diff --git a/src/HOL/Eisbach/method_closure.ML b/src/HOL/Eisbach/method_closure.ML --- a/src/HOL/Eisbach/method_closure.ML +++ b/src/HOL/Eisbach/method_closure.ML @@ -1,251 +1,251 @@ (* Title: HOL/Eisbach/method_closure.ML Author: Daniel Matichuk, NICTA/UNSW Facilities for treating method syntax as a closure, with abstraction over terms, facts and other methods. The 'method' command allows to define new proof methods by combining existing ones with their usual syntax. *) signature METHOD_CLOSURE = sig val apply_method: Proof.context -> string -> term list -> thm list list -> (Proof.context -> Method.method) list -> Proof.context -> thm list -> context_tactic val method: binding -> (binding * typ option * mixfix) list -> binding list -> binding list -> binding list -> Token.src -> local_theory -> string * local_theory val method_cmd: binding -> (binding * string option * mixfix) list -> binding list -> binding list -> binding list -> Token.src -> local_theory -> string * local_theory end; structure Method_Closure: METHOD_CLOSURE = struct (* auxiliary data for method definition *) structure Method_Definition = Proof_Data ( type T = (Proof.context -> Method.method) Symtab.table * (*dynamic methods*) (term list -> Proof.context -> Method.method) Symtab.table (*recursive methods*); fun init _ : T = (Symtab.empty, Symtab.empty); ); fun lookup_dynamic_method ctxt full_name = (case Symtab.lookup (#1 (Method_Definition.get ctxt)) full_name of SOME m => m ctxt | NONE => error ("Illegal use of internal Eisbach method: " ^ quote full_name)); val update_dynamic_method = Method_Definition.map o apfst o Symtab.update; fun get_recursive_method full_name ts ctxt = (case Symtab.lookup (#2 (Method_Definition.get ctxt)) full_name of SOME m => m ts ctxt | NONE => error ("Illegal use of internal Eisbach method: " ^ quote full_name)); val put_recursive_method = Method_Definition.map o apsnd o Symtab.update; (* stored method closures *) type closure = {vars: term list, named_thms: string list, methods: string list, body: Method.text}; structure Data = Generic_Data ( type T = closure Symtab.table; val empty: T = Symtab.empty; val extend = I; fun merge data : T = Symtab.merge (K true) data; ); fun get_closure ctxt name = (case Symtab.lookup (Data.get (Context.Proof ctxt)) name of SOME closure => closure | NONE => error ("Unknown Eisbach method: " ^ quote name)); fun put_closure binding (closure: closure) lthy = let val name = Local_Theory.full_name lthy binding; in lthy |> Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => Data.map (Symtab.update (name, {vars = map (Morphism.term phi) (#vars closure), named_thms = #named_thms closure, methods = #methods closure, body = (Method.map_source o map) (Token.transform phi) (#body closure)}))) end; (* instantiate and evaluate method text *) fun method_instantiate vars body ts ctxt = let val thy = Proof_Context.theory_of ctxt; val subst = fold (Pattern.match thy) (vars ~~ ts) (Vartab.empty, Vartab.empty); val morphism = Morphism.term_morphism "method_instantiate" (Envir.subst_term subst); val body' = (Method.map_source o map) (Token.transform morphism) body; in Method.evaluate_runtime body' ctxt end; (** apply method closure **) fun recursive_method full_name vars body ts = let val m = method_instantiate vars body in put_recursive_method (full_name, m) #> m ts end; fun apply_method ctxt method_name terms facts methods = let fun declare_facts (name :: names) (fact :: facts) = fold (Context.proof_map o Named_Theorems.add_thm name) fact #> declare_facts names facts | declare_facts _ [] = I | declare_facts [] (_ :: _) = error ("Excessive facts for method " ^ quote method_name); val {vars, named_thms, methods = method_args, body} = get_closure ctxt method_name; in declare_facts named_thms facts #> fold update_dynamic_method (method_args ~~ methods) #> recursive_method method_name vars body terms end; (** define method closure **) local fun setup_local_method binding lthy = let val full_name = Local_Theory.full_name lthy binding; fun dynamic_method ctxt = lookup_dynamic_method ctxt full_name; in lthy |> update_dynamic_method (full_name, K Method.fail) |> Method.local_setup binding (Scan.succeed dynamic_method) "(internal)" end; fun check_named_thm ctxt binding = let val bname = Binding.name_of binding; val pos = Binding.pos_of binding; val full_name = Named_Theorems.check ctxt (bname, pos); val parser: Method.modifier parser = Args.$$$ bname -- Args.colon >> K {init = I, attribute = Named_Theorems.add full_name, pos = pos}; in (full_name, parser) end; fun parse_term_args args = Args.context :|-- (fn ctxt => let val ctxt' = Proof_Context.set_mode (Proof_Context.mode_schematic) ctxt; fun parse T = (if T = propT then Syntax.parse_prop ctxt' else Syntax.parse_term ctxt') #> Type.constraint (Type_Infer.paramify_vars T); fun do_parse' T = Parse_Tools.name_term >> Parse_Tools.parse_val_cases (parse T); fun do_parse (Var (_, T)) = do_parse' T | do_parse (Free (_, T)) = do_parse' T | do_parse t = error ("Unexpected method parameter: " ^ Syntax.string_of_term ctxt' t); fun rep [] x = Scan.succeed [] x | rep (t :: ts) x = (do_parse t ::: rep ts) x; fun check ts = let val (ts, fs) = split_list ts; val ts' = Syntax.check_terms ctxt' ts |> Variable.polymorphic ctxt'; val _ = ListPair.app (fn (f, t) => f t) (fs, ts'); in ts' end; in Scan.lift (rep args) >> check end); fun parse_method_args method_args = let fun bind_method (name, text) ctxt = let val method = Method.evaluate_runtime text; val inner_update = method o update_dynamic_method (name, K (method ctxt)); in update_dynamic_method (name, inner_update) ctxt end; fun rep [] x = Scan.succeed [] x | rep (m :: ms) x = ((Method.text_closure >> pair m) ::: rep ms) x; in rep method_args >> fold bind_method end; fun gen_method add_fixes name vars uses declares methods source lthy = let val (uses_internal, lthy1) = lthy |> Proof_Context.concealed - |> Local_Theory.open_target |-> Proof_Context.private_scope + |> Local_Theory.begin_target |-> Proof_Context.private_scope |> Local_Theory.map_background_naming (Name_Space.add_path (Binding.name_of name)) |> Config.put Method.old_section_parser true |> fold setup_local_method methods |> fold_map (fn b => Named_Theorems.declare b "") uses; val (term_args, lthy2) = lthy1 |> add_fixes vars |-> fold_map Proof_Context.inferred_param |>> map Free; val (named_thms, modifiers) = map (check_named_thm lthy2) (declares @ uses) |> split_list; val method_args = map (Local_Theory.full_name lthy2) methods; fun parser args meth = apfst (Config.put_generic Method.old_section_parser true) #> (parse_term_args args -- parse_method_args method_args --| (Scan.depend (fn context => Scan.succeed (fold Named_Theorems.clear uses_internal context, ())) -- Method.sections modifiers)) >> (fn (ts, decl) => meth ts o decl); val full_name = Local_Theory.full_name lthy name; val lthy3 = lthy2 |> Method.local_setup (Binding.make (Binding.name_of name, Position.none)) (parser term_args (get_recursive_method full_name)) "(internal)" |> put_recursive_method (full_name, fn _ => fn _ => Method.fail); val (text, src) = Method.read_closure (Config.put Proof_Context.dynamic_facts_dummy true lthy3) source; val morphism = Variable.export_morphism lthy3 (lthy |> Proof_Context.transfer (Proof_Context.theory_of lthy3) |> fold Token.declare_maxidx src |> Variable.declare_maxidx (Variable.maxidx_of lthy3)); val text' = (Method.map_source o map) (Token.transform morphism) text; val term_args' = map (Morphism.term morphism) term_args; in lthy3 |> Local_Theory.close_target |> Proof_Context.restore_naming lthy |> put_closure name {vars = term_args', named_thms = named_thms, methods = method_args, body = text'} |> Method.local_setup name (Args.context :|-- (fn ctxt => let val {body, vars, ...} = get_closure ctxt full_name in parser vars (recursive_method full_name vars body) end)) "" |> pair full_name end; in val method = gen_method Proof_Context.add_fixes; val method_cmd = gen_method Proof_Context.add_fixes_cmd; end; val _ = Outer_Syntax.local_theory \<^command_keyword>\method\ "Eisbach method definition" (Parse.binding -- Parse.for_fixes -- ((Scan.optional (\<^keyword>\methods\ |-- Parse.!!! (Scan.repeat1 Parse.binding)) []) -- (Scan.optional (\<^keyword>\uses\ |-- Parse.!!! (Scan.repeat1 Parse.binding)) [])) -- (Scan.optional (\<^keyword>\declares\ |-- Parse.!!! (Scan.repeat1 Parse.binding)) []) -- Parse.!!! (\<^keyword>\=\ |-- Parse.args1 (K true)) >> (fn ((((name, vars), (methods, uses)), declares), src) => #2 o method_cmd name vars uses declares methods src)); end; diff --git a/src/HOL/Library/code_lazy.ML b/src/HOL/Library/code_lazy.ML --- a/src/HOL/Library/code_lazy.ML +++ b/src/HOL/Library/code_lazy.ML @@ -1,660 +1,660 @@ (* Author: Pascal Stoop, ETH Zurich Author: Andreas Lochbihler, Digital Asset *) signature CODE_LAZY = sig type lazy_info = {eagerT: typ, lazyT: typ, ctr: term, destr: term, lazy_ctrs: term list, case_lazy: term, active: bool, activate: theory -> theory, deactivate: theory -> theory}; val code_lazy_type: string -> theory -> theory val activate_lazy_type: string -> theory -> theory val deactivate_lazy_type: string -> theory -> theory val activate_lazy_types: theory -> theory val deactivate_lazy_types: theory -> theory val get_lazy_types: theory -> (string * lazy_info) list val print_lazy_types: theory -> unit val transform_code_eqs: Proof.context -> (thm * bool) list -> (thm * bool) list option end; structure Code_Lazy : CODE_LAZY = struct type lazy_info = {eagerT: typ, lazyT: typ, ctr: term, destr: term, lazy_ctrs: term list, case_lazy: term, active: bool, activate: theory -> theory, deactivate: theory -> theory}; fun map_active f {eagerT, lazyT, ctr, destr, lazy_ctrs, case_lazy, active, activate, deactivate} = { eagerT = eagerT, lazyT = lazyT, ctr = ctr, destr = destr, lazy_ctrs = lazy_ctrs, case_lazy = case_lazy, active = f active, activate = activate, deactivate = deactivate } structure Laziness_Data = Theory_Data( type T = lazy_info Symtab.table; val empty = Symtab.empty; val merge = Symtab.join (fn _ => fn (_, record) => record); val extend = I; ); fun fold_type type' tfree tvar typ = let fun go (Type (s, Ts)) = type' (s, map go Ts) | go (TFree T) = tfree T | go (TVar T) = tvar T in go typ end; fun read_typ lthy name = let val (s, Ts) = Proof_Context.read_type_name {proper = true, strict = true} lthy name |> dest_Type val (Ts', _) = Ctr_Sugar_Util.mk_TFrees (length Ts) lthy in Type (s, Ts') end fun mk_name prefix suffix name ctxt = Ctr_Sugar_Util.mk_fresh_names ctxt 1 (prefix ^ name ^ suffix) |>> hd; fun generate_typedef_name name ctxt = mk_name "" "_lazy" name ctxt; fun add_syntax_definition short_type_name eagerT lazyT lazy_ctr lthy = let val (name, _) = mk_name "lazy_" "" short_type_name lthy val freeT = HOLogic.unitT --> lazyT val lazyT' = Type (\<^type_name>\lazy\, [lazyT]) val def = Logic.all_const freeT $ absdummy freeT (Logic.mk_equals ( Free (name, (freeT --> eagerT)) $ Bound 0, lazy_ctr $ (Const (\<^const_name>\delay\, (freeT --> lazyT')) $ Bound 0))) - val (_, lthy') = Local_Theory.open_target lthy + val lthy' = Local_Theory.open_target lthy val ((t, (_, thm)), lthy') = Specification.definition NONE [] [] ((Thm.def_binding (Binding.name name), []), def) lthy' val lthy' = Specification.notation true ("", false) [(t, Mixfix.mixfix "_")] lthy' val lthy = Local_Theory.close_target lthy' val def_thm = singleton (Proof_Context.export lthy' lthy) thm in (def_thm, lthy) end; fun add_ctr_code raw_ctrs case_thms thy = let fun mk_case_certificate ctxt raw_thms = let val thms = raw_thms |> Conjunction.intr_balanced |> Thm.unvarify_global (Proof_Context.theory_of ctxt) |> Conjunction.elim_balanced (length raw_thms) |> map Simpdata.mk_meta_eq |> map Drule.zero_var_indexes val thm1 = case thms of thm :: _ => thm | _ => raise Empty val params = Term.add_free_names (Thm.prop_of thm1) []; val rhs = thm1 |> Thm.prop_of |> Logic.dest_equals |> fst |> strip_comb ||> fst o split_last |> list_comb val lhs = Free (singleton (Name.variant_list params) "case", fastype_of rhs); val assum = Thm.cterm_of ctxt (Logic.mk_equals (lhs, rhs)) in thms |> Conjunction.intr_balanced |> rewrite_rule ctxt [Thm.symmetric (Thm.assume assum)] |> Thm.implies_intr assum |> Thm.generalize ([], params) 0 |> Axclass.unoverload ctxt |> Thm.varifyT_global end val ctrs = map (apsnd (perhaps (try Logic.unvarifyT_global))) raw_ctrs val unover_ctrs = map (fn ctr as (_, fcT) => (Axclass.unoverload_const thy ctr, fcT)) ctrs in if can (Code.constrset_of_consts thy) unover_ctrs then thy |> Code.declare_datatype_global ctrs |> fold_rev (Code.add_eqn_global o rpair true) case_thms |> Code.declare_case_global (mk_case_certificate (Proof_Context.init_global thy) case_thms) else thy end; fun not_found s = error (s ^ " has not been added as lazy type"); fun validate_type_name thy type_name = let val lthy = Named_Target.theory_init thy val eager_type = read_typ lthy type_name val type_name = case eager_type of Type (s, _) => s | _ => raise Match in type_name end; fun set_active_lazy_type value eager_type_string thy = let val type_name = validate_type_name thy eager_type_string val x = case Symtab.lookup (Laziness_Data.get thy) type_name of NONE => not_found type_name | SOME x => x val new_x = map_active (K value) x val thy1 = if value = #active x then thy else if value then #activate x thy else #deactivate x thy in Laziness_Data.map (Symtab.update (type_name, new_x)) thy1 end; fun set_active_lazy_types value thy = let val lazy_type_names = Symtab.keys (Laziness_Data.get thy) fun fold_fun value type_name thy = let val x = case Symtab.lookup (Laziness_Data.get thy) type_name of SOME x => x | NONE => raise Match val new_x = map_active (K value) x val thy1 = if value = #active x then thy else if value then #activate x thy else #deactivate x thy in Laziness_Data.map (Symtab.update (type_name, new_x)) thy1 end in fold (fold_fun value) lazy_type_names thy end; (* code_lazy_type : string -> theory -> theory *) fun code_lazy_type eager_type_string thy = let val lthy = Named_Target.theory_init thy val eagerT = read_typ lthy eager_type_string val (type_name, type_args) = dest_Type eagerT val short_type_name = Long_Name.base_name type_name val _ = if Symtab.defined (Laziness_Data.get thy) type_name then error (type_name ^ " has already been added as lazy type.") else () val {case_thms, casex, ctrs, ...} = case Ctr_Sugar.ctr_sugar_of lthy type_name of SOME x => x | _ => error (type_name ^ " is not registered with free constructors.") fun substitute_ctr (old_T, new_T) ctr_T lthy = let val old_ctr_vars = map TVar (Term.add_tvarsT ctr_T []) val old_ctr_Ts = map TFree (Term.add_tfreesT ctr_T []) @ old_ctr_vars val (new_ctr_Ts, lthy') = Ctr_Sugar_Util.mk_TFrees (length old_ctr_Ts) lthy fun double_type_fold Ts = case Ts of (Type (_, Ts1), Type (_, Ts2)) => flat (map double_type_fold (Ts1 ~~ Ts2)) | (Type _, _) => raise Match | (_, Type _) => raise Match | Ts => [Ts] fun map_fun1 f = List.foldr (fn ((T1, T2), f) => fn T => if T = T1 then T2 else f T) f (double_type_fold (old_T, new_T)) val map_fun2 = AList.lookup (op =) (old_ctr_Ts ~~ new_ctr_Ts) #> the val map_fun = map_fun1 map_fun2 val new_ctr_type = fold_type Type (map_fun o TFree) (map_fun o TVar) ctr_T in (new_ctr_type, lthy') end val (short_lazy_type_name, lthy1) = generate_typedef_name short_type_name lthy fun mk_lazy_typedef (name, eager_type) lthy = let val binding = Binding.name name val (result, lthy1) = (Typedef.add_typedef { overloaded=false } (binding, rev (Term.add_tfreesT eager_type []), Mixfix.NoSyn) (Const (\<^const_name>\top\, Type (\<^type_name>\set\, [eager_type]))) NONE (fn ctxt => resolve_tac ctxt [@{thm UNIV_witness}] 1) - o (Local_Theory.open_target #> snd)) lthy + o Local_Theory.open_target) lthy in (binding, result, Local_Theory.close_target lthy1) end; val (typedef_binding, (_, info), lthy2) = mk_lazy_typedef (short_lazy_type_name, eagerT) lthy1 val lazy_type = Type (Local_Theory.full_name lthy2 typedef_binding, type_args) val (Abs_lazy, Rep_lazy) = let val info = fst info val Abs_name = #Abs_name info val Rep_name = #Rep_name info val Abs_type = eagerT --> lazy_type val Rep_type = lazy_type --> eagerT in (Const (Abs_name, Abs_type), Const (Rep_name, Rep_type)) end val Abs_inverse = #Abs_inverse (snd info) val Rep_inverse = #Rep_inverse (snd info) val (ctrs', lthy3) = List.foldr (fn (Const (s, T), (ctrs, lthy)) => let val (T', lthy') = substitute_ctr (body_type T, eagerT) T lthy in ((Const (s, T')) :: ctrs, lthy') end ) ([], lthy2) ctrs fun to_destr_eagerT typ = case typ of Type (\<^type_name>\fun\, [_, Type (\<^type_name>\fun\, Ts)]) => to_destr_eagerT (Type (\<^type_name>\fun\, Ts)) | Type (\<^type_name>\fun\, [T, _]) => T | _ => raise Match val (case', lthy4) = let val (eager_case, caseT) = dest_Const casex val (caseT', lthy') = substitute_ctr (to_destr_eagerT caseT, eagerT) caseT lthy3 in (Const (eager_case, caseT'), lthy') end val ctr_names = map (Long_Name.base_name o fst o dest_Const) ctrs val ((((lazy_ctr_name, lazy_sel_name), lazy_ctrs_name), lazy_case_name), ctxt) = lthy4 |> mk_name "Lazy_" "" short_type_name ||>> mk_name "unlazy_" "" short_type_name ||>> fold_map (mk_name "" "_Lazy") ctr_names ||>> mk_name "case_" "_lazy" short_type_name fun mk_def (name, T, rhs) lthy = let val binding = Binding.name name val ((_, (_, thm)), lthy1) = - Local_Theory.open_target lthy |> snd + Local_Theory.open_target lthy |> Specification.definition NONE [] [] ((Thm.def_binding binding, []), rhs) val lthy2 = Local_Theory.close_target lthy1 val def_thm = hd (Proof_Context.export lthy1 lthy2 [thm]) in ({binding = binding, const = Const (Local_Theory.full_name lthy2 binding, T), thm = def_thm}, lthy2) end; val lazy_datatype = Type (\<^type_name>\lazy\, [lazy_type]) val Lazy_type = lazy_datatype --> eagerT val unstr_type = eagerT --> lazy_datatype fun apply_bounds i n term = if n > i then apply_bounds i (n-1) (term $ Bound (n-1)) else term fun all_abs Ts t = Logic.list_all (map (pair Name.uu) Ts, t) fun mk_force t = Const (\<^const_name>\force\, lazy_datatype --> lazy_type) $ t fun mk_delay t = Const (\<^const_name>\delay\, (\<^typ>\unit\ --> lazy_type) --> lazy_datatype) $ t val lazy_ctr = all_abs [lazy_datatype] (Logic.mk_equals (Free (lazy_ctr_name, Lazy_type) $ Bound 0, Rep_lazy $ mk_force (Bound 0))) val (lazy_ctr_def, lthy5) = mk_def (lazy_ctr_name, Lazy_type, lazy_ctr) lthy4 val lazy_sel = all_abs [eagerT] (Logic.mk_equals (Free (lazy_sel_name, unstr_type) $ Bound 0, mk_delay (Abs (Name.uu, \<^typ>\unit\, Abs_lazy $ Bound 1)))) val (lazy_sel_def, lthy6) = mk_def (lazy_sel_name, unstr_type, lazy_sel) lthy5 fun mk_lazy_ctr (name, eager_ctr) = let val (_, ctrT) = dest_Const eager_ctr fun to_lazy_ctrT (Type (\<^type_name>\fun\, [T1, T2])) = T1 --> to_lazy_ctrT T2 | to_lazy_ctrT T = if T = eagerT then lazy_type else raise Match val lazy_ctrT = to_lazy_ctrT ctrT val argsT = binder_types ctrT val lhs = apply_bounds 0 (length argsT) (Free (name, lazy_ctrT)) val rhs = Abs_lazy $ apply_bounds 0 (length argsT) eager_ctr in mk_def (name, lazy_ctrT, all_abs argsT (Logic.mk_equals (lhs, rhs))) end val (lazy_ctrs_def, lthy7) = fold_map mk_lazy_ctr (lazy_ctrs_name ~~ ctrs') lthy6 val (lazy_case_def, lthy8) = let val (_, caseT) = dest_Const case' fun to_lazy_caseT (Type (\<^type_name>\fun\, [T1, T2])) = if T1 = eagerT then lazy_type --> T2 else T1 --> to_lazy_caseT T2 val lazy_caseT = to_lazy_caseT caseT val argsT = binder_types lazy_caseT val n = length argsT val lhs = apply_bounds 0 n (Free (lazy_case_name, lazy_caseT)) val rhs = apply_bounds 1 n case' $ (Rep_lazy $ Bound 0) in mk_def (lazy_case_name, lazy_caseT, all_abs argsT (Logic.mk_equals (lhs, rhs))) lthy7 end fun mk_thm ((name, term), thms) lthy = let val binding = Binding.name name fun tac {context, ...} = Simplifier.simp_tac (put_simpset HOL_basic_ss context addsimps thms) 1 val thm = Goal.prove lthy [] [] term tac val (_, lthy1) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> Local_Theory.note ((binding, []), [thm]) in (thm, Local_Theory.close_target lthy1) end fun mk_thms exec_list lthy = fold_map mk_thm exec_list lthy val mk_eq = HOLogic.mk_Trueprop o HOLogic.mk_eq val lazy_ctrs = map #const lazy_ctrs_def val eager_lazy_ctrs = ctrs' ~~ lazy_ctrs val (((((((Lazy_delay_eq_name, Rep_ctr_names), ctrs_lazy_names), force_sel_name), case_lazy_name), sel_lazy_name), case_ctrs_name), _) = lthy5 |> mk_name "Lazy_" "_delay" short_type_name ||>> fold_map (mk_name "Rep_" "_Lazy") ctr_names ||>> fold_map (mk_name "" "_conv_lazy") ctr_names ||>> mk_name "force_unlazy_" "" short_type_name ||>> mk_name "case_" "_conv_lazy" short_type_name ||>> mk_name "Lazy_" "_inverse" short_type_name ||>> fold_map (mk_name ("case_" ^ short_type_name ^ "_lazy_") "") ctr_names val mk_Lazy_delay_eq = (#const lazy_ctr_def $ mk_delay (Bound 0), Rep_lazy $ (Bound 0 $ \<^const>\Unity\)) |> mk_eq |> all_abs [\<^typ>\unit\ --> lazy_type] val (Lazy_delay_thm, lthy8a) = mk_thm ((Lazy_delay_eq_name, mk_Lazy_delay_eq), [#thm lazy_ctr_def, @{thm force_delay}]) lthy8 fun mk_lazy_ctr_eq (eager_ctr, lazy_ctr) = let val (_, ctrT) = dest_Const eager_ctr val argsT = binder_types ctrT val args = length argsT in (Rep_lazy $ apply_bounds 0 args lazy_ctr, apply_bounds 0 args eager_ctr) |> mk_eq |> all_abs argsT end val Rep_ctr_eqs = map mk_lazy_ctr_eq eager_lazy_ctrs val (Rep_ctr_thms, lthy8b) = mk_thms ((Rep_ctr_names ~~ Rep_ctr_eqs) ~~ (map (fn def => [#thm def, Abs_inverse, @{thm UNIV_I}]) lazy_ctrs_def) ) lthy8a fun mk_ctrs_lazy_eq (eager_ctr, lazy_ctr) = let val argsT = dest_Const eager_ctr |> snd |> binder_types val n = length argsT val lhs = apply_bounds 0 n eager_ctr val rhs = #const lazy_ctr_def $ (mk_delay (Abs (Name.uu, \<^typ>\unit\, apply_bounds 1 (n + 1) lazy_ctr))) in (lhs, rhs) |> mk_eq |> all_abs argsT end val ctrs_lazy_eq = map mk_ctrs_lazy_eq eager_lazy_ctrs val (ctrs_lazy_thms, lthy8c) = mk_thms ((ctrs_lazy_names ~~ ctrs_lazy_eq) ~~ map (fn thm => [Lazy_delay_thm, thm]) Rep_ctr_thms) lthy8b val force_sel_eq = (mk_force (#const lazy_sel_def $ Bound 0), Abs_lazy $ Bound 0) |> mk_eq |> all_abs [eagerT] val (force_sel_thm, lthy8d) = mk_thm ((force_sel_name, force_sel_eq), [#thm lazy_sel_def, @{thm force_delay}]) lthy8c val case_lazy_eq = let val (_, caseT) = case' |> dest_Const val argsT = binder_types caseT val n = length argsT val lhs = apply_bounds 0 n case' val rhs = apply_bounds 1 n (#const lazy_case_def) $ (mk_force (#const lazy_sel_def $ Bound 0)) in (lhs, rhs) |> mk_eq |> all_abs argsT end val (case_lazy_thm, lthy8e) = mk_thm ((case_lazy_name, case_lazy_eq), [#thm lazy_case_def, force_sel_thm, Abs_inverse, @{thm UNIV_I}]) lthy8d val sel_lazy_eq = (#const lazy_sel_def $ (#const lazy_ctr_def $ Bound 0), Bound 0) |> mk_eq |> all_abs [lazy_datatype] val (sel_lazy_thm, lthy8f) = mk_thm ((sel_lazy_name, sel_lazy_eq), [#thm lazy_sel_def, #thm lazy_ctr_def, Rep_inverse, @{thm delay_force}]) lthy8e fun mk_case_ctrs_eq (i, lazy_ctr) = let val lazy_case = #const lazy_case_def val (_, ctrT) = dest_Const lazy_ctr val ctr_argsT = binder_types ctrT val ctr_args_n = length ctr_argsT val (_, caseT) = dest_Const lazy_case val case_argsT = binder_types caseT fun n_bounds_from m n t = if n > 0 then n_bounds_from (m - 1) (n - 1) (t $ Bound (m - 1)) else t val case_argsT' = take (length case_argsT - 1) case_argsT val Ts = case_argsT' @ ctr_argsT val num_abs_types = length Ts val lhs = n_bounds_from num_abs_types (length case_argsT') lazy_case $ apply_bounds 0 ctr_args_n lazy_ctr val rhs = apply_bounds 0 ctr_args_n (Bound (num_abs_types - i - 1)) in (lhs, rhs) |> mk_eq |> all_abs Ts end val case_ctrs_eq = map_index mk_case_ctrs_eq lazy_ctrs val (case_ctrs_thms, lthy9) = mk_thms ((case_ctrs_name ~~ case_ctrs_eq) ~~ map2 (fn thm1 => fn thm2 => [#thm lazy_case_def, thm1, thm2]) Rep_ctr_thms case_thms ) lthy8f val (pat_def_thm, lthy10) = add_syntax_definition short_type_name eagerT lazy_type (#const lazy_ctr_def) lthy9 val add_lazy_ctrs = Code.declare_datatype_global [dest_Const (#const lazy_ctr_def)] val eager_ctrs = map (apsnd (perhaps (try Logic.unvarifyT_global)) o dest_Const) ctrs val add_eager_ctrs = fold Code.del_eqn_global ctrs_lazy_thms #> Code.declare_datatype_global eager_ctrs val add_code_eqs = fold (Code.add_eqn_global o rpair true) ([case_lazy_thm, sel_lazy_thm]) val add_lazy_ctr_thms = fold (Code.add_eqn_global o rpair true) ctrs_lazy_thms val add_lazy_case_thms = fold Code.del_eqn_global case_thms #> Code.add_eqn_global (case_lazy_thm, true) val add_eager_case_thms = Code.del_eqn_global case_lazy_thm #> fold (Code.add_eqn_global o rpair true) case_thms val delay_dummy_thm = (pat_def_thm RS @{thm symmetric}) |> Drule.infer_instantiate' lthy10 [SOME (Thm.cterm_of lthy10 (Const (\<^const_name>\Pure.dummy_pattern\, HOLogic.unitT --> lazy_type)))] |> Thm.generalize (map (fst o dest_TFree) type_args, []) (Variable.maxidx_of lthy10 + 1); val ctr_post = delay_dummy_thm :: map (fn thm => thm RS @{thm sym}) ctrs_lazy_thms val ctr_thms_abs = map (fn thm => Drule.abs_def (thm RS @{thm eq_reflection})) ctrs_lazy_thms val case_thm_abs = Drule.abs_def (case_lazy_thm RS @{thm eq_reflection}) val add_simps = Code_Preproc.map_pre (fn ctxt => ctxt addsimps (case_thm_abs :: ctr_thms_abs)) val del_simps = Code_Preproc.map_pre (fn ctxt => ctxt delsimps (case_thm_abs :: ctr_thms_abs)) val add_post = Code_Preproc.map_post (fn ctxt => ctxt addsimps ctr_post) val del_post = Code_Preproc.map_post (fn ctxt => ctxt delsimps ctr_post) in Local_Theory.exit_global lthy10 |> Laziness_Data.map (Symtab.update (type_name, {eagerT = eagerT, lazyT = lazy_type, ctr = #const lazy_ctr_def, destr = #const lazy_sel_def, lazy_ctrs = map #const lazy_ctrs_def, case_lazy = #const lazy_case_def, active = true, activate = add_lazy_ctrs #> add_lazy_ctr_thms #> add_lazy_case_thms #> add_simps #> add_post, deactivate = add_eager_ctrs #> add_eager_case_thms #> del_simps #> del_post})) |> add_lazy_ctrs |> add_ctr_code (map (dest_Const o #const) lazy_ctrs_def) case_ctrs_thms |> add_code_eqs |> add_lazy_ctr_thms |> add_simps |> add_post end; fun transform_code_eqs _ [] = NONE | transform_code_eqs ctxt eqs = let fun replace_ctr ctxt = let val thy = Proof_Context.theory_of ctxt val table = Laziness_Data.get thy in fn (s1, s2) => case Symtab.lookup table s1 of NONE => false | SOME x => #active x andalso s2 <> (#ctr x |> dest_Const |> fst) end val thy = Proof_Context.theory_of ctxt val table = Laziness_Data.get thy fun num_consts_fun (_, T) = let val s = body_type T |> dest_Type |> fst in if Symtab.defined table s then Ctr_Sugar.ctr_sugar_of ctxt s |> the |> #ctrs |> length else Code.get_type thy s |> fst |> snd |> length end val eqs = map (apfst (Thm.transfer thy)) eqs; val ((code_eqs, nbe_eqs), pure) = ((case hd eqs |> fst |> Thm.prop_of of Const (\<^const_name>\Pure.eq\, _) $ _ $ _ => (map (apfst (fn x => x RS @{thm meta_eq_to_obj_eq})) eqs, true) | _ => (eqs, false)) |> apfst (List.partition snd)) handle THM _ => (([], eqs), false) val to_original_eq = if pure then map (apfst (fn x => x RS @{thm eq_reflection})) else I in case Case_Converter.to_case ctxt (Case_Converter.replace_by_type replace_ctr) num_consts_fun (map fst code_eqs) of NONE => NONE | SOME thms => SOME (nbe_eqs @ map (rpair true) thms |> to_original_eq) end val activate_lazy_type = set_active_lazy_type true; val deactivate_lazy_type = set_active_lazy_type false; val activate_lazy_types = set_active_lazy_types true; val deactivate_lazy_types = set_active_lazy_types false; fun get_lazy_types thy = Symtab.dest (Laziness_Data.get thy) fun print_lazy_type thy (name, info : lazy_info) = let val ctxt = Proof_Context.init_global thy fun pretty_ctr ctr = let val argsT = dest_Const ctr |> snd |> binder_types in Pretty.block [ Syntax.pretty_term ctxt ctr, Pretty.brk 1, Pretty.block (Pretty.separate "" (map (Pretty.quote o Syntax.pretty_typ ctxt) argsT)) ] end in Pretty.block [ Pretty.str (name ^ (if #active info then "" else " (inactive)") ^ ":"), Pretty.brk 1, Pretty.block [ Syntax.pretty_typ ctxt (#eagerT info), Pretty.brk 1, Pretty.str "=", Pretty.brk 1, Syntax.pretty_term ctxt (#ctr info), Pretty.brk 1, Pretty.block [ Pretty.str "(", Syntax.pretty_term ctxt (#destr info), Pretty.str ":", Pretty.brk 1, Syntax.pretty_typ ctxt (Type (\<^type_name>\lazy\, [#lazyT info])), Pretty.str ")" ] ], Pretty.fbrk, Pretty.keyword2 "and", Pretty.brk 1, Pretty.block ([ Syntax.pretty_typ ctxt (#lazyT info), Pretty.brk 1, Pretty.str "=", Pretty.brk 1] @ Pretty.separate " |" (map pretty_ctr (#lazy_ctrs info)) @ [ Pretty.fbrk, Pretty.keyword2 "for", Pretty.brk 1, Pretty.str "case:", Pretty.brk 1, Syntax.pretty_term ctxt (#case_lazy info) ]) ] end; fun print_lazy_types thy = let fun cmp ((name1, _), (name2, _)) = string_ord (name1, name2) val infos = Laziness_Data.get thy |> Symtab.dest |> map (apfst Long_Name.base_name) |> sort cmp in Pretty.writeln_chunks (map (print_lazy_type thy) infos) end; val _ = Outer_Syntax.command \<^command_keyword>\code_lazy_type\ "make a lazy copy of the datatype and activate substitution" (Parse.binding >> (fn b => Toplevel.theory (Binding.name_of b |> code_lazy_type))); val _ = Outer_Syntax.command \<^command_keyword>\activate_lazy_type\ "activate substitution on a specific (lazy) type" (Parse.binding >> (fn b => Toplevel.theory (Binding.name_of b |> activate_lazy_type))); val _ = Outer_Syntax.command \<^command_keyword>\deactivate_lazy_type\ "deactivate substitution on a specific (lazy) type" (Parse.binding >> (fn b => Toplevel.theory (Binding.name_of b |> deactivate_lazy_type))); val _ = Outer_Syntax.command \<^command_keyword>\activate_lazy_types\ "activate substitution on all (lazy) types" (pair (Toplevel.theory activate_lazy_types)); val _ = Outer_Syntax.command \<^command_keyword>\deactivate_lazy_types\ "deactivate substitution on all (lazy) type" (pair (Toplevel.theory deactivate_lazy_types)); val _ = Outer_Syntax.command \<^command_keyword>\print_lazy_types\ "print the types that have been declared as lazy and their substitution state" (pair (Toplevel.theory (tap print_lazy_types))); end \ No newline at end of file diff --git a/src/HOL/Library/datatype_records.ML b/src/HOL/Library/datatype_records.ML --- a/src/HOL/Library/datatype_records.ML +++ b/src/HOL/Library/datatype_records.ML @@ -1,313 +1,312 @@ signature DATATYPE_RECORDS = sig type ctr_options = string -> bool type ctr_options_cmd = Proof.context -> string -> bool val default_ctr_options: ctr_options val default_ctr_options_cmd: ctr_options_cmd val mk_update_defs: string -> local_theory -> local_theory val record: binding -> ctr_options -> (binding option * (typ * sort)) list -> (binding * typ) list -> local_theory -> local_theory val record_cmd: binding -> ctr_options_cmd -> (binding option * (string * string option)) list -> (binding * string) list -> local_theory -> local_theory val setup: theory -> theory end structure Datatype_Records : DATATYPE_RECORDS = struct type ctr_options = string -> bool type ctr_options_cmd = Proof.context -> string -> bool val default_ctr_options = Plugin_Name.default_filter val default_ctr_options_cmd = K Plugin_Name.default_filter type data = string Symtab.table structure Data = Theory_Data ( type T = data val empty = Symtab.empty val merge = Symtab.merge op = val extend = I ) fun mk_eq_dummy (lhs, rhs) = Const (\<^const_name>\HOL.eq\, dummyT --> dummyT --> \<^typ>\bool\) $ lhs $ rhs val dummify = map_types (K dummyT) fun repeat_split_tac ctxt thm = REPEAT_ALL_NEW (CHANGED o Splitter.split_tac ctxt [thm]) fun mk_update_defs typ_name lthy = let val short_name = Long_Name.base_name typ_name val {ctrs, casex, selss, split, sel_thmss, injects, ...} = the (Ctr_Sugar.ctr_sugar_of lthy typ_name) val ctr = case ctrs of [ctr] => ctr | _ => error "Datatype_Records.mk_update_defs: expected only single constructor" val sels = case selss of [sels] => sels | _ => error "Datatype_Records.mk_update_defs: expected selectors" val sels_dummy = map dummify sels val ctr_dummy = dummify ctr val casex_dummy = dummify casex val len = length sels val simp_thms = flat sel_thmss @ injects fun mk_name sel = Binding.name ("update_" ^ Long_Name.base_name (fst (dest_Const sel))) val thms_binding = (Binding.name "record_simps", @{attributes [simp]}) fun mk_t idx = let val body = fold_rev (fn pos => fn t => t $ (if len - pos = idx + 1 then Bound len $ Bound pos else Bound pos)) (0 upto len - 1) ctr_dummy |> fold_rev (fn idx => fn t => Abs ("x" ^ Value.print_int idx, dummyT, t)) (1 upto len) in Abs ("f", dummyT, casex_dummy $ body) end fun simp_only_tac ctxt = REPEAT_ALL_NEW (resolve_tac ctxt @{thms impI allI}) THEN' asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps simp_thms) fun prove ctxt defs ts n = let val t = nth ts n val sel_dummy = nth sels_dummy n val t_dummy = dummify t fun tac {context = ctxt, ...} = Goal.conjunction_tac 1 THEN Local_Defs.unfold_tac ctxt defs THEN PARALLEL_ALLGOALS (repeat_split_tac ctxt split THEN' simp_only_tac ctxt) val sel_upd_same_thm = let val ([f, x], ctxt') = Variable.add_fixes ["f", "x"] ctxt val f = Free (f, dummyT) val x = Free (x, dummyT) val lhs = sel_dummy $ (t_dummy $ f $ x) val rhs = f $ (sel_dummy $ x) val prop = Syntax.check_term ctxt' (HOLogic.mk_Trueprop (mk_eq_dummy (lhs, rhs))) in [Goal.prove_future ctxt' [] [] prop tac] |> Variable.export ctxt' ctxt end val sel_upd_diff_thms = let val ([f, x], ctxt') = Variable.add_fixes ["f", "x"] ctxt val f = Free (f, dummyT) val x = Free (x, dummyT) fun lhs sel = sel $ (t_dummy $ f $ x) fun rhs sel = sel $ x fun eq sel = (lhs sel, rhs sel) fun is_n i = i = n val props = sels_dummy ~~ (0 upto len - 1) |> filter_out (is_n o snd) |> map (HOLogic.mk_Trueprop o mk_eq_dummy o eq o fst) |> Syntax.check_terms ctxt' in if length props > 0 then Goal.prove_common ctxt' (SOME ~1) [] [] props tac |> Variable.export ctxt' ctxt else [] end val upd_comp_thm = let val ([f, g, x], ctxt') = Variable.add_fixes ["f", "g", "x"] ctxt val f = Free (f, dummyT) val g = Free (g, dummyT) val x = Free (x, dummyT) val lhs = t_dummy $ f $ (t_dummy $ g $ x) val rhs = t_dummy $ Abs ("a", dummyT, f $ (g $ Bound 0)) $ x val prop = Syntax.check_term ctxt' (HOLogic.mk_Trueprop (mk_eq_dummy (lhs, rhs))) in [Goal.prove_future ctxt' [] [] prop tac] |> Variable.export ctxt' ctxt end val upd_comm_thms = let fun prop i ctxt = let val ([f, g, x], ctxt') = Variable.variant_fixes ["f", "g", "x"] ctxt val self = t_dummy $ Free (f, dummyT) val other = dummify (nth ts i) $ Free (g, dummyT) val lhs = other $ (self $ Free (x, dummyT)) val rhs = self $ (other $ Free (x, dummyT)) in (HOLogic.mk_Trueprop (mk_eq_dummy (lhs, rhs)), ctxt') end val (props, ctxt') = fold_map prop (0 upto n - 1) ctxt val props = Syntax.check_terms ctxt' props in if length props > 0 then Goal.prove_common ctxt' (SOME ~1) [] [] props tac |> Variable.export ctxt' ctxt else [] end val upd_sel_thm = let val ([x], ctxt') = Variable.add_fixes ["x"] ctxt val lhs = t_dummy $ Abs("_", dummyT, (sel_dummy $ Free(x, dummyT))) $ Free (x, dummyT) val rhs = Free (x, dummyT) val prop = Syntax.check_term ctxt (HOLogic.mk_Trueprop (mk_eq_dummy (lhs, rhs))) in [Goal.prove_future ctxt [] [] prop tac] |> Variable.export ctxt' ctxt end in sel_upd_same_thm @ sel_upd_diff_thms @ upd_comp_thm @ upd_comm_thms @ upd_sel_thm end fun define name t = Local_Theory.define ((name, NoSyn), ((Binding.empty, @{attributes [datatype_record_update, code]}),t)) #> apfst (apsnd snd) val (updates, (lthy'', lthy')) = lthy |> Local_Theory.open_target - |> snd |> Local_Theory.map_background_naming (Name_Space.qualified_path false (Binding.name short_name)) |> @{fold_map 2} define (map mk_name sels) (Syntax.check_terms lthy (map mk_t (0 upto len - 1))) ||> `Local_Theory.close_target val phi = Proof_Context.export_morphism lthy' lthy'' val (update_ts, update_defs) = split_list updates |>> map (Morphism.term phi) ||> map (Morphism.thm phi) val thms = flat (map (prove lthy'' update_defs update_ts) (0 upto len-1)) fun insert sel = Symtab.insert op = (fst (dest_Const sel), Local_Theory.full_name lthy' (mk_name sel)) in lthy'' |> Local_Theory.map_background_naming (Name_Space.mandatory_path short_name) |> Local_Theory.note (thms_binding, thms) |> snd |> Local_Theory.restore_background_naming lthy |> Local_Theory.background_theory (Data.map (fold insert sels)) end fun record binding opts tyargs args lthy = let val constructor = (((Binding.empty, Binding.map_name (fn c => "make_" ^ c) binding), args), NoSyn) val datatyp = ((tyargs, binding), NoSyn) val dtspec = ((opts, false), [(((datatyp, [constructor]), (Binding.empty, Binding.empty, Binding.empty)), [])]) val lthy' = BNF_FP_Def_Sugar.co_datatypes BNF_Util.Least_FP BNF_LFP.construct_lfp dtspec lthy |> mk_update_defs (Local_Theory.full_name lthy binding) in lthy' end fun record_cmd binding opts tyargs args lthy = record binding (opts lthy) (map (apsnd (apfst (Syntax.parse_typ lthy) o apsnd (Typedecl.read_constraint lthy))) tyargs) (map (apsnd (Syntax.parse_typ lthy)) args) lthy (* syntax *) (* copied and adapted from record.ML *) val read_const = dest_Const oo Proof_Context.read_const {proper = true, strict = true} fun field_tr ((Const (\<^syntax_const>\_datatype_field\, _) $ Const (name, _) $ arg)) = (name, arg) | field_tr t = raise TERM ("field_tr", [t]); fun fields_tr (Const (\<^syntax_const>\_datatype_fields\, _) $ t $ u) = field_tr t :: fields_tr u | fields_tr t = [field_tr t]; fun record_fields_tr ctxt t = let val assns = map (apfst (read_const ctxt)) (fields_tr t) val typ_name = snd (fst (hd assns)) |> domain_type |> dest_Type |> fst val assns' = map (apfst fst) assns val {ctrs, selss, ...} = the (Ctr_Sugar.ctr_sugar_of ctxt typ_name) val ctr = case ctrs of [ctr] => ctr | _ => error "BNF_Record.record_fields_tr: expected only single constructor" val sels = case selss of [sels] => sels | _ => error "BNF_Record.record_fields_tr: expected selectors" val ctr_dummy = Const (fst (dest_Const ctr), dummyT) fun mk_arg name = case AList.lookup op = assns' name of NONE => error ("BNF_Record.record_fields_tr: missing field " ^ name) | SOME t => t in if length assns = length sels then list_comb (ctr_dummy, map (mk_arg o fst o dest_Const) sels) else error ("BNF_Record.record_fields_tr: expected " ^ Value.print_int (length sels) ^ " field(s)") end fun field_update_tr ctxt (Const (\<^syntax_const>\_datatype_field_update\, _) $ Const (name, _) $ arg) = let val thy = Proof_Context.theory_of ctxt val (name, _) = read_const ctxt name in case Symtab.lookup (Data.get thy) name of NONE => raise Fail ("not a valid record field: " ^ name) | SOME s => Const (s, dummyT) $ Abs (Name.uu_, dummyT, arg) end | field_update_tr _ t = raise TERM ("field_update_tr", [t]); fun field_updates_tr ctxt (Const (\<^syntax_const>\_datatype_field_updates\, _) $ t $ u) = field_update_tr ctxt t :: field_updates_tr ctxt u | field_updates_tr ctxt t = [field_update_tr ctxt t]; fun record_tr ctxt [t] = record_fields_tr ctxt t | record_tr _ ts = raise TERM ("record_tr", ts); fun record_update_tr ctxt [t, u] = fold (curry op $) (field_updates_tr ctxt u) t | record_update_tr _ ts = raise TERM ("record_update_tr", ts); val parse_ctr_options = Scan.optional (\<^keyword>\(\ |-- Parse.list1 (Plugin_Name.parse_filter >> K) --| \<^keyword>\)\ >> (fn fs => fold I fs default_ctr_options_cmd)) default_ctr_options_cmd val parser = (parse_ctr_options -- BNF_Util.parse_type_args_named_constrained -- Parse.binding) -- (\<^keyword>\=\ |-- Scan.repeat1 (Parse.binding -- (\<^keyword>\::\ |-- Parse.!!! Parse.typ))) val _ = Outer_Syntax.local_theory \<^command_keyword>\datatype_record\ "Defines a record based on the BNF/datatype machinery" (parser >> (fn (((ctr_options, tyargs), binding), args) => record_cmd binding ctr_options tyargs args)) val setup = (Sign.parse_translation [(\<^syntax_const>\_datatype_record_update\, record_update_tr), (\<^syntax_const>\_datatype_record\, record_tr)]); end \ No newline at end of file diff --git a/src/HOL/Tools/BNF/bnf_def.ML b/src/HOL/Tools/BNF/bnf_def.ML --- a/src/HOL/Tools/BNF/bnf_def.ML +++ b/src/HOL/Tools/BNF/bnf_def.ML @@ -1,2161 +1,2161 @@ (* Title: HOL/Tools/BNF/bnf_def.ML Author: Dmitriy Traytel, TU Muenchen Author: Jasmin Blanchette, TU Muenchen Author: Martin Desharnais, TU Muenchen Copyright 2012, 2013, 2014 Definition of bounded natural functors. *) signature BNF_DEF = sig type bnf type nonemptiness_witness = {I: int list, wit: term, prop: thm list} val morph_bnf: morphism -> bnf -> bnf val morph_bnf_defs: morphism -> bnf -> bnf val permute_deads: (typ list -> typ list) -> bnf -> bnf val transfer_bnf: theory -> bnf -> bnf val bnf_of: Proof.context -> string -> bnf option val bnf_of_global: theory -> string -> bnf option val bnf_interpretation: string -> (bnf -> local_theory -> local_theory) -> theory -> theory val interpret_bnf: (string -> bool) -> bnf -> local_theory -> local_theory val register_bnf_raw: string -> bnf -> local_theory -> local_theory val register_bnf: (string -> bool) -> string -> bnf -> local_theory -> local_theory val name_of_bnf: bnf -> binding val T_of_bnf: bnf -> typ val live_of_bnf: bnf -> int val lives_of_bnf: bnf -> typ list val dead_of_bnf: bnf -> int val deads_of_bnf: bnf -> typ list val bd_of_bnf: bnf -> term val nwits_of_bnf: bnf -> int val mapN: string val predN: string val relN: string val setN: string val mk_setN: int -> string val mk_witN: int -> string val map_of_bnf: bnf -> term val pred_of_bnf: bnf -> term val rel_of_bnf: bnf -> term val sets_of_bnf: bnf -> term list val mk_T_of_bnf: typ list -> typ list -> bnf -> typ val mk_bd_of_bnf: typ list -> typ list -> bnf -> term val mk_map_of_bnf: typ list -> typ list -> typ list -> bnf -> term val mk_pred_of_bnf: typ list -> typ list -> bnf -> term val mk_rel_of_bnf: typ list -> typ list -> typ list -> bnf -> term val mk_sets_of_bnf: typ list list -> typ list list -> bnf -> term list val mk_wits_of_bnf: typ list list -> typ list list -> bnf -> (int list * term) list val bd_Card_order_of_bnf: bnf -> thm val bd_Cinfinite_of_bnf: bnf -> thm val bd_Cnotzero_of_bnf: bnf -> thm val bd_card_order_of_bnf: bnf -> thm val bd_cinfinite_of_bnf: bnf -> thm val collect_set_map_of_bnf: bnf -> thm val in_bd_of_bnf: bnf -> thm val in_cong_of_bnf: bnf -> thm val in_mono_of_bnf: bnf -> thm val in_rel_of_bnf: bnf -> thm val inj_map_of_bnf: bnf -> thm val inj_map_strong_of_bnf: bnf -> thm val le_rel_OO_of_bnf: bnf -> thm val map_comp0_of_bnf: bnf -> thm val map_comp_of_bnf: bnf -> thm val map_cong0_of_bnf: bnf -> thm val map_cong_of_bnf: bnf -> thm val map_cong_pred_of_bnf: bnf -> thm val map_cong_simp_of_bnf: bnf -> thm val map_def_of_bnf: bnf -> thm val map_id0_of_bnf: bnf -> thm val map_id_of_bnf: bnf -> thm val map_ident0_of_bnf: bnf -> thm val map_ident_of_bnf: bnf -> thm val map_transfer_of_bnf: bnf -> thm val pred_cong0_of_bnf: bnf -> thm val pred_cong_of_bnf: bnf -> thm val pred_cong_simp_of_bnf: bnf -> thm val pred_def_of_bnf: bnf -> thm val pred_map_of_bnf: bnf -> thm val pred_mono_strong0_of_bnf: bnf -> thm val pred_mono_strong_of_bnf: bnf -> thm val pred_mono_of_bnf: bnf -> thm val pred_set_of_bnf: bnf -> thm val pred_rel_of_bnf: bnf -> thm val pred_transfer_of_bnf: bnf -> thm val pred_True_of_bnf: bnf -> thm val rel_Grp_of_bnf: bnf -> thm val rel_OO_Grp_of_bnf: bnf -> thm val rel_OO_of_bnf: bnf -> thm val rel_cong0_of_bnf: bnf -> thm val rel_cong_of_bnf: bnf -> thm val rel_cong_simp_of_bnf: bnf -> thm val rel_conversep_of_bnf: bnf -> thm val rel_def_of_bnf: bnf -> thm val rel_eq_of_bnf: bnf -> thm val rel_flip_of_bnf: bnf -> thm val rel_map_of_bnf: bnf -> thm list val rel_mono_of_bnf: bnf -> thm val rel_mono_strong0_of_bnf: bnf -> thm val rel_mono_strong_of_bnf: bnf -> thm val rel_eq_onp_of_bnf: bnf -> thm val rel_refl_of_bnf: bnf -> thm val rel_refl_strong_of_bnf: bnf -> thm val rel_reflp_of_bnf: bnf -> thm val rel_symp_of_bnf: bnf -> thm val rel_transfer_of_bnf: bnf -> thm val rel_transp_of_bnf: bnf -> thm val set_bd_of_bnf: bnf -> thm list val set_defs_of_bnf: bnf -> thm list val set_map0_of_bnf: bnf -> thm list val set_map_of_bnf: bnf -> thm list val set_transfer_of_bnf: bnf -> thm list val wit_thms_of_bnf: bnf -> thm list val wit_thmss_of_bnf: bnf -> thm list list val mk_map: int -> typ list -> typ list -> term -> term val mk_pred: typ list -> term -> term val mk_rel: int -> typ list -> typ list -> term -> term val mk_set: typ list -> term -> term val build_map: Proof.context -> typ list -> typ list -> (typ * typ -> term) -> typ * typ -> term val build_rel: (string * (int * term)) list -> Proof.context -> typ list -> typ list -> (typ * typ -> term) -> typ * typ -> term val build_set: Proof.context -> typ -> typ -> term val flatten_type_args_of_bnf: bnf -> 'a -> 'a list -> 'a list val map_flattened_map_args: Proof.context -> string -> (term list -> 'a list) -> term list -> 'a list val mk_witness: int list * term -> thm list -> nonemptiness_witness val mk_wit_goals: term list -> term list -> term list -> int list * term -> term list val minimize_wits: (''a list * 'b) list -> (''a list * 'b) list val wits_of_bnf: bnf -> nonemptiness_witness list val zip_axioms: 'a -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a -> 'a list datatype inline_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline datatype fact_policy = Dont_Note | Note_Some | Note_All val bnf_internals: bool Config.T val bnf_timing: bool Config.T val user_policy: fact_policy -> Proof.context -> fact_policy val note_bnf_thms: fact_policy -> (binding -> binding) -> binding -> bnf -> local_theory -> bnf * local_theory val note_bnf_defs: bnf -> local_theory -> bnf * local_theory val print_bnfs: Proof.context -> unit val prepare_def: inline_policy -> (Proof.context -> fact_policy) -> bool -> (binding -> binding) -> (Proof.context -> 'a -> typ) -> (Proof.context -> 'b -> term) -> typ list option -> binding -> binding -> binding -> binding list -> ((((((binding * 'a) * 'b) * 'b list) * 'b) * 'b list) * 'b option) * 'b option -> Proof.context -> string * term list * ((Proof.context -> thm list -> tactic) option * term list list) * ((thm list -> thm list list) -> thm list list -> Proof.context -> bnf * local_theory) * local_theory * thm list val define_bnf_consts: inline_policy -> fact_policy -> bool -> typ list option -> binding -> binding -> binding -> binding list -> ((((((binding * typ) * term) * term list) * term) * term list) * term option) * term option -> local_theory -> ((typ list * typ list * typ list * typ) * (term * term list * term * (int list * term) list * term * term) * (thm * thm list * thm * thm list * thm * thm) * ((typ list -> typ list -> typ list -> term) * (typ list -> typ list -> term -> term) * (typ list -> typ list -> typ -> typ) * (typ list -> typ list -> typ list -> term) * (typ list -> typ list -> term) * (typ list -> typ list -> typ list -> term) * (typ list -> typ list -> term))) * local_theory val bnf_def: inline_policy -> (Proof.context -> fact_policy) -> bool -> (binding -> binding) -> (Proof.context -> tactic) list -> (Proof.context -> tactic) -> typ list option -> binding -> binding -> binding -> binding list -> ((((((binding * typ) * term) * term list) * term) * term list) * term option) * term option -> local_theory -> bnf * local_theory val bnf_cmd: (((((((binding * string) * string) * string list) * string) * string list) * string option) * string option) * (Proof.context -> Plugin_Name.filter) -> Proof.context -> Proof.state end; structure BNF_Def : BNF_DEF = struct open BNF_Util open BNF_Tactics open BNF_Def_Tactics val fundefcong_attrs = @{attributes [fundef_cong]}; val mono_attrs = @{attributes [mono]}; type axioms = { map_id0: thm, map_comp0: thm, map_cong0: thm, set_map0: thm list, bd_card_order: thm, bd_cinfinite: thm, set_bd: thm list, le_rel_OO: thm, rel_OO_Grp: thm, pred_set: thm }; fun mk_axioms' (((((((((id, comp), cong), map), c_o), cinf), set_bd), le_rel_OO), rel), pred) = {map_id0 = id, map_comp0 = comp, map_cong0 = cong, set_map0 = map, bd_card_order = c_o, bd_cinfinite = cinf, set_bd = set_bd, le_rel_OO = le_rel_OO, rel_OO_Grp = rel, pred_set = pred}; fun dest_cons [] = raise List.Empty | dest_cons (x :: xs) = (x, xs); fun mk_axioms n thms = thms |> map the_single |> dest_cons ||>> dest_cons ||>> dest_cons ||>> chop n ||>> dest_cons ||>> dest_cons ||>> chop n ||>> dest_cons ||>> dest_cons ||> the_single |> mk_axioms'; fun zip_axioms mid mcomp mcong smap bdco bdinf sbd le_rel_OO rel pred = [mid, mcomp, mcong] @ smap @ [bdco, bdinf] @ sbd @ [le_rel_OO, rel, pred]; fun map_axioms f {map_id0, map_comp0, map_cong0, set_map0, bd_card_order, bd_cinfinite, set_bd, le_rel_OO, rel_OO_Grp, pred_set} = {map_id0 = f map_id0, map_comp0 = f map_comp0, map_cong0 = f map_cong0, set_map0 = map f set_map0, bd_card_order = f bd_card_order, bd_cinfinite = f bd_cinfinite, set_bd = map f set_bd, le_rel_OO = f le_rel_OO, rel_OO_Grp = f rel_OO_Grp, pred_set = f pred_set}; val morph_axioms = map_axioms o Morphism.thm; type defs = { map_def: thm, set_defs: thm list, rel_def: thm, pred_def: thm } fun mk_defs map sets rel pred = {map_def = map, set_defs = sets, rel_def = rel, pred_def = pred}; fun map_defs f {map_def, set_defs, rel_def, pred_def} = {map_def = f map_def, set_defs = map f set_defs, rel_def = f rel_def, pred_def = f pred_def}; val morph_defs = map_defs o Morphism.thm; type facts = { bd_Card_order: thm, bd_Cinfinite: thm, bd_Cnotzero: thm, collect_set_map: thm lazy, in_bd: thm lazy, in_cong: thm lazy, in_mono: thm lazy, in_rel: thm lazy, inj_map: thm lazy, inj_map_strong: thm lazy, map_comp: thm lazy, map_cong: thm lazy, map_cong_simp: thm lazy, map_cong_pred: thm lazy, map_id: thm lazy, map_ident0: thm lazy, map_ident: thm lazy, map_transfer: thm lazy, rel_eq: thm lazy, rel_flip: thm lazy, set_map: thm lazy list, rel_cong0: thm lazy, rel_cong: thm lazy, rel_cong_simp: thm lazy, rel_map: thm list lazy, rel_mono: thm lazy, rel_mono_strong0: thm lazy, rel_mono_strong: thm lazy, set_transfer: thm list lazy, rel_Grp: thm lazy, rel_conversep: thm lazy, rel_OO: thm lazy, rel_refl: thm lazy, rel_refl_strong: thm lazy, rel_reflp: thm lazy, rel_symp: thm lazy, rel_transp: thm lazy, rel_transfer: thm lazy, rel_eq_onp: thm lazy, pred_transfer: thm lazy, pred_True: thm lazy, pred_map: thm lazy, pred_rel: thm lazy, pred_mono_strong0: thm lazy, pred_mono_strong: thm lazy, pred_mono: thm lazy, pred_cong0: thm lazy, pred_cong: thm lazy, pred_cong_simp: thm lazy }; fun mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_map in_bd in_cong in_mono in_rel inj_map inj_map_strong map_comp map_cong map_cong_simp map_cong_pred map_id map_ident0 map_ident map_transfer rel_eq rel_flip set_map rel_cong0 rel_cong rel_cong_simp rel_map rel_mono rel_mono_strong0 rel_mono_strong set_transfer rel_Grp rel_conversep rel_OO rel_refl rel_refl_strong rel_reflp rel_symp rel_transp rel_transfer rel_eq_onp pred_transfer pred_True pred_map pred_rel pred_mono_strong0 pred_mono_strong pred_mono pred_cong0 pred_cong pred_cong_simp = { bd_Card_order = bd_Card_order, bd_Cinfinite = bd_Cinfinite, bd_Cnotzero = bd_Cnotzero, collect_set_map = collect_set_map, in_bd = in_bd, in_cong = in_cong, in_mono = in_mono, in_rel = in_rel, inj_map = inj_map, inj_map_strong = inj_map_strong, map_comp = map_comp, map_cong = map_cong, map_cong_simp = map_cong_simp, map_cong_pred = map_cong_pred, map_id = map_id, map_ident0 = map_ident0, map_ident = map_ident, map_transfer = map_transfer, rel_eq = rel_eq, rel_flip = rel_flip, set_map = set_map, rel_cong0 = rel_cong0, rel_cong = rel_cong, rel_cong_simp = rel_cong_simp, rel_map = rel_map, rel_mono = rel_mono, rel_mono_strong0 = rel_mono_strong0, rel_mono_strong = rel_mono_strong, rel_transfer = rel_transfer, rel_Grp = rel_Grp, rel_conversep = rel_conversep, rel_OO = rel_OO, rel_refl = rel_refl, rel_refl_strong = rel_refl_strong, rel_reflp = rel_reflp, rel_symp = rel_symp, rel_transp = rel_transp, set_transfer = set_transfer, rel_eq_onp = rel_eq_onp, pred_transfer = pred_transfer, pred_True = pred_True, pred_map = pred_map, pred_rel = pred_rel, pred_mono_strong0 = pred_mono_strong0, pred_mono_strong = pred_mono_strong, pred_mono = pred_mono, pred_cong0 = pred_cong0, pred_cong = pred_cong, pred_cong_simp = pred_cong_simp}; fun map_facts f { bd_Card_order, bd_Cinfinite, bd_Cnotzero, collect_set_map, in_bd, in_cong, in_mono, in_rel, inj_map, inj_map_strong, map_comp, map_cong, map_cong_simp, map_cong_pred, map_id, map_ident0, map_ident, map_transfer, rel_eq, rel_flip, set_map, rel_cong0, rel_cong, rel_cong_simp, rel_map, rel_mono, rel_mono_strong0, rel_mono_strong, rel_transfer, rel_Grp, rel_conversep, rel_OO, rel_refl, rel_refl_strong, rel_reflp, rel_symp, rel_transp, set_transfer, rel_eq_onp, pred_transfer, pred_True, pred_map, pred_rel, pred_mono_strong0, pred_mono_strong, pred_mono, pred_cong0, pred_cong, pred_cong_simp} = {bd_Card_order = f bd_Card_order, bd_Cinfinite = f bd_Cinfinite, bd_Cnotzero = f bd_Cnotzero, collect_set_map = Lazy.map f collect_set_map, in_bd = Lazy.map f in_bd, in_cong = Lazy.map f in_cong, in_mono = Lazy.map f in_mono, in_rel = Lazy.map f in_rel, inj_map = Lazy.map f inj_map, inj_map_strong = Lazy.map f inj_map_strong, map_comp = Lazy.map f map_comp, map_cong = Lazy.map f map_cong, map_cong_simp = Lazy.map f map_cong_simp, map_cong_pred = Lazy.map f map_cong_pred, map_id = Lazy.map f map_id, map_ident0 = Lazy.map f map_ident0, map_ident = Lazy.map f map_ident, map_transfer = Lazy.map f map_transfer, rel_eq = Lazy.map f rel_eq, rel_flip = Lazy.map f rel_flip, set_map = map (Lazy.map f) set_map, rel_cong0 = Lazy.map f rel_cong0, rel_cong = Lazy.map f rel_cong, rel_cong_simp = Lazy.map f rel_cong_simp, rel_map = Lazy.map (map f) rel_map, rel_mono = Lazy.map f rel_mono, rel_mono_strong0 = Lazy.map f rel_mono_strong0, rel_mono_strong = Lazy.map f rel_mono_strong, rel_transfer = Lazy.map f rel_transfer, rel_Grp = Lazy.map f rel_Grp, rel_conversep = Lazy.map f rel_conversep, rel_OO = Lazy.map f rel_OO, rel_refl = Lazy.map f rel_refl, rel_refl_strong = Lazy.map f rel_refl_strong, rel_reflp = Lazy.map f rel_reflp, rel_symp = Lazy.map f rel_symp, rel_transp = Lazy.map f rel_transp, set_transfer = Lazy.map (map f) set_transfer, rel_eq_onp = Lazy.map f rel_eq_onp, pred_transfer = Lazy.map f pred_transfer, pred_True = Lazy.map f pred_True, pred_map = Lazy.map f pred_map, pred_rel = Lazy.map f pred_rel, pred_mono_strong0 = Lazy.map f pred_mono_strong0, pred_mono_strong = Lazy.map f pred_mono_strong, pred_mono = Lazy.map f pred_mono, pred_cong0 = Lazy.map f pred_cong0, pred_cong = Lazy.map f pred_cong, pred_cong_simp = Lazy.map f pred_cong_simp}; val morph_facts = map_facts o Morphism.thm; type nonemptiness_witness = { I: int list, wit: term, prop: thm list }; fun mk_witness (I, wit) prop = {I = I, wit = wit, prop = prop}; fun map_witness f g {I, wit, prop} = {I = I, wit = f wit, prop = map g prop}; fun morph_witness phi = map_witness (Morphism.term phi) (Morphism.thm phi); datatype bnf = BNF of { name: binding, T: typ, live: int, lives: typ list, (*source type variables of map*) lives': typ list, (*target type variables of map*) dead: int, deads: typ list, map: term, sets: term list, bd: term, axioms: axioms, defs: defs, facts: facts, nwits: int, wits: nonemptiness_witness list, rel: term, pred: term }; (* getters *) fun rep_bnf (BNF bnf) = bnf; val name_of_bnf = #name o rep_bnf; val T_of_bnf = #T o rep_bnf; fun mk_T_of_bnf Ds Ts bnf = let val bnf_rep = rep_bnf bnf in Term.typ_subst_atomic ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#T bnf_rep) end; val live_of_bnf = #live o rep_bnf; val lives_of_bnf = #lives o rep_bnf; val dead_of_bnf = #dead o rep_bnf; val deads_of_bnf = #deads o rep_bnf; val axioms_of_bnf = #axioms o rep_bnf; val facts_of_bnf = #facts o rep_bnf; val nwits_of_bnf = #nwits o rep_bnf; val wits_of_bnf = #wits o rep_bnf; fun flatten_type_args_of_bnf bnf dead_x xs = let val Type (_, Ts) = T_of_bnf bnf; val lives = lives_of_bnf bnf; val deads = deads_of_bnf bnf; in permute_like_unique (op =) (deads @ lives) Ts (replicate (length deads) dead_x @ xs) end; (*terms*) val map_of_bnf = #map o rep_bnf; val sets_of_bnf = #sets o rep_bnf; fun mk_map_of_bnf Ds Ts Us bnf = let val bnf_rep = rep_bnf bnf; in Term.subst_atomic_types ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#map bnf_rep) end; fun mk_sets_of_bnf Dss Tss bnf = let val bnf_rep = rep_bnf bnf; in map2 (fn (Ds, Ts) => Term.subst_atomic_types ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts))) (Dss ~~ Tss) (#sets bnf_rep) end; val bd_of_bnf = #bd o rep_bnf; fun mk_bd_of_bnf Ds Ts bnf = let val bnf_rep = rep_bnf bnf; in Term.subst_atomic_types ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#bd bnf_rep) end; fun mk_wits_of_bnf Dss Tss bnf = let val bnf_rep = rep_bnf bnf; val wits = map (fn x => (#I x, #wit x)) (#wits bnf_rep); in map2 (fn (Ds, Ts) => apsnd (Term.subst_atomic_types ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)))) (Dss ~~ Tss) wits end; val rel_of_bnf = #rel o rep_bnf; fun mk_rel_of_bnf Ds Ts Us bnf = let val bnf_rep = rep_bnf bnf; in Term.subst_atomic_types ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#rel bnf_rep) end; val pred_of_bnf = #pred o rep_bnf; fun mk_pred_of_bnf Ds Ts bnf = let val bnf_rep = rep_bnf bnf; in Term.subst_atomic_types ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#pred bnf_rep) end; (*thms*) val bd_Card_order_of_bnf = #bd_Card_order o #facts o rep_bnf; val bd_Cinfinite_of_bnf = #bd_Cinfinite o #facts o rep_bnf; val bd_Cnotzero_of_bnf = #bd_Cnotzero o #facts o rep_bnf; val bd_card_order_of_bnf = #bd_card_order o #axioms o rep_bnf; val bd_cinfinite_of_bnf = #bd_cinfinite o #axioms o rep_bnf; val collect_set_map_of_bnf = Lazy.force o #collect_set_map o #facts o rep_bnf; val in_bd_of_bnf = Lazy.force o #in_bd o #facts o rep_bnf; val in_cong_of_bnf = Lazy.force o #in_cong o #facts o rep_bnf; val in_mono_of_bnf = Lazy.force o #in_mono o #facts o rep_bnf; val in_rel_of_bnf = Lazy.force o #in_rel o #facts o rep_bnf; val inj_map_of_bnf = Lazy.force o #inj_map o #facts o rep_bnf; val inj_map_strong_of_bnf = Lazy.force o #inj_map_strong o #facts o rep_bnf; val le_rel_OO_of_bnf = #le_rel_OO o #axioms o rep_bnf; val map_comp0_of_bnf = #map_comp0 o #axioms o rep_bnf; val map_comp_of_bnf = Lazy.force o #map_comp o #facts o rep_bnf; val map_cong0_of_bnf = #map_cong0 o #axioms o rep_bnf; val map_cong_of_bnf = Lazy.force o #map_cong o #facts o rep_bnf; val map_cong_pred_of_bnf = Lazy.force o #map_cong_pred o #facts o rep_bnf; val map_cong_simp_of_bnf = Lazy.force o #map_cong_simp o #facts o rep_bnf; val map_def_of_bnf = #map_def o #defs o rep_bnf; val map_id0_of_bnf = #map_id0 o #axioms o rep_bnf; val map_id_of_bnf = Lazy.force o #map_id o #facts o rep_bnf; val map_ident0_of_bnf = Lazy.force o #map_ident0 o #facts o rep_bnf; val map_ident_of_bnf = Lazy.force o #map_ident o #facts o rep_bnf; val map_transfer_of_bnf = Lazy.force o #map_transfer o #facts o rep_bnf; val rel_eq_onp_of_bnf = Lazy.force o #rel_eq_onp o #facts o rep_bnf; val pred_def_of_bnf = #pred_def o #defs o rep_bnf; val pred_map_of_bnf = Lazy.force o #pred_map o #facts o rep_bnf; val pred_mono_strong0_of_bnf = Lazy.force o #pred_mono_strong0 o #facts o rep_bnf; val pred_mono_strong_of_bnf = Lazy.force o #pred_mono_strong o #facts o rep_bnf; val pred_mono_of_bnf = Lazy.force o #pred_mono o #facts o rep_bnf; val pred_cong0_of_bnf = Lazy.force o #pred_cong0 o #facts o rep_bnf; val pred_cong_of_bnf = Lazy.force o #pred_cong o #facts o rep_bnf; val pred_cong_simp_of_bnf = Lazy.force o #pred_cong_simp o #facts o rep_bnf; val pred_rel_of_bnf = Lazy.force o #pred_rel o #facts o rep_bnf; val pred_set_of_bnf = #pred_set o #axioms o rep_bnf; val pred_transfer_of_bnf = Lazy.force o #pred_transfer o #facts o rep_bnf; val pred_True_of_bnf = Lazy.force o #pred_True o #facts o rep_bnf; val rel_Grp_of_bnf = Lazy.force o #rel_Grp o #facts o rep_bnf; val rel_OO_Grp_of_bnf = #rel_OO_Grp o #axioms o rep_bnf; val rel_OO_of_bnf = Lazy.force o #rel_OO o #facts o rep_bnf; val rel_cong0_of_bnf = Lazy.force o #rel_cong0 o #facts o rep_bnf; val rel_cong_of_bnf = Lazy.force o #rel_cong o #facts o rep_bnf; val rel_cong_simp_of_bnf = Lazy.force o #rel_cong_simp o #facts o rep_bnf; val rel_conversep_of_bnf = Lazy.force o #rel_conversep o #facts o rep_bnf; val rel_def_of_bnf = #rel_def o #defs o rep_bnf; val rel_eq_of_bnf = Lazy.force o #rel_eq o #facts o rep_bnf; val rel_flip_of_bnf = Lazy.force o #rel_flip o #facts o rep_bnf; val rel_map_of_bnf = Lazy.force o #rel_map o #facts o rep_bnf; val rel_mono_of_bnf = Lazy.force o #rel_mono o #facts o rep_bnf; val rel_mono_strong0_of_bnf = Lazy.force o #rel_mono_strong0 o #facts o rep_bnf; val rel_mono_strong_of_bnf = Lazy.force o #rel_mono_strong o #facts o rep_bnf; val rel_refl_of_bnf = Lazy.force o #rel_refl o #facts o rep_bnf; val rel_refl_strong_of_bnf = Lazy.force o #rel_refl_strong o #facts o rep_bnf; val rel_reflp_of_bnf = Lazy.force o #rel_reflp o #facts o rep_bnf; val rel_symp_of_bnf = Lazy.force o #rel_symp o #facts o rep_bnf; val rel_transfer_of_bnf = Lazy.force o #rel_transfer o #facts o rep_bnf; val rel_transp_of_bnf = Lazy.force o #rel_transp o #facts o rep_bnf; val set_bd_of_bnf = #set_bd o #axioms o rep_bnf; val set_defs_of_bnf = #set_defs o #defs o rep_bnf; val set_map0_of_bnf = #set_map0 o #axioms o rep_bnf; val set_map_of_bnf = map Lazy.force o #set_map o #facts o rep_bnf; val set_transfer_of_bnf = Lazy.force o #set_transfer o #facts o rep_bnf; val wit_thms_of_bnf = maps #prop o wits_of_bnf; val wit_thmss_of_bnf = map #prop o wits_of_bnf; fun mk_bnf name T live lives lives' dead deads map sets bd axioms defs facts wits rel pred = BNF {name = name, T = T, live = live, lives = lives, lives' = lives', dead = dead, deads = deads, map = map, sets = sets, bd = bd, axioms = axioms, defs = defs, facts = facts, nwits = length wits, wits = wits, rel = rel, pred = pred}; fun map_bnf f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 (BNF {name = name, T = T, live = live, lives = lives, lives' = lives', dead = dead, deads = deads, map = map, sets = sets, bd = bd, axioms = axioms, defs = defs, facts = facts, nwits = nwits, wits = wits, rel = rel, pred = pred}) = BNF {name = f1 name, T = f2 T, live = f3 live, lives = f4 lives, lives' = f5 lives', dead = f6 dead, deads = f7 deads, map = f8 map, sets = f9 sets, bd = f10 bd, axioms = f11 axioms, defs = f12 defs, facts = f13 facts, nwits = f14 nwits, wits = f15 wits, rel = f16 rel, pred = f17 pred}; fun morph_bnf phi = let val Tphi = Morphism.typ phi; val tphi = Morphism.term phi; in map_bnf (Morphism.binding phi) Tphi I (map Tphi) (map Tphi) I (map Tphi) tphi (map tphi) tphi (morph_axioms phi) (morph_defs phi) (morph_facts phi) I (map (morph_witness phi)) tphi tphi end; fun morph_bnf_defs phi = map_bnf I I I I I I I I I I I (morph_defs phi) I I I I I; fun permute_deads perm = map_bnf I I I I I I perm I I I I I I I I I I; val transfer_bnf = morph_bnf o Morphism.transfer_morphism; structure Data = Generic_Data ( type T = bnf Symtab.table; val empty = Symtab.empty; val extend = I; fun merge data : T = Symtab.merge (K true) data; ); fun bnf_of_generic context = Option.map (transfer_bnf (Context.theory_of context)) o Symtab.lookup (Data.get context); val bnf_of = bnf_of_generic o Context.Proof; val bnf_of_global = bnf_of_generic o Context.Theory; (* Utilities *) fun normalize_set insts instA set = let val (T, T') = dest_funT (fastype_of set); val A = fst (Term.dest_TVar (HOLogic.dest_setT T')); val params = Term.add_tvar_namesT T []; in Term.subst_TVars ((A :: params) ~~ (instA :: insts)) set end; fun normalize_rel ctxt instTs instA instB rel = let val thy = Proof_Context.theory_of ctxt; val tyenv = Sign.typ_match thy (fastype_of rel, Library.foldr (op -->) (instTs, mk_pred2T instA instB)) Vartab.empty; in Envir.subst_term (tyenv, Vartab.empty) rel end handle Type.TYPE_MATCH => error "Bad relator"; fun normalize_pred ctxt instTs instA pred = let val thy = Proof_Context.theory_of ctxt; val tyenv = Sign.typ_match thy (fastype_of pred, Library.foldr (op -->) (instTs, mk_pred1T instA)) Vartab.empty; in Envir.subst_term (tyenv, Vartab.empty) pred end handle Type.TYPE_MATCH => error "Bad predicator"; fun normalize_wit insts CA As wit = let fun strip_param (Ts, T as Type (\<^type_name>\fun\, [T1, T2])) = if Type.raw_instance (CA, T) then (Ts, T) else strip_param (T1 :: Ts, T2) | strip_param x = x; val (Ts, T) = strip_param ([], fastype_of wit); val subst = Term.add_tvar_namesT T [] ~~ insts; fun find y = find_index (fn x => x = y) As; in (map (find o Term.typ_subst_TVars subst) (rev Ts), Term.subst_TVars subst wit) end; fun minimize_wits wits = let fun minimize done [] = done | minimize done ((I, wit) :: todo) = if exists (fn (J, _) => subset (op =) (J, I)) (done @ todo) then minimize done todo else minimize ((I, wit) :: done) todo; in minimize [] wits end; fun mk_map live Ts Us t = let val (Type (_, Ts0), Type (_, Us0)) = strip_typeN (live + 1) (fastype_of t) |>> List.last in Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t end; fun mk_pred Ts t = let val Type (_, Ts0) = domain_type (body_fun_type (fastype_of t)) in Term.subst_atomic_types (Ts0 ~~ Ts) t end; val mk_set = mk_pred; fun mk_rel live Ts Us t = let val [Type (_, Ts0), Type (_, Us0)] = binder_types (snd (strip_typeN live (fastype_of t))) in Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t end; fun build_map_or_rel mk const of_bnf dest pre_cst_table ctxt simple_Ts simple_Us build_simple = let fun build (TU as (T, U)) = if exists (curry (op =) T) simple_Ts orelse exists (curry (op =) U) simple_Us then build_simple TU else if T = U andalso not (exists_subtype_in simple_Ts T) andalso not (exists_subtype_in simple_Us U) then const T else (case TU of (Type (s, Ts), Type (s', Us)) => if s = s' then let fun recurse (live, cst0) = let val cst = mk live Ts Us cst0; val TUs' = map dest (fst (strip_typeN live (fastype_of cst))); in Term.list_comb (cst, map build TUs') end; in (case AList.lookup (op =) pre_cst_table s of NONE => (case bnf_of ctxt s of SOME bnf => recurse (live_of_bnf bnf, of_bnf bnf) | NONE => build_simple TU) | SOME entry => recurse entry) end else build_simple TU | _ => build_simple TU); in build end; val build_map = build_map_or_rel mk_map HOLogic.id_const map_of_bnf dest_funT [(\<^type_name>\set\, (1, \<^term>\image\))]; val build_rel = build_map_or_rel mk_rel HOLogic.eq_const rel_of_bnf dest_pred2T o append [(\<^type_name>\set\, (1, \<^term>\rel_set\)), (\<^type_name>\fun\, (2, \<^term>\rel_fun\))]; fun build_set ctxt A = let fun build T = Abs (Name.uu, T, if T = A then HOLogic.mk_set A [Bound 0] else (case T of Type (s, Ts) => let val sets = map (mk_set Ts) (sets_of_bnf (the (bnf_of ctxt s))) |> filter (exists_subtype_in [A] o range_type o fastype_of); val set_apps = map (fn set => Term.betapply (set, Bound 0)) sets; fun recurse set_app = let val Type (\<^type_name>\set\, [elemT]) = fastype_of set_app in if elemT = A then set_app else mk_UNION set_app (build elemT) end; in if null set_apps then HOLogic.mk_set A [] else Library.foldl1 mk_union (map recurse set_apps) end | _ => HOLogic.mk_set A [])); in build end; fun map_flattened_map_args ctxt s map_args fs = let val flat_fs = flatten_type_args_of_bnf (the (bnf_of ctxt s)) Term.dummy fs; val flat_fs' = map_args flat_fs; in permute_like_unique (op aconv) flat_fs fs flat_fs' end; (* Names *) val mapN = "map"; val setN = "set"; fun mk_setN i = setN ^ nonzero_string_of_int i; val bdN = "bd"; val witN = "wit"; fun mk_witN i = witN ^ nonzero_string_of_int i; val relN = "rel"; val predN = "pred"; val bd_Card_orderN = "bd_Card_order"; val bd_CinfiniteN = "bd_Cinfinite"; val bd_CnotzeroN = "bd_Cnotzero"; val bd_card_orderN = "bd_card_order"; val bd_cinfiniteN = "bd_cinfinite"; val collect_set_mapN = "collect_set_map"; val in_bdN = "in_bd"; val in_monoN = "in_mono"; val in_relN = "in_rel"; val inj_mapN = "inj_map"; val inj_map_strongN = "inj_map_strong"; val map_comp0N = "map_comp0"; val map_compN = "map_comp"; val map_cong0N = "map_cong0"; val map_congN = "map_cong"; val map_cong_simpN = "map_cong_simp"; val map_cong_predN = "map_cong_pred"; val map_id0N = "map_id0"; val map_idN = "map_id"; val map_identN = "map_ident"; val map_transferN = "map_transfer"; val pred_mono_strong0N = "pred_mono_strong0"; val pred_mono_strongN = "pred_mono_strong"; val pred_monoN = "pred_mono"; val pred_transferN = "pred_transfer"; val pred_TrueN = "pred_True"; val pred_mapN = "pred_map"; val pred_relN = "pred_rel"; val pred_setN = "pred_set"; val pred_congN = "pred_cong"; val pred_cong_simpN = "pred_cong_simp"; val rel_GrpN = "rel_Grp"; val rel_comppN = "rel_compp"; val rel_compp_GrpN = "rel_compp_Grp"; val rel_congN = "rel_cong"; val rel_cong_simpN = "rel_cong_simp"; val rel_conversepN = "rel_conversep"; val rel_eqN = "rel_eq"; val rel_eq_onpN = "rel_eq_onp"; val rel_flipN = "rel_flip"; val rel_mapN = "rel_map"; val rel_monoN = "rel_mono"; val rel_mono_strong0N = "rel_mono_strong0"; val rel_mono_strongN = "rel_mono_strong"; val rel_reflN = "rel_refl"; val rel_refl_strongN = "rel_refl_strong"; val rel_reflpN = "rel_reflp"; val rel_sympN = "rel_symp"; val rel_transferN = "rel_transfer"; val rel_transpN = "rel_transp"; val set_bdN = "set_bd"; val set_map0N = "set_map0"; val set_mapN = "set_map"; val set_transferN = "set_transfer"; datatype inline_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline; datatype fact_policy = Dont_Note | Note_Some | Note_All; val bnf_internals = Attrib.setup_config_bool \<^binding>\bnf_internals\ (K false); val bnf_timing = Attrib.setup_config_bool \<^binding>\bnf_timing\ (K false); fun user_policy policy ctxt = if Config.get ctxt bnf_internals then Note_All else policy; val smart_max_inline_term_size = 25; (*FUDGE*) fun note_bnf_thms fact_policy qualify0 bnf_b bnf lthy = let val axioms = axioms_of_bnf bnf; val facts = facts_of_bnf bnf; val wits = wits_of_bnf bnf; val qualify = let val qs = Binding.path_of bnf_b; in fold_rev (fn (s, mand) => Binding.qualify mand s) qs #> qualify0 end; fun note_if_note_all (noted0, lthy0) = let val witNs = if length wits = 1 then [witN] else map mk_witN (1 upto length wits); val notes = [(bd_card_orderN, [#bd_card_order axioms]), (bd_cinfiniteN, [#bd_cinfinite axioms]), (bd_Card_orderN, [#bd_Card_order facts]), (bd_CinfiniteN, [#bd_Cinfinite facts]), (bd_CnotzeroN, [#bd_Cnotzero facts]), (collect_set_mapN, [Lazy.force (#collect_set_map facts)]), (in_bdN, [Lazy.force (#in_bd facts)]), (in_monoN, [Lazy.force (#in_mono facts)]), (map_comp0N, [#map_comp0 axioms]), (rel_mono_strong0N, [Lazy.force (#rel_mono_strong0 facts)]), (pred_mono_strong0N, [Lazy.force (#pred_mono_strong0 facts)]), (set_map0N, #set_map0 axioms), (set_bdN, #set_bd axioms)] @ (witNs ~~ wit_thmss_of_bnf bnf) |> map (fn (thmN, thms) => ((qualify (Binding.qualify true (Binding.name_of bnf_b) (Binding.name thmN)), []), [(thms, [])])); in Local_Theory.notes notes lthy0 |>> append noted0 end; fun note_unless_dont_note (noted0, lthy0) = let val notes = [(in_relN, [Lazy.force (#in_rel facts)], []), (inj_mapN, [Lazy.force (#inj_map facts)], []), (inj_map_strongN, [Lazy.force (#inj_map_strong facts)], []), (map_compN, [Lazy.force (#map_comp facts)], []), (map_cong0N, [#map_cong0 axioms], []), (map_congN, [Lazy.force (#map_cong facts)], fundefcong_attrs), (map_cong_simpN, [Lazy.force (#map_cong_simp facts)], []), (map_cong_predN, [Lazy.force (#map_cong_pred facts)], []), (map_idN, [Lazy.force (#map_id facts)], []), (map_id0N, [#map_id0 axioms], []), (map_transferN, [Lazy.force (#map_transfer facts)], []), (map_identN, [Lazy.force (#map_ident facts)], []), (pred_mono_strongN, [Lazy.force (#pred_mono_strong facts)], []), (pred_monoN, [Lazy.force (#pred_mono facts)], []), (pred_congN, [Lazy.force (#pred_cong facts)], fundefcong_attrs), (pred_cong_simpN, [Lazy.force (#pred_cong_simp facts)], []), (pred_mapN, [Lazy.force (#pred_map facts)], []), (pred_relN, [Lazy.force (#pred_rel facts)], []), (pred_transferN, [Lazy.force (#pred_transfer facts)], []), (pred_TrueN, [Lazy.force (#pred_True facts)], []), (pred_setN, [#pred_set axioms], []), (rel_comppN, [Lazy.force (#rel_OO facts)], []), (rel_compp_GrpN, no_refl [#rel_OO_Grp axioms], []), (rel_conversepN, [Lazy.force (#rel_conversep facts)], []), (rel_eqN, [Lazy.force (#rel_eq facts)], []), (rel_eq_onpN, [Lazy.force (#rel_eq_onp facts)], []), (rel_flipN, [Lazy.force (#rel_flip facts)], []), (rel_GrpN, [Lazy.force (#rel_Grp facts)], []), (rel_mapN, Lazy.force (#rel_map facts), []), (rel_monoN, [Lazy.force (#rel_mono facts)], mono_attrs), (rel_mono_strongN, [Lazy.force (#rel_mono_strong facts)], []), (rel_congN, [Lazy.force (#rel_cong facts)], fundefcong_attrs), (rel_cong_simpN, [Lazy.force (#rel_cong_simp facts)], []), (rel_reflN, [Lazy.force (#rel_refl facts)], []), (rel_refl_strongN, [Lazy.force (#rel_refl_strong facts)], []), (rel_reflpN, [Lazy.force (#rel_reflp facts)], []), (rel_sympN, [Lazy.force (#rel_symp facts)], []), (rel_transpN, [Lazy.force (#rel_transp facts)], []), (rel_transferN, [Lazy.force (#rel_transfer facts)], []), (set_mapN, map Lazy.force (#set_map facts), []), (set_transferN, Lazy.force (#set_transfer facts), [])] |> filter_out (null o #2) |> map (fn (thmN, thms, attrs) => ((qualify (Binding.qualify true (Binding.name_of bnf_b) (Binding.name thmN)), attrs), [(thms, [])])); in Local_Theory.notes notes lthy0 |>> append noted0 end; in ([], lthy) |> fact_policy = Note_All ? note_if_note_all |> fact_policy <> Dont_Note ? note_unless_dont_note |>> (fn [] => bnf | noted => morph_bnf (substitute_noted_thm noted) bnf) end; fun note_bnf_defs bnf lthy = let fun mk_def_binding cst_of = Thm.def_binding (Binding.qualified_name (dest_Const (cst_of bnf) |> fst)); val notes = [(mk_def_binding map_of_bnf, map_def_of_bnf bnf), (mk_def_binding rel_of_bnf, rel_def_of_bnf bnf), (mk_def_binding pred_of_bnf, pred_def_of_bnf bnf)] @ @{map 2} (pair o mk_def_binding o K) (sets_of_bnf bnf) (set_defs_of_bnf bnf) |> map (fn (b, thm) => ((b, []), [([thm], [])])); in lthy |> Local_Theory.notes notes |>> (fn noted => morph_bnf (substitute_noted_thm noted) bnf) end; fun mk_wit_goals zs bs sets (I, wit) = let val xs = map (nth bs) I; fun wit_goal i = let val z = nth zs i; val set_wit = nth sets i $ Term.list_comb (wit, xs); val concl = HOLogic.mk_Trueprop (if member (op =) I i then HOLogic.mk_eq (z, nth bs i) else \<^term>\False\); in fold_rev Logic.all (z :: xs) (Logic.mk_implies (mk_Trueprop_mem (z, set_wit), concl)) end; in map wit_goal (0 upto length sets - 1) end; (* Define new BNFs *) fun define_bnf_consts const_policy fact_policy internal Ds_opt map_b rel_b pred_b set_bs (((((((bnf_b, T_rhs), map_rhs), set_rhss), bd_rhs), wit_rhss), rel_rhs_opt), pred_rhs_opt) no_defs_lthy = let val live = length set_rhss; val def_qualify = Binding.qualify false (Binding.name_of bnf_b); fun mk_prefix_binding pre = Binding.prefix_name (pre ^ "_") bnf_b; fun maybe_define user_specified (b, rhs) lthy = let val inline = (user_specified orelse fact_policy = Dont_Note) andalso (case const_policy of Dont_Inline => false | Hardly_Inline => Term.is_Free rhs orelse Term.is_Const rhs | Smart_Inline => Term.size_of_term rhs <= smart_max_inline_term_size | Do_Inline => true); in if inline then ((rhs, Drule.reflexive_thm), lthy) else let val b = b () in apfst (apsnd snd) ((if internal then Local_Theory.define_internal else Local_Theory.define) ((b, NoSyn), ((Binding.concealed (Thm.def_binding b), []), rhs)) lthy) end end; val map_bind_def = (fn () => def_qualify (if Binding.is_empty map_b then mk_prefix_binding mapN else map_b), map_rhs); val set_binds_defs = let fun set_name i get_b = (case try (nth set_bs) (i - 1) of SOME b => if Binding.is_empty b then get_b else K b | NONE => get_b) #> def_qualify; val bs = if live = 1 then [set_name 1 (fn () => mk_prefix_binding setN)] else map (fn i => set_name i (fn () => mk_prefix_binding (mk_setN i))) (1 upto live); in bs ~~ set_rhss end; val bd_bind_def = (fn () => def_qualify (mk_prefix_binding bdN), bd_rhs); val ((((bnf_map_term, raw_map_def), (bnf_set_terms, raw_set_defs)), (bnf_bd_term, raw_bd_def)), (lthy, lthy_old)) = no_defs_lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> maybe_define true map_bind_def ||>> apfst split_list o fold_map (maybe_define true) set_binds_defs ||>> maybe_define true bd_bind_def ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val bnf_map_def = Morphism.thm phi raw_map_def; val bnf_set_defs = map (Morphism.thm phi) raw_set_defs; val bnf_bd_def = Morphism.thm phi raw_bd_def; val bnf_map = Morphism.term phi bnf_map_term; (*TODO: handle errors*) (*simple shape analysis of a map function*) val ((alphas, betas), (Calpha, _)) = fastype_of bnf_map |> strip_typeN live |>> map_split dest_funT ||> dest_funT handle TYPE _ => error "Bad map function"; val Calpha_params = map TVar (Term.add_tvarsT Calpha []); val bnf_T = Morphism.typ phi T_rhs; val bad_args = Term.add_tfreesT bnf_T []; val _ = null bad_args orelse error ("Locally fixed type arguments " ^ commas_quote (map (Syntax.string_of_typ no_defs_lthy o TFree) bad_args)); val bnf_sets = map2 (normalize_set Calpha_params) alphas (map (Morphism.term phi) bnf_set_terms); val bnf_bd = Term.subst_TVars (Term.add_tvar_namesT bnf_T [] ~~ Calpha_params) (Morphism.term phi bnf_bd_term); (*TODO: assert Ds = (TVars of bnf_map) \ (alphas @ betas) as sets*) val deads = (case Ds_opt of NONE => subtract (op =) (alphas @ betas) (map TVar (Term.add_tvars bnf_map [])) | SOME Ds => map (Morphism.typ phi) Ds); (*TODO: further checks of type of bnf_map*) (*TODO: check types of bnf_sets*) (*TODO: check type of bnf_bd*) (*TODO: check type of bnf_rel*) fun mk_bnf_map Ds As' Bs' = Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As') @ (betas ~~ Bs')) bnf_map; fun mk_bnf_t Ds As' = Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As')); fun mk_bnf_T Ds As' = Term.typ_subst_atomic ((deads ~~ Ds) @ (alphas ~~ As')); val (((As, Bs), unsorted_Ds), names_lthy) = lthy |> mk_TFrees live ||>> mk_TFrees live ||>> mk_TFrees (length deads); val Ds = map2 (resort_tfree_or_tvar o Type.sort_of_atyp) deads unsorted_Ds; val RTs = map2 (curry HOLogic.mk_prodT) As Bs; val pred2RTs = map2 mk_pred2T As Bs; val (Rs, Rs') = names_lthy |> mk_Frees' "R" pred2RTs |> fst; val CA = mk_bnf_T Ds As Calpha; val CR = mk_bnf_T Ds RTs Calpha; val setRs = @{map 3} (fn R => fn T => fn U => HOLogic.Collect_const (HOLogic.mk_prodT (T, U)) $ HOLogic.mk_case_prod R) Rs As Bs; (*Grp (in (Collect (split R1) .. Collect (split Rn))) (map fst .. fst)^--1 OO Grp (in (Collect (split R1) .. Collect (split Rn))) (map snd .. snd)*) val rel_spec = let val map1 = Term.list_comb (mk_bnf_map Ds RTs As, map fst_const RTs); val map2 = Term.list_comb (mk_bnf_map Ds RTs Bs, map snd_const RTs); val bnf_in = mk_in setRs (map (mk_bnf_t Ds RTs) bnf_sets) CR; in mk_rel_compp (mk_conversep (mk_Grp bnf_in map1), mk_Grp bnf_in map2) |> fold_rev Term.absfree Rs' end; val rel_rhs = the_default rel_spec rel_rhs_opt; val rel_bind_def = (fn () => def_qualify (if Binding.is_empty rel_b then mk_prefix_binding relN else rel_b), rel_rhs); val pred_spec = if live = 0 then Term.absdummy (mk_bnf_T Ds As Calpha) \<^term>\True\ else let val sets = map (mk_bnf_t Ds As) bnf_sets; val argTs = map mk_pred1T As; val T = mk_bnf_T Ds As Calpha; val ((Ps, Ps'), x) = lthy |> mk_Frees' "P" argTs ||>> yield_singleton (mk_Frees "x") T |> fst; val conjs = map2 (fn set => fn P => mk_Ball (set $ x) P) sets Ps; in fold_rev Term.absfree Ps' (Term.absfree (dest_Free x) (Library.foldr1 HOLogic.mk_conj conjs)) end; val pred_rhs = the_default pred_spec pred_rhs_opt; val pred_bind_def = (fn () => def_qualify (if Binding.is_empty pred_b then mk_prefix_binding predN else pred_b), pred_rhs); val wit_rhss = if null wit_rhss then [fold_rev Term.absdummy As (Term.list_comb (mk_bnf_map Ds As As, map2 (fn T => fn i => Term.absdummy T (Bound i)) As (live downto 1)) $ Const (\<^const_name>\undefined\, CA))] else wit_rhss; val nwits = length wit_rhss; val wit_binds_defs = let val bs = if nwits = 1 then [fn () => def_qualify (mk_prefix_binding witN)] else map (fn i => fn () => def_qualify (mk_prefix_binding (mk_witN i))) (1 upto nwits); in bs ~~ wit_rhss end; val ((((bnf_rel_term, raw_rel_def), (bnf_pred_term, raw_pred_def)), (bnf_wit_terms, raw_wit_defs)), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> maybe_define (is_some rel_rhs_opt) rel_bind_def ||>> maybe_define (is_some pred_rhs_opt) pred_bind_def ||>> apfst split_list o fold_map (maybe_define (not (null wit_rhss))) wit_binds_defs ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val bnf_rel_def = Morphism.thm phi raw_rel_def; val bnf_rel = Morphism.term phi bnf_rel_term; fun mk_bnf_rel Ds As' Bs' = normalize_rel lthy (map2 mk_pred2T As' Bs') (mk_bnf_T Ds As' Calpha) (mk_bnf_T Ds Bs' Calpha) bnf_rel; val bnf_pred_def = Morphism.thm phi raw_pred_def; val bnf_pred = Morphism.term phi bnf_pred_term; fun mk_bnf_pred Ds As' = normalize_pred lthy (map mk_pred1T As') (mk_bnf_T Ds As' Calpha) bnf_pred; val bnf_wit_defs = map (Morphism.thm phi) raw_wit_defs; val bnf_wits = map (normalize_wit Calpha_params Calpha alphas o Morphism.term phi) bnf_wit_terms; fun mk_rel_spec Ds' As' Bs' = Term.subst_atomic_types ((Ds ~~ Ds') @ (As ~~ As') @ (Bs ~~ Bs')) rel_spec; fun mk_pred_spec Ds' As' = Term.subst_atomic_types ((Ds ~~ Ds') @ (As ~~ As')) pred_spec; in (((alphas, betas, deads, Calpha), (bnf_map, bnf_sets, bnf_bd, bnf_wits, bnf_rel, bnf_pred), (bnf_map_def, bnf_set_defs, bnf_bd_def, bnf_wit_defs, bnf_rel_def, bnf_pred_def), (mk_bnf_map, mk_bnf_t, mk_bnf_T, mk_bnf_rel, mk_bnf_pred, mk_rel_spec, mk_pred_spec)), lthy) end; fun prepare_def const_policy mk_fact_policy internal qualify prep_typ prep_term Ds_opt map_b rel_b pred_b set_bs (((((((raw_bnf_b, raw_bnf_T), raw_map), raw_sets), raw_bd), raw_wits), raw_rel_opt), raw_pred_opt) no_defs_lthy = let val fact_policy = mk_fact_policy no_defs_lthy; val bnf_b = qualify raw_bnf_b; val live = length raw_sets; val T_rhs = prep_typ no_defs_lthy raw_bnf_T; val map_rhs = prep_term no_defs_lthy raw_map; val set_rhss = map (prep_term no_defs_lthy) raw_sets; val bd_rhs = prep_term no_defs_lthy raw_bd; val wit_rhss = map (prep_term no_defs_lthy) raw_wits; val rel_rhs_opt = Option.map (prep_term no_defs_lthy) raw_rel_opt; val pred_rhs_opt = Option.map (prep_term no_defs_lthy) raw_pred_opt; fun err T = error ("Trying to register the type " ^ quote (Syntax.string_of_typ no_defs_lthy T) ^ " as unnamed BNF"); val (bnf_b, key) = if Binding.is_empty bnf_b then (case T_rhs of Type (C, Ts) => if forall (can dest_TFree) Ts andalso not (has_duplicates (op =) Ts) then (Binding.qualified_name C, C) else err T_rhs | T => err T) else (bnf_b, Local_Theory.full_name no_defs_lthy bnf_b); val (((alphas, betas, deads, Calpha), (bnf_map, bnf_sets, bnf_bd, bnf_wits, bnf_rel, bnf_pred), (bnf_map_def, bnf_set_defs, bnf_bd_def, bnf_wit_defs, bnf_rel_def, bnf_pred_def), (mk_bnf_map_Ds, mk_bnf_t_Ds, mk_bnf_T_Ds, _, _, mk_rel_spec, mk_pred_spec)), lthy) = define_bnf_consts const_policy fact_policy internal Ds_opt map_b rel_b pred_b set_bs (((((((bnf_b, T_rhs), map_rhs), set_rhss), bd_rhs), wit_rhss), rel_rhs_opt), pred_rhs_opt) no_defs_lthy; val dead = length deads; val (((((((As', Bs'), Cs), unsorted_Ds), Es), B1Ts), B2Ts), (Ts, T)) = lthy |> mk_TFrees live ||>> mk_TFrees live ||>> mk_TFrees live ||>> mk_TFrees dead ||>> mk_TFrees live ||>> mk_TFrees live ||>> mk_TFrees live ||> fst o mk_TFrees 1 ||> the_single ||> `(replicate live); val Ds = map2 (resort_tfree_or_tvar o Type.sort_of_atyp) deads unsorted_Ds; val mk_bnf_map = mk_bnf_map_Ds Ds; val mk_bnf_t = mk_bnf_t_Ds Ds; val mk_bnf_T = mk_bnf_T_Ds Ds; val pred1PTs = map mk_pred1T As'; val pred1QTs = map mk_pred1T Bs'; val pred2RTs = map2 mk_pred2T As' Bs'; val pred2RTsAsCs = map2 mk_pred2T As' Cs; val pred2RTsBsCs = map2 mk_pred2T Bs' Cs; val pred2RTsBsEs = map2 mk_pred2T Bs' Es; val pred2RTsCsBs = map2 mk_pred2T Cs Bs'; val pred2RTsCsEs = map2 mk_pred2T Cs Es; val pred2RT's = map2 mk_pred2T Bs' As'; val self_pred2RTs = map2 mk_pred2T As' As'; val transfer_domRTs = map2 mk_pred2T As' B1Ts; val transfer_ranRTs = map2 mk_pred2T Bs' B2Ts; val CA' = mk_bnf_T As' Calpha; val CB' = mk_bnf_T Bs' Calpha; val CC' = mk_bnf_T Cs Calpha; val CE' = mk_bnf_T Es Calpha; val CB1 = mk_bnf_T B1Ts Calpha; val CB2 = mk_bnf_T B2Ts Calpha; val bnf_map_AsAs = mk_bnf_map As' As'; val bnf_map_AsBs = mk_bnf_map As' Bs'; val bnf_map_AsCs = mk_bnf_map As' Cs; val bnf_map_BsCs = mk_bnf_map Bs' Cs; val bnf_sets_As = map (mk_bnf_t As') bnf_sets; val bnf_sets_Bs = map (mk_bnf_t Bs') bnf_sets; val bnf_bd_As = mk_bnf_t As' bnf_bd; fun mk_bnf_rel RTs CA CB = normalize_rel lthy RTs CA CB bnf_rel; fun mk_bnf_pred PTs CA = normalize_pred lthy PTs CA bnf_pred; val ((((((((((((((((((((((((((fs, fs'), gs), hs), is), x), x'), y), y'), zs), zs'), ys), As), As_copy), bs), (Ps, Ps')), Ps_copy), Qs), Rs), Rs_copy), Ss), S_AsCs), S_CsBs), S_BsEs), transfer_domRs), transfer_ranRs), _) = lthy |> mk_Frees "f" (map2 (curry op -->) As' Bs') ||>> mk_Frees "f" (map2 (curry op -->) As' Bs') ||>> mk_Frees "g" (map2 (curry op -->) Bs' Cs) ||>> mk_Frees "h" (map2 (curry op -->) As' Ts) ||>> mk_Frees "i" (map2 (curry op -->) As' Cs) ||>> yield_singleton (mk_Frees "x") CA' ||>> yield_singleton (mk_Frees "x") CA' ||>> yield_singleton (mk_Frees "y") CB' ||>> yield_singleton (mk_Frees "y") CB' ||>> mk_Frees "z" As' ||>> mk_Frees "z" As' ||>> mk_Frees "y" Bs' ||>> mk_Frees "A" (map HOLogic.mk_setT As') ||>> mk_Frees "A" (map HOLogic.mk_setT As') ||>> mk_Frees "b" As' ||>> mk_Frees' "P" pred1PTs ||>> mk_Frees "P" pred1PTs ||>> mk_Frees "Q" pred1QTs ||>> mk_Frees "R" pred2RTs ||>> mk_Frees "R" pred2RTs ||>> mk_Frees "S" pred2RTsBsCs ||>> mk_Frees "S" pred2RTsAsCs ||>> mk_Frees "S" pred2RTsCsBs ||>> mk_Frees "S" pred2RTsBsEs ||>> mk_Frees "R" transfer_domRTs ||>> mk_Frees "S" transfer_ranRTs; val fs_copy = map2 (retype_const_or_free o fastype_of) fs gs; val x_copy = retype_const_or_free CA' y'; val y_copy = retype_const_or_free CB' x'; val rel = mk_bnf_rel pred2RTs CA' CB'; val pred = mk_bnf_pred pred1PTs CA'; val pred' = mk_bnf_pred pred1QTs CB'; val relCsEs = mk_bnf_rel pred2RTsCsEs CC' CE'; val relAsAs = mk_bnf_rel self_pred2RTs CA' CA'; val bnf_wit_As = map (apsnd (mk_bnf_t As')) bnf_wits; val map_id0_goal = let val bnf_map_app_id = Term.list_comb (bnf_map_AsAs, map HOLogic.id_const As') in mk_Trueprop_eq (bnf_map_app_id, HOLogic.id_const CA') end; val map_comp0_goal = let val bnf_map_app_comp = Term.list_comb (bnf_map_AsCs, map2 (curry HOLogic.mk_comp) gs fs); val comp_bnf_map_app = HOLogic.mk_comp (Term.list_comb (bnf_map_BsCs, gs), Term.list_comb (bnf_map_AsBs, fs)); in fold_rev Logic.all (fs @ gs) (mk_Trueprop_eq (bnf_map_app_comp, comp_bnf_map_app)) end; fun mk_map_cong_prem mk_implies x z set f f_copy = Logic.all z (mk_implies (mk_Trueprop_mem (z, set $ x), mk_Trueprop_eq (f $ z, f_copy $ z))); val map_cong0_goal = let val prems = @{map 4} (mk_map_cong_prem Logic.mk_implies x) zs bnf_sets_As fs fs_copy; val eq = mk_Trueprop_eq (Term.list_comb (bnf_map_AsBs, fs) $ x, Term.list_comb (bnf_map_AsBs, fs_copy) $ x); in fold_rev Logic.all (x :: fs @ fs_copy) (Logic.list_implies (prems, eq)) end; val set_map0s_goal = let fun mk_goal setA setB f = let val set_comp_map = HOLogic.mk_comp (setB, Term.list_comb (bnf_map_AsBs, fs)); val image_comp_set = HOLogic.mk_comp (mk_image f, setA); in fold_rev Logic.all fs (mk_Trueprop_eq (set_comp_map, image_comp_set)) end; in @{map 3} mk_goal bnf_sets_As bnf_sets_Bs fs end; val card_order_bd_goal = HOLogic.mk_Trueprop (mk_card_order bnf_bd_As); val cinfinite_bd_goal = HOLogic.mk_Trueprop (mk_cinfinite bnf_bd_As); val set_bds_goal = let fun mk_goal set = Logic.all x (HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of (set $ x)) bnf_bd_As)); in map mk_goal bnf_sets_As end; val relAsCs = mk_bnf_rel pred2RTsAsCs CA' CC'; val relBsCs = mk_bnf_rel pred2RTsBsCs CB' CC'; val relCsBs = mk_bnf_rel pred2RTsCsBs CC' CB'; val rel_OO_lhs = Term.list_comb (relAsCs, map2 (curry mk_rel_compp) Rs Ss); val rel_OO_rhs = mk_rel_compp (Term.list_comb (rel, Rs), Term.list_comb (relBsCs, Ss)); val le_rel_OO_goal = fold_rev Logic.all (Rs @ Ss) (HOLogic.mk_Trueprop (mk_leq rel_OO_rhs rel_OO_lhs)); val rel_OO_Grp_goal = fold_rev Logic.all Rs (mk_Trueprop_eq (Term.list_comb (rel, Rs), Term.list_comb (mk_rel_spec Ds As' Bs', Rs))); val pred_set_goal = fold_rev Logic.all Ps (mk_Trueprop_eq (Term.list_comb (pred, Ps), Term.list_comb (mk_pred_spec Ds As', Ps))); val goals = zip_axioms map_id0_goal map_comp0_goal map_cong0_goal set_map0s_goal card_order_bd_goal cinfinite_bd_goal set_bds_goal le_rel_OO_goal rel_OO_Grp_goal pred_set_goal; val mk_wit_goals = mk_wit_goals bs zs bnf_sets_As; fun triv_wit_tac ctxt = mk_trivial_wit_tac ctxt bnf_wit_defs; val wit_goalss = (if null raw_wits then SOME triv_wit_tac else NONE, map mk_wit_goals bnf_wit_As); fun after_qed mk_wit_thms thms lthy = let val (axioms, nontriv_wit_thms) = apfst (mk_axioms live) (chop (length goals) thms); val bd_Card_order = #bd_card_order axioms RS @{thm conjunct2[OF card_order_on_Card_order]}; val bd_Cinfinite = @{thm conjI} OF [#bd_cinfinite axioms, bd_Card_order]; val bd_Cnotzero = bd_Cinfinite RS @{thm Cinfinite_Cnotzero}; fun mk_collect_set_map () = let val defT = mk_bnf_T Ts Calpha --> HOLogic.mk_setT T; val collect_map = HOLogic.mk_comp (mk_collect (map (mk_bnf_t Ts) bnf_sets) defT, Term.list_comb (mk_bnf_map As' Ts, hs)); val image_collect = mk_collect (map2 (fn h => fn set => HOLogic.mk_comp (mk_image h, set)) hs bnf_sets_As) defT; (*collect {set1 ... setm} o map f1 ... fm = collect {f1` o set1 ... fm` o setm}*) val goal = fold_rev Logic.all hs (mk_Trueprop_eq (collect_map, image_collect)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_collect_set_map_tac ctxt (#set_map0 axioms)) |> Thm.close_derivation \<^here> end; val collect_set_map = Lazy.lazy mk_collect_set_map; fun mk_in_mono () = let val prems_mono = map2 (HOLogic.mk_Trueprop oo mk_leq) As As_copy; val in_mono_goal = fold_rev Logic.all (As @ As_copy) (Logic.list_implies (prems_mono, HOLogic.mk_Trueprop (mk_leq (mk_in As bnf_sets_As CA') (mk_in As_copy bnf_sets_As CA')))); in Goal.prove_sorry lthy [] [] in_mono_goal (fn {context = ctxt, prems = _} => mk_in_mono_tac ctxt live) |> Thm.close_derivation \<^here> end; val in_mono = Lazy.lazy mk_in_mono; fun mk_in_cong () = let val prems_cong = map2 (curry mk_Trueprop_eq) As As_copy; val in_cong_goal = fold_rev Logic.all (As @ As_copy) (Logic.list_implies (prems_cong, mk_Trueprop_eq (mk_in As bnf_sets_As CA', mk_in As_copy bnf_sets_As CA'))); in Goal.prove_sorry lthy [] [] in_cong_goal (fn {context = ctxt, prems = _} => (TRY o hyp_subst_tac ctxt THEN' rtac ctxt refl) 1) |> Thm.close_derivation \<^here> end; val in_cong = Lazy.lazy mk_in_cong; val map_id = Lazy.lazy (fn () => mk_map_id (#map_id0 axioms)); val map_ident0 = Lazy.lazy (fn () => mk_map_ident lthy (#map_id0 axioms)); val map_ident = Lazy.lazy (fn () => mk_map_ident lthy (Lazy.force map_id)); val map_comp = Lazy.lazy (fn () => mk_map_comp (#map_comp0 axioms)); fun mk_map_cong mk_implies () = let val prem0 = mk_Trueprop_eq (x, x_copy); val prems = @{map 4} (mk_map_cong_prem mk_implies x_copy) zs bnf_sets_As fs fs_copy; val eq = mk_Trueprop_eq (Term.list_comb (bnf_map_AsBs, fs) $ x, Term.list_comb (bnf_map_AsBs, fs_copy) $ x_copy); val goal = fold_rev Logic.all (x :: x_copy :: fs @ fs_copy) (Logic.list_implies (prem0 :: prems, eq)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt @{thms simp_implies_def} THEN mk_map_cong_tac ctxt (#map_cong0 axioms)) |> Thm.close_derivation \<^here> end; val map_cong = Lazy.lazy (mk_map_cong Logic.mk_implies); val map_cong_simp = Lazy.lazy (mk_map_cong (fn (a, b) => \<^term>\simp_implies\ $ a $ b)); fun mk_inj_map () = let val prems = map (HOLogic.mk_Trueprop o mk_inj) fs; val concl = HOLogic.mk_Trueprop (mk_inj (Term.list_comb (bnf_map_AsBs, fs))); val goal = fold_rev Logic.all fs (Logic.list_implies (prems, concl)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_inj_map_tac ctxt live (Lazy.force map_id) (Lazy.force map_comp) (#map_cong0 axioms) (Lazy.force map_cong)) |> Thm.close_derivation \<^here> end; val inj_map = Lazy.lazy mk_inj_map; val set_map = map (fn thm => Lazy.lazy (fn () => mk_set_map thm)) (#set_map0 axioms); val wit_thms = if null nontriv_wit_thms then mk_wit_thms (map Lazy.force set_map) else nontriv_wit_thms; fun mk_in_bd () = let val bdT = fst (dest_relT (fastype_of bnf_bd_As)); val bdTs = replicate live bdT; val bd_bnfT = mk_bnf_T bdTs Calpha; val surj_imp_ordLeq_inst = (if live = 0 then TrueI else let val ranTs = map (fn AT => mk_sumT (AT, HOLogic.unitT)) As'; val funTs = map (fn T => bdT --> T) ranTs; val ran_bnfT = mk_bnf_T ranTs Calpha; val (revTs, Ts) = `rev (bd_bnfT :: funTs); val cTs = map (SOME o Thm.ctyp_of lthy) [ran_bnfT, Library.foldr1 HOLogic.mk_prodT Ts]; val tinst = fold (fn T => fn t => HOLogic.mk_case_prod (Term.absdummy T t)) (tl revTs) (Term.absdummy (hd revTs) (Term.list_comb (mk_bnf_map bdTs ranTs, map Bound (live - 1 downto 0)) $ Bound live)); val cts = [NONE, SOME (Thm.cterm_of lthy tinst)]; in Thm.instantiate' cTs cts @{thm surj_imp_ordLeq} end); val bd = mk_cexp (if live = 0 then ctwo else mk_csum (Library.foldr1 (uncurry mk_csum) (map mk_card_of As)) ctwo) (mk_csum bnf_bd_As (mk_card_of (HOLogic.mk_UNIV bd_bnfT))); val in_bd_goal = fold_rev Logic.all As (HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of (mk_in As bnf_sets_As CA')) bd)); in Goal.prove_sorry lthy [] [] in_bd_goal (fn {context = ctxt, prems = _} => mk_in_bd_tac ctxt live surj_imp_ordLeq_inst (Lazy.force map_comp) (Lazy.force map_id) (#map_cong0 axioms) (map Lazy.force set_map) (#set_bd axioms) (#bd_card_order axioms) bd_Card_order bd_Cinfinite bd_Cnotzero) |> Thm.close_derivation \<^here> end; val in_bd = Lazy.lazy mk_in_bd; val rel_OO_Grp = #rel_OO_Grp axioms; val rel_OO_Grps = no_refl [rel_OO_Grp]; fun mk_rel_Grp () = let val lhs = Term.list_comb (rel, map2 mk_Grp As fs); val rhs = mk_Grp (mk_in As bnf_sets_As CA') (Term.list_comb (bnf_map_AsBs, fs)); val goal = fold_rev Logic.all (As @ fs) (mk_Trueprop_eq (lhs, rhs)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_rel_Grp_tac ctxt rel_OO_Grps (#map_id0 axioms) (#map_cong0 axioms) (Lazy.force map_id) (Lazy.force map_comp) (map Lazy.force set_map)) |> Thm.close_derivation \<^here> end; val rel_Grp = Lazy.lazy mk_rel_Grp; fun mk_rel_prems f = map2 (HOLogic.mk_Trueprop oo f) Rs Rs_copy; fun mk_rel_concl f = HOLogic.mk_Trueprop (f (Term.list_comb (rel, Rs), Term.list_comb (rel, Rs_copy))); fun mk_rel_mono () = let val mono_prems = mk_rel_prems mk_leq; val mono_concl = mk_rel_concl (uncurry mk_leq); in Goal.prove_sorry lthy [] [] (fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (mono_prems, mono_concl))) (fn {context = ctxt, prems = _} => mk_rel_mono_tac ctxt rel_OO_Grps (Lazy.force in_mono)) |> Thm.close_derivation \<^here> end; fun mk_rel_cong0 () = let val cong_prems = mk_rel_prems (curry HOLogic.mk_eq); val cong_concl = mk_rel_concl HOLogic.mk_eq; in Goal.prove_sorry lthy [] [] (fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (cong_prems, cong_concl))) (fn {context = ctxt, prems = _} => (TRY o hyp_subst_tac ctxt THEN' rtac ctxt refl) 1) |> Thm.close_derivation \<^here> end; val rel_mono = Lazy.lazy mk_rel_mono; val rel_cong0 = Lazy.lazy mk_rel_cong0; fun mk_rel_eq () = Goal.prove_sorry lthy [] [] (mk_Trueprop_eq (Term.list_comb (relAsAs, map HOLogic.eq_const As'), HOLogic.eq_const CA')) (fn {context = ctxt, prems = _} => mk_rel_eq_tac ctxt live (Lazy.force rel_Grp) (Lazy.force rel_cong0) (#map_id0 axioms)) |> Thm.close_derivation \<^here>; val rel_eq = Lazy.lazy mk_rel_eq; fun mk_rel_conversep () = let val relBsAs = mk_bnf_rel pred2RT's CB' CA'; val lhs = Term.list_comb (relBsAs, map mk_conversep Rs); val rhs = mk_conversep (Term.list_comb (rel, Rs)); val le_goal = fold_rev Logic.all Rs (HOLogic.mk_Trueprop (mk_leq lhs rhs)); val le_thm = Goal.prove_sorry lthy [] [] le_goal (fn {context = ctxt, prems = _} => mk_rel_conversep_le_tac ctxt rel_OO_Grps (Lazy.force rel_eq) (#map_cong0 axioms) (Lazy.force map_comp) (map Lazy.force set_map)) |> Thm.close_derivation \<^here> val goal = fold_rev Logic.all Rs (mk_Trueprop_eq (lhs, rhs)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_rel_conversep_tac ctxt le_thm (Lazy.force rel_mono)) |> Thm.close_derivation \<^here> end; val rel_conversep = Lazy.lazy mk_rel_conversep; fun mk_rel_OO () = Goal.prove_sorry lthy [] [] (fold_rev Logic.all (Rs @ Ss) (HOLogic.mk_Trueprop (mk_leq rel_OO_lhs rel_OO_rhs))) (fn {context = ctxt, prems = _} => mk_rel_OO_le_tac ctxt rel_OO_Grps (Lazy.force rel_eq) (#map_cong0 axioms) (Lazy.force map_comp) (map Lazy.force set_map)) |> Thm.close_derivation \<^here> |> (fn thm => @{thm antisym} OF [thm, #le_rel_OO axioms]); val rel_OO = Lazy.lazy mk_rel_OO; fun mk_in_rel () = trans OF [rel_OO_Grp, @{thm OO_Grp_alt}] RS @{thm predicate2_eqD}; val in_rel = Lazy.lazy mk_in_rel; fun mk_rel_flip () = unfold_thms lthy @{thms conversep_iff} (Lazy.force rel_conversep RS @{thm predicate2_eqD}); val rel_flip = Lazy.lazy mk_rel_flip; fun mk_rel_mono_strong0 () = let fun mk_prem setA setB R S a b = HOLogic.mk_Trueprop (mk_Ball (setA $ x) (Term.absfree (dest_Free a) (mk_Ball (setB $ y) (Term.absfree (dest_Free b) (HOLogic.mk_imp (R $ a $ b, S $ a $ b)))))); val prems = HOLogic.mk_Trueprop (Term.list_comb (rel, Rs) $ x $ y) :: @{map 6} mk_prem bnf_sets_As bnf_sets_Bs Rs Rs_copy zs ys; val concl = HOLogic.mk_Trueprop (Term.list_comb (rel, Rs_copy) $ x $ y); in Goal.prove_sorry lthy [] [] (fold_rev Logic.all (x :: y :: Rs @ Rs_copy) (Logic.list_implies (prems, concl))) (fn {context = ctxt, prems = _} => mk_rel_mono_strong0_tac ctxt (Lazy.force in_rel) (map Lazy.force set_map)) |> Thm.close_derivation \<^here> end; val rel_mono_strong0 = Lazy.lazy mk_rel_mono_strong0; val rel_mono_strong = Lazy.map (Object_Logic.rulify lthy) rel_mono_strong0; fun mk_rel_cong_prem mk_implies x x' z z' set set' R R_copy = Logic.all z (Logic.all z' (mk_implies (mk_Trueprop_mem (z, set $ x), mk_implies (mk_Trueprop_mem (z', set' $ x'), mk_Trueprop_eq (R $ z $ z', R_copy $ z $ z'))))); fun mk_rel_cong mk_implies () = let val prem0 = mk_Trueprop_eq (x, x_copy); val prem1 = mk_Trueprop_eq (y, y_copy); val prems = @{map 6} (mk_rel_cong_prem mk_implies x_copy y_copy) zs ys bnf_sets_As bnf_sets_Bs Rs Rs_copy; val eq = mk_Trueprop_eq (Term.list_comb (rel, Rs) $ x $ y, Term.list_comb (rel, Rs_copy) $ x_copy $ y_copy); in fold (Variable.add_free_names lthy) (eq :: prem0 :: prem1 :: prems) [] |> (fn vars => Goal.prove_sorry lthy vars (prem0 :: prem1 :: prems) eq (fn {context = ctxt, prems} => mk_rel_cong_tac ctxt (chop 2 prems) (Lazy.force rel_mono_strong))) |> Thm.close_derivation \<^here> end; val rel_cong = Lazy.lazy (mk_rel_cong Logic.mk_implies); val rel_cong_simp = Lazy.lazy (mk_rel_cong (fn (a, b) => \<^term>\simp_implies\ $ a $ b)); fun mk_pred_prems f = map2 (HOLogic.mk_Trueprop oo f) Ps Ps_copy; fun mk_pred_concl f = HOLogic.mk_Trueprop (f (Term.list_comb (pred, Ps), Term.list_comb (pred, Ps_copy))); fun mk_pred_cong0 () = let val cong_prems = mk_pred_prems (curry HOLogic.mk_eq); val cong_concl = mk_pred_concl HOLogic.mk_eq; in Goal.prove_sorry lthy [] [] (fold_rev Logic.all (Ps @ Ps_copy) (Logic.list_implies (cong_prems, cong_concl))) (fn {context = ctxt, prems = _} => (TRY o hyp_subst_tac ctxt THEN' rtac ctxt refl) 1) |> Thm.close_derivation \<^here> end; val pred_cong0 = Lazy.lazy mk_pred_cong0; fun mk_rel_eq_onp () = let val lhs = Term.list_comb (relAsAs, map mk_eq_onp Ps); val rhs = mk_eq_onp (Term.list_comb (pred, Ps)); in Goal.prove_sorry lthy (map fst Ps') [] (mk_Trueprop_eq (lhs, rhs)) (fn {context = ctxt, prems = _} => mk_rel_eq_onp_tac ctxt (#pred_set axioms) (#map_id0 axioms) (Lazy.force rel_Grp)) |> Thm.close_derivation \<^here> end; val rel_eq_onp = Lazy.lazy mk_rel_eq_onp; val pred_rel = Lazy.map (fn thm => thm RS sym RS @{thm eq_onp_eqD}) rel_eq_onp; fun mk_pred_mono_strong0 () = let fun mk_prem setA P Q a = HOLogic.mk_Trueprop (mk_Ball (setA $ x) (Term.absfree (dest_Free a) (HOLogic.mk_imp (P $ a, Q $ a)))); val prems = HOLogic.mk_Trueprop (Term.list_comb (pred, Ps) $ x) :: @{map 4} mk_prem bnf_sets_As Ps Ps_copy zs; val concl = HOLogic.mk_Trueprop (Term.list_comb (pred, Ps_copy) $ x); in Goal.prove_sorry lthy [] [] (fold_rev Logic.all (x :: Ps @ Ps_copy) (Logic.list_implies (prems, concl))) (fn {context = ctxt, prems = _} => mk_pred_mono_strong0_tac ctxt (Lazy.force pred_rel) (Lazy.force rel_mono_strong0)) |> Thm.close_derivation \<^here> end; val pred_mono_strong0 = Lazy.lazy mk_pred_mono_strong0; val pred_mono_strong = Lazy.map (Object_Logic.rulify lthy) pred_mono_strong0; fun mk_pred_mono () = let val mono_prems = mk_pred_prems mk_leq; val mono_concl = mk_pred_concl (uncurry mk_leq); in Goal.prove_sorry lthy [] [] (fold_rev Logic.all (Ps @ Ps_copy) (Logic.list_implies (mono_prems, mono_concl))) (fn {context = ctxt, prems = _} => mk_pred_mono_tac ctxt (Lazy.force rel_eq_onp) (Lazy.force rel_mono)) |> Thm.close_derivation \<^here> end; val pred_mono = Lazy.lazy mk_pred_mono; fun mk_pred_cong_prem mk_implies x z set P P_copy = Logic.all z (mk_implies (mk_Trueprop_mem (z, set $ x), mk_Trueprop_eq (P $ z, P_copy $ z))); fun mk_pred_cong mk_implies () = let val prem0 = mk_Trueprop_eq (x, x_copy); val prems = @{map 4} (mk_pred_cong_prem mk_implies x_copy) zs bnf_sets_As Ps Ps_copy; val eq = mk_Trueprop_eq (Term.list_comb (pred, Ps) $ x, Term.list_comb (pred, Ps_copy) $ x_copy); in fold (Variable.add_free_names lthy) (eq :: prem0 :: prems) [] |> (fn vars => Goal.prove_sorry lthy vars (prem0 :: prems) eq (fn {context = ctxt, prems} => mk_rel_cong_tac ctxt (chop 1 prems) (Lazy.force pred_mono_strong))) |> Thm.close_derivation \<^here> end; val pred_cong = Lazy.lazy (mk_pred_cong Logic.mk_implies); val pred_cong_simp = Lazy.lazy (mk_pred_cong (fn (a, b) => \<^term>\simp_implies\ $ a $ b)); fun mk_map_cong_pred () = let val prem0 = mk_Trueprop_eq (x, x_copy); fun mk_eq f g z = Term.absfree (dest_Free z) (HOLogic.mk_eq (f $ z, g $ z)); val prem = HOLogic.mk_Trueprop (Term.list_comb (pred, @{map 3} mk_eq fs fs_copy zs) $ x_copy); val eq = mk_Trueprop_eq (Term.list_comb (bnf_map_AsBs, fs) $ x, Term.list_comb (bnf_map_AsBs, fs_copy) $ x_copy); val goal = fold_rev Logic.all (x :: x_copy :: fs @ fs_copy) (Logic.list_implies ([prem0, prem], eq)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt [#pred_set axioms] THEN HEADGOAL (EVERY' [REPEAT_DETERM o etac ctxt conjE, etac ctxt (Lazy.force map_cong) THEN_ALL_NEW (etac ctxt @{thm bspec} THEN' assume_tac ctxt)])) |> Thm.close_derivation \<^here> end; val map_cong_pred = Lazy.lazy mk_map_cong_pred; fun mk_rel_map () = let fun mk_goal lhs rhs = fold_rev Logic.all ([x, y] @ S_CsBs @ S_AsCs @ is @ gs) (mk_Trueprop_eq (lhs, rhs)); val lhss = [Term.list_comb (relCsBs, S_CsBs) $ (Term.list_comb (bnf_map_AsCs, is) $ x) $ y, Term.list_comb (relAsCs, S_AsCs) $ x $ (Term.list_comb (bnf_map_BsCs, gs) $ y)]; val rhss = [Term.list_comb (rel, @{map 3} (fn f => fn P => fn T => mk_vimage2p f (HOLogic.id_const T) $ P) is S_CsBs Bs') $ x $ y, Term.list_comb (rel, @{map 3} (fn f => fn P => fn T => mk_vimage2p (HOLogic.id_const T) f $ P) gs S_AsCs As') $ x $ y]; val goals = map2 mk_goal lhss rhss; in goals |> map (fn goal => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_rel_map0_tac ctxt live (Lazy.force rel_OO) (Lazy.force rel_conversep) (Lazy.force rel_Grp) (Lazy.force map_id))) |> map (unfold_thms lthy @{thms vimage2p_def[of id, simplified id_apply] vimage2p_def[of _ id, simplified id_apply]}) |> map (Thm.close_derivation \<^here>) end; val rel_map = Lazy.lazy mk_rel_map; fun mk_rel_refl () = @{thm ge_eq_refl[OF ord_eq_le_trans]} OF [Lazy.force rel_eq RS sym, Lazy.force rel_mono OF (replicate live @{thm refl_ge_eq})]; val rel_refl = Lazy.lazy mk_rel_refl; fun mk_rel_refl_strong () = (rule_by_tactic lthy (ALLGOALS (Object_Logic.full_atomize_tac lthy)) ((Lazy.force rel_eq RS @{thm predicate2_eqD}) RS @{thm iffD2[OF _ refl]} RS Lazy.force rel_mono_strong)) OF (replicate live @{thm diag_imp_eq_le}) val rel_refl_strong = Lazy.lazy mk_rel_refl_strong; fun mk_rel_preserves mk_prop prop_conv_thm thm () = let val Rs = map2 retype_const_or_free self_pred2RTs Rs; val prems = map (HOLogic.mk_Trueprop o mk_prop) Rs; val goal = HOLogic.mk_Trueprop (mk_prop (Term.list_comb (relAsAs, Rs))); val vars = fold (Variable.add_free_names lthy) (goal :: prems) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, goal)) (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt [prop_conv_thm] THEN HEADGOAL (rtac ctxt (Lazy.force thm RS sym RS @{thm ord_eq_le_trans}) THEN' rtac ctxt (Lazy.force rel_mono) THEN_ALL_NEW assume_tac ctxt)) |> Thm.close_derivation \<^here> end; val rel_reflp = Lazy.lazy (mk_rel_preserves mk_reflp @{thm reflp_eq} rel_eq); val rel_symp = Lazy.lazy (mk_rel_preserves mk_symp @{thm symp_conversep} rel_conversep); val rel_transp = Lazy.lazy (mk_rel_preserves mk_transp @{thm transp_relcompp} rel_OO); fun mk_pred_True () = let val lhs = Term.list_comb (pred, map (fn T => absdummy T \<^term>\True\) As'); val rhs = absdummy CA' \<^term>\True\; val goal = mk_Trueprop_eq (lhs, rhs); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => HEADGOAL (EVERY' (map (rtac ctxt) [ext, Lazy.force pred_rel RS trans, Lazy.force rel_cong0 RS fun_cong RS fun_cong RS trans OF replicate live @{thm eq_onp_True}, Lazy.force rel_eq RS fun_cong RS fun_cong RS trans, @{thm eqTrueI[OF refl]}]))) |> Thm.close_derivation \<^here> end; val pred_True = Lazy.lazy mk_pred_True; fun mk_pred_map () = let val lhs = Term.list_comb (pred', Qs) $ (Term.list_comb (bnf_map_AsBs, fs) $ x); val rhs = Term.list_comb (pred, @{map 2} (curry HOLogic.mk_comp) Qs fs) $ x; val goal = mk_Trueprop_eq (lhs, rhs); val vars = Variable.add_free_names lthy goal []; val pred_set = #pred_set axioms RS fun_cong RS sym; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => HEADGOAL (rtac ctxt (pred_set RSN (2, pred_set RSN (2, box_equals)))) THEN unfold_thms_tac ctxt (@{thms Ball_image_comp ball_empty} @ map Lazy.force set_map) THEN HEADGOAL (rtac ctxt refl)) |> Thm.close_derivation \<^here> end; val pred_map = Lazy.lazy mk_pred_map; fun mk_map_transfer () = let val rels = map2 mk_rel_fun transfer_domRs transfer_ranRs; val rel = mk_rel_fun (Term.list_comb (mk_bnf_rel transfer_domRTs CA' CB1, transfer_domRs)) (Term.list_comb (mk_bnf_rel transfer_ranRTs CB' CB2, transfer_ranRs)); val concl = HOLogic.mk_Trueprop (fold_rev mk_rel_fun rels rel $ bnf_map_AsBs $ mk_bnf_map B1Ts B2Ts); in Goal.prove_sorry lthy [] [] (fold_rev Logic.all (transfer_domRs @ transfer_ranRs) concl) (fn {context = ctxt, prems = _} => mk_map_transfer_tac ctxt (Lazy.force rel_mono) (Lazy.force in_rel) (map Lazy.force set_map) (#map_cong0 axioms) (Lazy.force map_comp)) |> Thm.close_derivation \<^here> end; val map_transfer = Lazy.lazy mk_map_transfer; fun mk_pred_transfer () = let val iff = HOLogic.eq_const HOLogic.boolT; val prem_rels = map (fn T => mk_rel_fun T iff) Rs; val prem_elems = mk_rel_fun (Term.list_comb (rel, Rs)) iff; val goal = HOLogic.mk_Trueprop (fold_rev mk_rel_fun prem_rels prem_elems $ pred $ pred'); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_pred_transfer_tac ctxt live (Lazy.force in_rel) (Lazy.force pred_map) (Lazy.force pred_cong)) |> Thm.close_derivation \<^here> end; val pred_transfer = Lazy.lazy mk_pred_transfer; fun mk_rel_transfer () = let val iff = HOLogic.eq_const HOLogic.boolT; val prem_rels = map2 (fn T1 => fn T2 => mk_rel_fun T1 (mk_rel_fun T2 iff)) S_AsCs S_BsEs; val prem_elems = mk_rel_fun (Term.list_comb (mk_bnf_rel pred2RTsAsCs CA' CC', S_AsCs)) (mk_rel_fun (Term.list_comb (mk_bnf_rel pred2RTsBsEs CB' CE', S_BsEs)) iff); val goal = HOLogic.mk_Trueprop (fold_rev mk_rel_fun prem_rels prem_elems $ rel $ relCsEs); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_rel_transfer_tac ctxt (Lazy.force in_rel) (Lazy.force rel_map) (Lazy.force rel_mono_strong)) |> Thm.close_derivation \<^here> end; val rel_transfer = Lazy.lazy mk_rel_transfer; fun mk_set_transfer () = let val rel_sets = map2 (fn A => fn B => mk_rel 1 [A] [B] \<^term>\rel_set\) As' Bs'; val rel_Rs = Term.list_comb (rel, Rs); val goals = @{map 4} (fn R => fn rel_set => fn setA => fn setB => HOLogic.mk_Trueprop (mk_rel_fun rel_Rs (rel_set $ R) $ setA $ setB)) Rs rel_sets bnf_sets_As bnf_sets_Bs; in if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_set_transfer_tac ctxt (Lazy.force in_rel) (map Lazy.force set_map)) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end end; val set_transfer = Lazy.lazy mk_set_transfer; fun mk_inj_map_strong () = let val assms = @{map 5} (fn setA => fn z => fn f => fn z' => fn f' => fold_rev Logic.all [z, z'] (Logic.mk_implies (mk_Trueprop_mem (z, setA $ x), Logic.mk_implies (mk_Trueprop_mem (z', setA $ x'), Logic.mk_implies (mk_Trueprop_eq (f $ z, f' $ z'), mk_Trueprop_eq (z, z')))))) bnf_sets_As zs fs zs' fs'; val concl = Logic.mk_implies (mk_Trueprop_eq (Term.list_comb (bnf_map_AsBs, fs) $ x, Term.list_comb (bnf_map_AsBs, fs') $ x'), mk_Trueprop_eq (x, x')); val goal = fold_rev Logic.all (x :: x' :: fs @ fs') (fold_rev (curry Logic.mk_implies) assms concl); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_inj_map_strong_tac ctxt (Lazy.force rel_eq) (Lazy.force rel_map) (Lazy.force rel_mono_strong)) |> Thm.close_derivation \<^here> end; val inj_map_strong = Lazy.lazy mk_inj_map_strong; val defs = mk_defs bnf_map_def bnf_set_defs bnf_rel_def bnf_pred_def; val facts = mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_map in_bd in_cong in_mono in_rel inj_map inj_map_strong map_comp map_cong map_cong_simp map_cong_pred map_id map_ident0 map_ident map_transfer rel_eq rel_flip set_map rel_cong0 rel_cong rel_cong_simp rel_map rel_mono rel_mono_strong0 rel_mono_strong set_transfer rel_Grp rel_conversep rel_OO rel_refl rel_refl_strong rel_reflp rel_symp rel_transp rel_transfer rel_eq_onp pred_transfer pred_True pred_map pred_rel pred_mono_strong0 pred_mono_strong pred_mono pred_cong0 pred_cong pred_cong_simp; val wits = map2 mk_witness bnf_wits wit_thms; val bnf_rel = Term.subst_atomic_types ((Ds ~~ deads) @ (As' ~~ alphas) @ (Bs' ~~ betas)) rel; val bnf_pred = Term.subst_atomic_types ((Ds ~~ deads) @ (As' ~~ alphas)) pred; val bnf = mk_bnf bnf_b Calpha live alphas betas dead deads bnf_map bnf_sets bnf_bd axioms defs facts wits bnf_rel bnf_pred; in note_bnf_thms fact_policy qualify bnf_b bnf lthy end; val one_step_defs = no_reflexive (bnf_map_def :: bnf_bd_def :: bnf_set_defs @ bnf_wit_defs @ [bnf_rel_def, bnf_pred_def]); in (key, goals, wit_goalss, after_qed, lthy, one_step_defs) end; structure BNF_Plugin = Plugin(type T = bnf); fun bnf_interpretation name f = BNF_Plugin.interpretation name (fn bnf => fn lthy => f (transfer_bnf (Proof_Context.theory_of lthy) bnf) lthy); val interpret_bnf = BNF_Plugin.data; fun register_bnf_raw key bnf = Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => Data.map (Symtab.update (key, morph_bnf phi bnf))); fun register_bnf plugins key bnf = register_bnf_raw key bnf #> interpret_bnf plugins bnf; fun bnf_def const_policy fact_policy internal qualify tacs wit_tac Ds map_b rel_b pred_b set_bs raw_csts = (fn (_, goals, (triv_tac_opt, wit_goalss), after_qed, lthy, one_step_defs) => let fun mk_wits_tac ctxt set_maps = TRYALL Goal.conjunction_tac THEN (case triv_tac_opt of SOME tac => tac ctxt set_maps | NONE => unfold_thms_tac ctxt one_step_defs THEN wit_tac ctxt); val wit_goals = map Logic.mk_conjunction_balanced wit_goalss; fun mk_wit_thms set_maps = Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced wit_goals) (fn {context = ctxt, prems = _} => mk_wits_tac ctxt set_maps) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length wit_goals) |> map2 (Conjunction.elim_balanced o length) wit_goalss |> (map o map) (Thm.forall_elim_vars 0); in map2 (Thm.close_derivation \<^here> oo Goal.prove_sorry lthy [] []) goals (map (fn tac => fn {context = ctxt, prems = _} => unfold_thms_tac ctxt one_step_defs THEN tac ctxt) tacs) |> (fn thms => after_qed mk_wit_thms (map single thms) lthy) end) o prepare_def const_policy fact_policy internal qualify (K I) (K I) Ds map_b rel_b pred_b set_bs raw_csts; fun bnf_cmd (raw_csts, raw_plugins) = (fn (key, goals, (triv_tac_opt, wit_goalss), after_qed, lthy, defs) => let val plugins = raw_plugins lthy; val wit_goals = map Logic.mk_conjunction_balanced wit_goalss; fun mk_triv_wit_thms tac set_maps = Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced wit_goals) (fn {context = ctxt, prems = _} => TRYALL Goal.conjunction_tac THEN tac ctxt set_maps) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length wit_goals) |> map2 (Conjunction.elim_balanced o length) wit_goalss |> (map o map) (Thm.forall_elim_vars 0); val (mk_wit_thms, nontriv_wit_goals) = (case triv_tac_opt of NONE => (fn _ => [], map (map (rpair [])) wit_goalss) | SOME tac => (mk_triv_wit_thms tac, [])); in lthy |> Proof.theorem NONE (uncurry (register_bnf plugins key) oo after_qed mk_wit_thms) (map (single o rpair []) goals @ nontriv_wit_goals) |> Proof.unfolding ([[(@{thm OO_Grp_alt} :: @{thm mem_Collect_eq} :: defs, [])]]) |> Proof.refine_singleton (Method.Basic (fn ctxt => Method.SIMPLE_METHOD (TRYALL (rtac ctxt refl)))) end) o prepare_def Do_Inline (user_policy Note_Some) false I Syntax.read_typ Syntax.read_term NONE Binding.empty Binding.empty Binding.empty [] raw_csts; fun print_bnfs ctxt = let fun pretty_set sets i = Pretty.block [Pretty.str (mk_setN (i + 1) ^ ":"), Pretty.brk 1, Pretty.quote (Syntax.pretty_term ctxt (nth sets i))]; fun pretty_bnf (key, BNF {T, map, sets, bd, live, lives, dead, deads, ...}) = Pretty.big_list (Pretty.string_of (Pretty.block [Pretty.str key, Pretty.str ":", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt T)])) ([Pretty.block [Pretty.str "live:", Pretty.brk 1, Pretty.str (string_of_int live), Pretty.brk 3, Pretty.list "[" "]" (List.map (Syntax.pretty_typ ctxt) lives)], Pretty.block [Pretty.str "dead:", Pretty.brk 1, Pretty.str (string_of_int dead), Pretty.brk 3, Pretty.list "[" "]" (List.map (Syntax.pretty_typ ctxt) deads)], Pretty.block [Pretty.str (mapN ^ ":"), Pretty.brk 1, Pretty.quote (Syntax.pretty_term ctxt map)]] @ List.map (pretty_set sets) (0 upto length sets - 1) @ [Pretty.block [Pretty.str (bdN ^ ":"), Pretty.brk 1, Pretty.quote (Syntax.pretty_term ctxt bd)]]); in Pretty.big_list "Registered bounded natural functors:" (map pretty_bnf (sort_by fst (Symtab.dest (Data.get (Context.Proof ctxt))))) |> Pretty.writeln end; val _ = Outer_Syntax.command \<^command_keyword>\print_bnfs\ "print all bounded natural functors" (Scan.succeed (Toplevel.keep (print_bnfs o Toplevel.context_of))); val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\bnf\ "register a type as a bounded natural functor" (parse_opt_binding_colon -- Parse.typ --| (Parse.reserved "map" -- \<^keyword>\:\) -- Parse.term -- Scan.optional ((Parse.reserved "sets" -- \<^keyword>\:\) |-- Scan.repeat1 (Scan.unless (Parse.reserved "bd") Parse.term)) [] --| (Parse.reserved "bd" -- \<^keyword>\:\) -- Parse.term -- Scan.optional ((Parse.reserved "wits" -- \<^keyword>\:\) |-- Scan.repeat1 (Scan.unless (Parse.reserved "rel" || Parse.reserved "plugins") Parse.term)) [] -- Scan.option ((Parse.reserved "rel" -- \<^keyword>\:\) |-- Parse.term) -- Scan.option ((Parse.reserved "pred" -- \<^keyword>\:\) |-- Parse.term) -- Scan.optional Plugin_Name.parse_filter (K Plugin_Name.default_filter) >> bnf_cmd); end; diff --git a/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML b/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML --- a/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML +++ b/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML @@ -1,2930 +1,2930 @@ (* Title: HOL/Tools/BNF/bnf_fp_def_sugar.ML Author: Jasmin Blanchette, TU Muenchen Author: Martin Desharnais, TU Muenchen Copyright 2012, 2013, 2014 Sugared datatype and codatatype constructions. *) signature BNF_FP_DEF_SUGAR = sig type fp_ctr_sugar = {ctrXs_Tss: typ list list, ctor_iff_dtor: thm, ctr_defs: thm list, ctr_sugar: Ctr_Sugar.ctr_sugar, ctr_transfers: thm list, case_transfers: thm list, disc_transfers: thm list, sel_transfers: thm list} type fp_bnf_sugar = {map_thms: thm list, map_disc_iffs: thm list, map_selss: thm list list, rel_injects: thm list, rel_distincts: thm list, rel_sels: thm list, rel_intros: thm list, rel_cases: thm list, pred_injects: thm list, set_thms: thm list, set_selssss: thm list list list list, set_introssss: thm list list list list, set_cases: thm list} type fp_co_induct_sugar = {co_rec: term, common_co_inducts: thm list, co_inducts: thm list, co_rec_def: thm, co_rec_thms: thm list, co_rec_discs: thm list, co_rec_disc_iffs: thm list, co_rec_selss: thm list list, co_rec_codes: thm list, co_rec_transfers: thm list, co_rec_o_maps: thm list, common_rel_co_inducts: thm list, rel_co_inducts: thm list, common_set_inducts: thm list, set_inducts: thm list} type fp_sugar = {T: typ, BT: typ, X: typ, fp: BNF_Util.fp_kind, fp_res_index: int, fp_res: BNF_FP_Util.fp_result, pre_bnf: BNF_Def.bnf, fp_bnf: BNF_Def.bnf, absT_info: BNF_Comp.absT_info, fp_nesting_bnfs: BNF_Def.bnf list, live_nesting_bnfs: BNF_Def.bnf list, fp_ctr_sugar: fp_ctr_sugar, fp_bnf_sugar: fp_bnf_sugar, fp_co_induct_sugar: fp_co_induct_sugar option} val co_induct_of: 'a list -> 'a val strong_co_induct_of: 'a list -> 'a val morph_fp_bnf_sugar: morphism -> fp_bnf_sugar -> fp_bnf_sugar val morph_fp_co_induct_sugar: morphism -> fp_co_induct_sugar -> fp_co_induct_sugar val morph_fp_ctr_sugar: morphism -> fp_ctr_sugar -> fp_ctr_sugar val morph_fp_sugar: morphism -> fp_sugar -> fp_sugar val transfer_fp_sugar: theory -> fp_sugar -> fp_sugar val fp_sugar_of: Proof.context -> string -> fp_sugar option val fp_sugar_of_global: theory -> string -> fp_sugar option val fp_sugars_of: Proof.context -> fp_sugar list val fp_sugars_of_global: theory -> fp_sugar list val fp_sugars_interpretation: string -> (fp_sugar list -> local_theory -> local_theory) -> theory -> theory val interpret_fp_sugars: (string -> bool) -> fp_sugar list -> local_theory -> local_theory val register_fp_sugars_raw: fp_sugar list -> local_theory -> local_theory val register_fp_sugars: (string -> bool) -> fp_sugar list -> local_theory -> local_theory val merge_type_args: BNF_Util.fp_kind -> ''a list * ''a list -> ''a list val type_args_named_constrained_of_spec: (((('a * 'b) * 'c) * 'd) * 'e) * 'f -> 'a val type_binding_of_spec: (((('a * 'b) * 'c) * 'd) * 'e) * 'f -> 'b val mixfix_of_spec: ((('a * 'b) * 'c) * 'd) * 'e -> 'b val mixfixed_ctr_specs_of_spec: (('a * 'b) * 'c) * 'd -> 'b val map_binding_of_spec: ('a * ('b * 'c * 'd)) * 'e -> 'b val rel_binding_of_spec: ('a * ('b * 'c * 'd)) * 'e -> 'c val pred_binding_of_spec: ('a * ('b * 'c * 'd)) * 'e -> 'd val sel_default_eqs_of_spec: 'a * 'b -> 'b val mk_parametricity_goal: Proof.context -> term list -> term -> term -> term val flat_corec_preds_predsss_gettersss: 'a list -> 'a list list list -> 'a list list list -> 'a list val mk_ctor: typ list -> term -> term val mk_dtor: typ list -> term -> term val mk_bnf_sets: BNF_Def.bnf -> string * term list val liveness_of_fp_bnf: int -> BNF_Def.bnf -> bool list val nesting_bnfs: Proof.context -> typ list list list -> typ list -> BNF_Def.bnf list val massage_simple_notes: string -> (bstring * 'a list * (int -> 'b)) list -> ((binding * 'c list) * ('a list * 'b) list) list val massage_multi_notes: string list -> typ list -> (string * 'a list list * (string -> 'b)) list -> ((binding * 'b) * ('a list * 'c list) list) list val define_ctrs_dtrs_for_type: string -> typ -> term -> term -> thm -> thm -> int -> int list -> term -> binding list -> mixfix list -> typ list list -> local_theory -> (term list list * term list * thm * thm list) * local_theory val wrap_ctrs: (string -> bool) -> BNF_Util.fp_kind -> bool -> string -> thm -> int -> int list -> thm -> thm -> binding list -> binding list list -> term list -> term list -> thm -> thm list -> local_theory -> Ctr_Sugar.ctr_sugar * local_theory val derive_map_set_rel_pred_thms: (string -> bool) -> BNF_Util.fp_kind -> int -> typ list -> typ list -> typ -> typ -> thm list -> thm list -> thm list -> thm list -> thm list -> thm list -> thm list -> thm list -> thm list -> string -> BNF_Def.bnf -> BNF_Def.bnf list -> typ -> term -> thm -> thm -> thm -> thm list -> thm -> thm -> thm list -> thm -> thm list -> thm list -> thm list -> typ list list -> Ctr_Sugar.ctr_sugar -> local_theory -> (thm list * thm list * thm list list * thm list * thm list * thm list * thm list * thm list * thm list * thm list * thm list list list list * thm list list list list * thm list * thm list * thm list * thm list * thm list) * local_theory type lfp_sugar_thms = (thm list * thm * Token.src list) * (thm list list * Token.src list) val morph_lfp_sugar_thms: morphism -> lfp_sugar_thms -> lfp_sugar_thms val transfer_lfp_sugar_thms: theory -> lfp_sugar_thms -> lfp_sugar_thms type gfp_sugar_thms = ((thm list * thm) list * (Token.src list * Token.src list)) * thm list list * thm list list * (thm list list * Token.src list) * (thm list list list * Token.src list) val morph_gfp_sugar_thms: morphism -> gfp_sugar_thms -> gfp_sugar_thms val transfer_gfp_sugar_thms: theory -> gfp_sugar_thms -> gfp_sugar_thms val mk_co_recs_prelims: Proof.context -> BNF_Util.fp_kind -> typ list list list -> typ list -> typ list -> typ list -> typ list -> int list -> int list list -> term list -> term list * (typ list list * typ list list list list * term list list * term list list list list) option * (string * term list * term list list * (((term list list * term list list * term list list list list * term list list list list) * term list list list) * typ list)) option val repair_nullary_single_ctr: typ list list -> typ list list val mk_corec_p_pred_types: typ list -> int list -> typ list list val mk_corec_fun_arg_types: typ list list list -> typ list -> typ list -> typ list -> int list -> int list list -> term -> typ list list * (typ list list list list * typ list list list * typ list list list list * typ list) val define_co_rec_as: BNF_Util.fp_kind -> typ list -> typ -> binding -> term -> local_theory -> (term * thm) * local_theory val define_rec: typ list list * typ list list list list * term list list * term list list list list -> (string -> binding) -> typ list -> typ list -> term list -> term -> Proof.context -> (term * thm) * Proof.context val define_corec: 'a * term list * term list list * (((term list list * term list list * term list list list list * term list list list list) * term list list list) * typ list) -> (string -> binding) -> 'b list -> typ list -> term list -> term -> local_theory -> (term * thm) * local_theory val mk_induct_raw_prem: (term -> term) -> Proof.context -> typ list list -> (string * term list) list -> term -> term -> typ list -> typ list -> term list * ((term * (term * term)) list * (int * term)) list * term val finish_induct_prem: Proof.context -> int -> term list -> term list * ((term * (term * term)) list * (int * term)) list * term -> term val mk_coinduct_prem: Proof.context -> typ list list -> typ list list -> term list -> term -> term -> term -> int -> term list -> term list list -> term list -> term list list -> typ list list -> term val mk_induct_attrs: term list list -> Token.src list val mk_coinduct_attrs: typ list -> term list list -> term list list -> int list list -> Token.src list * Token.src list val derive_induct_recs_thms_for_types: (string -> bool) -> BNF_Def.bnf list -> ('a * typ list list list list * term list list * 'b) option -> thm -> thm list -> BNF_Def.bnf list -> BNF_Def.bnf list -> typ list -> typ list -> typ list -> typ list list list -> thm list -> thm list -> thm list -> term list list -> thm list list -> term list -> thm list -> Proof.context -> lfp_sugar_thms val derive_coinduct_thms_for_types: Proof.context -> bool -> (term -> term) -> BNF_Def.bnf list -> thm -> thm list -> BNF_Def.bnf list -> typ list -> typ list -> typ list list list -> int list -> thm list -> thm list -> (thm -> thm) -> thm list list -> Ctr_Sugar.ctr_sugar list -> (thm list * thm) list val derive_coinduct_corecs_thms_for_types: Proof.context -> BNF_Def.bnf list -> string * term list * term list list * (((term list list * term list list * term list list list list * term list list list list) * term list list list) * typ list) -> thm -> thm list -> thm list -> thm list -> BNF_Def.bnf list -> typ list -> typ list -> typ list -> typ list list list -> int list list -> int list list -> int list -> thm list -> thm list -> (thm -> thm) -> thm list list -> Ctr_Sugar.ctr_sugar list -> term list -> thm list -> gfp_sugar_thms val co_datatypes: BNF_Util.fp_kind -> (mixfix list -> binding list -> binding list -> binding list -> binding list list -> binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list -> BNF_Comp.absT_info list -> local_theory -> BNF_FP_Util.fp_result * local_theory) -> Ctr_Sugar.ctr_options * ((((((binding option * (typ * sort)) list * binding) * mixfix) * ((binding, binding * typ) Ctr_Sugar.ctr_spec * mixfix) list) * (binding * binding * binding)) * term list) list -> local_theory -> local_theory val co_datatype_cmd: BNF_Util.fp_kind -> (mixfix list -> binding list -> binding list -> binding list -> binding list list -> binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list -> BNF_Comp.absT_info list -> local_theory -> BNF_FP_Util.fp_result * Proof.context) -> ((Proof.context -> Plugin_Name.filter) * bool) * ((((((binding option * (string * string option)) list * binding) * mixfix) * ((binding, binding * string) Ctr_Sugar.ctr_spec * mixfix) list) * (binding * binding * binding)) * string list) list -> Proof.context -> local_theory val parse_ctr_arg: (binding * string) parser val parse_ctr_specs: ((binding, binding * string) Ctr_Sugar.ctr_spec * mixfix) list parser val parse_spec: ((((((binding option * (string * string option)) list * binding) * mixfix) * ((binding, binding * string) Ctr_Sugar.ctr_spec * mixfix) list) * (binding * binding * binding)) * string list) parser val parse_co_datatype: (Ctr_Sugar.ctr_options_cmd * ((((((binding option * (string * string option)) list * binding) * mixfix) * ((binding, binding * string) Ctr_Sugar.ctr_spec * mixfix) list) * (binding * binding * binding)) * string list) list) parser val parse_co_datatype_cmd: BNF_Util.fp_kind -> (mixfix list -> binding list -> binding list -> binding list -> binding list list -> binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list -> BNF_Comp.absT_info list -> local_theory -> BNF_FP_Util.fp_result * local_theory) -> (local_theory -> local_theory) parser end; structure BNF_FP_Def_Sugar : BNF_FP_DEF_SUGAR = struct open Ctr_Sugar open BNF_FP_Rec_Sugar_Util open BNF_Util open BNF_Comp open BNF_Def open BNF_FP_Util open BNF_FP_Def_Sugar_Tactics val Eq_prefix = "Eq_"; val case_transferN = "case_transfer"; val ctor_iff_dtorN = "ctor_iff_dtor"; val ctr_transferN = "ctr_transfer"; val disc_transferN = "disc_transfer"; val sel_transferN = "sel_transfer"; val corec_codeN = "corec_code"; val corec_transferN = "corec_transfer"; val map_disc_iffN = "map_disc_iff"; val map_o_corecN = "map_o_corec"; val map_selN = "map_sel"; val pred_injectN = "pred_inject"; val rec_o_mapN = "rec_o_map"; val rec_transferN = "rec_transfer"; val set0N = "set0"; val set_casesN = "set_cases"; val set_introsN = "set_intros"; val set_inductN = "set_induct"; val set_selN = "set_sel"; type fp_ctr_sugar = {ctrXs_Tss: typ list list, ctor_iff_dtor: thm, ctr_defs: thm list, ctr_sugar: Ctr_Sugar.ctr_sugar, ctr_transfers: thm list, case_transfers: thm list, disc_transfers: thm list, sel_transfers: thm list}; type fp_bnf_sugar = {map_thms: thm list, map_disc_iffs: thm list, map_selss: thm list list, rel_injects: thm list, rel_distincts: thm list, rel_sels: thm list, rel_intros: thm list, rel_cases: thm list, pred_injects: thm list, set_thms: thm list, set_selssss: thm list list list list, set_introssss: thm list list list list, set_cases: thm list}; type fp_co_induct_sugar = {co_rec: term, common_co_inducts: thm list, co_inducts: thm list, co_rec_def: thm, co_rec_thms: thm list, co_rec_discs: thm list, co_rec_disc_iffs: thm list, co_rec_selss: thm list list, co_rec_codes: thm list, co_rec_transfers: thm list, co_rec_o_maps: thm list, common_rel_co_inducts: thm list, rel_co_inducts: thm list, common_set_inducts: thm list, set_inducts: thm list}; type fp_sugar = {T: typ, BT: typ, X: typ, fp: fp_kind, fp_res_index: int, fp_res: fp_result, pre_bnf: bnf, fp_bnf: bnf, absT_info: absT_info, fp_nesting_bnfs: bnf list, live_nesting_bnfs: bnf list, fp_ctr_sugar: fp_ctr_sugar, fp_bnf_sugar: fp_bnf_sugar, fp_co_induct_sugar: fp_co_induct_sugar option}; fun co_induct_of (i :: _) = i; fun strong_co_induct_of [_, s] = s; fun morph_fp_bnf_sugar phi ({map_thms, map_disc_iffs, map_selss, rel_injects, rel_distincts, rel_sels, rel_intros, rel_cases, pred_injects, set_thms, set_selssss, set_introssss, set_cases} : fp_bnf_sugar) = {map_thms = map (Morphism.thm phi) map_thms, map_disc_iffs = map (Morphism.thm phi) map_disc_iffs, map_selss = map (map (Morphism.thm phi)) map_selss, rel_injects = map (Morphism.thm phi) rel_injects, rel_distincts = map (Morphism.thm phi) rel_distincts, rel_sels = map (Morphism.thm phi) rel_sels, rel_intros = map (Morphism.thm phi) rel_intros, rel_cases = map (Morphism.thm phi) rel_cases, pred_injects = map (Morphism.thm phi) pred_injects, set_thms = map (Morphism.thm phi) set_thms, set_selssss = map (map (map (map (Morphism.thm phi)))) set_selssss, set_introssss = map (map (map (map (Morphism.thm phi)))) set_introssss, set_cases = map (Morphism.thm phi) set_cases}; fun morph_fp_co_induct_sugar phi ({co_rec, common_co_inducts, co_inducts, co_rec_def, co_rec_thms, co_rec_discs, co_rec_disc_iffs, co_rec_selss, co_rec_codes, co_rec_transfers, co_rec_o_maps, common_rel_co_inducts, rel_co_inducts, common_set_inducts, set_inducts} : fp_co_induct_sugar) = {co_rec = Morphism.term phi co_rec, common_co_inducts = map (Morphism.thm phi) common_co_inducts, co_inducts = map (Morphism.thm phi) co_inducts, co_rec_def = Morphism.thm phi co_rec_def, co_rec_thms = map (Morphism.thm phi) co_rec_thms, co_rec_discs = map (Morphism.thm phi) co_rec_discs, co_rec_disc_iffs = map (Morphism.thm phi) co_rec_disc_iffs, co_rec_selss = map (map (Morphism.thm phi)) co_rec_selss, co_rec_codes = map (Morphism.thm phi) co_rec_codes, co_rec_transfers = map (Morphism.thm phi) co_rec_transfers, co_rec_o_maps = map (Morphism.thm phi) co_rec_o_maps, common_rel_co_inducts = map (Morphism.thm phi) common_rel_co_inducts, rel_co_inducts = map (Morphism.thm phi) rel_co_inducts, common_set_inducts = map (Morphism.thm phi) common_set_inducts, set_inducts = map (Morphism.thm phi) set_inducts}; fun morph_fp_ctr_sugar phi ({ctrXs_Tss, ctor_iff_dtor, ctr_defs, ctr_sugar, ctr_transfers, case_transfers, disc_transfers, sel_transfers} : fp_ctr_sugar) = {ctrXs_Tss = map (map (Morphism.typ phi)) ctrXs_Tss, ctor_iff_dtor = Morphism.thm phi ctor_iff_dtor, ctr_defs = map (Morphism.thm phi) ctr_defs, ctr_sugar = morph_ctr_sugar phi ctr_sugar, ctr_transfers = map (Morphism.thm phi) ctr_transfers, case_transfers = map (Morphism.thm phi) case_transfers, disc_transfers = map (Morphism.thm phi) disc_transfers, sel_transfers = map (Morphism.thm phi) sel_transfers}; fun morph_fp_sugar phi ({T, BT, X, fp, fp_res, fp_res_index, pre_bnf, fp_bnf, absT_info, fp_nesting_bnfs, live_nesting_bnfs, fp_ctr_sugar, fp_bnf_sugar, fp_co_induct_sugar} : fp_sugar) = {T = Morphism.typ phi T, BT = Morphism.typ phi BT, X = Morphism.typ phi X, fp = fp, fp_res = morph_fp_result phi fp_res, fp_res_index = fp_res_index, pre_bnf = morph_bnf phi pre_bnf, fp_bnf = morph_bnf phi fp_bnf, absT_info = morph_absT_info phi absT_info, fp_nesting_bnfs = map (morph_bnf phi) fp_nesting_bnfs, live_nesting_bnfs = map (morph_bnf phi) live_nesting_bnfs, fp_ctr_sugar = morph_fp_ctr_sugar phi fp_ctr_sugar, fp_bnf_sugar = morph_fp_bnf_sugar phi fp_bnf_sugar, fp_co_induct_sugar = Option.map (morph_fp_co_induct_sugar phi) fp_co_induct_sugar}; val transfer_fp_sugar = morph_fp_sugar o Morphism.transfer_morphism; structure Data = Generic_Data ( type T = fp_sugar Symtab.table; val empty = Symtab.empty; val extend = I; fun merge data : T = Symtab.merge (K true) data; ); fun fp_sugar_of_generic context = Option.map (transfer_fp_sugar (Context.theory_of context)) o Symtab.lookup (Data.get context); fun fp_sugars_of_generic context = Symtab.fold (cons o transfer_fp_sugar (Context.theory_of context) o snd) (Data.get context) []; val fp_sugar_of = fp_sugar_of_generic o Context.Proof; val fp_sugar_of_global = fp_sugar_of_generic o Context.Theory; val fp_sugars_of = fp_sugars_of_generic o Context.Proof; val fp_sugars_of_global = fp_sugars_of_generic o Context.Theory; structure FP_Sugar_Plugin = Plugin(type T = fp_sugar list); fun fp_sugars_interpretation name f = FP_Sugar_Plugin.interpretation name (fn fp_sugars => fn lthy => f (map (transfer_fp_sugar (Proof_Context.theory_of lthy)) fp_sugars) lthy); val interpret_fp_sugars = FP_Sugar_Plugin.data; val register_fp_sugars_raw = fold (fn fp_sugar as {T = Type (s, _), ...} => Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => Data.map (Symtab.update (s, morph_fp_sugar phi fp_sugar)))); fun register_fp_sugars plugins fp_sugars = register_fp_sugars_raw fp_sugars #> interpret_fp_sugars plugins fp_sugars; fun interpret_bnfs_register_fp_sugars plugins Ts BTs Xs fp pre_bnfs absT_infos fp_nesting_bnfs live_nesting_bnfs fp_res ctrXs_Tsss ctor_iff_dtors ctr_defss ctr_sugars co_recs co_rec_defs map_thmss common_co_inducts co_inductss co_rec_thmss co_rec_discss co_rec_selsss rel_injectss rel_distinctss map_disc_iffss map_selsss rel_selss rel_intross rel_casess pred_injectss set_thmss set_selsssss set_introsssss set_casess ctr_transferss case_transferss disc_transferss sel_transferss co_rec_disc_iffss co_rec_codess co_rec_transferss common_rel_co_inducts rel_co_inductss common_set_inducts set_inductss co_rec_o_mapss noted = let val fp_sugars = map_index (fn (kk, T) => {T = T, BT = nth BTs kk, X = nth Xs kk, fp = fp, fp_res = fp_res, fp_res_index = kk, pre_bnf = nth pre_bnfs kk, absT_info = nth absT_infos kk, fp_bnf = nth (#bnfs fp_res) kk, fp_nesting_bnfs = fp_nesting_bnfs, live_nesting_bnfs = live_nesting_bnfs, fp_ctr_sugar = {ctrXs_Tss = nth ctrXs_Tsss kk, ctor_iff_dtor = nth ctor_iff_dtors kk, ctr_defs = nth ctr_defss kk, ctr_sugar = nth ctr_sugars kk, ctr_transfers = nth ctr_transferss kk, case_transfers = nth case_transferss kk, disc_transfers = nth disc_transferss kk, sel_transfers = nth sel_transferss kk}, fp_bnf_sugar = {map_thms = nth map_thmss kk, map_disc_iffs = nth map_disc_iffss kk, map_selss = nth map_selsss kk, rel_injects = nth rel_injectss kk, rel_distincts = nth rel_distinctss kk, rel_sels = nth rel_selss kk, rel_intros = nth rel_intross kk, rel_cases = nth rel_casess kk, pred_injects = nth pred_injectss kk, set_thms = nth set_thmss kk, set_selssss = nth set_selsssss kk, set_introssss = nth set_introsssss kk, set_cases = nth set_casess kk}, fp_co_induct_sugar = SOME {co_rec = nth co_recs kk, common_co_inducts = common_co_inducts, co_inducts = nth co_inductss kk, co_rec_def = nth co_rec_defs kk, co_rec_thms = nth co_rec_thmss kk, co_rec_discs = nth co_rec_discss kk, co_rec_disc_iffs = nth co_rec_disc_iffss kk, co_rec_selss = nth co_rec_selsss kk, co_rec_codes = nth co_rec_codess kk, co_rec_transfers = nth co_rec_transferss kk, co_rec_o_maps = nth co_rec_o_mapss kk, common_rel_co_inducts = common_rel_co_inducts, rel_co_inducts = nth rel_co_inductss kk, common_set_inducts = common_set_inducts, set_inducts = nth set_inductss kk}} |> morph_fp_sugar (substitute_noted_thm noted)) Ts; in register_fp_sugars_raw fp_sugars #> fold (interpret_bnf plugins) (#bnfs fp_res) #> interpret_fp_sugars plugins fp_sugars end; fun quasi_unambiguous_case_names names = let val ps = map (`Long_Name.base_name) names; val dups = Library.duplicates (op =) (map fst ps); fun underscore s = let val ss = Long_Name.explode s in space_implode "_" (drop (length ss - 2) ss) end; in map (fn (base, full) => if member (op =) dups base then underscore full else base) ps |> Name.variant_list [] end; fun zipper_map f = let fun zed _ [] = [] | zed xs (y :: ys) = f (xs, y, ys) :: zed (xs @ [y]) ys; in zed [] end; fun cannot_merge_types fp = error ("Mutually " ^ co_prefix fp ^ "recursive types must have the same type parameters"); fun merge_type_arg fp T T' = if T = T' then T else cannot_merge_types fp; fun merge_type_args fp (As, As') = if length As = length As' then map2 (merge_type_arg fp) As As' else cannot_merge_types fp; fun type_args_named_constrained_of_spec (((((ncAs, _), _), _), _), _) = ncAs; fun type_binding_of_spec (((((_, b), _), _), _), _) = b; fun mixfix_of_spec ((((_, mx), _), _), _) = mx; fun mixfixed_ctr_specs_of_spec (((_, mx_ctr_specs), _), _) = mx_ctr_specs; fun map_binding_of_spec ((_, (b, _, _)), _) = b; fun rel_binding_of_spec ((_, (_, b, _)), _) = b; fun pred_binding_of_spec ((_, (_, _, b)), _) = b; fun sel_default_eqs_of_spec (_, ts) = ts; fun ctr_sugar_kind_of_fp_kind Least_FP = Datatype | ctr_sugar_kind_of_fp_kind Greatest_FP = Codatatype; fun uncurry_thm 0 thm = thm | uncurry_thm 1 thm = thm | uncurry_thm n thm = rotate_prems ~1 (uncurry_thm (n - 1) (rotate_prems 1 (conjI RS thm))); fun choose_binary_fun fs AB = find_first (fastype_of #> binder_types #> (fn [A, B] => AB = (A, B))) fs; fun build_binary_fun_app fs t u = Option.map (rapp u o rapp t) (choose_binary_fun fs (fastype_of t, fastype_of u)); fun build_the_rel ctxt Rs Ts A B = build_rel [] ctxt Ts [] (the o choose_binary_fun Rs) (A, B); fun build_rel_app ctxt Rs Ts t u = build_the_rel ctxt Rs Ts (fastype_of t) (fastype_of u) $ t $ u; fun build_set_app ctxt A t = Term.betapply (build_set ctxt A (fastype_of t), t); fun mk_parametricity_goal ctxt Rs t u = let val prem = build_the_rel ctxt Rs [] (fastype_of t) (fastype_of u) in HOLogic.mk_Trueprop (prem $ t $ u) end; val name_of_set = name_of_const "set function" domain_type; val fundefcong_attrs = @{attributes [fundef_cong]}; val nitpicksimp_attrs = @{attributes [nitpick_simp]}; val simp_attrs = @{attributes [simp]}; val lists_bmoc = fold (fn xs => fn t => Term.list_comb (t, xs)); fun flat_corec_predss_getterss qss gss = maps (op @) (qss ~~ gss); fun flat_corec_preds_predsss_gettersss [] [qss] [gss] = flat_corec_predss_getterss qss gss | flat_corec_preds_predsss_gettersss (p :: ps) (qss :: qsss) (gss :: gsss) = p :: flat_corec_predss_getterss qss gss @ flat_corec_preds_predsss_gettersss ps qsss gsss; fun mk_flip (x, Type (_, [T1, Type (_, [T2, T3])])) = Abs ("x", T1, Abs ("y", T2, Var (x, T2 --> T1 --> T3) $ Bound 0 $ Bound 1)); fun flip_rels ctxt n thm = let val Rs = Term.add_vars (Thm.prop_of thm) []; val Rs' = rev (drop (length Rs - n) Rs); in infer_instantiate ctxt (map (fn f => (fst f, Thm.cterm_of ctxt (mk_flip f))) Rs') thm end; fun mk_ctor_or_dtor get_T Ts t = let val Type (_, Ts0) = get_T (fastype_of t) in Term.subst_atomic_types (Ts0 ~~ Ts) t end; val mk_ctor = mk_ctor_or_dtor range_type; val mk_dtor = mk_ctor_or_dtor domain_type; fun mk_bnf_sets bnf = let val Type (T_name, Us) = T_of_bnf bnf; val lives = lives_of_bnf bnf; val sets = sets_of_bnf bnf; fun mk_set U = (case find_index (curry (op =) U) lives of ~1 => Term.dummy | i => nth sets i); in (T_name, map mk_set Us) end; fun mk_xtor_co_recs thy fp fpTs Cs ts0 = let val nn = length fpTs; val (fpTs0, Cs0) = map ((fp = Greatest_FP ? swap) o dest_funT o snd o strip_typeN nn o fastype_of) ts0 |> split_list; val rho = tvar_subst thy (fpTs0 @ Cs0) (fpTs @ Cs); in map (Term.subst_TVars rho) ts0 end; fun liveness_of_fp_bnf n bnf = (case T_of_bnf bnf of Type (_, Ts) => map (not o member (op =) (deads_of_bnf bnf)) Ts | _ => replicate n false); fun add_nesting_bnf_names Us = let fun add (Type (s, Ts)) ss = let val (needs, ss') = fold_map add Ts ss in if exists I needs then (true, insert (op =) s ss') else (false, ss') end | add T ss = (member (op =) Us T, ss); in snd oo add end; fun nesting_bnfs ctxt ctr_Tsss Us = map_filter (bnf_of ctxt) (fold (fold (fold (add_nesting_bnf_names Us))) ctr_Tsss []); fun indexify proj xs f p = f (find_index (curry (op =) (proj p)) xs) p; fun massage_simple_notes base = filter_out (null o #2) #> map (fn (thmN, thms, f_attrs) => ((Binding.qualify true base (Binding.name thmN), []), map_index (fn (i, thm) => ([thm], f_attrs i)) thms)); fun massage_multi_notes b_names Ts = maps (fn (thmN, thmss, attrs) => @{map 3} (fn b_name => fn Type (T_name, _) => fn thms => ((Binding.qualify true b_name (Binding.name thmN), attrs T_name), [(thms, [])])) b_names Ts thmss) #> filter_out (null o fst o hd o snd); fun define_ctrs_dtrs_for_type fp_b_name fpT ctor dtor ctor_dtor dtor_ctor n ks abs ctr_bindings ctr_mixfixes ctr_Tss lthy = let val ctor_absT = domain_type (fastype_of ctor); val (((w, xss), u'), _) = lthy |> yield_singleton (mk_Frees "w") ctor_absT ||>> mk_Freess "x" ctr_Tss ||>> yield_singleton Variable.variant_fixes fp_b_name; val u = Free (u', fpT); val ctor_iff_dtor_thm = let val goal = fold_rev Logic.all [w, u] (mk_Trueprop_eq (HOLogic.mk_eq (u, ctor $ w), HOLogic.mk_eq (dtor $ u, w))); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_ctor_iff_dtor_tac ctxt (map (SOME o Thm.ctyp_of lthy) [ctor_absT, fpT]) (Thm.cterm_of lthy ctor) (Thm.cterm_of lthy dtor) ctor_dtor dtor_ctor) |> Thm.close_derivation \<^here> end; val ctr_rhss = map2 (fn k => fn xs => fold_rev Term.lambda xs (ctor $ mk_absumprod ctor_absT abs n k xs)) ks xss; val ((raw_ctrs, raw_ctr_defs), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> apfst split_list o @{fold_map 3} (fn b => fn mx => fn rhs => Local_Theory.define ((b, mx), ((Thm.make_def_binding (Config.get lthy bnf_internals) b, []), rhs)) #>> apsnd snd) ctr_bindings ctr_mixfixes ctr_rhss ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val ctr_defs = map (Morphism.thm phi) raw_ctr_defs; val ctrs0 = map (Morphism.term phi) raw_ctrs; in ((xss, ctrs0, ctor_iff_dtor_thm, ctr_defs), lthy) end; fun wrap_ctrs plugins fp discs_sels fp_b_name ctor_inject n ms abs_inject type_definition disc_bindings sel_bindingss sel_default_eqs ctrs0 ctor_iff_dtor_thm ctr_defs lthy = let val sumEN_thm' = unfold_thms lthy @{thms unit_all_eq1} (mk_absumprodE type_definition ms); fun exhaust_tac {context = ctxt, prems = _} = mk_exhaust_tac ctxt n ctr_defs ctor_iff_dtor_thm sumEN_thm'; val inject_tacss = map2 (fn ctr_def => fn 0 => [] | _ => [fn {context = ctxt, ...} => mk_inject_tac ctxt ctr_def ctor_inject abs_inject]) ctr_defs ms; val half_distinct_tacss = map (map (fn (def, def') => fn {context = ctxt, ...} => mk_half_distinct_tac ctxt ctor_inject abs_inject [def, def'])) (mk_half_pairss (`I ctr_defs)); val tacss = [exhaust_tac] :: inject_tacss @ half_distinct_tacss; fun ctr_spec_of disc_b ctr0 sel_bs = ((disc_b, ctr0), sel_bs); val ctr_specs = @{map 3} ctr_spec_of disc_bindings ctrs0 sel_bindingss; val (ctr_sugar as {case_cong, ...}, lthy) = free_constructors (ctr_sugar_kind_of_fp_kind fp) tacss ((((plugins, discs_sels), standard_binding), ctr_specs), sel_default_eqs) lthy; val anonymous_notes = [([case_cong], fundefcong_attrs)] |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])])); val notes = if Config.get lthy bnf_internals then [(ctor_iff_dtorN, [ctor_iff_dtor_thm], K [])] |> massage_simple_notes fp_b_name else []; in (ctr_sugar, lthy |> Local_Theory.notes (anonymous_notes @ notes) |> snd) end; fun derive_map_set_rel_pred_thms plugins fp live As Bs C E abs_inverses ctr_defs fp_nesting_set_maps fp_nesting_rel_eq_onps live_nesting_map_id0s live_nesting_set_maps live_nesting_rel_eqs live_nesting_rel_eq_onps fp_nested_rel_eq_onps fp_b_name fp_bnf fp_bnfs fpT ctor ctor_dtor dtor_ctor pre_map_def pre_set_defs pre_rel_def fp_map_thm fp_set_thms fp_rel_thm extra_unfolds_map extra_unfolds_set extra_unfolds_rel ctr_Tss ({casex, case_thms, discs, selss, sel_defs, ctrs, exhaust, exhaust_discs, disc_thmss, sel_thmss, injects, distincts, distinct_discsss, ...} : ctr_sugar) lthy = let val n = length ctr_Tss; val ms = map length ctr_Tss; val B_ify_T = Term.typ_subst_atomic (As ~~ Bs); val fpBT = B_ify_T fpT; val live_AsBs = filter (op <>) (As ~~ Bs); val live_As = map fst live_AsBs; val fTs = map (op -->) live_AsBs; val ((((((((xss, yss), fs), Ps), Rs), ta), tb), thesis), names_lthy) = lthy |> fold (fold Variable.declare_typ) [As, Bs] |> mk_Freess "x" ctr_Tss ||>> mk_Freess "y" (map (map B_ify_T) ctr_Tss) ||>> mk_Frees "f" fTs ||>> mk_Frees "P" (map mk_pred1T live_As) ||>> mk_Frees "R" (map (uncurry mk_pred2T) live_AsBs) ||>> yield_singleton (mk_Frees "a") fpT ||>> yield_singleton (mk_Frees "b") fpBT ||>> apfst HOLogic.mk_Trueprop o yield_singleton (mk_Frees "thesis") HOLogic.boolT; val ctrAs = map (mk_ctr As) ctrs; val ctrBs = map (mk_ctr Bs) ctrs; val ctr_defs' = map2 (fn m => fn def => mk_unabs_def m (HOLogic.mk_obj_eq def)) ms ctr_defs; val ABfs = live_AsBs ~~ fs; fun derive_rel_case relAsBs rel_inject_thms rel_distinct_thms = let val rel_Rs_a_b = list_comb (relAsBs, Rs) $ ta $ tb; fun mk_assms ctrA ctrB ctxt = let val argA_Ts = binder_types (fastype_of ctrA); val argB_Ts = binder_types (fastype_of ctrB); val ((argAs, argBs), names_ctxt) = ctxt |> mk_Frees "x" argA_Ts ||>> mk_Frees "y" argB_Ts; val ctrA_args = list_comb (ctrA, argAs); val ctrB_args = list_comb (ctrB, argBs); in (fold_rev Logic.all (argAs @ argBs) (Logic.list_implies (mk_Trueprop_eq (ta, ctrA_args) :: mk_Trueprop_eq (tb, ctrB_args) :: map2 (HOLogic.mk_Trueprop oo build_rel_app lthy Rs []) argAs argBs, thesis)), names_ctxt) end; val (assms, names_lthy) = @{fold_map 2} mk_assms ctrAs ctrBs names_lthy; val goal = Logic.list_implies (HOLogic.mk_Trueprop rel_Rs_a_b :: assms, thesis); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_rel_case_tac ctxt (Thm.cterm_of ctxt ta) (Thm.cterm_of ctxt tb) exhaust injects rel_inject_thms distincts rel_distinct_thms live_nesting_rel_eqs) |> singleton (Proof_Context.export names_lthy lthy) |> Thm.close_derivation \<^here> end; fun derive_case_transfer rel_case_thm = let val (S, names_lthy) = yield_singleton (mk_Frees "S") (mk_pred2T C E) names_lthy; val caseA = mk_case As C casex; val caseB = mk_case Bs E casex; val goal = mk_parametricity_goal names_lthy (S :: Rs) caseA caseB; in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_case_transfer_tac ctxt rel_case_thm case_thms) |> singleton (Proof_Context.export names_lthy lthy) |> Thm.close_derivation \<^here> end; in if live = 0 then if plugins transfer_plugin then let val relAsBs = HOLogic.eq_const fpT; val rel_case_thm = derive_rel_case relAsBs [] []; val case_transfer_thm = derive_case_transfer rel_case_thm; val notes = [(case_transferN, [case_transfer_thm], K [])] |> massage_simple_notes fp_b_name; val (noted, lthy') = lthy |> Local_Theory.notes notes; val subst = Morphism.thm (substitute_noted_thm noted); in (([], [], [], [], [], [], [], [], [], [], [], [], [], [], [subst case_transfer_thm], [], []), lthy') end else (([], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], []), lthy) else let val mapx = mk_map live As Bs (map_of_bnf fp_bnf); val relAsBs = mk_rel live As Bs (rel_of_bnf fp_bnf); val setAs = map (mk_set As) (sets_of_bnf fp_bnf); val discAs = map (mk_disc_or_sel As) discs; val discBs = map (mk_disc_or_sel Bs) discs; val selAss = map (map (mk_disc_or_sel As)) selss; val selBss = map (map (mk_disc_or_sel Bs)) selss; val map_ctor_thm = if fp = Least_FP then fp_map_thm else let val ctorA = mk_ctor As ctor; val ctorB = mk_ctor Bs ctor; val y_T = domain_type (fastype_of ctorA); val (y as Free (y_s, _), _) = lthy |> yield_singleton (mk_Frees "y") y_T; val ctor_cong = infer_instantiate' lthy [NONE, NONE, SOME (Thm.cterm_of lthy ctorB)] arg_cong; val fp_map_thm' = fp_map_thm |> infer_instantiate' lthy (replicate live NONE @ [SOME (Thm.cterm_of lthy (ctorA $ y))]) |> unfold_thms lthy [dtor_ctor]; in (fp_map_thm' RS ctor_cong RS (ctor_dtor RS sym RS trans)) |> Drule.generalize ([], [y_s]) end; val map_thms = let fun mk_goal ctrA ctrB xs ys = let val fmap = list_comb (mapx, fs); fun mk_arg (x as Free (_, T)) (Free (_, U)) = if T = U then x else build_map lthy [] [] (the o AList.lookup (op =) ABfs) (T, U) $ x; val xs' = map2 mk_arg xs ys; in mk_Trueprop_eq (fmap $ list_comb (ctrA, xs), list_comb (ctrB, xs')) end; val goals = @{map 4} mk_goal ctrAs ctrBs xss yss; val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_map_tac ctxt abs_inverses pre_map_def map_ctor_thm live_nesting_map_id0s ctr_defs' extra_unfolds_map) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end; val set0_thms = let fun mk_goal A setA ctrA xs = let val sets = map (build_set_app lthy A) (filter (exists_subtype_in [A] o fastype_of) xs); in mk_Trueprop_eq (setA $ list_comb (ctrA, xs), (if null sets then HOLogic.mk_set A [] else Library.foldl1 mk_union sets)) end; val goals = @{map 2} (fn live_A => fn setA => map2 (mk_goal live_A setA) ctrAs xss) live_As setAs |> flat; in if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_set0_tac ctxt abs_inverses pre_set_defs dtor_ctor fp_set_thms fp_nesting_set_maps live_nesting_set_maps ctr_defs' extra_unfolds_set) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end end; val set_thms = set0_thms |> map (unfold_thms lthy @{thms insert_is_Un[THEN sym] Un_empty_left Un_insert_left}); val rel_ctor_thm = if fp = Least_FP then fp_rel_thm else let val ctorA = mk_ctor As ctor; val ctorB = mk_ctor Bs ctor; val y_T = domain_type (fastype_of ctorA); val z_T = domain_type (fastype_of ctorB); val ((y as Free (y_s, _), z as Free (z_s, _)), _) = lthy |> yield_singleton (mk_Frees "y") y_T ||>> yield_singleton (mk_Frees "z") z_T; in fp_rel_thm |> infer_instantiate' lthy (replicate live NONE @ [SOME (Thm.cterm_of lthy (ctorA $ y)), SOME (Thm.cterm_of lthy (ctorB $ z))]) |> unfold_thms lthy [dtor_ctor] |> Drule.generalize ([], [y_s, z_s]) end; val rel_inject_thms = let fun mk_goal ctrA ctrB xs ys = let val lhs = list_comb (relAsBs, Rs) $ list_comb (ctrA, xs) $ list_comb (ctrB, ys); val conjuncts = map2 (build_rel_app lthy Rs []) xs ys; in HOLogic.mk_Trueprop (if null conjuncts then lhs else HOLogic.mk_eq (lhs, Library.foldr1 HOLogic.mk_conj conjuncts)) end; val goals = @{map 4} mk_goal ctrAs ctrBs xss yss; val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_rel_tac ctxt abs_inverses pre_rel_def rel_ctor_thm live_nesting_rel_eqs ctr_defs' extra_unfolds_rel) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end; val half_rel_distinct_thmss = let fun mk_goal ((ctrA, xs), (ctrB, ys)) = HOLogic.mk_Trueprop (HOLogic.mk_not (list_comb (relAsBs, Rs) $ list_comb (ctrA, xs) $ list_comb (ctrB, ys))); val rel_infos = (ctrAs ~~ xss, ctrBs ~~ yss); val goalss = map (map mk_goal) (mk_half_pairss rel_infos); val goals = flat goalss; in unflat goalss (if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_rel_tac ctxt abs_inverses pre_rel_def rel_ctor_thm live_nesting_rel_eqs ctr_defs' extra_unfolds_rel) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end) end; val rel_flip = rel_flip_of_bnf fp_bnf; fun mk_other_half_rel_distinct_thm thm = flip_rels lthy live thm RS (rel_flip RS sym RS @{thm arg_cong[of _ _ Not]} RS iffD2); val other_half_rel_distinct_thmss = map (map mk_other_half_rel_distinct_thm) half_rel_distinct_thmss; val (rel_distinct_thms, _) = join_halves n half_rel_distinct_thmss other_half_rel_distinct_thmss; fun mk_rel_intro_thm m thm = uncurry_thm m (thm RS iffD2) handle THM _ => thm; val rel_intro_thms = map2 mk_rel_intro_thm ms rel_inject_thms; val rel_code_thms = map (fn thm => thm RS @{thm eq_False[THEN iffD2]}) rel_distinct_thms @ map2 (fn thm => fn 0 => thm RS @{thm eq_True[THEN iffD2]} | _ => thm) rel_inject_thms ms; val ctr_transfer_thms = let val goals = map2 (mk_parametricity_goal names_lthy Rs) ctrAs ctrBs; val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_ctr_transfer_tac ctxt rel_intro_thms live_nesting_rel_eqs) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end; val (set_cases_thms, set_cases_attrss) = let fun mk_prems assms elem t ctxt = (case fastype_of t of Type (type_name, xs) => (case bnf_of ctxt type_name of NONE => ([], ctxt) | SOME bnf => apfst flat (fold_map (fn set => fn ctxt => let val T = HOLogic.dest_setT (range_type (fastype_of set)); val new_var = not (T = fastype_of elem); val (x, ctxt') = if new_var then yield_singleton (mk_Frees "x") T ctxt else (elem, ctxt); in mk_prems (mk_Trueprop_mem (x, set $ t) :: assms) elem x ctxt' |>> map (new_var ? Logic.all x) end) (map (mk_set xs) (sets_of_bnf bnf)) ctxt)) | T => rpair ctxt (if T = fastype_of elem then [fold (curry Logic.mk_implies) assms thesis] else [])); in split_list (map (fn set => let val A = HOLogic.dest_setT (range_type (fastype_of set)); val (elem, names_lthy) = yield_singleton (mk_Frees "e") A names_lthy; val premss = map (fn ctr => let val (args, names_lthy) = mk_Frees "z" (binder_types (fastype_of ctr)) names_lthy; in flat (zipper_map (fn (prev_args, arg, next_args) => let val (args_with_elem, args_without_elem) = if fastype_of arg = A then (prev_args @ [elem] @ next_args, prev_args @ next_args) else `I (prev_args @ [arg] @ next_args); in mk_prems [mk_Trueprop_eq (ta, Term.list_comb (ctr, args_with_elem))] elem arg names_lthy |> fst |> map (fold_rev Logic.all args_without_elem) end) args) end) ctrAs; val goal = Logic.mk_implies (mk_Trueprop_mem (elem, set $ ta), thesis); val vars = Variable.add_free_names lthy goal []; val thm = Goal.prove_sorry lthy vars (flat premss) goal (fn {context = ctxt, prems} => mk_set_cases_tac ctxt (Thm.cterm_of ctxt ta) prems exhaust set_thms) |> Thm.close_derivation \<^here> |> rotate_prems ~1; val cases_set_attr = Attrib.internal (K (Induct.cases_pred (name_of_set set))); val ctr_names = quasi_unambiguous_case_names (flat (map (uncurry mk_names o map_prod length name_of_ctr) (premss ~~ ctrAs))); in (* TODO: @{attributes [elim?]} *) (thm, [Attrib.consumes 1, cases_set_attr, Attrib.case_names ctr_names]) end) setAs) end; val (set_intros_thmssss, set_intros_thms) = let fun mk_goals A setA ctr_args t ctxt = (case fastype_of t of Type (type_name, innerTs) => (case bnf_of ctxt type_name of NONE => ([], ctxt) | SOME bnf => apfst flat (fold_map (fn set => fn ctxt => let val T = HOLogic.dest_setT (range_type (fastype_of set)); val (y, ctxt') = yield_singleton (mk_Frees "y") T ctxt; val assm = mk_Trueprop_mem (y, set $ t); in apfst (map (Logic.mk_implies o pair assm)) (mk_goals A setA ctr_args y ctxt') end) (map (mk_set innerTs) (sets_of_bnf bnf)) ctxt)) | T => (if T = A then [mk_Trueprop_mem (t, setA $ ctr_args)] else [], ctxt)); val (goalssss, _) = fold_map (fn set => let val A = HOLogic.dest_setT (range_type (fastype_of set)) in @{fold_map 2} (fn ctr => fn xs => fold_map (mk_goals A set (Term.list_comb (ctr, xs))) xs) ctrAs xss end) setAs lthy; val goals = flat (flat (flat goalssss)); in `(unflattt goalssss) (if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_set_intros_tac ctxt set0_thms) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end) end; val rel_sel_thms = let val n = length discAs; fun mk_conjunct n k discA selAs discB selBs = (if k = n then [] else [HOLogic.mk_eq (discA $ ta, discB $ tb)]) @ (if null selAs then [] else [Library.foldr HOLogic.mk_imp (if n = 1 then [] else [discA $ ta, discB $ tb], Library.foldr1 HOLogic.mk_conj (map2 (build_rel_app names_lthy Rs []) (map (rapp ta) selAs) (map (rapp tb) selBs)))]); val goals = if n = 0 then [] else [mk_Trueprop_eq (build_rel_app names_lthy Rs [] ta tb, (case flat (@{map 5} (mk_conjunct n) (1 upto n) discAs selAss discBs selBss) of [] => \<^term>\True\ | conjuncts => Library.foldr1 HOLogic.mk_conj conjuncts))]; fun prove goal = Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_rel_sel_tac ctxt (Thm.cterm_of ctxt ta) (Thm.cterm_of ctxt tb) exhaust (flat disc_thmss) (flat sel_thmss) rel_inject_thms distincts rel_distinct_thms live_nesting_rel_eqs)) |> Thm.close_derivation \<^here>; in map prove goals end; val (rel_case_thm, rel_case_attrs) = let val thm = derive_rel_case relAsBs rel_inject_thms rel_distinct_thms; val ctr_names = quasi_unambiguous_case_names (map name_of_ctr ctrAs); in (thm, [Attrib.case_names ctr_names, Attrib.consumes 1] @ @{attributes [cases pred]}) end; val case_transfer_thm = derive_case_transfer rel_case_thm; val sel_transfer_thms = if null selAss then [] else let val shared_sels = foldl1 (uncurry (inter (op =))) (map (op ~~) (selAss ~~ selBss)); val goals = map (uncurry (mk_parametricity_goal names_lthy Rs)) shared_sels; in if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_sel_transfer_tac ctxt n sel_defs case_transfer_thm) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end end; val disc_transfer_thms = let val goals = map2 (mk_parametricity_goal names_lthy Rs) discAs discBs in if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_disc_transfer_tac ctxt (the_single rel_sel_thms) (the_single exhaust_discs) (flat (flat distinct_discsss))) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end end; val map_disc_iff_thms = let val discsB = map (mk_disc_or_sel Bs) discs; val discsA_t = map (fn disc1 => Term.betapply (disc1, ta)) discAs; fun mk_goal (discA_t, discB) = if head_of discA_t aconv HOLogic.Not orelse is_refl_bool discA_t then NONE else SOME (mk_Trueprop_eq (betapply (discB, (Term.list_comb (mapx, fs) $ ta)), discA_t)); val goals = map_filter mk_goal (discsA_t ~~ discsB); in if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_map_disc_iff_tac ctxt (Thm.cterm_of ctxt ta) exhaust (flat disc_thmss) map_thms) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end end; val (map_sel_thmss, map_sel_thms) = let fun mk_goal discA selA selB = let val prem = Term.betapply (discA, ta); val lhs = selB $ (Term.list_comb (mapx, fs) $ ta); val lhsT = fastype_of lhs; val map_rhsT = map_atyps (perhaps (AList.lookup (op =) (map swap live_AsBs))) lhsT; val map_rhs = build_map lthy [] [] (the o (AList.lookup (op =) (live_AsBs ~~ fs))) (map_rhsT, lhsT); val rhs = (case map_rhs of Const (\<^const_name>\id\, _) => selA $ ta | _ => map_rhs $ (selA $ ta)); val concl = mk_Trueprop_eq (lhs, rhs); in if is_refl_bool prem then concl else Logic.mk_implies (HOLogic.mk_Trueprop prem, concl) end; val goalss = @{map 3} (map2 o mk_goal) discAs selAss selBss; val goals = flat goalss; in `(unflat goalss) (if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_map_sel_tac ctxt (Thm.cterm_of ctxt ta) exhaust (flat disc_thmss) map_thms (flat sel_thmss) live_nesting_map_id0s) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end) end; val (set_sel_thmssss, set_sel_thms) = let fun mk_goal setA discA selA ctxt = let val prem = Term.betapply (discA, ta); val sel_rangeT = range_type (fastype_of selA); val A = HOLogic.dest_setT (range_type (fastype_of setA)); fun travese_nested_types t ctxt = (case fastype_of t of Type (type_name, innerTs) => (case bnf_of ctxt type_name of NONE => ([], ctxt) | SOME bnf => let fun seq_assm a set ctxt = let val T = HOLogic.dest_setT (range_type (fastype_of set)); val (x, ctxt') = yield_singleton (mk_Frees "x") T ctxt; val assm = mk_Trueprop_mem (x, set $ a); in travese_nested_types x ctxt' |>> map (Logic.mk_implies o pair assm) end; in fold_map (seq_assm t o mk_set innerTs) (sets_of_bnf bnf) ctxt |>> flat end) | T => if T = A then ([mk_Trueprop_mem (t, setA $ ta)], ctxt) else ([], ctxt)); val (concls, ctxt') = if sel_rangeT = A then ([mk_Trueprop_mem (selA $ ta, setA $ ta)], ctxt) else travese_nested_types (selA $ ta) ctxt; in if exists_subtype_in [A] sel_rangeT then if is_refl_bool prem then (concls, ctxt') else (map (Logic.mk_implies o pair (HOLogic.mk_Trueprop prem)) concls, ctxt') else ([], ctxt) end; val (goalssss, _) = fold_map (fn set => @{fold_map 2} (fold_map o mk_goal set) discAs selAss) setAs names_lthy; val goals = flat (flat (flat goalssss)); in `(unflattt goalssss) (if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_set_sel_tac ctxt (Thm.cterm_of ctxt ta) exhaust (flat disc_thmss) (flat sel_thmss) set0_thms) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end) end; val pred_injects = let fun top_sweep_rewr_conv rewrs = Conv.top_sweep_conv (K (Conv.rewrs_conv rewrs)) \<^context>; val rel_eq_onp_with_tops_of = Conv.fconv_rule (HOLogic.Trueprop_conv (Conv.arg1_conv (top_sweep_rewr_conv @{thms eq_onp_top_eq_eq[symmetric, THEN eq_reflection]}))); val eq_onps = map rel_eq_onp_with_tops_of (map rel_eq_onp_of_bnf fp_bnfs @ fp_nesting_rel_eq_onps @ live_nesting_rel_eq_onps @ fp_nested_rel_eq_onps); val cTs = map (SOME o Thm.ctyp_of lthy) (maps (replicate 2) live_As); val cts = map (SOME o Thm.cterm_of lthy) (map mk_eq_onp Ps); val get_rhs = Thm.concl_of #> HOLogic.dest_Trueprop #> HOLogic.dest_eq #> snd; val pred_eq_onp_conj = List.foldr (fn (_, thm) => thm RS @{thm eq_onp_live_step}) @{thm refl[of True]}; fun predify_rel_inject rel_inject = let val conjuncts = try (get_rhs #> HOLogic.dest_conj) rel_inject |> the_default []; fun postproc thm = if null conjuncts then thm RS (@{thm eq_onp_same_args} RS iffD1) else @{thm box_equals} OF [thm, @{thm eq_onp_same_args}, pred_eq_onp_conj conjuncts |> unfold_thms lthy @{thms simp_thms(21)}]; in rel_inject |> Thm.instantiate' cTs cts |> Conv.fconv_rule (HOLogic.Trueprop_conv (Conv.arg_conv (Raw_Simplifier.rewrite lthy false @{thms eq_onp_top_eq_eq[symmetric, THEN eq_reflection]}))) |> unfold_thms lthy eq_onps |> postproc |> unfold_thms lthy @{thms top_conj} end; in rel_inject_thms |> map (unfold_thms lthy [@{thm conj_assoc}]) |> map predify_rel_inject |> Proof_Context.export names_lthy lthy end; val anonymous_notes = [(rel_code_thms, nitpicksimp_attrs)] |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])])); val notes = (if Config.get lthy bnf_internals then [(set0N, set0_thms, K [])] else []) @ [(case_transferN, [case_transfer_thm], K []), (ctr_transferN, ctr_transfer_thms, K []), (disc_transferN, disc_transfer_thms, K []), (sel_transferN, sel_transfer_thms, K []), (mapN, map_thms, K (nitpicksimp_attrs @ simp_attrs)), (map_disc_iffN, map_disc_iff_thms, K simp_attrs), (map_selN, map_sel_thms, K []), (pred_injectN, pred_injects, K simp_attrs), (rel_casesN, [rel_case_thm], K rel_case_attrs), (rel_distinctN, rel_distinct_thms, K simp_attrs), (rel_injectN, rel_inject_thms, K simp_attrs), (rel_introsN, rel_intro_thms, K []), (rel_selN, rel_sel_thms, K []), (setN, set_thms, K (case_fp fp nitpicksimp_attrs [] @ simp_attrs)), (set_casesN, set_cases_thms, nth set_cases_attrss), (set_introsN, set_intros_thms, K []), (set_selN, set_sel_thms, K [])] |> massage_simple_notes fp_b_name; val (noted, lthy') = lthy |> uncurry (Spec_Rules.add Binding.empty Spec_Rules.equational) (`(single o lhs_head_of o hd) map_thms) |> fp = Least_FP ? uncurry (Spec_Rules.add Binding.empty Spec_Rules.equational) (`(single o lhs_head_of o hd) rel_code_thms) |> uncurry (Spec_Rules.add Binding.empty Spec_Rules.equational) (`(single o lhs_head_of o hd) set0_thms) |> plugins code_plugin ? Code.declare_default_eqns (map (rpair true) (rel_code_thms @ map_thms @ set_thms)) |> Local_Theory.notes (anonymous_notes @ notes); val subst = Morphism.thm (substitute_noted_thm noted); in ((map subst map_thms, map subst map_disc_iff_thms, map (map subst) map_sel_thmss, map subst rel_inject_thms, map subst rel_distinct_thms, map subst rel_sel_thms, map subst rel_intro_thms, [subst rel_case_thm], map subst pred_injects, map subst set_thms, map (map (map (map subst))) set_sel_thmssss, map (map (map (map subst))) set_intros_thmssss, map subst set_cases_thms, map subst ctr_transfer_thms, [subst case_transfer_thm], map subst disc_transfer_thms, map subst sel_transfer_thms), lthy') end end; type lfp_sugar_thms = (thm list * thm * Token.src list) * (thm list list * Token.src list); fun morph_lfp_sugar_thms phi ((inducts, induct, induct_attrs), (recss, rec_attrs)) = ((map (Morphism.thm phi) inducts, Morphism.thm phi induct, induct_attrs), (map (map (Morphism.thm phi)) recss, rec_attrs)) : lfp_sugar_thms; val transfer_lfp_sugar_thms = morph_lfp_sugar_thms o Morphism.transfer_morphism; type gfp_sugar_thms = ((thm list * thm) list * (Token.src list * Token.src list)) * thm list list * thm list list * (thm list list * Token.src list) * (thm list list list * Token.src list); fun morph_gfp_sugar_thms phi ((coinducts_pairs, coinduct_attrs_pair), corecss, corec_discss, (corec_disc_iffss, corec_disc_iff_attrs), (corec_selsss, corec_sel_attrs)) = ((map (apfst (map (Morphism.thm phi)) o apsnd (Morphism.thm phi)) coinducts_pairs, coinduct_attrs_pair), map (map (Morphism.thm phi)) corecss, map (map (Morphism.thm phi)) corec_discss, (map (map (Morphism.thm phi)) corec_disc_iffss, corec_disc_iff_attrs), (map (map (map (Morphism.thm phi))) corec_selsss, corec_sel_attrs)) : gfp_sugar_thms; val transfer_gfp_sugar_thms = morph_gfp_sugar_thms o Morphism.transfer_morphism; fun unzip_recT (Type (\<^type_name>\prod\, [_, TFree x])) (T as Type (\<^type_name>\prod\, Ts as [_, TFree y])) = if x = y then [T] else Ts | unzip_recT _ (Type (\<^type_name>\prod\, Ts as [_, TFree _])) = Ts | unzip_recT _ T = [T]; fun mk_recs_args_types ctxt ctr_Tsss Cs absTs repTs ns mss ctor_rec_fun_Ts = let val Css = map2 replicate ns Cs; val x_Tssss = @{map 6} (fn absT => fn repT => fn n => fn ms => fn ctr_Tss => fn ctor_rec_fun_T => map2 (map2 unzip_recT) ctr_Tss (dest_absumprodT absT repT n ms (domain_type ctor_rec_fun_T))) absTs repTs ns mss ctr_Tsss ctor_rec_fun_Ts; val x_Tsss' = map (map flat_rec_arg_args) x_Tssss; val f_Tss = map2 (map2 (curry (op --->))) x_Tsss' Css; val ((fss, xssss), _) = ctxt |> mk_Freess "f" f_Tss ||>> mk_Freessss "x" x_Tssss; in (f_Tss, x_Tssss, fss, xssss) end; fun unzip_corecT (Type (\<^type_name>\sum\, _)) T = [T] | unzip_corecT _ (Type (\<^type_name>\sum\, Ts)) = Ts | unzip_corecT _ T = [T]; (*avoid "'a itself" arguments in corecursors*) fun repair_nullary_single_ctr [[]] = [[HOLogic.unitT]] | repair_nullary_single_ctr Tss = Tss; fun mk_corec_fun_arg_types0 ctr_Tsss Cs absTs repTs ns mss fun_Ts = let val ctr_Tsss' = map repair_nullary_single_ctr ctr_Tsss; val g_absTs = map range_type fun_Ts; val g_Tsss = map repair_nullary_single_ctr (@{map 5} dest_absumprodT absTs repTs ns mss g_absTs); val g_Tssss = @{map 3} (fn C => map2 (map2 (map (curry (op -->) C) oo unzip_corecT))) Cs ctr_Tsss' g_Tsss; val q_Tssss = map (map (map (fn [_] => [] | [_, T] => [mk_pred1T (domain_type T)]))) g_Tssss; in (q_Tssss, g_Tsss, g_Tssss, g_absTs) end; fun mk_corec_p_pred_types Cs ns = map2 (fn n => replicate (Int.max (0, n - 1)) o mk_pred1T) ns Cs; fun mk_corec_fun_arg_types ctr_Tsss Cs absTs repTs ns mss dtor_corec = (mk_corec_p_pred_types Cs ns, mk_corec_fun_arg_types0 ctr_Tsss Cs absTs repTs ns mss (binder_fun_types (fastype_of dtor_corec))); fun mk_corecs_args_types ctxt ctr_Tsss Cs absTs repTs ns mss dtor_corec_fun_Ts = let val p_Tss = mk_corec_p_pred_types Cs ns; val (q_Tssss, g_Tsss, g_Tssss, corec_types) = mk_corec_fun_arg_types0 ctr_Tsss Cs absTs repTs ns mss dtor_corec_fun_Ts; val (((((Free (x, _), cs), pss), qssss), gssss), _) = ctxt |> yield_singleton (mk_Frees "x") dummyT ||>> mk_Frees "a" Cs ||>> mk_Freess "p" p_Tss ||>> mk_Freessss "q" q_Tssss ||>> mk_Freessss "g" g_Tssss; val cpss = map2 (map o rapp) cs pss; fun build_sum_inj mk_inj = build_map ctxt [] [] (uncurry mk_inj o dest_sumT o snd); fun build_dtor_corec_arg _ [] [cg] = cg | build_dtor_corec_arg T [cq] [cg, cg'] = mk_If cq (build_sum_inj Inl_const (fastype_of cg, T) $ cg) (build_sum_inj Inr_const (fastype_of cg', T) $ cg'); val pgss = @{map 3} flat_corec_preds_predsss_gettersss pss qssss gssss; val cqssss = map2 (map o map o map o rapp) cs qssss; val cgssss = map2 (map o map o map o rapp) cs gssss; val cqgsss = @{map 3} (@{map 3} (@{map 3} build_dtor_corec_arg)) g_Tsss cqssss cgssss; in (x, cs, cpss, (((pgss, pss, qssss, gssss), cqgsss), corec_types)) end; fun mk_co_recs_prelims ctxt fp ctr_Tsss fpTs Cs absTs repTs ns mss xtor_co_recs0 = let val thy = Proof_Context.theory_of ctxt; val (xtor_co_rec_fun_Ts, xtor_co_recs) = mk_xtor_co_recs thy fp fpTs Cs xtor_co_recs0 |> `(binder_fun_types o fastype_of o hd); val (recs_args_types, corecs_args_types) = if fp = Least_FP then mk_recs_args_types ctxt ctr_Tsss Cs absTs repTs ns mss xtor_co_rec_fun_Ts |> (rpair NONE o SOME) else mk_corecs_args_types ctxt ctr_Tsss Cs absTs repTs ns mss xtor_co_rec_fun_Ts |> (pair NONE o SOME); in (xtor_co_recs, recs_args_types, corecs_args_types) end; fun mk_preds_getterss_join c cps absT abs cqgss = let val n = length cqgss; val ts = map2 (mk_absumprod absT abs n) (1 upto n) cqgss; in Term.lambda c (mk_IfN absT cps ts) end; fun define_co_rec_as fp Cs fpT b rhs lthy0 = let val thy = Proof_Context.theory_of lthy0; val ((cst, (_, def)), (lthy', lthy)) = lthy0 - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> Local_Theory.define ((b, NoSyn), ((Thm.make_def_binding (Config.get lthy0 bnf_internals) b, []), rhs)) ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy lthy'; val cst' = mk_co_rec thy fp Cs fpT (Morphism.term phi cst); val def' = Morphism.thm phi def; in ((cst', def'), lthy') end; fun define_rec (_, _, fss, xssss) mk_binding fpTs Cs reps ctor_rec = let val nn = length fpTs; val (ctor_rec_absTs, fpT) = strip_typeN nn (fastype_of ctor_rec) |>> map domain_type ||> domain_type; in define_co_rec_as Least_FP Cs fpT (mk_binding recN) (fold_rev (fold_rev Term.lambda) fss (Term.list_comb (ctor_rec, @{map 4} (fn ctor_rec_absT => fn rep => fn fs => fn xsss => mk_case_absumprod ctor_rec_absT rep fs (map (map HOLogic.mk_tuple) xsss) (map flat_rec_arg_args xsss)) ctor_rec_absTs reps fss xssss))) end; fun define_corec (_, cs, cpss, (((pgss, _, _, _), cqgsss), f_absTs)) mk_binding fpTs Cs abss dtor_corec = let val nn = length fpTs; val fpT = range_type (snd (strip_typeN nn (fastype_of dtor_corec))); in define_co_rec_as Greatest_FP Cs fpT (mk_binding corecN) (fold_rev (fold_rev Term.lambda) pgss (Term.list_comb (dtor_corec, @{map 5} mk_preds_getterss_join cs cpss f_absTs abss cqgsss))) end; fun mk_induct_raw_prem_prems names_ctxt Xss setss_fp_nesting (x as Free (s, Type (T_name, Ts0))) (Type (_, Xs_Ts0)) = (case AList.lookup (op =) setss_fp_nesting T_name of NONE => [] | SOME raw_sets0 => let val (Xs_Ts, (Ts, raw_sets)) = filter (exists_subtype_in (flat Xss) o fst) (Xs_Ts0 ~~ (Ts0 ~~ raw_sets0)) |> split_list ||> split_list; val sets = map (mk_set Ts0) raw_sets; val (ys, names_ctxt') = names_ctxt |> mk_Frees s Ts; val xysets = map (pair x) (ys ~~ sets); val ppremss = map2 (mk_induct_raw_prem_prems names_ctxt' Xss setss_fp_nesting) ys Xs_Ts; in flat (map2 (map o apfst o cons) xysets ppremss) end) | mk_induct_raw_prem_prems _ Xss _ (x as Free (_, Type _)) X = [([], (find_index (fn Xs => member (op =) Xs X) Xss + 1, x))] | mk_induct_raw_prem_prems _ _ _ _ _ = []; fun mk_induct_raw_prem alter_x names_ctxt Xss setss_fp_nesting p ctr ctr_Ts ctrXs_Ts = let val (xs, names_ctxt') = names_ctxt |> mk_Frees "x" ctr_Ts; val pprems = flat (map2 (mk_induct_raw_prem_prems names_ctxt' Xss setss_fp_nesting) xs ctrXs_Ts); val y = Term.list_comb (ctr, map alter_x xs); val p' = enforce_type names_ctxt domain_type (fastype_of y) p; in (xs, pprems, HOLogic.mk_Trueprop (p' $ y)) end; fun close_induct_prem_prem nn ps xs t = fold_rev Logic.all (map Free (drop (nn + length xs) (rev (Term.add_frees t (map dest_Free xs @ map_filter (try dest_Free) ps))))) t; fun finish_induct_prem_prem ctxt nn ps xs (xysets, (j, x)) = let val p' = enforce_type ctxt domain_type (fastype_of x) (nth ps (j - 1)) in close_induct_prem_prem nn ps xs (Logic.list_implies (map (fn (x', (y, set)) => mk_Trueprop_mem (y, set $ x')) xysets, HOLogic.mk_Trueprop (p' $ x))) end; fun finish_induct_prem ctxt nn ps (xs, raw_pprems, concl) = fold_rev Logic.all xs (Logic.list_implies (map (finish_induct_prem_prem ctxt nn ps xs) raw_pprems, concl)); fun mk_coinduct_prem_ctr_concls ctxt Xss fpTss rs' n k udisc usels vdisc vsels ctrXs_Ts = let fun build_the_rel T Xs_T = build_rel [] ctxt [] [] (fn (T, X) => nth rs' (find_index (fn Xs => member (op =) Xs X) Xss) |> enforce_type ctxt domain_type T) (T, Xs_T) |> Term.subst_atomic_types (flat Xss ~~ flat fpTss); fun build_rel_app usel vsel Xs_T = fold rapp [usel, vsel] (build_the_rel (fastype_of usel) Xs_T); in (if k = n then [] else [HOLogic.mk_eq (udisc, vdisc)]) @ (if null usels then [] else [Library.foldr HOLogic.mk_imp (if n = 1 then [] else [udisc, vdisc], Library.foldr1 HOLogic.mk_conj (@{map 3} build_rel_app usels vsels ctrXs_Ts))]) end; fun mk_coinduct_prem_concl ctxt Xss fpTss rs' n udiscs uselss vdiscs vselss ctrXs_Tss = @{map 6} (mk_coinduct_prem_ctr_concls ctxt Xss fpTss rs' n) (1 upto n) udiscs uselss vdiscs vselss ctrXs_Tss |> flat |> Library.foldr1 HOLogic.mk_conj handle List.Empty => \<^term>\True\; fun mk_coinduct_prem ctxt Xss fpTss rs' uvr u v n udiscs uselss vdiscs vselss ctrXs_Tss = fold_rev Logic.all [u, v] (Logic.mk_implies (HOLogic.mk_Trueprop uvr, HOLogic.mk_Trueprop (mk_coinduct_prem_concl ctxt Xss fpTss rs' n udiscs uselss vdiscs vselss ctrXs_Tss))); fun postproc_co_induct ctxt nn prop prop_conj = Drule.zero_var_indexes #> `(conj_dests nn) #>> map (fn thm => Thm.permute_prems 0 ~1 (thm RS prop)) ##> (fn thm => Thm.permute_prems 0 (~ nn) (if nn = 1 then thm RS prop else funpow nn (fn thm => unfold_thms ctxt @{thms conj_assoc} (thm RS prop_conj)) thm)); fun mk_induct_attrs ctrss = let val induct_cases = quasi_unambiguous_case_names (maps (map name_of_ctr) ctrss); in [Attrib.case_names induct_cases] end; fun derive_rel_induct_thms_for_types ctxt nn fpA_Ts As Bs ctrAss ctrAs_Tsss exhausts ctor_rel_induct ctor_defss ctor_injects pre_rel_defs abs_inverses live_nesting_rel_eqs = let val B_ify_T = Term.typ_subst_atomic (As ~~ Bs); val B_ify = Term.map_types B_ify_T; val fpB_Ts = map B_ify_T fpA_Ts; val ctrBs_Tsss = map (map (map B_ify_T)) ctrAs_Tsss; val ctrBss = map (map B_ify) ctrAss; val ((((Rs, IRs), ctrAsss), ctrBsss), names_ctxt) = ctxt |> mk_Frees "R" (map2 mk_pred2T As Bs) ||>> mk_Frees "IR" (map2 mk_pred2T fpA_Ts fpB_Ts) ||>> mk_Freesss "a" ctrAs_Tsss ||>> mk_Freesss "b" ctrBs_Tsss; val prems = let fun mk_prem ctrA ctrB argAs argBs = fold_rev Logic.all (argAs @ argBs) (fold_rev (curry Logic.mk_implies) (map2 (HOLogic.mk_Trueprop oo build_rel_app names_ctxt (Rs @ IRs) fpA_Ts) argAs argBs) (HOLogic.mk_Trueprop (build_rel_app names_ctxt (Rs @ IRs) fpA_Ts (Term.list_comb (ctrA, argAs)) (Term.list_comb (ctrB, argBs))))); in flat (@{map 4} (@{map 4} mk_prem) ctrAss ctrBss ctrAsss ctrBsss) end; val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_leq (map2 (build_the_rel ctxt (Rs @ IRs) []) fpA_Ts fpB_Ts) IRs)); val vars = Variable.add_free_names ctxt goal []; val rel_induct0_thm = Goal.prove_sorry ctxt vars prems goal (fn {context = ctxt, prems} => mk_rel_induct0_tac ctxt ctor_rel_induct prems (map (Thm.cterm_of ctxt) IRs) exhausts ctor_defss ctor_injects pre_rel_defs abs_inverses live_nesting_rel_eqs) |> Thm.close_derivation \<^here>; in (postproc_co_induct ctxt nn @{thm predicate2D} @{thm predicate2D_conj} rel_induct0_thm, mk_induct_attrs ctrAss) end; fun derive_induct_recs_thms_for_types plugins pre_bnfs rec_args_typess ctor_induct ctor_rec_thms live_nesting_bnfs fp_nesting_bnfs fpTs Cs Xs ctrXs_Tsss pre_abs_inverses pre_type_definitions abs_inverses ctrss ctr_defss recs rec_defs ctxt = let val ctr_Tsss = map (map (binder_types o fastype_of)) ctrss; val nn = length pre_bnfs; val ns = map length ctr_Tsss; val mss = map (map length) ctr_Tsss; val pre_map_defs = map map_def_of_bnf pre_bnfs; val pre_set_defss = map set_defs_of_bnf pre_bnfs; val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs; val fp_nesting_map_ident0s = map map_ident0_of_bnf fp_nesting_bnfs; val fp_nesting_set_maps = maps set_map_of_bnf fp_nesting_bnfs; val fp_b_names = map base_name_of_typ fpTs; val (((ps, xsss), us'), names_ctxt) = ctxt |> mk_Frees "P" (map mk_pred1T fpTs) ||>> mk_Freesss "x" ctr_Tsss ||>> Variable.variant_fixes fp_b_names; val us = map2 (curry Free) us' fpTs; val setss_fp_nesting = map mk_bnf_sets fp_nesting_bnfs; val (induct_thms, induct_thm) = let val raw_premss = @{map 4} (@{map 3} o mk_induct_raw_prem I names_ctxt (map single Xs) setss_fp_nesting) ps ctrss ctr_Tsss ctrXs_Tsss; val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (curry (op $)) ps us)); val goal = Library.foldr (Logic.list_implies o apfst (map (finish_induct_prem ctxt nn ps))) (raw_premss, concl); val vars = Variable.add_free_names ctxt goal []; val kksss = map (map (map (fst o snd) o #2)) raw_premss; val ctor_induct' = ctor_induct OF (map2 mk_absumprodE pre_type_definitions mss); val thm = Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, ...} => mk_induct_tac ctxt nn ns mss kksss (flat ctr_defss) ctor_induct' pre_abs_inverses abs_inverses fp_nesting_set_maps pre_set_defss) |> Thm.close_derivation \<^here>; in `(conj_dests nn) thm end; val xctrss = map2 (map2 (curry Term.list_comb)) ctrss xsss; fun mk_rec_thmss (_, x_Tssss, fss, _) recs rec_defs ctor_rec_thms = let val frecs = map (lists_bmoc fss) recs; fun mk_goal frec xctr f xs fxs = fold_rev (fold_rev Logic.all) (xs :: fss) (mk_Trueprop_eq (frec $ xctr, Term.list_comb (f, fxs))); fun maybe_tick (T, U) u f = if try (fst o HOLogic.dest_prodT) U = SOME T then Term.lambda u (HOLogic.mk_prod (u, f $ u)) else f; fun build_rec (x as Free (_, T)) U = if T = U then x else let val build_simple = indexify (perhaps (try (snd o HOLogic.dest_prodT)) o snd) Cs (fn kk => fn TU => maybe_tick TU (nth us kk) (nth frecs kk)); in build_map ctxt [] [] build_simple (T, U) $ x end; val fxsss = map2 (map2 (flat_rec_arg_args oo map2 (map o build_rec))) xsss x_Tssss; val goalss = @{map 5} (@{map 4} o mk_goal) frecs xctrss fss xsss fxsss; val tacss = @{map 4} (map ooo mk_rec_tac pre_map_defs (fp_nesting_map_ident0s @ live_nesting_map_ident0s) rec_defs) ctor_rec_thms pre_abs_inverses abs_inverses ctr_defss; fun prove goal tac = Goal.prove_sorry ctxt [] [] goal (tac o #context) |> Thm.close_derivation \<^here>; in map2 (map2 prove) goalss tacss end; val rec_thmss = mk_rec_thmss (the rec_args_typess) recs rec_defs ctor_rec_thms; in ((induct_thms, induct_thm, mk_induct_attrs ctrss), (rec_thmss, nitpicksimp_attrs @ simp_attrs)) end; fun mk_coinduct_attrs fpTs ctrss discss mss = let val fp_b_names = map base_name_of_typ fpTs; fun mk_coinduct_concls ms discs ctrs = let fun mk_disc_concl disc = [name_of_disc disc]; fun mk_ctr_concl 0 _ = [] | mk_ctr_concl _ ctr = [name_of_ctr ctr]; val disc_concls = map mk_disc_concl (fst (split_last discs)) @ [[]]; val ctr_concls = map2 mk_ctr_concl ms ctrs; in flat (map2 append disc_concls ctr_concls) end; val coinduct_cases = quasi_unambiguous_case_names (map (prefix Eq_prefix) fp_b_names); val coinduct_conclss = @{map 3} (quasi_unambiguous_case_names ooo mk_coinduct_concls) mss discss ctrss; val coinduct_case_names_attr = Attrib.case_names coinduct_cases; val coinduct_case_concl_attrs = map2 (fn casex => fn concls => Attrib.case_conclusion (casex, concls)) coinduct_cases coinduct_conclss; val common_coinduct_attrs = coinduct_case_names_attr :: coinduct_case_concl_attrs; val coinduct_attrs = Attrib.consumes 1 :: coinduct_case_names_attr :: coinduct_case_concl_attrs; in (coinduct_attrs, common_coinduct_attrs) end; fun derive_rel_coinduct_thms_for_types ctxt nn fpA_Ts ns As Bs mss (ctr_sugars : ctr_sugar list) abs_inverses abs_injects ctor_injects dtor_ctors rel_pre_defs ctor_defss dtor_rel_coinduct live_nesting_rel_eqs = let val B_ify_T = Term.typ_subst_atomic (As ~~ Bs); val fpB_Ts = map B_ify_T fpA_Ts; val (Rs, IRs, fpAs, fpBs, _) = let val fp_names = map base_name_of_typ fpA_Ts; val ((((Rs, IRs), fpAs_names), fpBs_names), names_ctxt) = ctxt |> mk_Frees "R" (map2 mk_pred2T As Bs) ||>> mk_Frees "IR" (map2 mk_pred2T fpA_Ts fpB_Ts) ||>> Variable.variant_fixes fp_names ||>> Variable.variant_fixes (map (suffix "'") fp_names); in (Rs, IRs, map2 (curry Free) fpAs_names fpA_Ts, map2 (curry Free) fpBs_names fpB_Ts, names_ctxt) end; val ((discA_tss, selA_tsss), (discB_tss, selB_tsss)) = let val discss = map #discs ctr_sugars; val selsss = map #selss ctr_sugars; fun mk_discss ts Ts = map2 (map o rapp) ts (map (map (mk_disc_or_sel Ts)) discss); fun mk_selsss ts Ts = map2 (map o map o rapp) ts (map (map (map (mk_disc_or_sel Ts))) selsss); in ((mk_discss fpAs As, mk_selsss fpAs As), (mk_discss fpBs Bs, mk_selsss fpBs Bs)) end; val prems = let fun mk_prem_ctr_concls n k discA_t selA_ts discB_t selB_ts = (if k = n then [] else [HOLogic.mk_eq (discA_t, discB_t)]) @ (case (selA_ts, selB_ts) of ([], []) => [] | (_ :: _, _ :: _) => [Library.foldr HOLogic.mk_imp (if n = 1 then [] else [discA_t, discB_t], Library.foldr1 HOLogic.mk_conj (map2 (build_rel_app ctxt (Rs @ IRs) fpA_Ts) selA_ts selB_ts))]); fun mk_prem_concl n discA_ts selA_tss discB_ts selB_tss = Library.foldr1 HOLogic.mk_conj (flat (@{map 5} (mk_prem_ctr_concls n) (1 upto n) discA_ts selA_tss discB_ts selB_tss)) handle List.Empty => \<^term>\True\; fun mk_prem IR tA tB n discA_ts selA_tss discB_ts selB_tss = fold_rev Logic.all [tA, tB] (Logic.mk_implies (HOLogic.mk_Trueprop (IR $ tA $ tB), HOLogic.mk_Trueprop (mk_prem_concl n discA_ts selA_tss discB_ts selB_tss))); in @{map 8} mk_prem IRs fpAs fpBs ns discA_tss selA_tsss discB_tss selB_tsss end; val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_leq IRs (map2 (build_the_rel ctxt (Rs @ IRs) []) fpA_Ts fpB_Ts))); val vars = Variable.add_free_names ctxt goal []; val rel_coinduct0_thm = Goal.prove_sorry ctxt vars prems goal (fn {context = ctxt, prems} => mk_rel_coinduct0_tac ctxt dtor_rel_coinduct (map (Thm.cterm_of ctxt) IRs) prems (map #exhaust ctr_sugars) (map (flat o #disc_thmss) ctr_sugars) (map (flat o #sel_thmss) ctr_sugars) ctor_defss dtor_ctors ctor_injects abs_injects rel_pre_defs abs_inverses live_nesting_rel_eqs) |> Thm.close_derivation \<^here>; in (postproc_co_induct ctxt nn @{thm predicate2D} @{thm predicate2D_conj} rel_coinduct0_thm, mk_coinduct_attrs fpA_Ts (map #ctrs ctr_sugars) (map #discs ctr_sugars) mss) end; fun derive_set_induct_thms_for_types ctxt nn fpTs ctrss setss dtor_set_inducts exhausts set_pre_defs ctor_defs dtor_ctors Abs_pre_inverses = let fun mk_prems A Ps ctr_args t ctxt = (case fastype_of t of Type (type_name, innerTs) => (case bnf_of ctxt type_name of NONE => ([], ctxt) | SOME bnf => let fun seq_assm a set ctxt = let val X = HOLogic.dest_setT (range_type (fastype_of set)); val (x, ctxt') = yield_singleton (mk_Frees "x") X ctxt; val assm = mk_Trueprop_mem (x, set $ a); in (case build_binary_fun_app Ps x a of NONE => mk_prems A Ps ctr_args x ctxt' |>> map (Logic.all x o Logic.mk_implies o pair assm) | SOME f => ([Logic.all x (Logic.mk_implies (assm, Logic.mk_implies (HOLogic.mk_Trueprop f, HOLogic.mk_Trueprop (the (build_binary_fun_app Ps x ctr_args)))))], ctxt')) end; in fold_map (seq_assm t o mk_set innerTs) (sets_of_bnf bnf) ctxt |>> flat end) | T => if T = A then ([HOLogic.mk_Trueprop (the (build_binary_fun_app Ps t ctr_args))], ctxt) else ([], ctxt)); fun mk_prems_for_ctr A Ps ctr ctxt = let val (args, ctxt') = mk_Frees "z" (binder_types (fastype_of ctr)) ctxt; in fold_map (mk_prems A Ps (list_comb (ctr, args))) args ctxt' |>> map (fold_rev Logic.all args) o flat |>> (fn prems => (prems, mk_names (length prems) (name_of_ctr ctr))) end; fun mk_prems_and_concl_for_type A Ps ((fpT, ctrs), set) ctxt = let val ((x, fp), ctxt') = ctxt |> yield_singleton (mk_Frees "x") A ||>> yield_singleton (mk_Frees "a") fpT; val concl = mk_Ball (set $ fp) (Term.absfree (dest_Free x) (the (build_binary_fun_app Ps x fp))); in fold_map (mk_prems_for_ctr A Ps) ctrs ctxt' |>> split_list |>> map_prod flat flat |>> apfst (rpair concl) end; fun mk_thm ctxt fpTs ctrss sets = let val A = HOLogic.dest_setT (range_type (fastype_of (hd sets))); val (Ps, ctxt') = mk_Frees "P" (map (fn fpT => A --> fpT --> HOLogic.boolT) fpTs) ctxt; val (((prems, concl), case_names), ctxt'') = fold_map (mk_prems_and_concl_for_type A Ps) (fpTs ~~ ctrss ~~ sets) ctxt' |>> apfst split_list o split_list |>> apfst (apfst flat) |>> apfst (apsnd (Library.foldr1 HOLogic.mk_conj)) |>> apsnd flat; val vars = fold (Variable.add_free_names ctxt) (concl :: prems) []; val thm = Goal.prove_sorry ctxt vars prems (HOLogic.mk_Trueprop concl) (fn {context = ctxt, prems} => mk_set_induct0_tac ctxt (map (Thm.cterm_of ctxt'') Ps) prems dtor_set_inducts exhausts set_pre_defs ctor_defs dtor_ctors Abs_pre_inverses) |> Thm.close_derivation \<^here>; val case_names_attr = Attrib.case_names (quasi_unambiguous_case_names case_names); val induct_set_attrs = map (Attrib.internal o K o Induct.induct_pred o name_of_set) sets; in (thm, case_names_attr :: induct_set_attrs) end val consumes_attr = Attrib.consumes 1; in map (mk_thm ctxt fpTs ctrss #> nn = 1 ? map_prod (fn thm => rotate_prems ~1 (thm RS bspec)) (cons consumes_attr)) (transpose setss) end; fun mk_coinduct_strong_thm coind rel_eqs rel_monos mk_vimage2p ctxt = let val n = Thm.nprems_of coind; val m = Thm.nprems_of (hd rel_monos) - n; fun mk_inst phi = (phi, Thm.cterm_of ctxt (mk_union (Var phi, HOLogic.eq_const (fst (dest_pred2T (#2 phi)))))); val insts = Term.add_vars (Thm.prop_of coind) [] |> rev |> take n |> map mk_inst; fun mk_unfold rel_eq rel_mono = let val eq = iffD2 OF [rel_eq RS @{thm predicate2_eqD}, refl]; val mono = rel_mono OF (replicate m @{thm order_refl} @ replicate n @{thm eq_subset}); in mk_vimage2p (eq RS (mono RS @{thm predicate2D})) RS eqTrueI end; val unfolds = map2 mk_unfold rel_eqs rel_monos @ @{thms sup_fun_def sup_bool_def imp_disjL all_conj_distrib subst_eq_imp simp_thms(18,21,35)}; in Thm.instantiate ([], insts) coind |> unfold_thms ctxt unfolds end; fun derive_coinduct_thms_for_types ctxt strong alter_r pre_bnfs dtor_coinduct dtor_ctors live_nesting_bnfs fpTs Xs ctrXs_Tsss ns pre_abs_inverses abs_inverses mk_vimage2p ctr_defss (ctr_sugars : ctr_sugar list) = let val nn = length pre_bnfs; val pre_rel_defs = map rel_def_of_bnf pre_bnfs; val live_nesting_rel_eqs = map rel_eq_of_bnf live_nesting_bnfs; val fp_b_names = map base_name_of_typ fpTs; val discss = map #discs ctr_sugars; val selsss = map #selss ctr_sugars; val exhausts = map #exhaust ctr_sugars; val disc_thmsss = map #disc_thmss ctr_sugars; val sel_thmsss = map #sel_thmss ctr_sugars; val (((rs, us'), vs'), _) = ctxt |> mk_Frees "R" (map (fn T => mk_pred2T T T) fpTs) ||>> Variable.variant_fixes fp_b_names ||>> Variable.variant_fixes (map (suffix "'") fp_b_names); val us = map2 (curry Free) us' fpTs; val udiscss = map2 (map o rapp) us discss; val uselsss = map2 (map o map o rapp) us selsss; val vs = map2 (curry Free) vs' fpTs; val vdiscss = map2 (map o rapp) vs discss; val vselsss = map2 (map o map o rapp) vs selsss; val uvrs = @{map 3} (fn r => fn u => fn v => r $ u $ v) rs us vs; val uv_eqs = map2 (curry HOLogic.mk_eq) us vs; val strong_rs = @{map 4} (fn u => fn v => fn uvr => fn uv_eq => fold_rev Term.lambda [u, v] (HOLogic.mk_disj (uvr, uv_eq))) us vs uvrs uv_eqs; val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (@{map 3} (fn uvr => fn u => fn v => HOLogic.mk_imp (uvr, HOLogic.mk_eq (u, v))) uvrs us vs)) fun mk_goal rs0' = Logic.list_implies (@{map 9} (mk_coinduct_prem ctxt (map single Xs) (map single fpTs) (map alter_r rs0')) uvrs us vs ns udiscss uselsss vdiscss vselsss ctrXs_Tsss, concl); val goals = map mk_goal ([rs] @ (if strong then [strong_rs] else [])); fun prove dtor_coinduct' goal = Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, ...} => mk_coinduct_tac ctxt live_nesting_rel_eqs nn ns dtor_coinduct' pre_rel_defs pre_abs_inverses abs_inverses dtor_ctors exhausts ctr_defss disc_thmsss sel_thmsss)) |> Thm.close_derivation \<^here>; val rel_eqs = map rel_eq_of_bnf pre_bnfs; val rel_monos = map rel_mono_of_bnf pre_bnfs; val dtor_coinducts = [dtor_coinduct] @ (if strong then [mk_coinduct_strong_thm dtor_coinduct rel_eqs rel_monos mk_vimage2p ctxt] else []); in map2 (postproc_co_induct ctxt nn mp @{thm conj_commute[THEN iffD1]} oo prove) dtor_coinducts goals end; fun derive_coinduct_corecs_thms_for_types ctxt pre_bnfs (x, cs, cpss, (((pgss, _, _, _), cqgsss), _)) dtor_coinduct dtor_injects dtor_ctors dtor_corec_thms live_nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss mss ns pre_abs_inverses abs_inverses mk_vimage2p ctr_defss (ctr_sugars : ctr_sugar list) corecs corec_defs = let fun mk_ctor_dtor_corec_thm dtor_inject dtor_ctor corec = iffD1 OF [dtor_inject, trans OF [corec, dtor_ctor RS sym]]; val ctor_dtor_corec_thms = @{map 3} mk_ctor_dtor_corec_thm dtor_injects dtor_ctors dtor_corec_thms; val pre_map_defs = map map_def_of_bnf pre_bnfs; val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs; val fp_b_names = map base_name_of_typ fpTs; val ctrss = map #ctrs ctr_sugars; val discss = map #discs ctr_sugars; val selsss = map #selss ctr_sugars; val disc_thmsss = map #disc_thmss ctr_sugars; val discIss = map #discIs ctr_sugars; val sel_thmsss = map #sel_thmss ctr_sugars; val coinduct_thms_pairs = derive_coinduct_thms_for_types ctxt true I pre_bnfs dtor_coinduct dtor_ctors live_nesting_bnfs fpTs Xs ctrXs_Tsss ns pre_abs_inverses abs_inverses mk_vimage2p ctr_defss ctr_sugars; fun mk_maybe_not pos = not pos ? HOLogic.mk_not; val gcorecs = map (lists_bmoc pgss) corecs; val corec_thmss = let val (us', _) = ctxt |> Variable.variant_fixes fp_b_names; val us = map2 (curry Free) us' fpTs; fun mk_goal c cps gcorec n k ctr m cfs' = fold_rev (fold_rev Logic.all) ([c] :: pgss) (Logic.list_implies (seq_conds (HOLogic.mk_Trueprop oo mk_maybe_not) n k cps, mk_Trueprop_eq (gcorec $ c, Term.list_comb (ctr, take m cfs')))); val mk_U = typ_subst_nonatomic (map2 (fn C => fn fpT => (mk_sumT (fpT, C), fpT)) Cs fpTs); fun tack (c, u) f = let val x' = Free (x, mk_sumT (fastype_of u, fastype_of c)) in Term.lambda x' (mk_case_sum (Term.lambda u u, Term.lambda c (f $ c)) $ x') end; fun build_corec cqg = let val T = fastype_of cqg in if exists_subtype_in Cs T then let val U = mk_U T; val build_simple = indexify fst (map2 (curry mk_sumT) fpTs Cs) (fn kk => fn _ => tack (nth cs kk, nth us kk) (nth gcorecs kk)); in build_map ctxt [] [] build_simple (T, U) $ cqg end else cqg end; val cqgsss' = map (map (map build_corec)) cqgsss; val goalss = @{map 8} (@{map 4} oooo mk_goal) cs cpss gcorecs ns kss ctrss mss cqgsss'; val tacss = @{map 4} (map ooo mk_corec_tac corec_defs live_nesting_map_ident0s) ctor_dtor_corec_thms pre_map_defs abs_inverses ctr_defss; fun prove goal tac = Goal.prove_sorry ctxt [] [] goal (tac o #context) |> Thm.close_derivation \<^here>; in map2 (map2 prove) goalss tacss |> map (map (unfold_thms ctxt @{thms case_sum_if})) end; val corec_disc_iff_thmss = let fun mk_goal c cps gcorec n k disc = mk_Trueprop_eq (disc $ (gcorec $ c), if n = 1 then \<^const>\True\ else Library.foldr1 HOLogic.mk_conj (seq_conds mk_maybe_not n k cps)); val goalss = @{map 6} (map2 oooo mk_goal) cs cpss gcorecs ns kss discss; fun mk_case_split' cp = Thm.instantiate' [] [SOME (Thm.cterm_of ctxt cp)] @{thm case_split}; val case_splitss' = map (map mk_case_split') cpss; val tacss = @{map 3} (map oo mk_corec_disc_iff_tac) case_splitss' corec_thmss disc_thmsss; fun prove goal tac = Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (tac o #context)) |> Thm.close_derivation \<^here>; fun proves [_] [_] = [] | proves goals tacs = map2 prove goals tacs; in map2 proves goalss tacss end; fun mk_corec_disc_thms corecs discIs = map (op RS) (corecs ~~ discIs); val corec_disc_thmss = map2 mk_corec_disc_thms corec_thmss discIss; fun mk_corec_sel_thm corec_thm sel sel_thm = let val (domT, ranT) = dest_funT (fastype_of sel); val arg_cong' = Thm.instantiate' (map (SOME o Thm.ctyp_of ctxt) [domT, ranT]) [NONE, NONE, SOME (Thm.cterm_of ctxt sel)] arg_cong |> Thm.varifyT_global; val sel_thm' = sel_thm RSN (2, trans); in corec_thm RS arg_cong' RS sel_thm' end; fun mk_corec_sel_thms corec_thmss = @{map 3} (@{map 3} (map2 o mk_corec_sel_thm)) corec_thmss selsss sel_thmsss; val corec_sel_thmsss = mk_corec_sel_thms corec_thmss; in ((coinduct_thms_pairs, mk_coinduct_attrs fpTs (map #ctrs ctr_sugars) (map #discs ctr_sugars) mss), corec_thmss, corec_disc_thmss, (corec_disc_iff_thmss, simp_attrs), (corec_sel_thmsss, simp_attrs)) end; fun define_co_datatypes prepare_plugins prepare_constraint prepare_typ prepare_term fp construct_fp ((raw_plugins, discs_sels0), specs) lthy = let val plugins = prepare_plugins lthy raw_plugins; val discs_sels = discs_sels0 orelse fp = Greatest_FP; val nn = length specs; val fp_bs = map type_binding_of_spec specs; val fp_b_names = map Binding.name_of fp_bs; val fp_common_name = mk_common_name fp_b_names; val map_bs = map map_binding_of_spec specs; val rel_bs = map rel_binding_of_spec specs; val pred_bs = map pred_binding_of_spec specs; fun prepare_type_arg (_, (ty, c)) = let val TFree (s, _) = prepare_typ lthy ty in TFree (s, prepare_constraint lthy c) end; val Ass0 = map (map prepare_type_arg o type_args_named_constrained_of_spec) specs; val unsorted_Ass0 = map (map (resort_tfree_or_tvar \<^sort>\type\)) Ass0; val unsorted_As = Library.foldr1 (merge_type_args fp) unsorted_Ass0; val num_As = length unsorted_As; val set_boss = map (map fst o type_args_named_constrained_of_spec) specs; val set_bss = map (map (the_default Binding.empty)) set_boss; fun add_fake_type spec = Typedecl.basic_typedecl {final = true} (type_binding_of_spec spec, num_As, Mixfix.reset_pos (mixfix_of_spec spec)); val (fake_T_names, fake_lthy) = fold_map add_fake_type specs lthy; val qsoty = quote o Syntax.string_of_typ fake_lthy; val _ = (case Library.duplicates (op =) unsorted_As of [] => () | A :: _ => error ("Duplicate type parameter " ^ qsoty A ^ " in " ^ co_prefix fp ^ "datatype specification")); val bad_args = map (Logic.type_map (singleton (Variable.polymorphic lthy))) unsorted_As |> filter_out Term.is_TVar; val _ = null bad_args orelse error ("Locally fixed type argument " ^ qsoty (hd bad_args) ^ " in " ^ co_prefix fp ^ "datatype specification"); val mixfixes = map mixfix_of_spec specs; val _ = (case Library.duplicates Binding.eq_name fp_bs of [] => () | b :: _ => error ("Duplicate type name declaration " ^ quote (Binding.name_of b))); val mx_ctr_specss = map mixfixed_ctr_specs_of_spec specs; val ctr_specss = map (map fst) mx_ctr_specss; val ctr_mixfixess = map (map snd) mx_ctr_specss; val disc_bindingss = map (map disc_of_ctr_spec) ctr_specss; val ctr_bindingss = map2 (fn fp_b_name => map (Binding.qualify false fp_b_name o ctr_of_ctr_spec)) fp_b_names ctr_specss; val ctr_argsss = map (map args_of_ctr_spec) ctr_specss; val sel_bindingsss = map (map (map fst)) ctr_argsss; val fake_ctr_Tsss0 = map (map (map (prepare_typ fake_lthy o snd))) ctr_argsss; val raw_sel_default_eqss = map sel_default_eqs_of_spec specs; val (As :: _) :: fake_ctr_Tsss = burrow (burrow (Syntax.check_typs fake_lthy)) (Ass0 :: fake_ctr_Tsss0); val As' = map dest_TFree As; val rhs_As' = fold (fold (fold Term.add_tfreesT)) fake_ctr_Tsss []; val _ = (case subtract (op =) As' rhs_As' of [] => () | extras => error ("Extra type variables on right-hand side: " ^ commas (map (qsoty o TFree) extras))); val fake_Ts = map (fn s => Type (s, As)) fake_T_names; val ((((Bs0, Cs), Es), Xs), _) = lthy |> fold (Variable.declare_typ o resort_tfree_or_tvar dummyS) unsorted_As |> mk_TFrees num_As ||>> mk_TFrees nn ||>> mk_TFrees nn ||>> variant_tfrees fp_b_names; fun eq_fpT_check (T as Type (s, Ts)) (T' as Type (s', Ts')) = s = s' andalso (Ts = Ts' orelse error ("Wrong type arguments in " ^ co_prefix fp ^ "recursive type " ^ qsoty T ^ " (expected " ^ qsoty T' ^ ")")) | eq_fpT_check _ _ = false; fun freeze_fp (T as Type (s, Ts)) = (case find_index (eq_fpT_check T) fake_Ts of ~1 => Type (s, map freeze_fp Ts) | kk => nth Xs kk) | freeze_fp T = T; val unfreeze_fp = Term.typ_subst_atomic (Xs ~~ fake_Ts); val ctrXs_Tsss = map (map (map freeze_fp)) fake_ctr_Tsss; val ctrXs_repTs = map mk_sumprodT_balanced ctrXs_Tsss; val _ = let fun add_deps i = fold (fn T => fold_index (fn (j, X) => (i <> j andalso exists_subtype_in [X] T) ? insert (op =) (i, j)) Xs); val add_missing_nodes = fold (AList.default (op =) o rpair []) (0 upto nn - 1); val deps = fold_index (uncurry (fold o add_deps)) ctrXs_Tsss [] |> AList.group (op =) |> add_missing_nodes; val G = Int_Graph.make (map (apfst (rpair ())) deps); val sccs = map (sort int_ord) (Int_Graph.strong_conn G); val str_of_scc = prefix (co_prefix fp ^ "datatype ") o space_implode " and " o map (suffix " = \" o Long_Name.base_name); fun warn [_] = () | warn sccs = warning ("Defined types not fully mutually " ^ co_prefix fp ^ "recursive\n\ \Alternative specification:\n" ^ cat_lines (map (prefix " " o str_of_scc o map (nth fp_b_names)) sccs)); in warn (order_strong_conn (op =) Int_Graph.make Int_Graph.topological_order deps sccs) end; val killed_As = map_filter (fn (A, set_bos) => if exists is_none set_bos then SOME A else NONE) (As ~~ transpose set_boss); val (((pre_bnfs, absT_infos), _), (fp_res as {bnfs = fp_bnfs as any_fp_bnf :: _, ctors = ctors0, dtors = dtors0, xtor_co_recs = xtor_co_recs0, xtor_co_induct, dtor_ctors, ctor_dtors, ctor_injects, dtor_injects, xtor_maps, xtor_setss, xtor_rels, xtor_co_rec_thms, xtor_rel_co_induct, dtor_set_inducts, xtor_co_rec_transfers, xtor_co_rec_o_maps, ...}, lthy)) = fixpoint_bnf false I (construct_fp mixfixes map_bs rel_bs pred_bs set_bss) fp_bs (map dest_TFree As) (map dest_TFree killed_As) (map dest_TFree Xs) ctrXs_repTs empty_comp_cache lthy handle BAD_DEAD (X, X_backdrop) => (case X_backdrop of Type (bad_tc, _) => let val fake_T = qsoty (unfreeze_fp X); val fake_T_backdrop = qsoty (unfreeze_fp X_backdrop); fun register_hint () = "\nUse the " ^ quote (#1 \<^command_keyword>\bnf\) ^ " command to register " ^ quote bad_tc ^ " as a bounded natural functor to allow nested (co)recursion through \ \it"; in if is_some (bnf_of lthy bad_tc) orelse is_some (fp_sugar_of lthy bad_tc) then error ("Inadmissible " ^ co_prefix fp ^ "recursive occurrence of type " ^ fake_T ^ " in type expression " ^ fake_T_backdrop) else if is_some (Old_Datatype_Data.get_info (Proof_Context.theory_of lthy) bad_tc) then error ("Unsupported " ^ co_prefix fp ^ "recursive occurrence of type " ^ fake_T ^ " via the old-style datatype " ^ quote bad_tc ^ " in type expression " ^ fake_T_backdrop ^ register_hint ()) else error ("Unsupported " ^ co_prefix fp ^ "recursive occurrence of type " ^ fake_T ^ " via type constructor " ^ quote bad_tc ^ " in type expression " ^ fake_T_backdrop ^ register_hint ()) end); val time = time lthy; val timer = time (Timer.startRealTimer ()); val fp_nesting_bnfs = nesting_bnfs lthy ctrXs_Tsss Xs; val live_nesting_bnfs = nesting_bnfs lthy ctrXs_Tsss As; val pre_map_defs = map map_def_of_bnf pre_bnfs; val pre_set_defss = map set_defs_of_bnf pre_bnfs; val pre_rel_defs = map rel_def_of_bnf pre_bnfs; val fp_nesting_set_maps = maps set_map_of_bnf fp_nesting_bnfs; val fp_nesting_rel_eq_onps = map rel_eq_onp_of_bnf fp_nesting_bnfs; val live_nesting_map_id0s = map map_id0_of_bnf live_nesting_bnfs; val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs; val live_nesting_set_maps = maps set_map_of_bnf live_nesting_bnfs; val live_nesting_rel_eqs = map rel_eq_of_bnf live_nesting_bnfs; val live_nesting_rel_eq_onps = map rel_eq_onp_of_bnf live_nesting_bnfs; val liveness = liveness_of_fp_bnf num_As any_fp_bnf; val live = live_of_bnf any_fp_bnf; val _ = if live = 0 andalso exists (not o Binding.is_empty) (map_bs @ rel_bs @ pred_bs) then warning "Map function, relator, and predicator names ignored" else (); val Bs = @{map 3} (fn alive => fn A as TFree (_, S) => fn B => if alive then resort_tfree_or_tvar S B else A) liveness As Bs0; val B_ify_T = Term.typ_subst_atomic (As ~~ Bs); val B_ify = Term.map_types B_ify_T; val live_AsBs = filter (op <>) (As ~~ Bs); val abss = map #abs absT_infos; val reps = map #rep absT_infos; val absTs = map #absT absT_infos; val repTs = map #repT absT_infos; val abs_injects = map #abs_inject absT_infos; val abs_inverses = map #abs_inverse absT_infos; val type_definitions = map #type_definition absT_infos; val ctors = map (mk_ctor As) ctors0; val dtors = map (mk_dtor As) dtors0; val fpTs = map (domain_type o fastype_of) dtors; val fpBTs = map B_ify_T fpTs; val real_unfreeze_fp = Term.typ_subst_atomic (Xs ~~ fpTs); val ctr_Tsss = map (map (map real_unfreeze_fp)) ctrXs_Tsss; val ns = map length ctr_Tsss; val kss = map (fn n => 1 upto n) ns; val mss = map (map length) ctr_Tsss; val (xtor_co_recs, recs_args_types, corecs_args_types) = mk_co_recs_prelims lthy fp ctr_Tsss fpTs Cs absTs repTs ns mss xtor_co_recs0; fun define_ctrs_dtrs_for_type_etc fp_bnf fp_b fpT C E ctor dtor xtor_co_rec ctor_dtor dtor_ctor ctor_inject pre_map_def pre_set_defs pre_rel_def fp_map_thm fp_set_thms fp_rel_thm n ks ms abs abs_inject type_definition ctr_bindings ctr_mixfixes ctr_Tss disc_bindings sel_bindingss raw_sel_default_eqs lthy = let val fp_b_name = Binding.name_of fp_b; val ((xss, ctrs0, ctor_iff_dtor_thm, ctr_defs), lthy) = define_ctrs_dtrs_for_type fp_b_name fpT ctor dtor ctor_dtor dtor_ctor n ks abs ctr_bindings ctr_mixfixes ctr_Tss lthy; val ctrs = map (mk_ctr As) ctrs0; val sel_default_eqs = let val sel_Tss = map (map (curry (op -->) fpT)) ctr_Tss; val sel_bTs = flat sel_bindingss ~~ flat sel_Tss |> filter_out (Binding.is_empty o fst) |> distinct (Binding.eq_name o apply2 fst); val sel_default_lthy = fake_local_theory_for_sel_defaults sel_bTs lthy in map (prepare_term sel_default_lthy) raw_sel_default_eqs end; fun mk_binding pre = Binding.qualify false fp_b_name (Binding.prefix_name (pre ^ "_") fp_b); fun massage_res (ctr_sugar, maps_sets_rels) = (maps_sets_rels, (ctrs, xss, ctor_iff_dtor_thm, ctr_defs, ctr_sugar)); in (wrap_ctrs plugins fp discs_sels fp_b_name ctor_inject n ms abs_inject type_definition disc_bindings sel_bindingss sel_default_eqs ctrs0 ctor_iff_dtor_thm ctr_defs #> (fn (ctr_sugar, lthy) => derive_map_set_rel_pred_thms plugins fp live As Bs C E abs_inverses ctr_defs fp_nesting_set_maps fp_nesting_rel_eq_onps live_nesting_map_id0s live_nesting_set_maps live_nesting_rel_eqs live_nesting_rel_eq_onps [] fp_b_name fp_bnf fp_bnfs fpT ctor ctor_dtor dtor_ctor pre_map_def pre_set_defs pre_rel_def fp_map_thm fp_set_thms fp_rel_thm [] [] [] ctr_Tss ctr_sugar lthy |>> pair ctr_sugar) ##>> (if fp = Least_FP then define_rec (the recs_args_types) mk_binding fpTs Cs reps else define_corec (the corecs_args_types) mk_binding fpTs Cs abss) xtor_co_rec #>> apfst massage_res, lthy) end; fun wrap_ctrs_derive_map_set_rel_pred_thms_define_co_rec_for_types (wrap_one_etcs, lthy) = fold_map I wrap_one_etcs lthy |>> apsnd split_list o apfst (apsnd @{split_list 5} o apfst @{split_list 17} o split_list) o split_list; fun mk_simp_thms ({injects, distincts, case_thms, ...} : ctr_sugar) co_recs map_thms rel_injects rel_distincts set_thmss = injects @ distincts @ case_thms @ co_recs @ map_thms @ rel_injects @ rel_distincts @ set_thmss; fun mk_co_rec_transfer_goals lthy co_recs = let val BE_ify = Term.subst_atomic_types (live_AsBs @ (Cs ~~ Es)); val ((Rs, Ss), names_lthy) = lthy |> mk_Frees "R" (map (uncurry mk_pred2T) live_AsBs) ||>> mk_Frees "S" (map2 mk_pred2T Cs Es); val co_recBs = map BE_ify co_recs; in (Rs, Ss, map2 (mk_parametricity_goal lthy (Rs @ Ss)) co_recs co_recBs, names_lthy) end; fun derive_rec_transfer_thms lthy recs rec_defs (SOME (_, _, _, xsssss)) = let val (Rs, Ss, goals, _) = mk_co_rec_transfer_goals lthy recs; val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_rec_transfer_tac ctxt nn ns (map (Thm.cterm_of ctxt) Ss) (map (Thm.cterm_of ctxt) Rs) xsssss rec_defs xtor_co_rec_transfers pre_rel_defs live_nesting_rel_eqs) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced nn end; fun derive_rec_o_map_thmss lthy recs rec_defs = if live = 0 then replicate nn [] else let fun variant_names n pre = fst (Variable.variant_fixes (replicate n pre) lthy); val maps0 = map map_of_bnf fp_bnfs; val f_names = variant_names num_As "f"; val fs = map2 (curry Free) f_names (map (op -->) (As ~~ Bs)); val live_gs = AList.find (op =) (fs ~~ liveness) true; val gmaps = map (fn map0 => Term.list_comb (mk_map live As Bs map0, live_gs)) maps0; val rec_arg_Ts = binder_fun_types (hd (map fastype_of recs)); val num_rec_args = length rec_arg_Ts; val g_Ts = map B_ify_T rec_arg_Ts; val g_names = variant_names num_rec_args "g"; val gs = map2 (curry Free) g_names g_Ts; val grecs = map (fn recx => Term.list_comb (Term.map_types B_ify_T recx, gs)) recs; val rec_o_map_lhss = map2 (curry HOLogic.mk_comp) grecs gmaps; val ABfs = (As ~~ Bs) ~~ fs; fun mk_rec_arg_arg (x as Free (_, T)) = let val U = B_ify_T T in if T = U then x else build_map lthy [] [] (the o AList.lookup (op =) ABfs) (T, U) $ x end; fun mk_rec_o_map_arg rec_arg_T h = let val x_Ts = binder_types rec_arg_T; val m = length x_Ts; val x_names = variant_names m "x"; val xs = map2 (curry Free) x_names x_Ts; val xs' = map mk_rec_arg_arg xs; in fold_rev Term.lambda xs (Term.list_comb (h, xs')) end; fun mk_rec_o_map_rhs recx = let val args = map2 mk_rec_o_map_arg rec_arg_Ts gs in Term.list_comb (recx, args) end; val rec_o_map_rhss = map mk_rec_o_map_rhs recs; val rec_o_map_goals = map2 (fold_rev (fold_rev Logic.all) [fs, gs] o HOLogic.mk_Trueprop oo curry HOLogic.mk_eq) rec_o_map_lhss rec_o_map_rhss; val rec_o_map_thms = @{map 3} (fn goal => fn rec_def => fn ctor_rec_o_map => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => mk_co_rec_o_map_tac ctxt rec_def pre_map_defs live_nesting_map_ident0s abs_inverses ctor_rec_o_map) |> Thm.close_derivation \<^here>) rec_o_map_goals rec_defs xtor_co_rec_o_maps; in map single rec_o_map_thms end; fun derive_note_induct_recs_thms_for_types ((((map_thmss, map_disc_iffss, map_selsss, rel_injectss, rel_distinctss, rel_selss, rel_intross, rel_casess, pred_injectss, set_thmss, set_selsssss, set_introsssss, set_casess, ctr_transferss, case_transferss, disc_transferss, sel_transferss), (ctrss, _, ctor_iff_dtors, ctr_defss, ctr_sugars)), (recs, rec_defs)), lthy) = let val ((induct_thms, induct_thm, induct_attrs), (rec_thmss, rec_attrs)) = derive_induct_recs_thms_for_types plugins pre_bnfs recs_args_types xtor_co_induct xtor_co_rec_thms live_nesting_bnfs fp_nesting_bnfs fpTs Cs Xs ctrXs_Tsss abs_inverses type_definitions abs_inverses ctrss ctr_defss recs rec_defs lthy; val rec_transfer_thmss = map single (derive_rec_transfer_thms lthy recs rec_defs recs_args_types); val induct_type_attr = Attrib.internal o K o Induct.induct_type; val induct_pred_attr = Attrib.internal o K o Induct.induct_pred; val ((rel_induct_thmss, common_rel_induct_thms), (rel_induct_attrs, common_rel_induct_attrs)) = if live = 0 then ((replicate nn [], []), ([], [])) else let val ((rel_induct_thms, common_rel_induct_thm), rel_induct_attrs) = derive_rel_induct_thms_for_types lthy nn fpTs As Bs ctrss ctr_Tsss (map #exhaust ctr_sugars) xtor_rel_co_induct ctr_defss ctor_injects pre_rel_defs abs_inverses live_nesting_rel_eqs; in ((map single rel_induct_thms, single common_rel_induct_thm), (rel_induct_attrs, rel_induct_attrs)) end; val rec_o_map_thmss = derive_rec_o_map_thmss lthy recs rec_defs; val simp_thmss = @{map 6} mk_simp_thms ctr_sugars rec_thmss map_thmss rel_injectss rel_distinctss set_thmss; val common_notes = (if nn > 1 then [(inductN, [induct_thm], K induct_attrs), (rel_inductN, common_rel_induct_thms, K common_rel_induct_attrs)] else []) |> massage_simple_notes fp_common_name; val notes = [(inductN, map single induct_thms, fn T_name => induct_attrs @ [induct_type_attr T_name]), (recN, rec_thmss, K rec_attrs), (rec_o_mapN, rec_o_map_thmss, K []), (rec_transferN, rec_transfer_thmss, K []), (rel_inductN, rel_induct_thmss, K (rel_induct_attrs @ [induct_pred_attr ""])), (simpsN, simp_thmss, K [])] |> massage_multi_notes fp_b_names fpTs; in lthy |> Spec_Rules.add Binding.empty Spec_Rules.equational recs (flat rec_thmss) |> plugins code_plugin ? Code.declare_default_eqns (map (rpair true) (flat rec_thmss)) |> Local_Theory.notes (common_notes @ notes) (* for "datatype_realizer.ML": *) |>> name_noted_thms (fst (dest_Type (hd fpTs)) ^ implode (map (prefix "_") (tl fp_b_names))) inductN |-> interpret_bnfs_register_fp_sugars plugins fpTs fpBTs Xs Least_FP pre_bnfs absT_infos fp_nesting_bnfs live_nesting_bnfs fp_res ctrXs_Tsss ctor_iff_dtors ctr_defss ctr_sugars recs rec_defs map_thmss [induct_thm] (map single induct_thms) rec_thmss (replicate nn []) (replicate nn []) rel_injectss rel_distinctss map_disc_iffss map_selsss rel_selss rel_intross rel_casess pred_injectss set_thmss set_selsssss set_introsssss set_casess ctr_transferss case_transferss disc_transferss sel_transferss (replicate nn []) (replicate nn []) rec_transfer_thmss common_rel_induct_thms rel_induct_thmss [] (replicate nn []) rec_o_map_thmss end; fun derive_corec_transfer_thms lthy corecs corec_defs = let val (Rs, Ss, goals, _) = mk_co_rec_transfer_goals lthy corecs; val (_, _, _, (((pgss, pss, qssss, gssss), _), _)) = the corecs_args_types; val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_corec_transfer_tac ctxt (map (Thm.cterm_of ctxt) Ss) (map (Thm.cterm_of ctxt) Rs) type_definitions corec_defs xtor_co_rec_transfers pre_rel_defs live_nesting_rel_eqs (flat pgss) pss qssss gssss) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end; fun derive_map_o_corec_thmss lthy0 lthy2 corecs corec_defs = if live = 0 then replicate nn [] else let fun variant_names n pre = fst (Variable.variant_fixes (replicate n pre) lthy0); val maps0 = map map_of_bnf fp_bnfs; val f_names = variant_names num_As "f"; val fs = map2 (curry Free) f_names (map (op -->) (As ~~ Bs)); val live_fs = AList.find (op =) (fs ~~ liveness) true; val fmaps = map (fn map0 => Term.list_comb (mk_map live As Bs map0, live_fs)) maps0; val corec_arg_Ts = binder_fun_types (hd (map fastype_of corecs)); val num_rec_args = length corec_arg_Ts; val g_names = variant_names num_rec_args "g"; val gs = map2 (curry Free) g_names corec_arg_Ts; val gcorecs = map (fn corecx => Term.list_comb (corecx, gs)) corecs; val map_o_corec_lhss = map2 (curry HOLogic.mk_comp) fmaps gcorecs; val ABfs = (As ~~ Bs) ~~ fs; fun mk_map_o_corec_arg corec_argB_T g = let val T = range_type (fastype_of g); val U = range_type corec_argB_T; in if T = U then g else HOLogic.mk_comp (build_map lthy2 [] [] (the o AList.lookup (op =) ABfs) (T, U), g) end; fun mk_map_o_corec_rhs corecx = let val args = map2 (mk_map_o_corec_arg o B_ify_T) corec_arg_Ts gs in Term.list_comb (B_ify corecx, args) end; val map_o_corec_rhss = map mk_map_o_corec_rhs corecs; val map_o_corec_goals = map2 (fold_rev (fold_rev Logic.all) [fs, gs] o HOLogic.mk_Trueprop oo curry HOLogic.mk_eq) map_o_corec_lhss map_o_corec_rhss; val map_o_corec_thms = @{map 3} (fn goal => fn corec_def => fn dtor_map_o_corec => Goal.prove_sorry lthy2 [] [] goal (fn {context = ctxt, ...} => mk_co_rec_o_map_tac ctxt corec_def pre_map_defs live_nesting_map_ident0s abs_inverses dtor_map_o_corec) |> Thm.close_derivation \<^here>) map_o_corec_goals corec_defs xtor_co_rec_o_maps; in map single map_o_corec_thms end; fun derive_note_coinduct_corecs_thms_for_types ((((map_thmss, map_disc_iffss, map_selsss, rel_injectss, rel_distinctss, rel_selss, rel_intross, rel_casess, pred_injectss, set_thmss, set_selsssss, set_introsssss, set_casess, ctr_transferss, case_transferss, disc_transferss, sel_transferss), (_, _, ctor_iff_dtors, ctr_defss, ctr_sugars)), (corecs, corec_defs)), lthy) = let val (([(coinduct_thms, coinduct_thm), (coinduct_strong_thms, coinduct_strong_thm)], (coinduct_attrs, common_coinduct_attrs)), corec_thmss, corec_disc_thmss, (corec_disc_iff_thmss, corec_disc_iff_attrs), (corec_sel_thmsss, corec_sel_attrs)) = derive_coinduct_corecs_thms_for_types lthy pre_bnfs (the corecs_args_types) xtor_co_induct dtor_injects dtor_ctors xtor_co_rec_thms live_nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss mss ns abs_inverses abs_inverses I ctr_defss ctr_sugars corecs corec_defs; fun distinct_prems ctxt thm = Goal.prove (*no sorry*) ctxt [] [] (thm |> Thm.prop_of |> Logic.strip_horn |>> distinct (op aconv) |> Logic.list_implies) (fn _ => HEADGOAL (cut_tac thm THEN' assume_tac ctxt) THEN ALLGOALS eq_assume_tac); fun eq_ifIN _ [thm] = thm | eq_ifIN ctxt (thm :: thms) = distinct_prems ctxt (@{thm eq_ifI} OF (map (unfold_thms ctxt @{thms atomize_imp[of _ "t = u" for t u]}) [thm, eq_ifIN ctxt thms])); val corec_code_thms = map (eq_ifIN lthy) corec_thmss; val corec_sel_thmss = map flat corec_sel_thmsss; val coinduct_type_attr = Attrib.internal o K o Induct.coinduct_type; val coinduct_pred_attr = Attrib.internal o K o Induct.coinduct_pred; val flat_corec_thms = append oo append; val corec_transfer_thmss = map single (derive_corec_transfer_thms lthy corecs corec_defs); val ((rel_coinduct_thmss, common_rel_coinduct_thms), (rel_coinduct_attrs, common_rel_coinduct_attrs)) = if live = 0 then ((replicate nn [], []), ([], [])) else let val ((rel_coinduct_thms, common_rel_coinduct_thm), (rel_coinduct_attrs, common_rel_coinduct_attrs)) = derive_rel_coinduct_thms_for_types lthy nn fpTs ns As Bs mss ctr_sugars abs_inverses abs_injects ctor_injects dtor_ctors pre_rel_defs ctr_defss xtor_rel_co_induct live_nesting_rel_eqs; in ((map single rel_coinduct_thms, single common_rel_coinduct_thm), (rel_coinduct_attrs, common_rel_coinduct_attrs)) end; val map_o_corec_thmss = derive_map_o_corec_thmss lthy lthy corecs corec_defs; val (set_induct_thms, set_induct_attrss) = derive_set_induct_thms_for_types lthy nn fpTs (map #ctrs ctr_sugars) (map (map (mk_set As)) (map sets_of_bnf fp_bnfs)) dtor_set_inducts (map #exhaust ctr_sugars) (flat pre_set_defss) (flat ctr_defss) dtor_ctors abs_inverses |> split_list; val simp_thmss = @{map 6} mk_simp_thms ctr_sugars (@{map 3} flat_corec_thms corec_disc_thmss corec_disc_iff_thmss corec_sel_thmss) map_thmss rel_injectss rel_distinctss set_thmss; val common_notes = (set_inductN, set_induct_thms, nth set_induct_attrss) :: (if nn > 1 then [(coinductN, [coinduct_thm], K common_coinduct_attrs), (coinduct_strongN, [coinduct_strong_thm], K common_coinduct_attrs), (rel_coinductN, common_rel_coinduct_thms, K common_rel_coinduct_attrs)] else []) |> massage_simple_notes fp_common_name; val notes = [(coinductN, map single coinduct_thms, fn T_name => coinduct_attrs @ [coinduct_type_attr T_name]), (coinduct_strongN, map single coinduct_strong_thms, K coinduct_attrs), (corecN, corec_thmss, K []), (corec_codeN, map single corec_code_thms, K (nitpicksimp_attrs)), (corec_discN, corec_disc_thmss, K []), (corec_disc_iffN, corec_disc_iff_thmss, K corec_disc_iff_attrs), (corec_selN, corec_sel_thmss, K corec_sel_attrs), (corec_transferN, corec_transfer_thmss, K []), (map_o_corecN, map_o_corec_thmss, K []), (rel_coinductN, rel_coinduct_thmss, K (rel_coinduct_attrs @ [coinduct_pred_attr ""])), (simpsN, simp_thmss, K [])] |> massage_multi_notes fp_b_names fpTs; in lthy |> fold (Spec_Rules.add Binding.empty Spec_Rules.equational corecs) [flat corec_sel_thmss, flat corec_thmss] |> plugins code_plugin ? Code.declare_default_eqns (map (rpair true) corec_code_thms) |> Local_Theory.notes (common_notes @ notes) |-> interpret_bnfs_register_fp_sugars plugins fpTs fpBTs Xs Greatest_FP pre_bnfs absT_infos fp_nesting_bnfs live_nesting_bnfs fp_res ctrXs_Tsss ctor_iff_dtors ctr_defss ctr_sugars corecs corec_defs map_thmss [coinduct_thm, coinduct_strong_thm] (transpose [coinduct_thms, coinduct_strong_thms]) corec_thmss corec_disc_thmss corec_sel_thmsss rel_injectss rel_distinctss map_disc_iffss map_selsss rel_selss rel_intross rel_casess pred_injectss set_thmss set_selsssss set_introsssss set_casess ctr_transferss case_transferss disc_transferss sel_transferss corec_disc_iff_thmss (map single corec_code_thms) corec_transfer_thmss common_rel_coinduct_thms rel_coinduct_thmss set_induct_thms (replicate nn (if nn = 1 then set_induct_thms else [])) map_o_corec_thmss end; val lthy = lthy |> live > 0 ? fold2 (fn Type (s, _) => fn bnf => register_bnf_raw s bnf) fpTs fp_bnfs |> @{fold_map 29} define_ctrs_dtrs_for_type_etc fp_bnfs fp_bs fpTs Cs Es ctors dtors xtor_co_recs ctor_dtors dtor_ctors ctor_injects pre_map_defs pre_set_defss pre_rel_defs xtor_maps xtor_setss xtor_rels ns kss mss abss abs_injects type_definitions ctr_bindingss ctr_mixfixess ctr_Tsss disc_bindingss sel_bindingsss raw_sel_default_eqss |> wrap_ctrs_derive_map_set_rel_pred_thms_define_co_rec_for_types |> case_fp fp derive_note_induct_recs_thms_for_types derive_note_coinduct_corecs_thms_for_types; val timer = time (timer ("Constructors, discriminators, selectors, etc., for the new " ^ co_prefix fp ^ "datatype")); in lthy end; fun co_datatypes fp = define_co_datatypes (K I) (K I) (K I) (K I) fp; fun co_datatype_cmd fp construct_fp options lthy = define_co_datatypes Plugin_Name.make_filter Typedecl.read_constraint Syntax.parse_typ Syntax.parse_term fp construct_fp options lthy handle EMPTY_DATATYPE s => error ("Cannot define empty datatype " ^ quote s); val parse_ctr_arg = \<^keyword>\(\ |-- parse_binding_colon -- Parse.typ --| \<^keyword>\)\ || Parse.typ >> pair Binding.empty; val parse_ctr_specs = Parse.enum1 "|" (parse_ctr_spec Parse.binding parse_ctr_arg -- Parse.opt_mixfix); val parse_spec = parse_type_args_named_constrained -- Parse.binding -- Parse.opt_mixfix -- (\<^keyword>\=\ |-- parse_ctr_specs) -- parse_map_rel_pred_bindings -- parse_sel_default_eqs; val parse_co_datatype = parse_ctr_options -- Parse.and_list1 parse_spec; fun parse_co_datatype_cmd fp construct_fp = parse_co_datatype >> co_datatype_cmd fp construct_fp; end; diff --git a/src/HOL/Tools/BNF/bnf_fp_n2m.ML b/src/HOL/Tools/BNF/bnf_fp_n2m.ML --- a/src/HOL/Tools/BNF/bnf_fp_n2m.ML +++ b/src/HOL/Tools/BNF/bnf_fp_n2m.ML @@ -1,620 +1,620 @@ (* Title: HOL/Tools/BNF/bnf_fp_n2m.ML Author: Dmitriy Traytel, TU Muenchen Copyright 2013 Flattening of nested to mutual (co)recursion. *) signature BNF_FP_N2M = sig val construct_mutualized_fp: BNF_Util.fp_kind -> int list -> typ list -> (int * BNF_FP_Util.fp_result) list -> binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list -> BNF_Comp.absT_info list -> local_theory -> BNF_FP_Util.fp_result * local_theory end; structure BNF_FP_N2M : BNF_FP_N2M = struct open BNF_Def open BNF_Util open BNF_Comp open BNF_FP_Util open BNF_FP_Def_Sugar open BNF_Tactics open BNF_FP_N2M_Tactics fun mk_arg_cong ctxt n t = let val Us = fastype_of t |> strip_typeN n |> fst; val ((xs, ys), _) = ctxt |> mk_Frees "x" Us ||>> mk_Frees "y" Us; val goal = Logic.list_implies (@{map 2} (curry mk_Trueprop_eq) xs ys, mk_Trueprop_eq (list_comb (t, xs), list_comb (t, ys))); val vars = Variable.add_free_names ctxt goal []; in Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => HEADGOAL (hyp_subst_tac ctxt THEN' rtac ctxt refl)) |> Thm.close_derivation \<^here> end; val cacheN = "cache" fun mk_cacheN i = cacheN ^ string_of_int i ^ "_"; val cache_threshold = Attrib.setup_config_int \<^binding>\bnf_n2m_cache_threshold\ (K 200); type cache = int * (term * thm) Typtab.table val empty_cache = (0, Typtab.empty) fun update_cache b0 TU t (cache as (i, tab), lthy) = if size_of_term t < Config.get lthy cache_threshold then (t, (cache, lthy)) else let val b = Binding.prefix_name (mk_cacheN i) b0; val ((c, thm), lthy') = Local_Theory.define ((b, NoSyn), ((Binding.concealed (Thm.def_binding b), []), t)) lthy |>> apsnd snd; in (c, ((i + 1, Typtab.update (TU, (c, thm)) tab), lthy')) end; fun lookup_cache (SOME _) _ _ = NONE | lookup_cache NONE TU ((_, tab), _) = Typtab.lookup tab TU |> Option.map fst; fun construct_mutualized_fp fp mutual_cliques fpTs indexed_fp_ress bs resBs (resDs, Dss) bnfs (absT_infos : absT_info list) lthy = let val time = time lthy; val timer = time (Timer.startRealTimer ()); val b_names = map Binding.name_of bs; val b_name = mk_common_name b_names; val b = Binding.name b_name; fun of_fp_res get = map (uncurry nth o swap o apsnd get) indexed_fp_ress; fun mk_co_algT T U = case_fp fp (T --> U) (U --> T); fun co_swap pair = case_fp fp I swap pair; val mk_co_comp = curry (HOLogic.mk_comp o co_swap); val dest_co_algT = co_swap o dest_funT; val co_alg_argT = case_fp fp range_type domain_type; val co_alg_funT = case_fp fp domain_type range_type; val rewrite_comp_comp = case_fp fp @{thm rewriteL_comp_comp} @{thm rewriteR_comp_comp}; val fp_absT_infos = of_fp_res #absT_infos; val fp_bnfs = of_fp_res #bnfs; val fp_pre_bnfs = of_fp_res #pre_bnfs; val fp_absTs = map #absT fp_absT_infos; val fp_repTs = map #repT fp_absT_infos; val fp_abss = map #abs fp_absT_infos; val fp_reps = map #rep fp_absT_infos; val fp_type_definitions = map #type_definition fp_absT_infos; val absTs = map #absT absT_infos; val repTs = map #repT absT_infos; val absTs' = map (Logic.type_map (singleton (Variable.polymorphic lthy))) absTs; val repTs' = map (Logic.type_map (singleton (Variable.polymorphic lthy))) repTs; val abss = map #abs absT_infos; val reps = map #rep absT_infos; val abs_inverses = map #abs_inverse absT_infos; val type_definitions = map #type_definition absT_infos; val n = length bnfs; val deads = fold (union (op =)) Dss resDs; val As = subtract (op =) deads (map TFree resBs); val names_lthy = fold Variable.declare_typ (As @ deads) lthy; val m = length As; val live = m + n; val (((Xs, Ys), Bs), names_lthy) = names_lthy |> mk_TFrees n ||>> mk_TFrees n ||>> mk_TFrees m; val allAs = As @ Xs; val allBs = Bs @ Xs; val phiTs = map2 mk_pred2T As Bs; val thetaBs = As ~~ Bs; val fpTs' = map (Term.typ_subst_atomic thetaBs) fpTs; val fold_thetaAs = Xs ~~ fpTs; val fold_thetaBs = Xs ~~ fpTs'; val pre_phiTs = map2 mk_pred2T fpTs fpTs'; val ((ctors, dtors), (xtor's, xtors)) = let val ctors = map2 (force_typ names_lthy o (fn T => dummyT --> T)) fpTs (of_fp_res #ctors); val dtors = map2 (force_typ names_lthy o (fn T => T --> dummyT)) fpTs (of_fp_res #dtors); in ((ctors, dtors), `(map (Term.subst_atomic_types thetaBs)) (case_fp fp ctors dtors)) end; val absATs = map (domain_type o fastype_of) ctors; val absBTs = map (Term.typ_subst_atomic thetaBs) absATs; val xTs = map (domain_type o fastype_of) xtors; val yTs = map (domain_type o fastype_of) xtor's; val absAs = @{map 3} (fn Ds => mk_abs o mk_T_of_bnf Ds allAs) Dss bnfs abss; val absBs = @{map 3} (fn Ds => mk_abs o mk_T_of_bnf Ds allBs) Dss bnfs abss; val fp_repAs = map2 mk_rep absATs fp_reps; val fp_repBs = map2 mk_rep absBTs fp_reps; val typ_subst_nonatomic_sorted = fold_rev (typ_subst_nonatomic o single); val sorted_theta = sort (int_ord o apply2 (Term.size_of_typ o fst)) (fpTs ~~ Xs) val sorted_fpTs = map fst sorted_theta; val nesting_bnfs = nesting_bnfs lthy [[map (typ_subst_nonatomic_sorted (rev sorted_theta) o range_type o fastype_of) fp_repAs]] allAs; val fp_or_nesting_bnfs = distinct (op = o apply2 T_of_bnf) (fp_bnfs @ nesting_bnfs); val (((((phis, phis'), pre_phis), xs), ys), names_lthy) = names_lthy |> mk_Frees' "R" phiTs ||>> mk_Frees "S" pre_phiTs ||>> mk_Frees "x" xTs ||>> mk_Frees "y" yTs; val rels = let fun find_rel T As Bs = fp_or_nesting_bnfs |> filter_out (curry (op = o apply2 name_of_bnf) BNF_Comp.DEADID_bnf) |> find_first (fn bnf => Type.could_unify (T_of_bnf bnf, T)) |> Option.map (fn bnf => let val live = live_of_bnf bnf; in (mk_rel live As Bs (rel_of_bnf bnf), live) end) |> the_default (HOLogic.eq_const T, 0); fun mk_rel (T as Type (_, Ts)) (Type (_, Us)) = let val (rel, live) = find_rel T Ts Us; val (Ts', Us') = fastype_of rel |> strip_typeN live |> fst |> map_split dest_pred2T; val rels = map2 mk_rel Ts' Us'; in Term.list_comb (rel, rels) end | mk_rel (T as TFree _) _ = (nth phis (find_index (curry op = T) As) handle General.Subscript => HOLogic.eq_const T) | mk_rel _ _ = raise Fail "fpTs contains schematic type variables"; in map2 (fold_rev Term.absfree phis' oo mk_rel) fpTs fpTs' end; val pre_rels = map2 (fn Ds => mk_rel_of_bnf Ds (As @ fpTs) (Bs @ fpTs')) Dss bnfs; val rel_unfolds = maps (no_refl o single o rel_def_of_bnf) fp_pre_bnfs; val rel_xtor_co_inducts = of_fp_res (split_conj_thm o #xtor_rel_co_induct) |> map (zero_var_indexes o unfold_thms lthy (id_apply :: rel_unfolds)); val rel_defs = map rel_def_of_bnf bnfs; val rel_monos = map rel_mono_of_bnf bnfs; fun cast castA castB pre_rel = let val castAB = mk_vimage2p (Term.subst_atomic_types fold_thetaAs castA) (Term.subst_atomic_types fold_thetaBs castB); in fold_rev (fold_rev Term.absdummy) [phiTs, pre_phiTs] (castAB $ Term.list_comb (pre_rel, map Bound (live - 1 downto 0))) end; val castAs = map2 (curry HOLogic.mk_comp) absAs fp_repAs; val castBs = map2 (curry HOLogic.mk_comp) absBs fp_repBs; val fp_or_nesting_rel_eqs = no_refl (map rel_eq_of_bnf fp_or_nesting_bnfs); val fp_or_nesting_rel_monos = map rel_mono_of_bnf fp_or_nesting_bnfs; fun mutual_instantiate ctxt inst = let val thetas = AList.group (op =) (mutual_cliques ~~ inst); in map2 (infer_instantiate ctxt o the o AList.lookup (op =) thetas) mutual_cliques end; val rel_xtor_co_inducts_inst = let val extract = case_fp fp (snd o Term.dest_comb) (snd o Term.dest_comb o fst o Term.dest_comb); val raw_phis = map (extract o HOLogic.dest_Trueprop o Thm.concl_of) rel_xtor_co_inducts; val inst = map (fn (t, u) => (#1 (dest_Var t), Thm.cterm_of lthy u)) (raw_phis ~~ pre_phis); in mutual_instantiate lthy inst rel_xtor_co_inducts end val xtor_rel_co_induct = mk_xtor_rel_co_induct_thm fp (@{map 3} cast castAs castBs pre_rels) pre_phis rels phis xs ys xtors xtor's (mk_rel_xtor_co_induct_tac fp abs_inverses rel_xtor_co_inducts_inst rel_defs rel_monos fp_or_nesting_rel_eqs fp_or_nesting_rel_monos) lthy; val map_id0s = no_refl (map map_id0_of_bnf bnfs); val xtor_co_induct_thm = (case fp of Least_FP => let val (Ps, names_lthy) = names_lthy |> mk_Frees "P" (map (fn T => T --> HOLogic.boolT) fpTs); fun mk_Grp_id P = let val T = domain_type (fastype_of P); in mk_Grp (HOLogic.Collect_const T $ P) (HOLogic.id_const T) end; val cts = map (SOME o Thm.cterm_of names_lthy) (map HOLogic.eq_const As @ map mk_Grp_id Ps); fun mk_fp_type_copy_thms thm = map (curry op RS thm) @{thms type_copy_Abs_o_Rep type_copy_vimage2p_Grp_Rep}; fun mk_type_copy_thms thm = map (curry op RS thm) @{thms type_copy_Rep_o_Abs type_copy_vimage2p_Grp_Abs}; in infer_instantiate' names_lthy cts xtor_rel_co_induct |> singleton (Proof_Context.export names_lthy lthy) |> unfold_thms lthy (@{thms eq_le_Grp_id_iff all_simps(1,2)[symmetric]} @ fp_or_nesting_rel_eqs) |> funpow n (fn thm => thm RS spec) |> unfold_thms lthy (@{thm eq_alt} :: map rel_Grp_of_bnf bnfs @ map_id0s) |> unfold_thms lthy (@{thms vimage2p_id vimage2p_comp comp_apply comp_id Grp_id_mono_subst eqTrueI[OF subset_UNIV] simp_thms(22)} @ maps mk_fp_type_copy_thms fp_type_definitions @ maps mk_type_copy_thms type_definitions) |> unfold_thms lthy @{thms subset_iff mem_Collect_eq atomize_conjL[symmetric] atomize_all[symmetric] atomize_imp[symmetric]} end | Greatest_FP => let val cts = NONE :: map (SOME o Thm.cterm_of lthy) (map HOLogic.eq_const As); in infer_instantiate' lthy cts xtor_rel_co_induct |> unfold_thms lthy (@{thms le_fun_def le_bool_def all_simps(1,2)[symmetric]} @ fp_or_nesting_rel_eqs) |> funpow (2 * n) (fn thm => thm RS spec) |> Conv.fconv_rule (Object_Logic.atomize lthy) |> funpow n (fn thm => thm RS mp) end); val timer = time (timer "Nested-to-mutual (co)induction"); val fold_preTs = map2 (fn Ds => mk_T_of_bnf Ds allAs) Dss bnfs; val fold_strTs = map2 mk_co_algT fold_preTs Xs; val resTs = map2 mk_co_algT fpTs Xs; val fp_un_folds = of_fp_res #xtor_un_folds; val ns = map (length o #Ts o snd) indexed_fp_ress; fun force_fold i TU raw_un_fold = let val thy = Proof_Context.theory_of lthy; val approx_un_fold = raw_un_fold |> force_typ names_lthy (replicate (nth ns i) dummyT ---> TU); val subst = Term.typ_subst_atomic fold_thetaAs; fun mk_fp_absT_repT fp_repT fp_absT = mk_absT thy fp_repT fp_absT ooo mk_repT; val mk_fp_absT_repTs = @{map 5} mk_fp_absT_repT fp_repTs fp_absTs absTs repTs; val fold_preTs' = mk_fp_absT_repTs (map subst fold_preTs); val fold_pre_deads_only_Ts = map (typ_subst_nonatomic_sorted (map (rpair dummyT) (As @ sorted_fpTs))) fold_preTs'; val (mutual_clique, TUs) = map_split dest_co_algT (binder_fun_types (fastype_of approx_un_fold)) |>> map subst |> `(fn (_, Ys) => nth mutual_cliques (find_index (fn X => X = the (find_first (can dest_TFree) Ys)) Xs)) ||> uncurry (map2 mk_co_algT); val cands = mutual_cliques ~~ map2 mk_co_algT fold_preTs' Xs; val js = find_indices (fn ((cl, cand), TU) => cl = mutual_clique andalso Type.could_unify (TU, cand)) TUs cands; val Tpats = map (fn j => mk_co_algT (nth fold_pre_deads_only_Ts j) (nth Xs j)) js; in force_typ names_lthy (Tpats ---> TU) raw_un_fold end; fun mk_co_comp_abs_rep fp_absT absT fp_abs fp_rep abs rep t = case_fp fp (HOLogic.mk_comp (HOLogic.mk_comp (t, mk_abs absT abs), mk_rep fp_absT fp_rep)) (HOLogic.mk_comp (mk_abs fp_absT fp_abs, HOLogic.mk_comp (mk_rep absT rep, t))); val thy = Proof_Context.theory_of lthy; fun mk_absT_fp_repT repT absT = mk_absT thy repT absT ooo mk_repT; fun mk_un_fold b_opt ss un_folds cache_lthy TU = (case lookup_cache b_opt TU cache_lthy of SOME t => ((t, Drule.dummy_thm), cache_lthy) | NONE => let val x = co_alg_argT TU; val i = find_index (fn T => x = T) Xs; val TUfold = (case find_first (fn f => body_fun_type (fastype_of f) = TU) un_folds of NONE => force_fold i TU (nth fp_un_folds i) | SOME f => f); val TUs = binder_fun_types (fastype_of TUfold); fun mk_s TU' cache_lthy = let val i = find_index (fn T => co_alg_argT TU' = T) Xs; val fp_abs = nth fp_abss i; val fp_rep = nth fp_reps i; val abs = nth abss i; val rep = nth reps i; val sF = co_alg_funT TU'; val sF' = mk_absT_fp_repT (nth repTs' i) (nth absTs' i) (nth fp_absTs i) (nth fp_repTs i) sF handle Term.TYPE _ => sF; val F = nth fold_preTs i; val s = nth ss i; in if sF = F then (s, cache_lthy) else if sF' = F then (mk_co_comp_abs_rep sF sF' fp_abs fp_rep abs rep s, cache_lthy) else let val smapT = replicate live dummyT ---> mk_co_algT sF' F; fun hidden_to_unit t = Term.subst_TVars (map (rpair HOLogic.unitT) (Term.add_tvar_names t [])) t; val smap = map_of_bnf (nth bnfs i) |> force_typ names_lthy smapT |> hidden_to_unit; val smap_argTs = strip_typeN live (fastype_of smap) |> fst; fun mk_smap_arg T_to_U cache_lthy = (if domain_type T_to_U = range_type T_to_U then (HOLogic.id_const (domain_type T_to_U), cache_lthy) else mk_un_fold NONE ss un_folds cache_lthy T_to_U |>> fst); val (smap_args, cache_lthy') = fold_map mk_smap_arg smap_argTs cache_lthy; in (mk_co_comp_abs_rep sF sF' fp_abs fp_rep abs rep (mk_co_comp s (Term.list_comb (smap, smap_args))), cache_lthy') end end; val (args, cache_lthy) = fold_map mk_s TUs cache_lthy; val t = Term.list_comb (TUfold, args); in (case b_opt of NONE => update_cache b TU t cache_lthy |>> rpair Drule.dummy_thm | SOME b => cache_lthy |-> (fn cache => let val S = HOLogic.mk_tupleT fold_strTs; val s = HOLogic.mk_tuple ss; val u = Const (\<^const_name>\Let\, S --> (S --> TU) --> TU) $ s $ absdummy S t; in Local_Theory.define ((b, NoSyn), ((Binding.concealed (Thm.def_binding b), []), u)) #>> apsnd snd ##> pair cache end)) end); val un_foldN = case_fp fp ctor_foldN dtor_unfoldN; fun mk_un_folds (ss_names, lthy) = let val ss = map2 (curry Free) ss_names fold_strTs; in fold2 (fn TU => fn b => fn ((un_folds, defs), cache_lthy) => mk_un_fold (SOME b) (map2 (curry Free) ss_names fold_strTs) un_folds cache_lthy TU |>> (fn (f, d) => (f :: un_folds, d :: defs))) resTs (map (Binding.suffix_name ("_" ^ un_foldN)) bs) (([], []), (empty_cache, lthy)) |>> map_prod rev rev |>> pair ss end; val ((ss, (un_folds, un_fold_defs0)), (cache, (lthy, raw_lthy))) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> Variable.add_fixes (mk_names n "s") |> mk_un_folds ||> apsnd (`(Local_Theory.close_target)); val un_fold_defs = map (unfold_thms raw_lthy @{thms Let_const}) un_fold_defs0; val cache_defs = snd cache |> Typtab.dest |> map (snd o snd); val phi = Proof_Context.export_morphism raw_lthy lthy; val xtor_un_folds = map (head_of o Morphism.term phi) un_folds; val xtor_un_fold_defs = map (Drule.abs_def o Morphism.thm phi) un_fold_defs; val xtor_cache_defs = map (Drule.abs_def o Morphism.thm phi) cache_defs; val xtor_un_folds' = map2 (fn raw => fn t => Const (fst (dest_Const t), fold_strTs ---> fastype_of raw)) un_folds xtor_un_folds; val fp_un_fold_o_maps = of_fp_res #xtor_un_fold_o_maps |> maps (fn thm => [thm, thm RS rewrite_comp_comp]); val fold_mapTs = co_swap (As @ fpTs, As @ Xs); val pre_fold_maps = @{map 2} (fn Ds => uncurry (mk_map_of_bnf Ds) fold_mapTs) Dss bnfs fun mk_pre_fold_maps fs = map (fn mapx => Term.list_comb (mapx, map HOLogic.id_const As @ fs)) pre_fold_maps; val pre_map_defs = no_refl (map map_def_of_bnf bnfs); val fp_map_defs = no_refl (map map_def_of_bnf fp_pre_bnfs); val map_defs = pre_map_defs @ fp_map_defs; val pre_rel_defs = no_refl (map rel_def_of_bnf bnfs); val fp_rel_defs = no_refl (map rel_def_of_bnf fp_pre_bnfs); val rel_defs = pre_rel_defs @ fp_rel_defs; fun mk_Rep_o_Abs thm = (thm RS @{thm type_copy_Rep_o_Abs}) |> (fn thm => [thm, thm RS rewrite_comp_comp]); val fp_Rep_o_Abss = maps mk_Rep_o_Abs fp_type_definitions; val pre_Rep_o_Abss = maps mk_Rep_o_Abs type_definitions; val Rep_o_Abss = fp_Rep_o_Abss @ pre_Rep_o_Abss; val unfold_map = map (unfold_thms lthy (id_apply :: pre_map_defs)); val simp_thms = case_fp fp @{thm comp_assoc} @{thm comp_assoc[symmetric]} :: @{thms id_apply comp_id id_comp}; val eq_thm_prop_untyped = Term.aconv_untyped o apply2 Thm.full_prop_of; val map_thms = no_refl (maps (fn bnf => let val map_comp0 = map_comp0_of_bnf bnf RS sym in [map_comp0, map_comp0 RS rewrite_comp_comp, map_id0_of_bnf bnf] end) fp_or_nesting_bnfs) @ remove eq_thm_prop_untyped (case_fp fp @{thm comp_assoc[symmetric]} @{thm comp_assoc}) (map2 (fn thm => fn bnf => @{thm type_copy_map_comp0_undo} OF (replicate 3 thm @ unfold_map [map_comp0_of_bnf bnf]) RS rewrite_comp_comp) type_definitions bnfs); val xtor_un_fold_thms = let val pre_fold_maps = mk_pre_fold_maps un_folds; fun mk_goals f xtor s smap fp_abs fp_rep abs rep = let val lhs = mk_co_comp f xtor; val rhs = mk_co_comp s smap; in HOLogic.mk_eq (lhs, mk_co_comp_abs_rep (co_alg_funT (fastype_of lhs)) (co_alg_funT (fastype_of rhs)) fp_abs fp_rep abs rep rhs) end; val goals = @{map 8} mk_goals un_folds xtors ss pre_fold_maps fp_abss fp_reps abss reps; val fp_un_folds = map (mk_pointfree2 lthy) (of_fp_res #xtor_un_fold_thms); val simps = flat [simp_thms, un_fold_defs, map_defs, fp_un_folds, fp_un_fold_o_maps, map_thms, Rep_o_Abss]; in Library.foldr1 HOLogic.mk_conj goals |> HOLogic.mk_Trueprop |> (fn goal => Goal.prove_sorry raw_lthy [] [] goal (fn {context = ctxt, prems = _} => mk_xtor_un_fold_tac ctxt n simps cache_defs)) |> Thm.close_derivation \<^here> |> Morphism.thm phi |> split_conj_thm |> map (fn thm => thm RS @{thm comp_eq_dest}) end; val xtor_un_fold_o_maps = of_fp_res #xtor_un_fold_o_maps |> maps (fn thm => [thm, thm RS rewrite_comp_comp]); val xtor_un_fold_unique_thm = let val (fs, _) = names_lthy |> mk_Frees "f" resTs; val fold_maps = mk_pre_fold_maps fs; fun mk_prem f s mapx xtor fp_abs fp_rep abs rep = let val lhs = mk_co_comp f xtor; val rhs = mk_co_comp s mapx; in mk_Trueprop_eq (lhs, mk_co_comp_abs_rep (co_alg_funT (fastype_of lhs)) (co_alg_funT (fastype_of rhs)) fp_abs fp_rep abs rep rhs) end; val prems = @{map 8} mk_prem fs ss fold_maps xtors fp_abss fp_reps abss reps; val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_eq) fs un_folds)); val vars = Variable.add_free_names raw_lthy concl []; val fp_un_fold_uniques0 = of_fp_res (split_conj_thm o #xtor_un_fold_unique) |> map (Drule.zero_var_indexes o unfold_thms lthy fp_map_defs); val names = fp_un_fold_uniques0 |> map (Thm.concl_of #> HOLogic.dest_Trueprop #> HOLogic.dest_eq #> fst #> dest_Var #> fst); val inst = names ~~ map (Thm.cterm_of lthy) fs; val fp_un_fold_uniques = mutual_instantiate lthy inst fp_un_fold_uniques0; val map_arg_congs = map (fn bnf => mk_arg_cong lthy (live_of_bnf bnf) (map_of_bnf bnf) |> unfold_thms lthy (pre_map_defs @ simp_thms)) nesting_bnfs; in Goal.prove_sorry raw_lthy vars prems concl (mk_xtor_un_fold_unique_tac fp un_fold_defs map_arg_congs xtor_un_fold_o_maps Rep_o_Abss fp_un_fold_uniques simp_thms map_thms map_defs cache_defs) |> Thm.close_derivation \<^here> |> case_fp fp I (fn thm => thm OF replicate n sym) |> Morphism.thm phi end; val ABs = As ~~ Bs; val XYs = Xs ~~ Ys; val ABphiTs = @{map 2} mk_pred2T As Bs; val XYphiTs = @{map 2} mk_pred2T Xs Ys; val ((ABphis, XYphis), _) = names_lthy |> mk_Frees "R" ABphiTs ||>> mk_Frees "S" XYphiTs; val pre_rels = @{map 2} (fn Ds => mk_rel_of_bnf Ds (As @ Xs) (Bs @ Ys)) Dss bnfs; val ns = map (fn i => length (filter (fn c => i = c) mutual_cliques)) mutual_cliques; val map_transfers = map (funpow live (fn thm => thm RS @{thm rel_funD}) #> unfold_thms lthy pre_rel_defs) (map map_transfer_of_bnf bnfs); val fp_un_fold_transfers = map2 (fn n => funpow n (fn thm => thm RS @{thm rel_funD}) #> unfold_thms lthy fp_rel_defs) ns (of_fp_res #xtor_un_fold_transfers); val pre_Abs_transfers = map (fn thm => @{thm Abs_transfer} OF [thm, thm]) type_definitions; val fp_Abs_transfers = map (fn thm => @{thm Abs_transfer} OF [thm, thm]) fp_type_definitions; val Abs_transfers = pre_Abs_transfers @ fp_Abs_transfers; fun tac {context = ctxt, prems = _} = mk_xtor_un_fold_transfer_tac ctxt n xtor_un_fold_defs rel_defs fp_un_fold_transfers map_transfers Abs_transfers fp_or_nesting_rel_eqs xtor_cache_defs; val xtor_un_fold_transfer_thms = mk_xtor_co_iter_transfer_thms fp pre_rels XYphis XYphis rels ABphis xtor_un_folds' (map (subst_atomic_types (ABs @ XYs)) xtor_un_folds') tac lthy; val timer = time (timer "Nested-to-mutual (co)iteration"); val xtor_maps = of_fp_res #xtor_maps; val xtor_rels = of_fp_res #xtor_rels; fun mk_Ts Cs = map (typ_subst_atomic (As ~~ Cs)) fpTs; val phi = Local_Theory.target_morphism lthy; val export = map (Morphism.term phi); val ((xtor_co_recs, (xtor_co_rec_thms, xtor_co_rec_unique_thm, xtor_co_rec_o_map_thms, xtor_co_rec_transfer_thms)), lthy) = lthy |> derive_xtor_co_recs fp bs mk_Ts (Dss, resDs) bnfs (export xtors) (export xtor_un_folds) xtor_un_fold_unique_thm xtor_un_fold_thms xtor_un_fold_transfer_thms xtor_maps xtor_rels (@{map 2} (curry (SOME o @{apply 2} (morph_absT_info phi))) fp_absT_infos absT_infos); val timer = time (timer "Nested-to-mutual (co)recursion"); val common_notes = (case fp of Least_FP => [(ctor_inductN, [xtor_co_induct_thm]), (ctor_rel_inductN, [xtor_rel_co_induct])] | Greatest_FP => [(dtor_coinductN, [xtor_co_induct_thm]), (dtor_rel_coinductN, [xtor_rel_co_induct])]) |> map (fn (thmN, thms) => ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])])); val notes = (case fp of Least_FP => [(ctor_foldN, xtor_un_fold_thms)] | Greatest_FP => [(dtor_unfoldN, xtor_un_fold_thms)]) |> map (apsnd (map single)) |> maps (fn (thmN, thmss) => map2 (fn b => fn thms => ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])])) bs thmss); val lthy = lthy |> Config.get lthy bnf_internals ? snd o Local_Theory.notes (common_notes @ notes); (* These results are half broken. This is deliberate. We care only about those fields that are used by "primrec", "primcorecursive", and "datatype_compat". *) val fp_res = ({Ts = fpTs, bnfs = of_fp_res #bnfs, pre_bnfs = bnfs, absT_infos = absT_infos, dtors = dtors, ctors = ctors, xtor_un_folds = xtor_un_folds, xtor_co_recs = xtor_co_recs, xtor_co_induct = xtor_co_induct_thm, dtor_ctors = of_fp_res #dtor_ctors (*too general types*), ctor_dtors = of_fp_res #ctor_dtors (*too general types*), ctor_injects = of_fp_res #ctor_injects (*too general types*), dtor_injects = of_fp_res #dtor_injects (*too general types*), xtor_maps = of_fp_res #xtor_maps (*too general types and terms*), xtor_map_unique = xtor_un_fold_unique_thm (*wrong*), xtor_setss = of_fp_res #xtor_setss (*too general types and terms*), xtor_rels = of_fp_res #xtor_rels (*too general types and terms*), xtor_un_fold_thms = xtor_un_fold_thms, xtor_co_rec_thms = xtor_co_rec_thms, xtor_un_fold_unique = xtor_un_fold_unique_thm, xtor_co_rec_unique = xtor_co_rec_unique_thm, xtor_un_fold_o_maps = fp_un_fold_o_maps (*wrong*), xtor_co_rec_o_maps = xtor_co_rec_o_map_thms (*wrong*), xtor_un_fold_transfers = xtor_un_fold_transfer_thms, xtor_co_rec_transfers = xtor_co_rec_transfer_thms (*wrong*), xtor_rel_co_induct = xtor_rel_co_induct, dtor_set_inducts = []} |> morph_fp_result (Morphism.term_morphism "BNF" (singleton (Variable.polymorphic lthy)))); in timer; (fp_res, lthy) end; end; diff --git a/src/HOL/Tools/BNF/bnf_fp_util.ML b/src/HOL/Tools/BNF/bnf_fp_util.ML --- a/src/HOL/Tools/BNF/bnf_fp_util.ML +++ b/src/HOL/Tools/BNF/bnf_fp_util.ML @@ -1,984 +1,984 @@ (* Title: HOL/Tools/BNF/bnf_fp_util.ML Author: Dmitriy Traytel, TU Muenchen Author: Jasmin Blanchette, TU Muenchen Author: Martin Desharnais, TU Muenchen Author: Stefan Berghofer, TU Muenchen Copyright 2012, 2013, 2014 Shared library for the datatype and codatatype constructions. *) signature BNF_FP_UTIL = sig exception EMPTY_DATATYPE of string type fp_result = {Ts: typ list, bnfs: BNF_Def.bnf list, pre_bnfs: BNF_Def.bnf list, absT_infos: BNF_Comp.absT_info list, ctors: term list, dtors: term list, xtor_un_folds: term list, xtor_co_recs: term list, xtor_co_induct: thm, dtor_ctors: thm list, ctor_dtors: thm list, ctor_injects: thm list, dtor_injects: thm list, xtor_maps: thm list, xtor_map_unique: thm, xtor_setss: thm list list, xtor_rels: thm list, xtor_un_fold_thms: thm list, xtor_co_rec_thms: thm list, xtor_un_fold_unique: thm, xtor_co_rec_unique: thm, xtor_un_fold_o_maps: thm list, xtor_co_rec_o_maps: thm list, xtor_un_fold_transfers: thm list, xtor_co_rec_transfers: thm list, xtor_rel_co_induct: thm, dtor_set_inducts: thm list} val morph_fp_result: morphism -> fp_result -> fp_result val time: Proof.context -> Timer.real_timer -> string -> Timer.real_timer val fixpoint: ('a * 'a -> bool) -> ('a list -> 'a list) -> 'a list -> 'a list val IITN: string val LevN: string val algN: string val behN: string val bisN: string val carTN: string val caseN: string val coN: string val coinductN: string val coinduct_strongN: string val corecN: string val corec_discN: string val corec_disc_iffN: string val ctorN: string val ctor_dtorN: string val ctor_exhaustN: string val ctor_induct2N: string val ctor_inductN: string val ctor_injectN: string val ctor_foldN: string val ctor_fold_o_mapN: string val ctor_fold_transferN: string val ctor_fold_uniqueN: string val ctor_mapN: string val ctor_map_uniqueN: string val ctor_recN: string val ctor_rec_o_mapN: string val ctor_rec_transferN: string val ctor_rec_uniqueN: string val ctor_relN: string val ctor_rel_inductN: string val ctor_set_inclN: string val ctor_set_set_inclN: string val dtorN: string val dtor_coinductN: string val dtor_corecN: string val dtor_corec_o_mapN: string val dtor_corec_transferN: string val dtor_corec_uniqueN: string val dtor_ctorN: string val dtor_exhaustN: string val dtor_injectN: string val dtor_mapN: string val dtor_map_coinductN: string val dtor_map_coinduct_strongN: string val dtor_map_uniqueN: string val dtor_relN: string val dtor_rel_coinductN: string val dtor_set_inclN: string val dtor_set_set_inclN: string val dtor_coinduct_strongN: string val dtor_unfoldN: string val dtor_unfold_o_mapN: string val dtor_unfold_transferN: string val dtor_unfold_uniqueN: string val exhaustN: string val colN: string val inductN: string val injectN: string val isNodeN: string val lsbisN: string val mapN: string val map_uniqueN: string val min_algN: string val morN: string val nchotomyN: string val recN: string val rel_casesN: string val rel_coinductN: string val rel_inductN: string val rel_injectN: string val rel_introsN: string val rel_distinctN: string val rel_selN: string val rvN: string val corec_selN: string val set_inclN: string val set_set_inclN: string val setN: string val simpsN: string val strTN: string val str_initN: string val sum_bdN: string val sum_bdTN: string val uniqueN: string (* TODO: Don't index set facts. Isabelle packages traditionally generate uniform names. *) val mk_ctor_setN: int -> string val mk_dtor_setN: int -> string val mk_dtor_set_inductN: int -> string val mk_set_inductN: int -> string val co_prefix: BNF_Util.fp_kind -> string val split_conj_thm: thm -> thm list val split_conj_prems: int -> thm -> thm val mk_sumTN: typ list -> typ val mk_sumTN_balanced: typ list -> typ val mk_tupleT_balanced: typ list -> typ val mk_sumprodT_balanced: typ list list -> typ val mk_proj: typ -> int -> int -> term val mk_convol: term * term -> term val mk_rel_prod: term -> term -> term val mk_rel_sum: term -> term -> term val Inl_const: typ -> typ -> term val Inr_const: typ -> typ -> term val mk_tuple_balanced: term list -> term val mk_tuple1_balanced: typ list -> term list -> term val abs_curried_balanced: typ list -> term -> term val mk_tupled_fun: term -> term -> term list -> term val mk_case_sum: term * term -> term val mk_case_sumN: term list -> term val mk_case_absumprod: typ -> term -> term list -> term list list -> term list list -> term val mk_Inl: typ -> term -> term val mk_Inr: typ -> term -> term val mk_sumprod_balanced: typ -> int -> int -> term list -> term val mk_absumprod: typ -> term -> int -> int -> term list -> term val dest_sumT: typ -> typ * typ val dest_sumTN_balanced: int -> typ -> typ list val dest_tupleT_balanced: int -> typ -> typ list val dest_absumprodT: typ -> typ -> int -> int list -> typ -> typ list list val If_const: typ -> term val mk_Field: term -> term val mk_If: term -> term -> term -> term val mk_absumprodE: thm -> int list -> thm val mk_sum_caseN: int -> int -> thm val mk_sum_caseN_balanced: int -> int -> thm val mk_sum_Cinfinite: thm list -> thm val mk_sum_card_order: thm list -> thm val force_typ: Proof.context -> typ -> term -> term val mk_xtor_rel_co_induct_thm: BNF_Util.fp_kind -> term list -> term list -> term list -> term list -> term list -> term list -> term list -> term list -> ({prems: thm list, context: Proof.context} -> tactic) -> Proof.context -> thm val mk_xtor_co_iter_transfer_thms: BNF_Util.fp_kind -> term list -> term list -> term list -> term list -> term list -> term list -> term list -> ({prems: thm list, context: Proof.context} -> tactic) -> Proof.context -> thm list val mk_xtor_co_iter_o_map_thms: BNF_Util.fp_kind -> bool -> int -> thm -> thm list -> thm list -> thm list -> thm list -> thm list val derive_xtor_co_recs: BNF_Util.fp_kind -> binding list -> (typ list -> typ list) -> (typ list list * typ list) -> BNF_Def.bnf list -> term list -> term list -> thm -> thm list -> thm list -> thm list -> thm list -> (BNF_Comp.absT_info * BNF_Comp.absT_info) option list -> local_theory -> (term list * (thm list * thm * thm list * thm list)) * local_theory val raw_qualify: (binding -> binding) -> binding -> binding -> binding val fixpoint_bnf: bool -> (binding -> binding) -> (binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list -> BNF_Comp.absT_info list -> local_theory -> 'a) -> binding list -> (string * sort) list -> (string * sort) list -> (string * sort) list -> typ list -> BNF_Comp.comp_cache -> local_theory -> ((BNF_Def.bnf list * BNF_Comp.absT_info list) * BNF_Comp.comp_cache) * 'a end; structure BNF_FP_Util : BNF_FP_UTIL = struct open Ctr_Sugar open BNF_Comp open BNF_Def open BNF_Util open BNF_FP_Util_Tactics exception EMPTY_DATATYPE of string; type fp_result = {Ts: typ list, bnfs: bnf list, pre_bnfs: BNF_Def.bnf list, absT_infos: BNF_Comp.absT_info list, ctors: term list, dtors: term list, xtor_un_folds: term list, xtor_co_recs: term list, xtor_co_induct: thm, dtor_ctors: thm list, ctor_dtors: thm list, ctor_injects: thm list, dtor_injects: thm list, xtor_maps: thm list, xtor_map_unique: thm, xtor_setss: thm list list, xtor_rels: thm list, xtor_un_fold_thms: thm list, xtor_co_rec_thms: thm list, xtor_un_fold_unique: thm, xtor_co_rec_unique: thm, xtor_un_fold_o_maps: thm list, xtor_co_rec_o_maps: thm list, xtor_un_fold_transfers: thm list, xtor_co_rec_transfers: thm list, xtor_rel_co_induct: thm, dtor_set_inducts: thm list}; fun morph_fp_result phi {Ts, bnfs, pre_bnfs, absT_infos, ctors, dtors, xtor_un_folds, xtor_co_recs, xtor_co_induct, dtor_ctors, ctor_dtors, ctor_injects, dtor_injects, xtor_maps, xtor_map_unique, xtor_setss, xtor_rels, xtor_un_fold_thms, xtor_co_rec_thms, xtor_un_fold_unique, xtor_co_rec_unique, xtor_un_fold_o_maps, xtor_co_rec_o_maps, xtor_un_fold_transfers, xtor_co_rec_transfers, xtor_rel_co_induct, dtor_set_inducts} = {Ts = map (Morphism.typ phi) Ts, bnfs = map (morph_bnf phi) bnfs, pre_bnfs = map (morph_bnf phi) pre_bnfs, absT_infos = map (morph_absT_info phi) absT_infos, ctors = map (Morphism.term phi) ctors, dtors = map (Morphism.term phi) dtors, xtor_un_folds = map (Morphism.term phi) xtor_un_folds, xtor_co_recs = map (Morphism.term phi) xtor_co_recs, xtor_co_induct = Morphism.thm phi xtor_co_induct, dtor_ctors = map (Morphism.thm phi) dtor_ctors, ctor_dtors = map (Morphism.thm phi) ctor_dtors, ctor_injects = map (Morphism.thm phi) ctor_injects, dtor_injects = map (Morphism.thm phi) dtor_injects, xtor_maps = map (Morphism.thm phi) xtor_maps, xtor_map_unique = Morphism.thm phi xtor_map_unique, xtor_setss = map (map (Morphism.thm phi)) xtor_setss, xtor_rels = map (Morphism.thm phi) xtor_rels, xtor_un_fold_thms = map (Morphism.thm phi) xtor_un_fold_thms, xtor_co_rec_thms = map (Morphism.thm phi) xtor_co_rec_thms, xtor_un_fold_unique = Morphism.thm phi xtor_un_fold_unique, xtor_co_rec_unique = Morphism.thm phi xtor_co_rec_unique, xtor_un_fold_o_maps = map (Morphism.thm phi) xtor_un_fold_o_maps, xtor_co_rec_o_maps = map (Morphism.thm phi) xtor_co_rec_o_maps, xtor_un_fold_transfers = map (Morphism.thm phi) xtor_un_fold_transfers, xtor_co_rec_transfers = map (Morphism.thm phi) xtor_co_rec_transfers, xtor_rel_co_induct = Morphism.thm phi xtor_rel_co_induct, dtor_set_inducts = map (Morphism.thm phi) dtor_set_inducts}; fun time ctxt timer msg = (if Config.get ctxt bnf_timing then warning (msg ^ ": " ^ string_of_int (Time.toMilliseconds (Timer.checkRealTimer timer)) ^ " ms") else (); Timer.startRealTimer ()); val preN = "pre_" val rawN = "raw_" val coN = "co" val unN = "un" val algN = "alg" val IITN = "IITN" val foldN = "fold" val unfoldN = unN ^ foldN val uniqueN = "unique" val transferN = "transfer" val simpsN = "simps" val ctorN = "ctor" val dtorN = "dtor" val ctor_foldN = ctorN ^ "_" ^ foldN val dtor_unfoldN = dtorN ^ "_" ^ unfoldN val ctor_fold_uniqueN = ctor_foldN ^ "_" ^ uniqueN val ctor_fold_o_mapN = ctor_foldN ^ "_o_" ^ mapN val dtor_unfold_uniqueN = dtor_unfoldN ^ "_" ^ uniqueN val dtor_unfold_o_mapN = dtor_unfoldN ^ "_o_" ^ mapN val ctor_fold_transferN = ctor_foldN ^ "_" ^ transferN val dtor_unfold_transferN = dtor_unfoldN ^ "_" ^ transferN val ctor_mapN = ctorN ^ "_" ^ mapN val dtor_mapN = dtorN ^ "_" ^ mapN val map_uniqueN = mapN ^ "_" ^ uniqueN val ctor_map_uniqueN = ctorN ^ "_" ^ map_uniqueN val dtor_map_uniqueN = dtorN ^ "_" ^ map_uniqueN val min_algN = "min_alg" val morN = "mor" val bisN = "bis" val lsbisN = "lsbis" val sum_bdTN = "sbdT" val sum_bdN = "sbd" val carTN = "carT" val strTN = "strT" val isNodeN = "isNode" val LevN = "Lev" val rvN = "recover" val behN = "beh" val setN = "set" val mk_ctor_setN = prefix (ctorN ^ "_") o mk_setN val mk_dtor_setN = prefix (dtorN ^ "_") o mk_setN fun mk_set_inductN i = mk_setN i ^ "_induct" val mk_dtor_set_inductN = prefix (dtorN ^ "_") o mk_set_inductN val str_initN = "str_init" val recN = "rec" val corecN = coN ^ recN val ctor_recN = ctorN ^ "_" ^ recN val ctor_rec_o_mapN = ctor_recN ^ "_o_" ^ mapN val ctor_rec_transferN = ctor_recN ^ "_" ^ transferN val ctor_rec_uniqueN = ctor_recN ^ "_" ^ uniqueN val dtor_corecN = dtorN ^ "_" ^ corecN val dtor_corec_o_mapN = dtor_corecN ^ "_o_" ^ mapN val dtor_corec_transferN = dtor_corecN ^ "_" ^ transferN val dtor_corec_uniqueN = dtor_corecN ^ "_" ^ uniqueN val ctor_dtorN = ctorN ^ "_" ^ dtorN val dtor_ctorN = dtorN ^ "_" ^ ctorN val nchotomyN = "nchotomy" val injectN = "inject" val exhaustN = "exhaust" val ctor_injectN = ctorN ^ "_" ^ injectN val ctor_exhaustN = ctorN ^ "_" ^ exhaustN val dtor_injectN = dtorN ^ "_" ^ injectN val dtor_exhaustN = dtorN ^ "_" ^ exhaustN val ctor_relN = ctorN ^ "_" ^ relN val dtor_relN = dtorN ^ "_" ^ relN val inductN = "induct" val coinductN = coN ^ inductN val ctor_inductN = ctorN ^ "_" ^ inductN val ctor_induct2N = ctor_inductN ^ "2" val dtor_map_coinductN = dtor_mapN ^ "_" ^ coinductN val dtor_coinductN = dtorN ^ "_" ^ coinductN val coinduct_strongN = coinductN ^ "_strong" val dtor_map_coinduct_strongN = dtor_mapN ^ "_" ^ coinduct_strongN val dtor_coinduct_strongN = dtorN ^ "_" ^ coinduct_strongN val colN = "col" val set_inclN = "set_incl" val ctor_set_inclN = ctorN ^ "_" ^ set_inclN val dtor_set_inclN = dtorN ^ "_" ^ set_inclN val set_set_inclN = "set_set_incl" val ctor_set_set_inclN = ctorN ^ "_" ^ set_set_inclN val dtor_set_set_inclN = dtorN ^ "_" ^ set_set_inclN val caseN = "case" val discN = "disc" val corec_discN = corecN ^ "_" ^ discN val iffN = "_iff" val corec_disc_iffN = corec_discN ^ iffN val distinctN = "distinct" val rel_distinctN = relN ^ "_" ^ distinctN val injectN = "inject" val rel_casesN = relN ^ "_cases" val rel_injectN = relN ^ "_" ^ injectN val rel_introsN = relN ^ "_intros" val rel_coinductN = relN ^ "_" ^ coinductN val rel_selN = relN ^ "_sel" val dtor_rel_coinductN = dtorN ^ "_" ^ rel_coinductN val rel_inductN = relN ^ "_" ^ inductN val ctor_rel_inductN = ctorN ^ "_" ^ rel_inductN val selN = "sel" val corec_selN = corecN ^ "_" ^ selN fun co_prefix fp = case_fp fp "" "co"; fun dest_sumT (Type (\<^type_name>\sum\, [T, T'])) = (T, T'); val dest_sumTN_balanced = Balanced_Tree.dest dest_sumT; fun dest_tupleT_balanced 0 \<^typ>\unit\ = [] | dest_tupleT_balanced n T = Balanced_Tree.dest HOLogic.dest_prodT n T; fun dest_absumprodT absT repT n ms = map2 dest_tupleT_balanced ms o dest_sumTN_balanced n o mk_repT absT repT; val mk_sumTN = Library.foldr1 mk_sumT; val mk_sumTN_balanced = Balanced_Tree.make mk_sumT; fun mk_tupleT_balanced [] = HOLogic.unitT | mk_tupleT_balanced Ts = Balanced_Tree.make HOLogic.mk_prodT Ts; val mk_sumprodT_balanced = mk_sumTN_balanced o map mk_tupleT_balanced; fun mk_proj T n k = let val (binders, _) = strip_typeN n T in fold_rev (fn T => fn t => Abs (Name.uu, T, t)) binders (Bound (n - k - 1)) end; fun mk_convol (f, g) = let val (fU, fTU) = `range_type (fastype_of f); val ((gT, gU), gTU) = `dest_funT (fastype_of g); val convolT = fTU --> gTU --> gT --> HOLogic.mk_prodT (fU, gU); in Const (\<^const_name>\convol\, convolT) $ f $ g end; fun mk_rel_prod R S = let val ((A1, A2), RT) = `dest_pred2T (fastype_of R); val ((B1, B2), ST) = `dest_pred2T (fastype_of S); val rel_prodT = RT --> ST --> mk_pred2T (HOLogic.mk_prodT (A1, B1)) (HOLogic.mk_prodT (A2, B2)); in Const (\<^const_name>\rel_prod\, rel_prodT) $ R $ S end; fun mk_rel_sum R S = let val ((A1, A2), RT) = `dest_pred2T (fastype_of R); val ((B1, B2), ST) = `dest_pred2T (fastype_of S); val rel_sumT = RT --> ST --> mk_pred2T (mk_sumT (A1, B1)) (mk_sumT (A2, B2)); in Const (\<^const_name>\rel_sum\, rel_sumT) $ R $ S end; fun Inl_const LT RT = Const (\<^const_name>\Inl\, LT --> mk_sumT (LT, RT)); fun mk_Inl RT t = Inl_const (fastype_of t) RT $ t; fun Inr_const LT RT = Const (\<^const_name>\Inr\, RT --> mk_sumT (LT, RT)); fun mk_Inr LT t = Inr_const LT (fastype_of t) $ t; fun mk_prod1 bound_Ts (t, u) = HOLogic.pair_const (fastype_of1 (bound_Ts, t)) (fastype_of1 (bound_Ts, u)) $ t $ u; fun mk_tuple1_balanced _ [] = HOLogic.unit | mk_tuple1_balanced bound_Ts ts = Balanced_Tree.make (mk_prod1 bound_Ts) ts; val mk_tuple_balanced = mk_tuple1_balanced []; fun abs_curried_balanced Ts t = t $ mk_tuple1_balanced (List.rev Ts) (map Bound (length Ts - 1 downto 0)) |> fold_rev (Term.abs o pair Name.uu) Ts; fun mk_sumprod_balanced T n k ts = Sum_Tree.mk_inj T n k (mk_tuple_balanced ts); fun mk_absumprod absT abs0 n k ts = let val abs = mk_abs absT abs0; in abs $ mk_sumprod_balanced (domain_type (fastype_of abs)) n k ts end; fun mk_case_sum (f, g) = let val (fT, T') = dest_funT (fastype_of f); val (gT, _) = dest_funT (fastype_of g); in Sum_Tree.mk_sumcase fT gT T' f g end; val mk_case_sumN = Library.foldr1 mk_case_sum; val mk_case_sumN_balanced = Balanced_Tree.make mk_case_sum; fun mk_tupled_fun f x xs = if xs = [x] then f else HOLogic.tupled_lambda x (Term.list_comb (f, xs)); fun mk_case_absumprod absT rep fs xss xss' = HOLogic.mk_comp (mk_case_sumN_balanced (@{map 3} mk_tupled_fun fs (map mk_tuple_balanced xss) xss'), mk_rep absT rep); fun If_const T = Const (\<^const_name>\If\, HOLogic.boolT --> T --> T --> T); fun mk_If p t f = let val T = fastype_of t in If_const T $ p $ t $ f end; fun mk_Field r = let val T = fst (dest_relT (fastype_of r)); in Const (\<^const_name>\Field\, mk_relT (T, T) --> HOLogic.mk_setT T) $ r end; (*dangerous; use with monotonic, converging functions only!*) fun fixpoint eq f X = if subset eq (f X, X) then X else fixpoint eq f (f X); (* stolen from "~~/src/HOL/Tools/Datatype/datatype_aux.ML" *) fun split_conj_thm th = ((th RS conjunct1) :: split_conj_thm (th RS conjunct2)) handle THM _ => [th]; fun split_conj_prems limit th = let fun split n i th = if i = n then th else split n (i + 1) (conjI RSN (i, th)) handle THM _ => th; in split limit 1 th end; fun mk_obj_sumEN_balanced n = Balanced_Tree.make (fn (thm1, thm2) => thm1 RSN (1, thm2 RSN (2, @{thm obj_sumE_f}))) (replicate n asm_rl); fun mk_tupled_allIN_balanced 0 = @{thm unit_all_impI} | mk_tupled_allIN_balanced n = let val (tfrees, _) = BNF_Util.mk_TFrees n \<^context>; val T = mk_tupleT_balanced tfrees; in @{thm asm_rl[of "\x. P x \ Q x" for P Q]} |> Thm.instantiate' [SOME (Thm.ctyp_of \<^context> T)] [] |> Raw_Simplifier.rewrite_goals_rule \<^context> @{thms split_paired_All[THEN eq_reflection]} |> (fn thm => impI RS funpow n (fn th => allI RS th) thm) |> Thm.varifyT_global end; fun mk_absumprodE type_definition ms = let val n = length ms in mk_obj_sumEN_balanced n OF map mk_tupled_allIN_balanced ms RS (type_definition RS @{thm type_copy_obj_one_point_absE}) end; fun mk_sum_caseN 1 1 = refl | mk_sum_caseN _ 1 = @{thm sum.case(1)} | mk_sum_caseN 2 2 = @{thm sum.case(2)} | mk_sum_caseN n k = trans OF [@{thm case_sum_step(2)}, mk_sum_caseN (n - 1) (k - 1)]; fun mk_sum_step base step thm = if Thm.eq_thm_prop (thm, refl) then base else trans OF [step, thm]; fun mk_sum_caseN_balanced 1 1 = refl | mk_sum_caseN_balanced n k = Balanced_Tree.access {left = mk_sum_step @{thm sum.case(1)} @{thm case_sum_step(1)}, right = mk_sum_step @{thm sum.case(2)} @{thm case_sum_step(2)}, init = refl} n k; fun mk_sum_Cinfinite [thm] = thm | mk_sum_Cinfinite (thm :: thms) = @{thm Cinfinite_csum_weak} OF [thm, mk_sum_Cinfinite thms]; fun mk_sum_card_order [thm] = thm | mk_sum_card_order (thm :: thms) = @{thm card_order_csum} OF [thm, mk_sum_card_order thms]; fun mk_xtor_rel_co_induct_thm fp pre_rels pre_phis rels phis xs ys xtors xtor's tac lthy = let val pre_relphis = map (fn rel => Term.list_comb (rel, phis @ pre_phis)) pre_rels; val relphis = map (fn rel => Term.list_comb (rel, phis)) rels; fun mk_xtor fp' xtor x = if fp = fp' then xtor $ x else x; val dtor = mk_xtor Greatest_FP; val ctor = mk_xtor Least_FP; fun flip f x y = if fp = Greatest_FP then f y x else f x y; fun mk_prem pre_relphi phi x y xtor xtor' = HOLogic.mk_Trueprop (list_all_free [x, y] (flip (curry HOLogic.mk_imp) (pre_relphi $ (dtor xtor x) $ (dtor xtor' y)) (phi $ (ctor xtor x) $ (ctor xtor' y)))); val prems = @{map 6} mk_prem pre_relphis pre_phis xs ys xtors xtor's; val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (flip mk_leq) relphis pre_phis)); in Goal.prove_sorry lthy (map (fst o dest_Free) (phis @ pre_phis)) prems concl tac |> Thm.close_derivation \<^here> |> (fn thm => thm OF (replicate (length pre_rels) @{thm allI[OF allI[OF impI]]})) end; fun mk_xtor_co_iter_transfer_thms fp pre_rels pre_iphis pre_ophis rels phis un_folds un_folds' tac lthy = let val pre_relphis = map (fn rel => Term.list_comb (rel, phis @ pre_iphis)) pre_rels; val relphis = map (fn rel => Term.list_comb (rel, phis)) rels; fun flip f x y = if fp = Greatest_FP then f y x else f x y; val arg_rels = map2 (flip mk_rel_fun) pre_relphis pre_ophis; fun mk_transfer relphi pre_phi un_fold un_fold' = fold_rev mk_rel_fun arg_rels (flip mk_rel_fun relphi pre_phi) $ un_fold $ un_fold'; val transfers = @{map 4} mk_transfer relphis pre_ophis un_folds un_folds'; val goal = fold_rev Logic.all (phis @ pre_ophis) (HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj transfers)); in Goal.prove_sorry lthy [] [] goal tac |> Thm.close_derivation \<^here> |> split_conj_thm end; fun mk_xtor_co_iter_o_map_thms fp is_rec m un_fold_unique xtor_maps xtor_un_folds sym_map_comps map_cong0s = let val n = length sym_map_comps; val rewrite_comp_comp2 = case_fp fp @{thm rewriteR_comp_comp2} @{thm rewriteL_comp_comp2}; val rewrite_comp_comp = case_fp fp @{thm rewriteR_comp_comp} @{thm rewriteL_comp_comp}; val map_cong_passive_args1 = replicate m (case_fp fp @{thm id_comp} @{thm comp_id} RS fun_cong); val map_cong_active_args1 = replicate n (if is_rec then case_fp fp @{thm convol_o} @{thm o_case_sum} RS fun_cong else refl); val map_cong_passive_args2 = replicate m (case_fp fp @{thm comp_id} @{thm id_comp} RS fun_cong); val map_cong_active_args2 = replicate n (if is_rec then case_fp fp @{thm map_prod_o_convol_id} @{thm case_sum_o_map_sum_id} else case_fp fp @{thm id_comp} @{thm comp_id} RS fun_cong); fun mk_map_congs passive active = map (fn thm => thm OF (passive @ active) RS @{thm ext}) map_cong0s; val map_cong1s = mk_map_congs map_cong_passive_args1 map_cong_active_args1; val map_cong2s = mk_map_congs map_cong_passive_args2 map_cong_active_args2; fun mk_rewrites map_congs = map2 (fn sym_map_comp => fn map_cong => mk_trans sym_map_comp map_cong RS rewrite_comp_comp) sym_map_comps map_congs; val rewrite1s = mk_rewrites map_cong1s; val rewrite2s = mk_rewrites map_cong2s; val unique_prems = @{map 4} (fn xtor_map => fn un_fold => fn rewrite1 => fn rewrite2 => mk_trans (rewrite_comp_comp2 OF [xtor_map, un_fold]) (mk_trans rewrite1 (mk_sym rewrite2))) xtor_maps xtor_un_folds rewrite1s rewrite2s; in split_conj_thm (un_fold_unique OF map (case_fp fp I mk_sym) unique_prems) end; fun force_typ ctxt T = Term.map_types Type_Infer.paramify_vars #> Type.constraint T #> Syntax.check_term ctxt #> singleton (Variable.polymorphic ctxt); fun absT_info_encodeT thy (SOME (src : absT_info, dst : absT_info)) src_absT = let val src_repT = mk_repT (#absT src) (#repT src) src_absT; val dst_absT = mk_absT thy (#repT dst) (#absT dst) src_repT; in dst_absT end | absT_info_encodeT _ NONE T = T; fun absT_info_decodeT thy = absT_info_encodeT thy o Option.map swap; fun absT_info_encode thy fp (opt as SOME (src : absT_info, dst : absT_info)) t = let val co_alg_funT = case_fp fp domain_type range_type; fun co_swap pair = case_fp fp I swap pair; val mk_co_comp = curry (HOLogic.mk_comp o co_swap); val mk_co_abs = case_fp fp mk_abs mk_rep; val mk_co_rep = case_fp fp mk_rep mk_abs; val co_abs = case_fp fp #abs #rep; val co_rep = case_fp fp #rep #abs; val src_absT = co_alg_funT (fastype_of t); val dst_absT = absT_info_encodeT thy opt src_absT; val co_src_abs = mk_co_abs src_absT (co_abs src); val co_dst_rep = mk_co_rep dst_absT (co_rep dst); in mk_co_comp (mk_co_comp t co_src_abs) co_dst_rep end | absT_info_encode _ _ NONE t = t; fun absT_info_decode thy fp = absT_info_encode thy fp o Option.map swap; fun mk_xtor_un_fold_xtor_thms ctxt fp un_folds xtors xtor_un_fold_unique map_id0s absT_info_opts = let val thy = Proof_Context.theory_of ctxt; fun mk_goal un_fold = let val rhs = list_comb (un_fold, @{map 2} (absT_info_encode thy fp) absT_info_opts xtors); val T = range_type (fastype_of rhs); in HOLogic.mk_eq (HOLogic.id_const T, rhs) end; val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map mk_goal un_folds)); fun mk_inverses NONE = [] | mk_inverses (SOME (src, dst)) = [#type_definition dst RS @{thm type_definition.Abs_inverse[OF _ UNIV_I]}, #type_definition src RS @{thm type_definition.Rep_inverse}]; val inverses = maps mk_inverses absT_info_opts; in Goal.prove_sorry ctxt [] [] goal (fn {context = ctxt, prems = _} => mk_xtor_un_fold_xtor_tac ctxt xtor_un_fold_unique map_id0s inverses) |> split_conj_thm |> map mk_sym end; fun derive_xtor_co_recs fp bs mk_Ts (Dss, resDs) pre_bnfs xtors0 un_folds0 xtor_un_fold_unique xtor_un_folds xtor_un_fold_transfers xtor_maps xtor_rels absT_info_opts lthy = let val thy = Proof_Context.theory_of lthy; fun co_swap pair = case_fp fp I swap pair; val mk_co_comp = curry (HOLogic.mk_comp o co_swap); fun mk_co_algT T U = case_fp fp (T --> U) (U --> T); val co_alg_funT = case_fp fp domain_type range_type; val mk_co_product = curry (case_fp fp mk_convol mk_case_sum); val co_proj1_const = case_fp fp fst_const (uncurry Inl_const o dest_sumT) o co_alg_funT; val co_proj2_const = case_fp fp snd_const (uncurry Inr_const o dest_sumT) o co_alg_funT; val mk_co_productT = curry (case_fp fp HOLogic.mk_prodT mk_sumT); val rewrite_comp_comp = case_fp fp @{thm rewriteL_comp_comp} @{thm rewriteR_comp_comp}; val n = length pre_bnfs; val live = live_of_bnf (hd pre_bnfs); val m = live - n; val ks = 1 upto n; val map_id0s = map map_id0_of_bnf pre_bnfs; val map_comps = map map_comp_of_bnf pre_bnfs; val map_cong0s = map map_cong0_of_bnf pre_bnfs; val map_transfers = map map_transfer_of_bnf pre_bnfs; val sym_map_comp0s = map (mk_sym o map_comp0_of_bnf) pre_bnfs; val deads = fold (union (op =)) Dss resDs; val ((((As, Bs), Xs), Ys), names_lthy) = lthy |> fold Variable.declare_typ deads |> mk_TFrees m ||>> mk_TFrees m ||>> mk_TFrees n ||>> mk_TFrees n; val XFTs = @{map 2} (fn Ds => mk_T_of_bnf Ds (As @ Xs)) Dss pre_bnfs; val co_algXFTs = @{map 2} mk_co_algT XFTs Xs; val Ts = mk_Ts As; val un_foldTs = @{map 2} (fn T => fn X => co_algXFTs ---> mk_co_algT T X) Ts Xs; val un_folds = @{map 2} (force_typ names_lthy) un_foldTs un_folds0; val ABs = As ~~ Bs; val XYs = Xs ~~ Ys; val Us = map (typ_subst_atomic ABs) Ts; val TFTs = @{map 2} (fn Ds => mk_T_of_bnf Ds (As @ Ts)) Dss pre_bnfs; val TFTs' = @{map 2} (absT_info_decodeT thy) absT_info_opts TFTs; val xtors = @{map 3} (force_typ names_lthy oo mk_co_algT) TFTs' Ts xtors0; val ids = map HOLogic.id_const As; val co_rec_Xs = @{map 2} mk_co_productT Ts Xs; val co_rec_Ys = @{map 2} mk_co_productT Us Ys; val co_rec_algXs = @{map 2} mk_co_algT co_rec_Xs Xs; val co_proj1s = map co_proj1_const co_rec_algXs; val co_rec_maps = @{map 2} (fn Ds => mk_map_of_bnf Ds (As @ case_fp fp co_rec_Xs Ts) (As @ case_fp fp Ts co_rec_Xs)) Dss pre_bnfs; val co_rec_Ts = @{map 2} (fn Ds => mk_T_of_bnf Ds (As @ co_rec_Xs)) Dss pre_bnfs val co_rec_argTs = @{map 2} mk_co_algT co_rec_Ts Xs; val co_rec_resTs = @{map 2} mk_co_algT Ts Xs; val (((co_rec_ss, fs), xs), names_lthy) = names_lthy |> mk_Frees "s" co_rec_argTs ||>> mk_Frees "f" co_rec_resTs ||>> mk_Frees "x" (case_fp fp TFTs' Xs); val co_rec_strs = @{map 4} (fn xtor => fn s => fn mapx => fn absT_info_opt => mk_co_product (mk_co_comp (absT_info_encode thy fp absT_info_opt xtor) (list_comb (mapx, ids @ co_proj1s))) s) xtors co_rec_ss co_rec_maps absT_info_opts; val theta = Xs ~~ co_rec_Xs; val co_rec_un_folds = map (subst_atomic_types theta) un_folds; val co_rec_spec0s = map (fn un_fold => list_comb (un_fold, co_rec_strs)) co_rec_un_folds; val co_rec_ids = @{map 2} (mk_co_comp o co_proj1_const) co_rec_algXs co_rec_spec0s; val co_rec_specs = @{map 2} (mk_co_comp o co_proj2_const) co_rec_algXs co_rec_spec0s; val co_recN = case_fp fp ctor_recN dtor_corecN; fun co_rec_bind i = nth bs (i - 1) |> Binding.prefix_name (co_recN ^ "_"); val co_rec_def_bind = rpair [] o Binding.concealed o Thm.def_binding o co_rec_bind; fun co_rec_spec i = fold_rev (Term.absfree o Term.dest_Free) co_rec_ss (nth co_rec_specs (i - 1)); val ((co_rec_frees, (_, co_rec_def_frees)), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> fold_map (fn i => Local_Theory.define ((co_rec_bind i, NoSyn), (co_rec_def_bind i, co_rec_spec i))) ks |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val co_rec_names = map (fst o dest_Const o Morphism.term phi) co_rec_frees; val co_recs = @{map 2} (fn name => fn resT => Const (name, co_rec_argTs ---> resT)) co_rec_names co_rec_resTs; val co_rec_defs = map (fn def => mk_unabs_def n (HOLogic.mk_obj_eq (Morphism.thm phi def))) co_rec_def_frees; val xtor_un_fold_xtor_thms = mk_xtor_un_fold_xtor_thms lthy fp (map (Term.subst_atomic_types (Xs ~~ Ts)) un_folds) xtors xtor_un_fold_unique map_id0s absT_info_opts; val co_rec_id_thms = let val goal = @{map 2} (fn T => fn t => HOLogic.mk_eq (t, HOLogic.id_const T)) Ts co_rec_ids |> Library.foldr1 HOLogic.mk_conj |> HOLogic.mk_Trueprop; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_xtor_co_rec_id_tac ctxt xtor_un_fold_xtor_thms xtor_un_fold_unique xtor_un_folds map_comps) |> Thm.close_derivation \<^here> |> split_conj_thm end; val co_rec_app_ss = map (fn co_rec => list_comb (co_rec, co_rec_ss)) co_recs; val co_products = @{map 2} (fn T => mk_co_product (HOLogic.id_const T)) Ts co_rec_app_ss; val co_rec_maps_rev = @{map 2} (fn Ds => mk_map_of_bnf Ds (As @ case_fp fp Ts co_rec_Xs) (As @ case_fp fp co_rec_Xs Ts)) Dss pre_bnfs; fun mk_co_app f g x = case_fp fp (f $ (g $ x)) (g $ (f $ x)); val co_rec_expand_thms = map (fn thm => thm RS case_fp fp @{thm convol_expand_snd} @{thm case_sum_expand_Inr_pointfree}) co_rec_id_thms; val xtor_co_rec_thms = let fun mk_goal co_rec s mapx xtor x absT_info_opt = let val lhs = mk_co_app co_rec xtor x; val rhs = mk_co_app s (list_comb (mapx, ids @ co_products) |> absT_info_decode thy fp absT_info_opt) x; in mk_Trueprop_eq (lhs, rhs) end; val goals = @{map 6} mk_goal co_rec_app_ss co_rec_ss co_rec_maps_rev xtors xs absT_info_opts; in map2 (fn goal => fn un_fold => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_xtor_co_rec_tac ctxt un_fold co_rec_defs co_rec_expand_thms)) |> Thm.close_derivation \<^here>) goals xtor_un_folds end; val co_product_fs = @{map 2} (fn T => mk_co_product (HOLogic.id_const T)) Ts fs; val co_rec_expand'_thms = map (fn thm => thm RS case_fp fp @{thm convol_expand_snd'} @{thm case_sum_expand_Inr'}) co_rec_id_thms; val xtor_co_rec_unique_thm = let fun mk_prem f s mapx xtor absT_info_opt = let val lhs = mk_co_comp f xtor; val rhs = mk_co_comp s (list_comb (mapx, ids @ co_product_fs)) |> absT_info_decode thy fp absT_info_opt; in mk_Trueprop_eq (co_swap (lhs, rhs)) end; val prems = @{map 5} mk_prem fs co_rec_ss co_rec_maps_rev xtors absT_info_opts; val concl = @{map 2} (curry HOLogic.mk_eq) fs co_rec_app_ss |> Library.foldr1 HOLogic.mk_conj |> HOLogic.mk_Trueprop; val goal = Logic.list_implies (prems, concl); val vars = Variable.add_free_names lthy goal []; fun mk_inverses NONE = [] | mk_inverses (SOME (src, dst)) = [#type_definition dst RS @{thm type_copy_Rep_o_Abs} RS rewrite_comp_comp, #type_definition src RS @{thm type_copy_Abs_o_Rep}]; val inverses = maps mk_inverses absT_info_opts; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_xtor_co_rec_unique_tac ctxt fp co_rec_defs co_rec_expand'_thms xtor_un_fold_unique map_id0s sym_map_comp0s inverses) |> Thm.close_derivation \<^here> end; val xtor_co_rec_o_map_thms = if forall is_none absT_info_opts then mk_xtor_co_iter_o_map_thms fp true m xtor_co_rec_unique_thm (map (mk_pointfree2 lthy) xtor_maps) (map (mk_pointfree2 lthy) xtor_co_rec_thms) sym_map_comp0s map_cong0s else replicate n refl (* FIXME *); val ABphiTs = @{map 2} mk_pred2T As Bs; val XYphiTs = @{map 2} mk_pred2T Xs Ys; val ((ABphis, XYphis), names_lthy) = names_lthy |> mk_Frees "R" ABphiTs ||>> mk_Frees "S" XYphiTs; val xtor_co_rec_transfer_thms = if forall is_none absT_info_opts then let val pre_rels = @{map 2} (fn Ds => mk_rel_of_bnf Ds (As @ co_rec_Xs) (Bs @ co_rec_Ys)) Dss pre_bnfs; val rels = @{map 3} (fn T => fn T' => Thm.prop_of #> HOLogic.dest_Trueprop #> fst o dest_comb #> fst o dest_comb #> funpow n (snd o dest_comb) #> case_fp fp (fst o dest_comb #> snd o dest_comb) (snd o dest_comb) #> head_of #> force_typ names_lthy (ABphiTs ---> mk_pred2T T T')) Ts Us xtor_un_fold_transfers; fun tac {context = ctxt, prems = _} = mk_xtor_co_rec_transfer_tac ctxt fp n m co_rec_defs xtor_un_fold_transfers map_transfers xtor_rels; val mk_rel_co_product = case_fp fp mk_rel_prod mk_rel_sum; val rec_phis = map2 (fn rel => mk_rel_co_product (Term.list_comb (rel, ABphis))) rels XYphis; in mk_xtor_co_iter_transfer_thms fp pre_rels rec_phis XYphis rels ABphis co_recs (map (subst_atomic_types (ABs @ XYs)) co_recs) tac lthy end else replicate n TrueI (* FIXME *); val notes = [(case_fp fp ctor_recN dtor_corecN, xtor_co_rec_thms), (case_fp fp ctor_rec_uniqueN dtor_corec_uniqueN, split_conj_thm xtor_co_rec_unique_thm), (case_fp fp ctor_rec_o_mapN dtor_corec_o_mapN, xtor_co_rec_o_map_thms), (case_fp fp ctor_rec_transferN dtor_corec_transferN, xtor_co_rec_transfer_thms)] |> map (apsnd (map single)) |> maps (fn (thmN, thmss) => map2 (fn b => fn thms => ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])])) bs thmss); val lthy = lthy |> Config.get lthy bnf_internals ? snd o Local_Theory.notes notes; in ((co_recs, (xtor_co_rec_thms, xtor_co_rec_unique_thm, xtor_co_rec_o_map_thms, xtor_co_rec_transfer_thms)), lthy) end; fun raw_qualify extra_qualify base_b = let val qs = Binding.path_of base_b; val n = Binding.name_of base_b; in Binding.prefix_name rawN #> fold_rev (fn (s, mand) => Binding.qualify mand s) (qs @ [(n, true)]) #> extra_qualify #> Binding.concealed end; fun fixpoint_bnf force_out_of_line extra_qualify construct_fp bs resBs Ds0 Xs rhsXs comp_cache0 lthy = let val time = time lthy; val timer = time (Timer.startRealTimer ()); fun flatten_tyargs Ass = subtract (op =) Xs (filter (fn T => exists (fn Ts => member (op =) Ts T) Ass) resBs) @ Xs; val ((bnfs, (deadss, livess)), (comp_cache_unfold_set, lthy')) = apfst (apsnd split_list o split_list) (@{fold_map 2} (fn b => bnf_of_typ true Smart_Inline (raw_qualify extra_qualify b) flatten_tyargs Xs Ds0) bs rhsXs ((comp_cache0, empty_unfolds), lthy)); fun norm_qualify i = Binding.qualify true (Binding.name_of (nth bs (Int.max (0, i - 1)))) #> extra_qualify #> Binding.concealed; val Ass = map (map dest_TFree) livess; val Ds' = fold (fold Term.add_tfreesT) deadss []; val Ds = union (op =) Ds' Ds0; val missing = resBs |> fold (subtract (op =)) (Ds' :: Ass); val (dead_phantoms, live_phantoms) = List.partition (member (op =) Ds0) missing; val resBs' = resBs |> fold (subtract (op =)) [dead_phantoms, Ds]; val timer = time (timer "Construction of BNFs"); val ((kill_posss, _), (bnfs', ((comp_cache', unfold_set'), lthy''))) = normalize_bnfs norm_qualify Ass Ds (K (resBs' @ Xs)) bnfs (comp_cache_unfold_set, lthy'); val Dss = @{map 3} (fn lives => fn kill_posss => fn deads => deads @ map (nth lives) kill_posss) livess kill_posss deadss; val all_Dss = Dss |> force_out_of_line ? map (fn Ds' => union (op =) Ds' (map TFree Ds0)); fun pre_qualify b = Binding.qualify false (Binding.name_of b) #> extra_qualify #> not (Config.get lthy'' bnf_internals) ? Binding.concealed; val ((pre_bnfs, (deadss, absT_infos)), lthy''') = lthy'' |> @{fold_map 5} (fn b => seal_bnf (pre_qualify b) unfold_set' (Binding.prefix_name preN b)) bs (replicate (length rhsXs) (force_out_of_line orelse not (null live_phantoms))) Dss all_Dss bnfs' |>> split_list |>> apsnd split_list; val timer = time (timer "Normalization & sealing of BNFs"); val res = construct_fp bs resBs (map TFree dead_phantoms, deadss) pre_bnfs absT_infos lthy'''; val timer = time (timer "FP construction in total"); in (((pre_bnfs, absT_infos), comp_cache'), res) end; end; diff --git a/src/HOL/Tools/BNF/bnf_gfp.ML b/src/HOL/Tools/BNF/bnf_gfp.ML --- a/src/HOL/Tools/BNF/bnf_gfp.ML +++ b/src/HOL/Tools/BNF/bnf_gfp.ML @@ -1,2609 +1,2609 @@ (* Title: HOL/Tools/BNF/bnf_gfp.ML Author: Dmitriy Traytel, TU Muenchen Author: Andrei Popescu, TU Muenchen Author: Jasmin Blanchette, TU Muenchen Copyright 2012 Codatatype construction. *) signature BNF_GFP = sig val construct_gfp: mixfix list -> binding list -> binding list -> binding list -> binding list list -> binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list -> BNF_Comp.absT_info list -> local_theory -> BNF_FP_Util.fp_result * local_theory end; structure BNF_GFP : BNF_GFP = struct open BNF_Def open BNF_Util open BNF_Tactics open BNF_Comp open BNF_FP_Util open BNF_FP_Def_Sugar open BNF_GFP_Util open BNF_GFP_Tactics datatype wit_tree = Wit_Leaf of int | Wit_Node of (int * int * int list) * wit_tree list; fun mk_tree_args (I, T) (I', Ts) = (sort_distinct int_ord (I @ I'), T :: Ts); fun finish Iss m seen i (nwit, I) = let val treess = map (fn j => if j < m orelse member (op =) seen j then [([j], Wit_Leaf j)] else map_index (finish Iss m (insert (op =) j seen) j) (nth Iss (j - m)) |> flat |> minimize_wits) I; in map (fn (I, t) => (I, Wit_Node ((i - m, nwit, filter (fn i => i < m) I), t))) (fold_rev (map_product mk_tree_args) treess [([], [])]) |> minimize_wits end; fun tree_to_ctor_wit vars _ _ (Wit_Leaf j) = ([j], nth vars j) | tree_to_ctor_wit vars ctors witss (Wit_Node ((i, nwit, I), subtrees)) = (I, nth ctors i $ (Term.list_comb (snd (nth (nth witss i) nwit), map (snd o tree_to_ctor_wit vars ctors witss) subtrees))); fun tree_to_coind_wits _ (Wit_Leaf _) = [] | tree_to_coind_wits lwitss (Wit_Node ((i, nwit, I), subtrees)) = ((i, I), nth (nth lwitss i) nwit) :: maps (tree_to_coind_wits lwitss) subtrees; (*all BNFs have the same lives*) fun construct_gfp mixfixes map_bs rel_bs pred_bs set_bss0 bs resBs (resDs, Dss) bnfs absT_infos lthy = let val time = time lthy; val timer = time (Timer.startRealTimer ()); val live = live_of_bnf (hd bnfs); val n = length bnfs; (*active*) val ks = 1 upto n; val m = live - n; (*passive, if 0 don't generate a new BNF*) val ls = 1 upto m; val internals = Config.get lthy bnf_internals; val b_names = map Binding.name_of bs; val b_name = mk_common_name b_names; val b = Binding.name b_name; fun mk_internal_of_b name = Binding.prefix_name (name ^ "_") #> Binding.prefix true b_name #> Binding.concealed; fun mk_internal_b name = mk_internal_of_b name b; fun mk_internal_bs name = map (mk_internal_of_b name) bs; val external_bs = map2 (Binding.prefix false) b_names bs |> not internals ? map Binding.concealed; val deads = fold (union (op =)) Dss resDs; val names_lthy = fold Variable.declare_typ deads lthy; val passives = map fst (subtract (op = o apsnd TFree) deads resBs); (* tvars *) val ((((((passiveAs, activeAs), passiveBs), activeBs), passiveCs), activeCs), idxT) = names_lthy |> variant_tfrees passives ||>> mk_TFrees n ||>> variant_tfrees passives ||>> mk_TFrees n ||>> mk_TFrees m ||>> mk_TFrees n ||> fst o mk_TFrees 1 ||> the_single; val allAs = passiveAs @ activeAs; val allBs' = passiveBs @ activeBs; val Ass = replicate n allAs; val allBs = passiveAs @ activeBs; val Bss = replicate n allBs; val allCs = passiveAs @ activeCs; val allCs' = passiveBs @ activeCs; val Css' = replicate n allCs'; (* types *) val dead_poss = map (fn x => if member (op =) deads (TFree x) then SOME (TFree x) else NONE) resBs; fun mk_param NONE passive = (hd passive, tl passive) | mk_param (SOME a) passive = (a, passive); val mk_params = fold_map mk_param dead_poss #> fst; fun mk_FTs Ts = map2 (fn Ds => mk_T_of_bnf Ds Ts) Dss bnfs; val (params, params') = `(map Term.dest_TFree) (mk_params passiveAs); val FTsAs = mk_FTs allAs; val FTsBs = mk_FTs allBs; val FTsCs = mk_FTs allCs; val ATs = map HOLogic.mk_setT passiveAs; val BTs = map HOLogic.mk_setT activeAs; val B'Ts = map HOLogic.mk_setT activeBs; val B''Ts = map HOLogic.mk_setT activeCs; val sTs = map2 (fn T => fn U => T --> U) activeAs FTsAs; val s'Ts = map2 (fn T => fn U => T --> U) activeBs FTsBs; val s''Ts = map2 (fn T => fn U => T --> U) activeCs FTsCs; val fTs = map2 (fn T => fn U => T --> U) activeAs activeBs; val self_fTs = map (fn T => T --> T) activeAs; val gTs = map2 (fn T => fn U => T --> U) activeBs activeCs; val all_gTs = map2 (fn T => fn U => T --> U) allBs allCs'; val RTs = map2 (fn T => fn U => HOLogic.mk_prodT (T, U)) activeAs activeBs; val sRTs = map2 (fn T => fn U => HOLogic.mk_prodT (T, U)) activeAs activeAs; val R'Ts = map2 (fn T => fn U => HOLogic.mk_prodT (T, U)) activeBs activeCs; val setsRTs = map HOLogic.mk_setT sRTs; val setRTs = map HOLogic.mk_setT RTs; val all_sbisT = HOLogic.mk_tupleT setsRTs; val setR'Ts = map HOLogic.mk_setT R'Ts; val FRTs = mk_FTs (passiveAs @ RTs); (* terms *) val mapsAsAs = @{map 4} mk_map_of_bnf Dss Ass Ass bnfs; val mapsAsBs = @{map 4} mk_map_of_bnf Dss Ass Bss bnfs; val mapsBsCs' = @{map 4} mk_map_of_bnf Dss Bss Css' bnfs; val mapsAsCs' = @{map 4} mk_map_of_bnf Dss Ass Css' bnfs; val map_fsts = @{map 4} mk_map_of_bnf Dss (replicate n (passiveAs @ RTs)) Ass bnfs; val map_snds = @{map 4} mk_map_of_bnf Dss (replicate n (passiveAs @ RTs)) Bss bnfs; fun mk_setss Ts = @{map 3} mk_sets_of_bnf (map (replicate live) Dss) (map (replicate live) (replicate n Ts)) bnfs; val setssAs = mk_setss allAs; val setssAs' = transpose setssAs; val bis_setss = mk_setss (passiveAs @ RTs); val relsAsBs = @{map 4} mk_rel_of_bnf Dss Ass Bss bnfs; val bds = @{map 3} mk_bd_of_bnf Dss Ass bnfs; val sum_bd = Library.foldr1 (uncurry mk_csum) bds; val sum_bdT = fst (dest_relT (fastype_of sum_bd)); val (sum_bdT_params, sum_bdT_params') = `(map TFree) (Term.add_tfreesT sum_bdT []); val ((((((((((zs, zs'), Bs), ss), fs), self_fs), all_gs), xFs), yFs), yFs_copy), _) = lthy |> mk_Frees' "b" activeAs ||>> mk_Frees "B" BTs ||>> mk_Frees "s" sTs ||>> mk_Frees "f" fTs ||>> mk_Frees "f" self_fTs ||>> mk_Frees "g" all_gTs ||>> mk_Frees "x" FTsAs ||>> mk_Frees "y" FTsBs ||>> mk_Frees "y" FTsBs; val passive_UNIVs = map HOLogic.mk_UNIV passiveAs; val passive_eqs = map HOLogic.eq_const passiveAs; val active_UNIVs = map HOLogic.mk_UNIV activeAs; val passive_ids = map HOLogic.id_const passiveAs; val active_ids = map HOLogic.id_const activeAs; val fsts = map fst_const RTs; val snds = map snd_const RTs; (* thms *) val bd_card_orders = map bd_card_order_of_bnf bnfs; val bd_card_order = hd bd_card_orders val bd_Card_orders = map bd_Card_order_of_bnf bnfs; val bd_Card_order = hd bd_Card_orders; val bd_Cinfinites = map bd_Cinfinite_of_bnf bnfs; val bd_Cinfinite = hd bd_Cinfinites; val in_monos = map in_mono_of_bnf bnfs; val map_comp0s = map map_comp0_of_bnf bnfs; val sym_map_comps = map mk_sym map_comp0s; val map_comps = map map_comp_of_bnf bnfs; val map_cong0s = map map_cong0_of_bnf bnfs; val map_id0s = map map_id0_of_bnf bnfs; val map_ids = map map_id_of_bnf bnfs; val set_bdss = map set_bd_of_bnf bnfs; val set_mapss = map set_map_of_bnf bnfs; val rel_congs = map rel_cong0_of_bnf bnfs; val rel_converseps = map rel_conversep_of_bnf bnfs; val rel_Grps = map rel_Grp_of_bnf bnfs; val le_rel_OOs = map le_rel_OO_of_bnf bnfs; val in_rels = map in_rel_of_bnf bnfs; val timer = time (timer "Extracted terms & thms"); (* derived thms *) (*map g1 ... gm g(m+1) ... g(m+n) (map id ... id f(m+1) ... f(m+n) x) = map g1 ... gm (g(m+1) o f(m+1)) ... (g(m+n) o f(m+n)) x*) fun mk_map_comp_id x mapAsBs mapBsCs mapAsCs map_comp0 = let val lhs = Term.list_comb (mapBsCs, all_gs) $ (Term.list_comb (mapAsBs, passive_ids @ fs) $ x); val rhs = Term.list_comb (mapAsCs, take m all_gs @ map HOLogic.mk_comp (drop m all_gs ~~ fs)) $ x; val goal = mk_Trueprop_eq (lhs, rhs); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_map_comp_id_tac ctxt map_comp0) |> Thm.close_derivation \<^here> end; val map_comp_id_thms = @{map 5} mk_map_comp_id xFs mapsAsBs mapsBsCs' mapsAsCs' map_comps; (*forall a : set(m+1) x. f(m+1) a = a; ...; forall a : set(m+n) x. f(m+n) a = a ==> map id ... id f(m+1) ... f(m+n) x = x*) fun mk_map_cong0L x mapAsAs sets map_cong0 map_id = let fun mk_prem set f z z' = HOLogic.mk_Trueprop (mk_Ball (set $ x) (Term.absfree z' (HOLogic.mk_eq (f $ z, z)))); val prems = @{map 4} mk_prem (drop m sets) self_fs zs zs'; val goal = mk_Trueprop_eq (Term.list_comb (mapAsAs, passive_ids @ self_fs) $ x, x); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, goal)) (fn {context = ctxt, prems = _} => mk_map_cong0L_tac ctxt m map_cong0 map_id) |> Thm.close_derivation \<^here> end; val map_cong0L_thms = @{map 5} mk_map_cong0L xFs mapsAsAs setssAs map_cong0s map_ids; val in_mono'_thms = map (fn thm => (thm OF (replicate m subset_refl)) RS @{thm set_mp}) in_monos; val map_arg_cong_thms = let val prems = map2 (curry mk_Trueprop_eq) yFs yFs_copy; val maps = map (fn mapx => Term.list_comb (mapx, all_gs)) mapsBsCs'; val concls = @{map 3} (fn x => fn y => fn mapx => mk_Trueprop_eq (mapx $ x, mapx $ y)) yFs yFs_copy maps; val goals = map2 (fn prem => fn concl => Logic.mk_implies (prem, concl)) prems concls; in map (fn goal => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => (hyp_subst_tac ctxt THEN' rtac ctxt refl) 1)) |> Thm.close_derivation \<^here>) goals end; val timer = time (timer "Derived simple theorems"); (* coalgebra *) val coalg_bind = mk_internal_b (coN ^ algN) ; val coalg_def_bind = (Thm.def_binding coalg_bind, []); (*forall i = 1 ... n: (\x \ Bi. si \ Fi_in UNIV .. UNIV B1 ... Bn)*) val coalg_spec = let val ins = @{map 3} mk_in (replicate n (passive_UNIVs @ Bs)) setssAs FTsAs; fun mk_coalg_conjunct B s X z z' = mk_Ball B (Term.absfree z' (HOLogic.mk_mem (s $ z, X))); val rhs = Library.foldr1 HOLogic.mk_conj (@{map 5} mk_coalg_conjunct Bs ss ins zs zs') in fold_rev (Term.absfree o Term.dest_Free) (Bs @ ss) rhs end; val ((coalg_free, (_, coalg_def_free)), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> Local_Theory.define ((coalg_bind, NoSyn), (coalg_def_bind, coalg_spec)) ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val coalg = fst (Term.dest_Const (Morphism.term phi coalg_free)); val coalg_def = mk_unabs_def (2 * n) (HOLogic.mk_obj_eq (Morphism.thm phi coalg_def_free)); fun mk_coalg Bs ss = let val args = Bs @ ss; val Ts = map fastype_of args; val coalgT = Library.foldr (op -->) (Ts, HOLogic.boolT); in Term.list_comb (Const (coalg, coalgT), args) end; val((((((zs, zs'), Bs), B's), ss), s's), _) = lthy |> mk_Frees' "b" activeAs ||>> mk_Frees "B" BTs ||>> mk_Frees "B'" B'Ts ||>> mk_Frees "s" sTs ||>> mk_Frees "s'" s'Ts; val coalg_prem = HOLogic.mk_Trueprop (mk_coalg Bs ss); val coalg_in_thms = map (fn i => coalg_def RS iffD1 RS mk_conjunctN n i RS bspec) ks val coalg_set_thmss = let val coalg_prem = HOLogic.mk_Trueprop (mk_coalg Bs ss); fun mk_prem x B = mk_Trueprop_mem (x, B); fun mk_concl s x B set = HOLogic.mk_Trueprop (mk_leq (set $ (s $ x)) B); val prems = map2 mk_prem zs Bs; val conclss = @{map 3} (fn s => fn x => fn sets => map2 (mk_concl s x) Bs (drop m sets)) ss zs setssAs; val goalss = map2 (fn prem => fn concls => map (fn concl => Logic.list_implies (coalg_prem :: [prem], concl)) concls) prems conclss; in map (fn goals => map (fn goal => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_coalg_set_tac ctxt coalg_def)) |> Thm.close_derivation \<^here>) goals) goalss end; fun mk_tcoalg BTs = mk_coalg (map HOLogic.mk_UNIV BTs); val tcoalg_thm = let val goal = HOLogic.mk_Trueprop (mk_tcoalg activeAs ss); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => (rtac ctxt (coalg_def RS iffD2) 1 THEN CONJ_WRAP (K (EVERY' [rtac ctxt ballI, rtac ctxt CollectI, CONJ_WRAP' (K (EVERY' [rtac ctxt @{thm subset_UNIV}])) allAs] 1)) ss)) |> Thm.close_derivation \<^here> end; val timer = time (timer "Coalgebra definition & thms"); (* morphism *) val mor_bind = mk_internal_b morN; val mor_def_bind = (Thm.def_binding mor_bind, []); (*fbetw) forall i = 1 ... n: (\x \ Bi. fi x \ B'i)*) (*mor) forall i = 1 ... n: (\x \ Bi. Fi_map id ... id f1 ... fn (si x) = si' (fi x)*) val mor_spec = let fun mk_fbetw f B1 B2 z z' = mk_Ball B1 (Term.absfree z' (HOLogic.mk_mem (f $ z, B2))); fun mk_mor B mapAsBs f s s' z z' = mk_Ball B (Term.absfree z' (HOLogic.mk_eq (Term.list_comb (mapAsBs, passive_ids @ fs @ [s $ z]), s' $ (f $ z)))); val rhs = HOLogic.mk_conj (Library.foldr1 HOLogic.mk_conj (@{map 5} mk_fbetw fs Bs B's zs zs'), Library.foldr1 HOLogic.mk_conj (@{map 7} mk_mor Bs mapsAsBs fs ss s's zs zs')) in fold_rev (Term.absfree o Term.dest_Free) (Bs @ ss @ B's @ s's @ fs) rhs end; val ((mor_free, (_, mor_def_free)), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> Local_Theory.define ((mor_bind, NoSyn), (mor_def_bind, mor_spec)) ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val mor = fst (Term.dest_Const (Morphism.term phi mor_free)); val mor_def = mk_unabs_def (5 * n) (HOLogic.mk_obj_eq (Morphism.thm phi mor_def_free)); fun mk_mor Bs1 ss1 Bs2 ss2 fs = let val args = Bs1 @ ss1 @ Bs2 @ ss2 @ fs; val Ts = map fastype_of (Bs1 @ ss1 @ Bs2 @ ss2 @ fs); val morT = Library.foldr (op -->) (Ts, HOLogic.boolT); in Term.list_comb (Const (mor, morT), args) end; val ((((((((((((((zs, z's), Bs), Bs_copy), B's), B''s), ss), s's), s''s), fs), fs_copy), gs), RFs), Rs), _) = lthy |> mk_Frees "b" activeAs ||>> mk_Frees "b" activeBs ||>> mk_Frees "B" BTs ||>> mk_Frees "B" BTs ||>> mk_Frees "B'" B'Ts ||>> mk_Frees "B''" B''Ts ||>> mk_Frees "s" sTs ||>> mk_Frees "s'" s'Ts ||>> mk_Frees "s''" s''Ts ||>> mk_Frees "f" fTs ||>> mk_Frees "f" fTs ||>> mk_Frees "g" gTs ||>> mk_Frees "x" FRTs ||>> mk_Frees "R" setRTs; val mor_prem = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs); val (mor_image_thms, morE_thms) = let val prem = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs); fun mk_image_goal f B1 B2 = Logic.mk_implies (prem, HOLogic.mk_Trueprop (mk_leq (mk_image f $ B1) B2)); val image_goals = @{map 3} mk_image_goal fs Bs B's; fun mk_elim_goal B mapAsBs f s s' x = Logic.list_implies ([prem, mk_Trueprop_mem (x, B)], mk_Trueprop_eq (Term.list_comb (mapAsBs, passive_ids @ fs @ [s $ x]), s' $ (f $ x))); val elim_goals = @{map 6} mk_elim_goal Bs mapsAsBs fs ss s's zs; fun prove goal = Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_mor_elim_tac ctxt mor_def)) |> Thm.close_derivation \<^here>; in (map prove image_goals, map prove elim_goals) end; val mor_image'_thms = map (fn thm => @{thm set_mp} OF [thm, imageI]) mor_image_thms; val mor_incl_thm = let val prems = map2 (HOLogic.mk_Trueprop oo mk_leq) Bs Bs_copy; val concl = HOLogic.mk_Trueprop (mk_mor Bs ss Bs_copy ss active_ids); val vars = fold (Variable.add_free_names lthy) (concl :: prems) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, concl)) (fn {context = ctxt, prems = _} => mk_mor_incl_tac ctxt mor_def map_ids) |> Thm.close_derivation \<^here> end; val mor_id_thm = mor_incl_thm OF (replicate n subset_refl); val mor_comp_thm = let val prems = [HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs), HOLogic.mk_Trueprop (mk_mor B's s's B''s s''s gs)]; val concl = HOLogic.mk_Trueprop (mk_mor Bs ss B''s s''s (map2 (curry HOLogic.mk_comp) gs fs)); val vars = fold (Variable.add_free_names lthy) (concl :: prems) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, concl)) (fn {context = ctxt, prems = _} => mk_mor_comp_tac ctxt mor_def mor_image'_thms morE_thms map_comp_id_thms) |> Thm.close_derivation \<^here> end; val mor_cong_thm = let val prems = map HOLogic.mk_Trueprop (map2 (curry HOLogic.mk_eq) fs_copy fs @ [mk_mor Bs ss B's s's fs]) val concl = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs_copy); val vars = fold (Variable.add_free_names lthy) (concl :: prems) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, concl)) (fn {context = ctxt, prems = _} => (hyp_subst_tac ctxt THEN' assume_tac ctxt) 1) |> Thm.close_derivation \<^here> end; val mor_UNIV_thm = let fun mk_conjunct mapAsBs f s s' = HOLogic.mk_eq (HOLogic.mk_comp (Term.list_comb (mapAsBs, passive_ids @ fs), s), HOLogic.mk_comp (s', f)); val lhs = mk_mor active_UNIVs ss (map HOLogic.mk_UNIV activeBs) s's fs; val rhs = Library.foldr1 HOLogic.mk_conj (@{map 4} mk_conjunct mapsAsBs fs ss s's); val vars = fold (Variable.add_free_names lthy) [lhs, rhs] []; in Goal.prove_sorry lthy vars [] (mk_Trueprop_eq (lhs, rhs)) (fn {context = ctxt, prems = _} => mk_mor_UNIV_tac ctxt morE_thms mor_def) |> Thm.close_derivation \<^here> end; val mor_str_thm = let val maps = map2 (fn Ds => fn bnf => Term.list_comb (mk_map_of_bnf Ds allAs (passiveAs @ FTsAs) bnf, passive_ids @ ss)) Dss bnfs; val goal = HOLogic.mk_Trueprop (mk_mor active_UNIVs ss (map HOLogic.mk_UNIV FTsAs) maps ss); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_mor_str_tac ctxt ks mor_UNIV_thm) |> Thm.close_derivation \<^here> end; val timer = time (timer "Morphism definition & thms"); (* bisimulation *) val bis_bind = mk_internal_b bisN; val bis_def_bind = (Thm.def_binding bis_bind, []); fun mk_bis_le_conjunct R B1 B2 = mk_leq R (mk_Times (B1, B2)); val bis_le = Library.foldr1 HOLogic.mk_conj (@{map 3} mk_bis_le_conjunct Rs Bs B's) val bis_spec = let val fst_args = passive_ids @ fsts; val snd_args = passive_ids @ snds; fun mk_bis R s s' b1 b2 RF map1 map2 sets = list_all_free [b1, b2] (HOLogic.mk_imp (HOLogic.mk_mem (HOLogic.mk_prod (b1, b2), R), mk_Bex (mk_in (passive_UNIVs @ Rs) sets (snd (dest_Free RF))) (Term.absfree (dest_Free RF) (HOLogic.mk_conj (HOLogic.mk_eq (Term.list_comb (map1, fst_args) $ RF, s $ b1), HOLogic.mk_eq (Term.list_comb (map2, snd_args) $ RF, s' $ b2)))))); val rhs = HOLogic.mk_conj (bis_le, Library.foldr1 HOLogic.mk_conj (@{map 9} mk_bis Rs ss s's zs z's RFs map_fsts map_snds bis_setss)) in fold_rev (Term.absfree o Term.dest_Free) (Bs @ ss @ B's @ s's @ Rs) rhs end; val ((bis_free, (_, bis_def_free)), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> Local_Theory.define ((bis_bind, NoSyn), (bis_def_bind, bis_spec)) ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val bis = fst (Term.dest_Const (Morphism.term phi bis_free)); val bis_def = mk_unabs_def (5 * n) (HOLogic.mk_obj_eq (Morphism.thm phi bis_def_free)); fun mk_bis Bs1 ss1 Bs2 ss2 Rs = let val args = Bs1 @ ss1 @ Bs2 @ ss2 @ Rs; val Ts = map fastype_of args; val bisT = Library.foldr (op -->) (Ts, HOLogic.boolT); in Term.list_comb (Const (bis, bisT), args) end; val (((((((((((((((((zs, z's), Bs), B's), B''s), ss), s's), s''s), fs), (Rtuple, Rtuple')), Rs), Rs_copy), R's), sRs), (idx, idx')), Idx), Ris), _) = lthy |> mk_Frees "b" activeAs ||>> mk_Frees "b" activeBs ||>> mk_Frees "B" BTs ||>> mk_Frees "B'" B'Ts ||>> mk_Frees "B''" B''Ts ||>> mk_Frees "s" sTs ||>> mk_Frees "s'" s'Ts ||>> mk_Frees "s''" s''Ts ||>> mk_Frees "f" fTs ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "Rtuple") all_sbisT ||>> mk_Frees "R" setRTs ||>> mk_Frees "R" setRTs ||>> mk_Frees "R'" setR'Ts ||>> mk_Frees "R" setsRTs ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "i") idxT ||>> yield_singleton (mk_Frees "I") (HOLogic.mk_setT idxT) ||>> mk_Frees "Ri" (map (fn T => idxT --> T) setRTs); val bis_cong_thm = let val prems = map HOLogic.mk_Trueprop (mk_bis Bs ss B's s's Rs :: map2 (curry HOLogic.mk_eq) Rs_copy Rs) val concl = HOLogic.mk_Trueprop (mk_bis Bs ss B's s's Rs_copy); val vars = fold (Variable.add_free_names lthy) (concl :: prems) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, concl)) (fn {context = ctxt, prems = _} => (hyp_subst_tac ctxt THEN' assume_tac ctxt) 1) |> Thm.close_derivation \<^here> end; val bis_rel_thm = let fun mk_conjunct R s s' b1 b2 rel = list_all_free [b1, b2] (HOLogic.mk_imp (HOLogic.mk_mem (HOLogic.mk_prod (b1, b2), R), Term.list_comb (rel, passive_eqs @ map mk_in_rel Rs) $ (s $ b1) $ (s' $ b2))); val rhs = HOLogic.mk_conj (bis_le, Library.foldr1 HOLogic.mk_conj (@{map 6} mk_conjunct Rs ss s's zs z's relsAsBs)) val goal = mk_Trueprop_eq (mk_bis Bs ss B's s's Rs, rhs); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_bis_rel_tac ctxt m bis_def in_rels map_comps map_cong0s set_mapss) |> Thm.close_derivation \<^here> end; val bis_converse_thm = let val goal = Logic.mk_implies (HOLogic.mk_Trueprop (mk_bis Bs ss B's s's Rs), HOLogic.mk_Trueprop (mk_bis B's s's Bs ss (map mk_converse Rs))); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_bis_converse_tac ctxt m bis_rel_thm rel_congs rel_converseps) |> Thm.close_derivation \<^here> end; val bis_O_thm = let val prems = [HOLogic.mk_Trueprop (mk_bis Bs ss B's s's Rs), HOLogic.mk_Trueprop (mk_bis B's s's B''s s''s R's)]; val concl = HOLogic.mk_Trueprop (mk_bis Bs ss B''s s''s (map2 (curry mk_rel_comp) Rs R's)); val vars = fold (Variable.add_free_names lthy) (concl :: prems) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, concl)) (fn {context = ctxt, prems = _} => mk_bis_O_tac ctxt m bis_rel_thm rel_congs le_rel_OOs) |> Thm.close_derivation \<^here> end; val bis_Gr_thm = let val concl = HOLogic.mk_Trueprop (mk_bis Bs ss B's s's (map2 mk_Gr Bs fs)); val vars = fold (Variable.add_free_names lthy) ([coalg_prem, mor_prem, concl]) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies ([coalg_prem, mor_prem], concl)) (fn {context = ctxt, prems = _} => mk_bis_Gr_tac ctxt bis_rel_thm rel_Grps mor_image_thms morE_thms coalg_in_thms) |> Thm.close_derivation \<^here> end; val bis_image2_thm = bis_cong_thm OF ((bis_O_thm OF [bis_Gr_thm RS bis_converse_thm, bis_Gr_thm]) :: replicate n @{thm image2_Gr}); val bis_Id_on_thm = bis_cong_thm OF ((mor_id_thm RSN (2, bis_Gr_thm)) :: replicate n @{thm Id_on_Gr}); val bis_Union_thm = let val prem = HOLogic.mk_Trueprop (mk_Ball Idx (Term.absfree idx' (mk_bis Bs ss B's s's (map (fn R => R $ idx) Ris)))); val concl = HOLogic.mk_Trueprop (mk_bis Bs ss B's s's (map (mk_UNION Idx) Ris)); val vars = fold (Variable.add_free_names lthy) [prem, concl] []; in Goal.prove_sorry lthy vars [] (Logic.mk_implies (prem, concl)) (fn {context = ctxt, prems = _} => mk_bis_Union_tac ctxt bis_def in_mono'_thms) |> Thm.close_derivation \<^here> end; (* self-bisimulation *) fun mk_sbis Bs ss Rs = mk_bis Bs ss Bs ss Rs; (* largest self-bisimulation *) val lsbis_binds = mk_internal_bs lsbisN; fun lsbis_bind i = nth lsbis_binds (i - 1); val lsbis_def_bind = rpair [] o Thm.def_binding o lsbis_bind; val all_sbis = HOLogic.mk_Collect (fst Rtuple', snd Rtuple', list_exists_free sRs (HOLogic.mk_conj (HOLogic.mk_eq (Rtuple, HOLogic.mk_tuple sRs), mk_sbis Bs ss sRs))); fun lsbis_spec i = fold_rev (Term.absfree o Term.dest_Free) (Bs @ ss) (mk_UNION all_sbis (Term.absfree Rtuple' (mk_nthN n Rtuple i))); val ((lsbis_frees, (_, lsbis_def_frees)), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> fold_map (fn i => Local_Theory.define ((lsbis_bind i, NoSyn), (lsbis_def_bind i, lsbis_spec i))) ks |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val lsbis_defs = map (fn def => mk_unabs_def (2 * n) (HOLogic.mk_obj_eq (Morphism.thm phi def))) lsbis_def_frees; val lsbiss = map (fst o Term.dest_Const o Morphism.term phi) lsbis_frees; fun mk_lsbis Bs ss i = let val args = Bs @ ss; val Ts = map fastype_of args; val RT = mk_relT (`I (HOLogic.dest_setT (fastype_of (nth Bs (i - 1))))); val lsbisT = Library.foldr (op -->) (Ts, RT); in Term.list_comb (Const (nth lsbiss (i - 1), lsbisT), args) end; val (((((zs, zs'), Bs), ss), sRs), _) = lthy |> mk_Frees' "b" activeAs ||>> mk_Frees "B" BTs ||>> mk_Frees "s" sTs ||>> mk_Frees "R" setsRTs; val sbis_prem = HOLogic.mk_Trueprop (mk_sbis Bs ss sRs); val coalg_prem = HOLogic.mk_Trueprop (mk_coalg Bs ss); val sbis_lsbis_thm = let val goal = HOLogic.mk_Trueprop (mk_sbis Bs ss (map (mk_lsbis Bs ss) ks)); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_sbis_lsbis_tac ctxt lsbis_defs bis_Union_thm bis_cong_thm) |> Thm.close_derivation \<^here> end; val lsbis_incl_thms = map (fn i => sbis_lsbis_thm RS (bis_def RS iffD1 RS conjunct1 RS mk_conjunctN n i)) ks; val lsbisE_thms = map (fn i => (mk_specN 2 (sbis_lsbis_thm RS (bis_def RS iffD1 RS conjunct2 RS mk_conjunctN n i))) RS mp) ks; val incl_lsbis_thms = let fun mk_concl i R = HOLogic.mk_Trueprop (mk_leq R (mk_lsbis Bs ss i)); val goals = map2 (fn i => fn R => Logic.mk_implies (sbis_prem, mk_concl i R)) ks sRs; in @{map 3} (fn goal => fn i => fn def => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_incl_lsbis_tac ctxt n i def)) |> Thm.close_derivation \<^here>) goals ks lsbis_defs end; val equiv_lsbis_thms = let fun mk_concl i B = HOLogic.mk_Trueprop (mk_equiv B (mk_lsbis Bs ss i)); val goals = map2 (fn i => fn B => Logic.mk_implies (coalg_prem, mk_concl i B)) ks Bs; in @{map 3} (fn goal => fn l_incl => fn incl_l => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_equiv_lsbis_tac ctxt sbis_lsbis_thm l_incl incl_l bis_Id_on_thm bis_converse_thm bis_O_thm) |> Thm.close_derivation \<^here>)) goals lsbis_incl_thms incl_lsbis_thms end; val timer = time (timer "Bisimulations"); (* bounds *) val (lthy, sbd, sbdT, sbd_card_order, sbd_Cinfinite, sbd_Card_order, set_sbdss) = if n = 1 then (lthy, sum_bd, sum_bdT, bd_card_order, bd_Cinfinite, bd_Card_order, set_bdss) else let val sbdT_bind = mk_internal_b sum_bdTN; val ((sbdT_name, (sbdT_glob_info, sbdT_loc_info)), lthy) = typedef (sbdT_bind, sum_bdT_params', NoSyn) (HOLogic.mk_UNIV sum_bdT) NONE (fn ctxt => EVERY' [rtac ctxt exI, rtac ctxt UNIV_I] 1) lthy; val sbdT = Type (sbdT_name, sum_bdT_params); val Abs_sbdT = Const (#Abs_name sbdT_glob_info, sum_bdT --> sbdT); val sbd_bind = mk_internal_b sum_bdN; val sbd_def_bind = (Thm.def_binding sbd_bind, []); val sbd_spec = mk_dir_image sum_bd Abs_sbdT; val ((sbd_free, (_, sbd_def_free)), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> Local_Theory.define ((sbd_bind, NoSyn), (sbd_def_bind, sbd_spec)) ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val sbd_def = HOLogic.mk_obj_eq (Morphism.thm phi sbd_def_free); val sbd = Const (fst (Term.dest_Const (Morphism.term phi sbd_free)), mk_relT (`I sbdT)); val Abs_sbdT_inj = mk_Abs_inj_thm (#Abs_inject sbdT_loc_info); val Abs_sbdT_bij = mk_Abs_bij_thm lthy Abs_sbdT_inj (#Abs_cases sbdT_loc_info); val sum_Cinfinite = mk_sum_Cinfinite bd_Cinfinites; val sum_Card_order = sum_Cinfinite RS conjunct2; val sum_card_order = mk_sum_card_order bd_card_orders; val sbd_ordIso = @{thm ssubst_Pair_rhs} OF [@{thm dir_image} OF [Abs_sbdT_inj, sum_Card_order], sbd_def]; val sbd_card_order = @{thm iffD2[OF arg_cong[of _ _ card_order]]} OF [sbd_def, @{thm card_order_dir_image} OF [Abs_sbdT_bij, sum_card_order]]; val sbd_Cinfinite = @{thm Cinfinite_cong} OF [sbd_ordIso, sum_Cinfinite]; val sbd_Card_order = sbd_Cinfinite RS conjunct2; fun mk_set_sbd i bd_Card_order bds = map (fn thm => @{thm ordLeq_ordIso_trans} OF [bd_Card_order RS mk_ordLeq_csum n i thm, sbd_ordIso]) bds; val set_sbdss = @{map 3} mk_set_sbd ks bd_Card_orders set_bdss; in (lthy, sbd, sbdT, sbd_card_order, sbd_Cinfinite, sbd_Card_order, set_sbdss) end; val sbdTs = replicate n sbdT; val sum_sbdT = mk_sumTN sbdTs; val sum_sbd_listT = HOLogic.listT sum_sbdT; val sum_sbd_list_setT = HOLogic.mk_setT sum_sbd_listT; val bdTs = passiveAs @ replicate n sbdT; val to_sbd_maps = @{map 4} mk_map_of_bnf Dss Ass (replicate n bdTs) bnfs; val bdFTs = mk_FTs bdTs; val sbdFT = mk_sumTN bdFTs; val treeT = HOLogic.mk_prodT (sum_sbd_list_setT, sum_sbd_listT --> sbdFT); val treeQT = HOLogic.mk_setT treeT; val treeTs = passiveAs @ replicate n treeT; val treeQTs = passiveAs @ replicate n treeQT; val treeFTs = mk_FTs treeTs; val tree_maps = @{map 4} mk_map_of_bnf Dss (replicate n bdTs) (replicate n treeTs) bnfs; val final_maps = @{map 4} mk_map_of_bnf Dss (replicate n treeTs) (replicate n treeQTs) bnfs; val isNode_setss = mk_setss (passiveAs @ replicate n sbdT); val root = HOLogic.mk_set sum_sbd_listT [HOLogic.mk_list sum_sbdT []]; val Zero = HOLogic.mk_tuple (map (fn U => absdummy U root) activeAs); val Lev_recT = fastype_of Zero; val Nil = HOLogic.mk_tuple (@{map 3} (fn i => fn z => fn z'=> Term.absfree z' (mk_InN activeAs z i)) ks zs zs'); val rv_recT = fastype_of Nil; val (((((((((((((((zs, zs'), zs_copy), ss), (nat, nat')), (sumx, sumx')), (kks, kks')), (kl, kl')), (kl_copy, kl'_copy)), (Kl, Kl')), (lab, lab')), (Kl_lab, Kl_lab')), xs), (Lev_rec, Lev_rec')), (rv_rec, rv_rec')), _) = lthy |> mk_Frees' "b" activeAs ||>> mk_Frees "b" activeAs ||>> mk_Frees "s" sTs ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "n") HOLogic.natT ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "sumx") sum_sbdT ||>> mk_Frees' "k" sbdTs ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "kl") sum_sbd_listT ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "kl") sum_sbd_listT ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "Kl") sum_sbd_list_setT ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "lab") (sum_sbd_listT --> sbdFT) ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "Kl_lab") treeT ||>> mk_Frees "x" bdFTs ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "rec") Lev_recT ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "rec") rv_recT; val (k, k') = (hd kks, hd kks') val timer = time (timer "Bounds"); (* tree coalgebra *) val isNode_binds = mk_internal_bs isNodeN; fun isNode_bind i = nth isNode_binds (i - 1); val isNode_def_bind = rpair [] o Thm.def_binding o isNode_bind; val isNodeT = Library.foldr (op -->) (map fastype_of [Kl, lab, kl], HOLogic.boolT); val Succs = @{map 3} (fn i => fn k => fn k' => HOLogic.mk_Collect (fst k', snd k', HOLogic.mk_mem (mk_InN sbdTs k i, mk_Succ Kl kl))) ks kks kks'; fun isNode_spec sets x i = let val active_sets = drop m (map (fn set => set $ x) sets); val rhs = list_exists_free [x] (Library.foldr1 HOLogic.mk_conj (HOLogic.mk_eq (lab $ kl, mk_InN bdFTs x i) :: map2 (curry HOLogic.mk_eq) active_sets Succs)); in fold_rev (Term.absfree o Term.dest_Free) [Kl, lab, kl] rhs end; val ((isNode_frees, (_, isNode_def_frees)), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> @{fold_map 3} (fn i => fn x => fn sets => Local_Theory.define ((isNode_bind i, NoSyn), (isNode_def_bind i, isNode_spec sets x i))) ks xs isNode_setss |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val isNode_defs = map (fn def => mk_unabs_def 3 (HOLogic.mk_obj_eq (Morphism.thm phi def))) isNode_def_frees; val isNodes = map (fst o Term.dest_Const o Morphism.term phi) isNode_frees; fun mk_isNode kl i = Term.list_comb (Const (nth isNodes (i - 1), isNodeT), [Kl, lab, kl]); val isTree = let val empty = HOLogic.mk_mem (HOLogic.mk_list sum_sbdT [], Kl); val tree = mk_Ball Kl (Term.absfree kl' (Library.foldr1 HOLogic.mk_conj (@{map 4} (fn Succ => fn i => fn k => fn k' => mk_Ball Succ (Term.absfree k' (mk_isNode (mk_append (kl, HOLogic.mk_list sum_sbdT [mk_InN sbdTs k i])) i))) Succs ks kks kks'))); in HOLogic.mk_conj (empty, tree) end; val carT_binds = mk_internal_bs carTN; fun carT_bind i = nth carT_binds (i - 1); val carT_def_bind = rpair [] o Thm.def_binding o carT_bind; fun carT_spec i = HOLogic.mk_Collect (fst Kl_lab', snd Kl_lab', list_exists_free [Kl, lab] (HOLogic.mk_conj (HOLogic.mk_eq (Kl_lab, HOLogic.mk_prod (Kl, lab)), HOLogic.mk_conj (isTree, mk_isNode (HOLogic.mk_list sum_sbdT []) i)))); val ((carT_frees, (_, carT_def_frees)), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> fold_map (fn i => Local_Theory.define ((carT_bind i, NoSyn), (carT_def_bind i, carT_spec i))) ks |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val carT_defs = map (fn def => HOLogic.mk_obj_eq (Morphism.thm phi def)) carT_def_frees; val carTs = map (fst o Term.dest_Const o Morphism.term phi) carT_frees; fun mk_carT i = Const (nth carTs (i - 1), HOLogic.mk_setT treeT); val strT_binds = mk_internal_bs strTN; fun strT_bind i = nth strT_binds (i - 1); val strT_def_bind = rpair [] o Thm.def_binding o strT_bind; fun strT_spec mapFT FT i = let fun mk_f i k k' = let val in_k = mk_InN sbdTs k i; in Term.absfree k' (HOLogic.mk_prod (mk_Shift Kl in_k, mk_shift lab in_k)) end; val f = Term.list_comb (mapFT, passive_ids @ @{map 3} mk_f ks kks kks'); val (fTs1, fTs2) = apsnd tl (chop (i - 1) (map (fn T => T --> FT) bdFTs)); val fs = map mk_undefined fTs1 @ (f :: map mk_undefined fTs2); in HOLogic.mk_case_prod (Term.absfree Kl' (Term.absfree lab' (mk_case_sumN fs $ (lab $ HOLogic.mk_list sum_sbdT [])))) end; val ((strT_frees, (_, strT_def_frees)), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> @{fold_map 3} (fn i => fn mapFT => fn FT => Local_Theory.define ((strT_bind i, NoSyn), (strT_def_bind i, strT_spec mapFT FT i))) ks tree_maps treeFTs |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val strT_defs = map (fn def => trans OF [HOLogic.mk_obj_eq (Morphism.thm phi def) RS fun_cong, @{thm prod.case}]) strT_def_frees; val strTs = map (fst o Term.dest_Const o Morphism.term phi) strT_frees; fun mk_strT FT i = Const (nth strTs (i - 1), treeT --> FT); val carTAs = map mk_carT ks; val strTAs = map2 mk_strT treeFTs ks; val coalgT_thm = Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (mk_coalg carTAs strTAs)) (fn {context = ctxt, prems = _} => mk_coalgT_tac ctxt m (coalg_def :: isNode_defs @ carT_defs) strT_defs set_mapss) |> Thm.close_derivation \<^here>; val timer = time (timer "Tree coalgebra"); fun mk_to_sbd s x i i' = mk_toCard (nth (nth setssAs (i - 1)) (m + i' - 1) $ (s $ x)) sbd; fun mk_from_sbd s x i i' = mk_fromCard (nth (nth setssAs (i - 1)) (m + i' - 1) $ (s $ x)) sbd; fun mk_to_sbd_thmss thm = map (map (fn set_sbd => thm OF [set_sbd, sbd_Card_order]) o drop m) set_sbdss; val to_sbd_inj_thmss = mk_to_sbd_thmss @{thm toCard_inj}; val from_to_sbd_thmss = mk_to_sbd_thmss @{thm fromCard_toCard}; val Lev_bind = mk_internal_b LevN; val Lev_def_bind = rpair [] (Thm.def_binding Lev_bind); val Lev_spec = let fun mk_Suc i s setsAs a a' = let val sets = drop m setsAs; fun mk_set i' set b = let val Cons = HOLogic.mk_eq (kl_copy, mk_Cons (mk_InN sbdTs (mk_to_sbd s a i i' $ b) i') kl) val b_set = HOLogic.mk_mem (b, set $ (s $ a)); val kl_rec = HOLogic.mk_mem (kl, mk_nthN n Lev_rec i' $ b); in HOLogic.mk_Collect (fst kl'_copy, snd kl'_copy, list_exists_free [b, kl] (HOLogic.mk_conj (Cons, HOLogic.mk_conj (b_set, kl_rec)))) end; in Term.absfree a' (Library.foldl1 mk_union (@{map 3} mk_set ks sets zs_copy)) end; val Suc = Term.absdummy HOLogic.natT (Term.absfree Lev_rec' (HOLogic.mk_tuple (@{map 5} mk_Suc ks ss setssAs zs zs'))); val rhs = mk_rec_nat Zero Suc; in fold_rev (Term.absfree o Term.dest_Free) ss rhs end; val ((Lev_free, (_, Lev_def_free)), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> Local_Theory.define ((Lev_bind, NoSyn), (Lev_def_bind, Lev_spec)) ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val Lev_def = mk_unabs_def n (HOLogic.mk_obj_eq (Morphism.thm phi Lev_def_free)); val Lev = fst (Term.dest_Const (Morphism.term phi Lev_free)); fun mk_Lev ss nat i = let val Ts = map fastype_of ss; val LevT = Library.foldr (op -->) (Ts, HOLogic.natT --> HOLogic.mk_tupleT (map (fn U => domain_type U --> sum_sbd_list_setT) Ts)); in mk_nthN n (Term.list_comb (Const (Lev, LevT), ss) $ nat) i end; val Lev_0s = flat (mk_rec_simps n @{thm rec_nat_0_imp} [Lev_def]); val Lev_Sucs = flat (mk_rec_simps n @{thm rec_nat_Suc_imp} [Lev_def]); val rv_bind = mk_internal_b rvN; val rv_def_bind = rpair [] (Thm.def_binding rv_bind); val rv_spec = let fun mk_Cons i s b b' = let fun mk_case i' = Term.absfree k' (mk_nthN n rv_rec i' $ (mk_from_sbd s b i i' $ k)); in Term.absfree b' (mk_case_sumN (map mk_case ks) $ sumx) end; val Cons = Term.absfree sumx' (Term.absdummy sum_sbd_listT (Term.absfree rv_rec' (HOLogic.mk_tuple (@{map 4} mk_Cons ks ss zs zs')))); val rhs = mk_rec_list Nil Cons; in fold_rev (Term.absfree o Term.dest_Free) ss rhs end; val ((rv_free, (_, rv_def_free)), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> Local_Theory.define ((rv_bind, NoSyn), (rv_def_bind, rv_spec)) ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val rv_def = mk_unabs_def n (HOLogic.mk_obj_eq (Morphism.thm phi rv_def_free)); val rv = fst (Term.dest_Const (Morphism.term phi rv_free)); fun mk_rv ss kl i = let val Ts = map fastype_of ss; val As = map domain_type Ts; val rvT = Library.foldr (op -->) (Ts, fastype_of kl --> HOLogic.mk_tupleT (map (fn U => U --> mk_sumTN As) As)); in mk_nthN n (Term.list_comb (Const (rv, rvT), ss) $ kl) i end; val rv_Nils = flat (mk_rec_simps n @{thm rec_list_Nil_imp} [rv_def]); val rv_Conss = flat (mk_rec_simps n @{thm rec_list_Cons_imp} [rv_def]); val beh_binds = mk_internal_bs behN; fun beh_bind i = nth beh_binds (i - 1); val beh_def_bind = rpair [] o Thm.def_binding o beh_bind; fun beh_spec i z = let fun mk_case i to_sbd_map s k k' = Term.absfree k' (mk_InN bdFTs (Term.list_comb (to_sbd_map, passive_ids @ map (mk_to_sbd s k i) ks) $ (s $ k)) i); val Lab = Term.absfree kl' (mk_case_sumN (@{map 5} mk_case ks to_sbd_maps ss zs zs') $ (mk_rv ss kl i $ z)); val rhs = HOLogic.mk_prod (mk_UNION (HOLogic.mk_UNIV HOLogic.natT) (Term.absfree nat' (mk_Lev ss nat i $ z)), Lab); in fold_rev (Term.absfree o Term.dest_Free) (ss @ [z]) rhs end; val ((beh_frees, (_, beh_def_frees)), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> @{fold_map 2} (fn i => fn z => Local_Theory.define ((beh_bind i, NoSyn), (beh_def_bind i, beh_spec i z))) ks zs |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val beh_defs = map (fn def => mk_unabs_def (n + 1) (HOLogic.mk_obj_eq (Morphism.thm phi def))) beh_def_frees; val behs = map (fst o Term.dest_Const o Morphism.term phi) beh_frees; fun mk_beh ss i = let val Ts = map fastype_of ss; val behT = Library.foldr (op -->) (Ts, nth activeAs (i - 1) --> treeT); in Term.list_comb (Const (nth behs (i - 1), behT), ss) end; val ((((((zs, zs_copy), zs_copy2), ss), (nat, nat')), (kl, kl')), _) = lthy |> mk_Frees "b" activeAs ||>> mk_Frees "b" activeAs ||>> mk_Frees "b" activeAs ||>> mk_Frees "s" sTs ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "n") HOLogic.natT ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "kl") sum_sbd_listT; val (length_Lev_thms, length_Lev'_thms) = let fun mk_conjunct i z = HOLogic.mk_imp (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z), HOLogic.mk_eq (mk_size kl, nat)); val goal = list_all_free (kl :: zs) (Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct ks zs)); val vars = Variable.add_free_names lthy goal []; val cts = map (SOME o Thm.cterm_of lthy) [Term.absfree nat' goal, nat]; val length_Lev = Goal.prove_sorry lthy vars [] (HOLogic.mk_Trueprop goal) (fn {context = ctxt, prems = _} => mk_length_Lev_tac ctxt cts Lev_0s Lev_Sucs) |> Thm.close_derivation \<^here>; val length_Lev' = mk_specN (n + 1) length_Lev; val length_Levs = map (fn i => length_Lev' RS mk_conjunctN n i RS mp) ks; fun mk_goal i z = Logic.mk_implies (mk_Trueprop_mem (kl, mk_Lev ss nat i $ z), mk_Trueprop_mem (kl, mk_Lev ss (mk_size kl) i $ z)); val goals = map2 mk_goal ks zs; val length_Levs' = map2 (fn goal => fn length_Lev => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_length_Lev'_tac ctxt length_Lev)) |> Thm.close_derivation \<^here>) goals length_Levs; in (length_Levs, length_Levs') end; val rv_last_thmss = let fun mk_conjunct i z i' z_copy = list_exists_free [z_copy] (HOLogic.mk_eq (mk_rv ss (mk_append (kl, HOLogic.mk_list sum_sbdT [mk_InN sbdTs k i'])) i $ z, mk_InN activeAs z_copy i')); val goal = list_all_free (k :: zs) (Library.foldr1 HOLogic.mk_conj (map2 (fn i => fn z => Library.foldr1 HOLogic.mk_conj (map2 (mk_conjunct i z) ks zs_copy)) ks zs)); val vars = Variable.add_free_names lthy goal []; val cTs = [SOME (Thm.ctyp_of lthy sum_sbdT)]; val cts = map (SOME o Thm.cterm_of lthy) [Term.absfree kl' goal, kl]; val rv_last = Goal.prove_sorry lthy vars [] (HOLogic.mk_Trueprop goal) (fn {context = ctxt, prems = _} => mk_rv_last_tac ctxt cTs cts rv_Nils rv_Conss) |> Thm.close_derivation \<^here>; val rv_last' = mk_specN (n + 1) rv_last; in map (fn i => map (fn i' => rv_last' RS mk_conjunctN n i RS mk_conjunctN n i') ks) ks end; val set_Lev_thmsss = let fun mk_conjunct i z = let fun mk_conjunct' i' sets s z' = let fun mk_conjunct'' i'' set z'' = HOLogic.mk_imp (HOLogic.mk_mem (z'', set $ (s $ z')), HOLogic.mk_mem (mk_append (kl, HOLogic.mk_list sum_sbdT [mk_InN sbdTs (mk_to_sbd s z' i' i'' $ z'') i'']), mk_Lev ss (HOLogic.mk_Suc nat) i $ z)); in HOLogic.mk_imp (HOLogic.mk_eq (mk_rv ss kl i $ z, mk_InN activeAs z' i'), (Library.foldr1 HOLogic.mk_conj (@{map 3} mk_conjunct'' ks (drop m sets) zs_copy2))) end; in HOLogic.mk_imp (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z), Library.foldr1 HOLogic.mk_conj (@{map 4} mk_conjunct' ks setssAs ss zs_copy)) end; val goal = list_all_free (kl :: zs @ zs_copy @ zs_copy2) (Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct ks zs)); val vars = Variable.add_free_names lthy goal []; val cts = map (SOME o Thm.cterm_of lthy) [Term.absfree nat' goal, nat]; val set_Lev = Goal.prove_sorry lthy vars [] (HOLogic.mk_Trueprop goal) (fn {context = ctxt, prems = _} => mk_set_Lev_tac ctxt cts Lev_0s Lev_Sucs rv_Nils rv_Conss from_to_sbd_thmss) |> Thm.close_derivation \<^here>; val set_Lev' = mk_specN (3 * n + 1) set_Lev; in map (fn i => map (fn i' => map (fn i'' => set_Lev' RS mk_conjunctN n i RS mp RS mk_conjunctN n i' RS mp RS mk_conjunctN n i'' RS mp) ks) ks) ks end; val set_image_Lev_thmsss = let fun mk_conjunct i z = let fun mk_conjunct' i' sets = let fun mk_conjunct'' i'' set s z'' = HOLogic.mk_imp (HOLogic.mk_eq (mk_rv ss kl i $ z, mk_InN activeAs z'' i''), HOLogic.mk_mem (k, mk_image (mk_to_sbd s z'' i'' i') $ (set $ (s $ z'')))); in HOLogic.mk_imp (HOLogic.mk_mem (mk_append (kl, HOLogic.mk_list sum_sbdT [mk_InN sbdTs k i']), mk_Lev ss (HOLogic.mk_Suc nat) i $ z), (Library.foldr1 HOLogic.mk_conj (@{map 4} mk_conjunct'' ks sets ss zs_copy))) end; in HOLogic.mk_imp (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z), Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct' ks (drop m setssAs'))) end; val goal = list_all_free (kl :: k :: zs @ zs_copy) (Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct ks zs)); val vars = Variable.add_free_names lthy goal []; val cts = map (SOME o Thm.cterm_of lthy) [Term.absfree nat' goal, nat]; val set_image_Lev = Goal.prove_sorry lthy vars [] (HOLogic.mk_Trueprop goal) (fn {context = ctxt, prems = _} => mk_set_image_Lev_tac ctxt cts Lev_0s Lev_Sucs rv_Nils rv_Conss from_to_sbd_thmss to_sbd_inj_thmss) |> Thm.close_derivation \<^here>; val set_image_Lev' = mk_specN (2 * n + 2) set_image_Lev; in map (fn i => map (fn i' => map (fn i'' => set_image_Lev' RS mk_conjunctN n i RS mp RS mk_conjunctN n i'' RS mp RS mk_conjunctN n i' RS mp) ks) ks) ks end; val mor_beh_thm = let val goal = HOLogic.mk_Trueprop (mk_mor active_UNIVs ss carTAs strTAs (map (mk_beh ss) ks)); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_mor_beh_tac ctxt m mor_def mor_cong_thm beh_defs carT_defs strT_defs isNode_defs to_sbd_inj_thmss from_to_sbd_thmss Lev_0s Lev_Sucs rv_Nils rv_Conss length_Lev_thms length_Lev'_thms rv_last_thmss set_Lev_thmsss set_image_Lev_thmsss set_mapss map_comp_id_thms map_cong0s) |> Thm.close_derivation \<^here> end; val timer = time (timer "Behavioral morphism"); val lsbisAs = map (mk_lsbis carTAs strTAs) ks; fun mk_str_final i = mk_univ (HOLogic.mk_comp (Term.list_comb (nth final_maps (i - 1), passive_ids @ map mk_proj lsbisAs), nth strTAs (i - 1))); val car_finals = map2 mk_quotient carTAs lsbisAs; val str_finals = map mk_str_final ks; val coalgT_set_thmss = map (map (fn thm => coalgT_thm RS thm)) coalg_set_thmss; val equiv_LSBIS_thms = map (fn thm => coalgT_thm RS thm) equiv_lsbis_thms; val congruent_str_final_thms = let fun mk_goal R final_map strT = HOLogic.mk_Trueprop (mk_congruent R (HOLogic.mk_comp (Term.list_comb (final_map, passive_ids @ map mk_proj lsbisAs), strT))); val goals = @{map 3} mk_goal lsbisAs final_maps strTAs; in @{map 4} (fn goal => fn lsbisE => fn map_comp_id => fn map_cong0 => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_congruent_str_final_tac ctxt m lsbisE map_comp_id map_cong0 equiv_LSBIS_thms) |> Thm.close_derivation \<^here>) goals lsbisE_thms map_comp_id_thms map_cong0s end; val coalg_final_thm = Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (mk_coalg car_finals str_finals)) (fn {context = ctxt, prems = _} => mk_coalg_final_tac ctxt m coalg_def congruent_str_final_thms equiv_LSBIS_thms set_mapss coalgT_set_thmss) |> Thm.close_derivation \<^here>; val mor_T_final_thm = Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (mk_mor carTAs strTAs car_finals str_finals (map mk_proj lsbisAs))) (fn {context = ctxt, prems = _} => mk_mor_T_final_tac ctxt mor_def congruent_str_final_thms equiv_LSBIS_thms) |> Thm.close_derivation \<^here>; val mor_final_thm = mor_comp_thm OF [mor_beh_thm, mor_T_final_thm]; val in_car_final_thms = map (fn thm => thm OF [mor_final_thm, UNIV_I]) mor_image'_thms; val timer = time (timer "Final coalgebra"); val ((T_names, (T_glob_infos, T_loc_infos)), lthy) = lthy |> @{fold_map 4} (fn b => fn mx => fn car_final => fn in_car_final => typedef (b, params, mx) car_final NONE (fn ctxt => EVERY' [rtac ctxt exI, rtac ctxt in_car_final] 1)) bs mixfixes car_finals in_car_final_thms |>> apsnd split_list o split_list; val Ts = map (fn name => Type (name, params')) T_names; fun mk_Ts passive = map (Term.typ_subst_atomic (passiveAs ~~ passive)) Ts; val Ts' = mk_Ts passiveBs; val Rep_Ts = map2 (fn info => fn T => Const (#Rep_name info, T --> treeQT)) T_glob_infos Ts; val Abs_Ts = map2 (fn info => fn T => Const (#Abs_name info, treeQT --> T)) T_glob_infos Ts; val Reps = map #Rep T_loc_infos; val Rep_injects = map #Rep_inject T_loc_infos; val Abs_inverses = map #Abs_inverse T_loc_infos; val timer = time (timer "THE TYPEDEFs & Rep/Abs thms"); val UNIVs = map HOLogic.mk_UNIV Ts; val FTs = mk_FTs (passiveAs @ Ts); val FTs_setss = mk_setss (passiveAs @ Ts); val map_FTs = map2 (fn Ds => mk_map_of_bnf Ds treeQTs (passiveAs @ Ts)) Dss bnfs; val unfold_fTs = map2 (curry op -->) activeAs Ts; val emptys = map (fn T => HOLogic.mk_set T []) passiveAs; val Zeros = map (fn empty => HOLogic.mk_tuple (map (fn U => absdummy U empty) Ts)) emptys; val hrecTs = map fastype_of Zeros; val (((zs, ss), (Jzs, Jzs')), _) = lthy |> mk_Frees "b" activeAs ||>> mk_Frees "s" sTs ||>> mk_Frees' "z" Ts; fun dtor_bind i = nth external_bs (i - 1) |> Binding.prefix_name (dtorN ^ "_"); val dtor_def_bind = rpair [] o Binding.concealed o Thm.def_binding o dtor_bind; fun dtor_spec rep str map_FT Jz Jz' = Term.absfree Jz' (Term.list_comb (map_FT, map HOLogic.id_const passiveAs @ Abs_Ts) $ (str $ (rep $ Jz))); val ((dtor_frees, (_, dtor_def_frees)), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> @{fold_map 6} (fn i => fn rep => fn str => fn mapx => fn Jz => fn Jz' => Local_Theory.define ((dtor_bind i, NoSyn), (dtor_def_bind i, dtor_spec rep str mapx Jz Jz'))) ks Rep_Ts str_finals map_FTs Jzs Jzs' |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; fun mk_dtors passive = map (Term.subst_atomic_types (map (Morphism.typ phi) params' ~~ (mk_params passive)) o Morphism.term phi) dtor_frees; val dtors = mk_dtors passiveAs; val dtor's = mk_dtors passiveBs; val dtor_defs = map (fn def => HOLogic.mk_obj_eq (Morphism.thm phi def) RS fun_cong) dtor_def_frees; val coalg_final_set_thmss = map (map (fn thm => coalg_final_thm RS thm)) coalg_set_thmss; val (mor_Rep_thm, mor_Abs_thm) = let val mor_Rep = Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (mk_mor UNIVs dtors car_finals str_finals Rep_Ts)) (fn {context = ctxt, prems = _} => mk_mor_Rep_tac ctxt (mor_def :: dtor_defs) Reps Abs_inverses coalg_final_set_thmss map_comp_id_thms map_cong0L_thms) |> Thm.close_derivation \<^here>; val mor_Abs = Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (mk_mor car_finals str_finals UNIVs dtors Abs_Ts)) (fn {context = ctxt, prems = _} => mk_mor_Abs_tac ctxt (mor_def :: dtor_defs) Abs_inverses) |> Thm.close_derivation \<^here>; in (mor_Rep, mor_Abs) end; val timer = time (timer "dtor definitions & thms"); fun unfold_bind i = nth external_bs (i - 1) |> Binding.prefix_name (dtor_unfoldN ^ "_"); val unfold_def_bind = rpair [] o Binding.concealed o Thm.def_binding o unfold_bind; fun unfold_spec abs f z = fold_rev (Term.absfree o Term.dest_Free) (ss @ [z]) (abs $ (f $ z)); val ((unfold_frees, (_, unfold_def_frees)), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> @{fold_map 4} (fn i => fn abs => fn f => fn z => Local_Theory.define ((unfold_bind i, NoSyn), (unfold_def_bind i, unfold_spec abs f z))) ks Abs_Ts (map (fn i => HOLogic.mk_comp (mk_proj (nth lsbisAs (i - 1)), mk_beh ss i)) ks) zs |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val unfolds = map (Morphism.term phi) unfold_frees; val unfold_names = map (fst o dest_Const) unfolds; fun mk_unfolds passives actives = @{map 3} (fn name => fn T => fn active => Const (name, Library.foldr (op -->) (map2 (curry op -->) actives (mk_FTs (passives @ actives)), active --> T))) unfold_names (mk_Ts passives) actives; fun mk_unfold Ts ss i = Term.list_comb (Const (nth unfold_names (i - 1), Library.foldr (op -->) (map fastype_of ss, domain_type (fastype_of (nth ss (i - 1))) --> nth Ts (i - 1))), ss); val unfold_defs = map (fn def => mk_unabs_def (n + 1) (HOLogic.mk_obj_eq (Morphism.thm phi def))) unfold_def_frees; val (((ss, TRs), unfold_fs), _) = lthy |> mk_Frees "s" sTs ||>> mk_Frees "r" (map (mk_relT o `I) Ts) ||>> mk_Frees "f" unfold_fTs; val mor_unfold_thm = let val Abs_inverses' = map2 (curry op RS) in_car_final_thms Abs_inverses; val morEs' = map (fn thm => (thm OF [mor_final_thm, UNIV_I]) RS sym) morE_thms; val goal = HOLogic.mk_Trueprop (mk_mor active_UNIVs ss UNIVs dtors (map (mk_unfold Ts ss) ks)); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_mor_unfold_tac ctxt m mor_UNIV_thm dtor_defs unfold_defs Abs_inverses' morEs' map_comp_id_thms map_cong0s) |> Thm.close_derivation \<^here> end; val dtor_unfold_thms = map (fn thm => (thm OF [mor_unfold_thm, UNIV_I]) RS sym) morE_thms; val (raw_coind_thms, raw_coind_thm) = let val prem = HOLogic.mk_Trueprop (mk_sbis UNIVs dtors TRs); val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (fn R => fn T => mk_leq R (Id_const T)) TRs Ts)); val vars = fold (Variable.add_free_names lthy) [prem, concl] []; in `split_conj_thm (Goal.prove_sorry lthy vars [] (Logic.mk_implies (prem, concl)) (fn {context = ctxt, prems = _} => mk_raw_coind_tac ctxt bis_def bis_cong_thm bis_O_thm bis_converse_thm bis_Gr_thm tcoalg_thm coalgT_thm mor_T_final_thm sbis_lsbis_thm lsbis_incl_thms incl_lsbis_thms equiv_LSBIS_thms mor_Rep_thm Rep_injects) |> Thm.close_derivation \<^here>) end; val (unfold_unique_mor_thms, unfold_unique_mor_thm) = let val prem = HOLogic.mk_Trueprop (mk_mor active_UNIVs ss UNIVs dtors unfold_fs); fun mk_fun_eq f i = HOLogic.mk_eq (f, mk_unfold Ts ss i); val unique = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_fun_eq unfold_fs ks)); val vars = fold (Variable.add_free_names lthy) [prem, unique] []; val bis_thm = tcoalg_thm RSN (2, tcoalg_thm RS bis_image2_thm); val mor_thm = mor_comp_thm OF [mor_final_thm, mor_Abs_thm]; val unique_mor = Goal.prove_sorry lthy vars [] (Logic.mk_implies (prem, unique)) (fn {context = ctxt, prems = _} => mk_unfold_unique_mor_tac ctxt raw_coind_thms bis_thm mor_thm unfold_defs) |> Thm.close_derivation \<^here>; in `split_conj_thm unique_mor end; val (dtor_unfold_unique_thms, dtor_unfold_unique_thm) = `split_conj_thm (split_conj_prems n (mor_UNIV_thm RS iffD2 RS unfold_unique_mor_thm)); val unfold_dtor_thms = map (fn thm => mor_id_thm RS thm RS sym) unfold_unique_mor_thms; val unfold_o_dtor_thms = let val mor = mor_comp_thm OF [mor_str_thm, mor_unfold_thm]; in map2 (fn unique => fn unfold_ctor => trans OF [mor RS unique, unfold_ctor]) unfold_unique_mor_thms unfold_dtor_thms end; val timer = time (timer "unfold definitions & thms"); val map_dtors = map2 (fn Ds => fn bnf => Term.list_comb (mk_map_of_bnf Ds (passiveAs @ Ts) (passiveAs @ FTs) bnf, map HOLogic.id_const passiveAs @ dtors)) Dss bnfs; fun ctor_bind i = nth external_bs (i - 1) |> Binding.prefix_name (ctorN ^ "_"); val ctor_def_bind = rpair [] o Binding.concealed o Thm.def_binding o ctor_bind; fun ctor_spec i = mk_unfold Ts map_dtors i; val ((ctor_frees, (_, ctor_def_frees)), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> fold_map (fn i => Local_Theory.define ((ctor_bind i, NoSyn), (ctor_def_bind i, ctor_spec i))) ks |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; fun mk_ctors params = map (Term.subst_atomic_types (map (Morphism.typ phi) params' ~~ params) o Morphism.term phi) ctor_frees; val ctors = mk_ctors params'; val ctor_defs = map (fn def => HOLogic.mk_obj_eq (Morphism.thm phi def)) ctor_def_frees; val ctor_o_dtor_thms = map2 (Local_Defs.fold lthy o single) ctor_defs unfold_o_dtor_thms; val dtor_o_ctor_thms = let fun mk_goal dtor ctor FT = mk_Trueprop_eq (HOLogic.mk_comp (dtor, ctor), HOLogic.id_const FT); val goals = @{map 3} mk_goal dtors ctors FTs; in @{map 5} (fn goal => fn ctor_def => fn unfold => fn map_comp_id => fn map_cong0L => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_dtor_o_ctor_tac ctxt ctor_def unfold map_comp_id map_cong0L unfold_o_dtor_thms) |> Thm.close_derivation \<^here>) goals ctor_defs dtor_unfold_thms map_comp_id_thms map_cong0L_thms end; val dtor_ctor_thms = map (fn thm => thm RS @{thm pointfree_idE}) dtor_o_ctor_thms; val ctor_dtor_thms = map (fn thm => thm RS @{thm pointfree_idE}) ctor_o_dtor_thms; val bij_dtor_thms = map2 (fn thm1 => fn thm2 => @{thm o_bij} OF [thm1, thm2]) ctor_o_dtor_thms dtor_o_ctor_thms; val inj_dtor_thms = map (fn thm => thm RS @{thm bij_is_inj}) bij_dtor_thms; val surj_dtor_thms = map (fn thm => thm RS @{thm bij_is_surj}) bij_dtor_thms; val dtor_nchotomy_thms = map (fn thm => thm RS @{thm surjD}) surj_dtor_thms; val dtor_inject_thms = map (fn thm => thm RS @{thm inj_eq}) inj_dtor_thms; val dtor_exhaust_thms = map (fn thm => thm RS exE) dtor_nchotomy_thms; val bij_ctor_thms = map2 (fn thm1 => fn thm2 => @{thm o_bij} OF [thm1, thm2]) dtor_o_ctor_thms ctor_o_dtor_thms; val inj_ctor_thms = map (fn thm => thm RS @{thm bij_is_inj}) bij_ctor_thms; val surj_ctor_thms = map (fn thm => thm RS @{thm bij_is_surj}) bij_ctor_thms; val ctor_nchotomy_thms = map (fn thm => thm RS @{thm surjD}) surj_ctor_thms; val ctor_inject_thms = map (fn thm => thm RS @{thm inj_eq}) inj_ctor_thms; val ctor_exhaust_thms = map (fn thm => thm RS exE) ctor_nchotomy_thms; val timer = time (timer "ctor definitions & thms"); val (((((Jzs, Jzs_copy), Jzs1), Jzs2), phis), _) = lthy |> mk_Frees "z" Ts ||>> mk_Frees "z'" Ts ||>> mk_Frees "z1" Ts ||>> mk_Frees "z2" Ts ||>> mk_Frees "P" (map2 mk_pred2T Ts Ts); val (coinduct_params, dtor_coinduct_thm) = let val rels = map (Term.subst_atomic_types ((activeAs ~~ Ts) @ (activeBs ~~ Ts))) relsAsBs; fun mk_concl phi z1 z2 = HOLogic.mk_imp (phi $ z1 $ z2, HOLogic.mk_eq (z1, z2)); val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (@{map 3} mk_concl phis Jzs1 Jzs2)); fun mk_rel_prem phi dtor rel Jz Jz_copy = let val concl = Term.list_comb (rel, passive_eqs @ phis) $ (dtor $ Jz) $ (dtor $ Jz_copy); in HOLogic.mk_Trueprop (list_all_free [Jz, Jz_copy] (HOLogic.mk_imp (phi $ Jz $ Jz_copy, concl))) end; val rel_prems = @{map 5} mk_rel_prem phis dtors rels Jzs Jzs_copy; val dtor_coinduct_goal = Logic.list_implies (rel_prems, concl); val dtor_coinduct = Variable.add_free_names lthy dtor_coinduct_goal [] |> (fn vars => Goal.prove_sorry lthy vars [] dtor_coinduct_goal (fn {context = ctxt, prems = _} => mk_dtor_coinduct_tac ctxt m raw_coind_thm bis_rel_thm rel_congs)) |> Thm.close_derivation \<^here>; in (rev (Term.add_tfrees dtor_coinduct_goal []), dtor_coinduct) end; val timer = time (timer "coinduction"); fun mk_dtor_map_DEADID_thm dtor_inject map_id0 = trans OF [iffD2 OF [dtor_inject, id_apply], map_id0 RS sym]; fun mk_dtor_map_unique_DEADID_thm () = let val (funs, algs) = HOLogic.conjuncts (HOLogic.dest_Trueprop (Thm.concl_of dtor_unfold_unique_thm)) |> map_split HOLogic.dest_eq ||> snd o strip_comb o hd |> @{apply 2} (map (fst o dest_Var)); fun mk_fun_insts T ix = Thm.cterm_of lthy (Var (ix, T --> T)); val theta = (funs ~~ @{map 2} mk_fun_insts Ts funs) @ (algs ~~ map (Thm.cterm_of lthy) dtors); val dtor_unfold_dtors = (dtor_unfold_unique_thm OF map (fn thm => mk_trans (thm RS @{thm arg_cong2[of _ _ _ _ "(\)", OF _ refl]}) @{thm trans[OF id_o o_id[symmetric]]}) map_id0s) |> split_conj_thm |> map mk_sym; in infer_instantiate lthy theta dtor_unfold_unique_thm |> Morphism.thm (Local_Theory.target_morphism lthy) |> unfold_thms lthy dtor_unfold_dtors |> (fn thm => thm OF replicate n sym) end; (* thm trans[OF x.dtor_unfold_unique x.dtor_unfold_unique[symmetric, OF trans[OF arg_cong2[of _ _ _ _ "(o)", OF pre_x.map_id0 refl] trans[OF id_o o_id[symmetric]]]], OF sym] *) fun mk_dtor_Jrel_DEADID_thm dtor_inject bnf = trans OF [rel_eq_of_bnf bnf RS @{thm predicate2_eqD}, dtor_inject] RS sym; val JphiTs = map2 mk_pred2T passiveAs passiveBs; val Jpsi1Ts = map2 mk_pred2T passiveAs passiveCs; val Jpsi2Ts = map2 mk_pred2T passiveCs passiveBs; val prodTsTs' = map2 (curry HOLogic.mk_prodT) Ts Ts'; val fstsTsTs' = map fst_const prodTsTs'; val sndsTsTs' = map snd_const prodTsTs'; val activephiTs = map2 mk_pred2T activeAs activeBs; val activeJphiTs = map2 mk_pred2T Ts Ts'; val rels = map2 (fn Ds => mk_rel_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs; val ((((Jzs, Jz's), Jphis), activeJphis), _) = lthy |> mk_Frees "z" Ts ||>> mk_Frees "y" Ts' ||>> mk_Frees "R" JphiTs ||>> mk_Frees "JR" activeJphiTs; fun mk_Jrel_DEADID_coinduct_thm () = mk_xtor_rel_co_induct_thm Greatest_FP rels activeJphis (map HOLogic.eq_const Ts) Jphis Jzs Jz's dtors dtor's (fn {context = ctxt, prems} => (unfold_thms_tac ctxt @{thms le_fun_def le_bool_def all_simps(1,2)[symmetric]} THEN REPEAT_DETERM (rtac ctxt allI 1) THEN rtac ctxt (dtor_coinduct_thm OF prems) 1)) lthy; (*register new codatatypes as BNFs*) val (timer, Jbnfs, (dtor_Jmap_o_thms, dtor_Jmap_thms), dtor_Jmap_unique_thm, dtor_Jset_thmss', dtor_Jrel_thms, Jrel_coinduct_thm, Jbnf_notes, dtor_Jset_induct_thms, lthy) = if m = 0 then (timer, replicate n DEADID_bnf, map_split (`(mk_pointfree2 lthy)) (map2 mk_dtor_map_DEADID_thm dtor_inject_thms map_ids), mk_dtor_map_unique_DEADID_thm (), replicate n [], map2 mk_dtor_Jrel_DEADID_thm dtor_inject_thms bnfs, mk_Jrel_DEADID_coinduct_thm (), [], [], lthy) else let val fTs = map2 (curry op -->) passiveAs passiveBs; val gTs = map2 (curry op -->) passiveBs passiveCs; val uTs = map2 (curry op -->) Ts Ts'; val (((((nat, nat'), (Jzs, Jzs')), (hrecs, hrecs')), (fs, fs')), _) = lthy |> yield_singleton (apfst (op ~~) oo mk_Frees' "n") HOLogic.natT ||>> mk_Frees' "z" Ts ||>> mk_Frees' "rec" hrecTs ||>> mk_Frees' "f" fTs; val map_FTFT's = map2 (fn Ds => mk_map_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs; fun mk_maps ATs BTs Ts mk_T = map2 (fn Ds => mk_map_of_bnf Ds (ATs @ Ts) (BTs @ map mk_T Ts)) Dss bnfs; fun mk_Fmap mk_const fs Ts Fmap = Term.list_comb (Fmap, fs @ map mk_const Ts); fun mk_map mk_const mk_T Ts fs Ts' dtors mk_maps = mk_unfold Ts' (map2 (fn dtor => fn Fmap => HOLogic.mk_comp (mk_Fmap mk_const fs Ts Fmap, dtor)) dtors (mk_maps Ts mk_T)); val mk_map_id = mk_map HOLogic.id_const I; val mk_mapsAB = mk_maps passiveAs passiveBs; val fs_maps = map (mk_map_id Ts fs Ts' dtors mk_mapsAB) ks; val set_bss = map (flat o map2 (fn B => fn b => if member (op =) resDs (TFree B) then [] else [b]) resBs) set_bss0; fun col_bind j = mk_internal_b (colN ^ (if m = 1 then "" else string_of_int j)); val col_def_bind = rpair [] o Thm.def_binding o col_bind; fun col_spec j Zero hrec hrec' = let fun mk_Suc dtor sets z z' = let val (set, sets) = apfst (fn xs => nth xs (j - 1)) (chop m sets); fun mk_UN set k = mk_UNION (set $ (dtor $ z)) (mk_nthN n hrec k); in Term.absfree z' (mk_union (set $ (dtor $ z), Library.foldl1 mk_union (map2 mk_UN sets ks))) end; val Suc = Term.absdummy HOLogic.natT (Term.absfree hrec' (HOLogic.mk_tuple (@{map 4} mk_Suc dtors FTs_setss Jzs Jzs'))); in mk_rec_nat Zero Suc end; val ((col_frees, (_, col_def_frees)), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> @{fold_map 4} (fn j => fn Zero => fn hrec => fn hrec' => Local_Theory.define ((col_bind j, NoSyn), (col_def_bind j, col_spec j Zero hrec hrec'))) ls Zeros hrecs hrecs' |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val col_defs = map (fn def => HOLogic.mk_obj_eq (Morphism.thm phi def)) col_def_frees; val cols = map (fst o Term.dest_Const o Morphism.term phi) col_frees; fun mk_col Ts nat i j T = let val hrecT = HOLogic.mk_tupleT (map (fn U => U --> HOLogic.mk_setT T) Ts) val colT = HOLogic.natT --> hrecT; in mk_nthN n (Term.list_comb (Const (nth cols (j - 1), colT), [nat])) i end; val col_0ss = mk_rec_simps n @{thm rec_nat_0_imp} col_defs; val col_Sucss = mk_rec_simps n @{thm rec_nat_Suc_imp} col_defs; val col_0ss' = transpose col_0ss; val col_Sucss' = transpose col_Sucss; fun mk_set Ts i j T = Abs (Name.uu, nth Ts (i - 1), mk_UNION (HOLogic.mk_UNIV HOLogic.natT) (Term.absfree nat' (mk_col Ts nat i j T $ Bound 1))); val setss = map (fn i => map2 (mk_set Ts i) ls passiveAs) ks; val (Jbnf_consts, lthy) = @{fold_map 8} (fn b => fn map_b => fn rel_b => fn pred_b => fn set_bs => fn mapx => fn sets => fn T => fn lthy => define_bnf_consts Hardly_Inline (user_policy Note_Some lthy) false (SOME deads) map_b rel_b pred_b set_bs (((((((b, T), fold_rev Term.absfree fs' mapx), sets), sbd), [Const (\<^const_name>\undefined\, T)]), NONE), NONE) lthy) bs map_bs rel_bs pred_bs set_bss fs_maps setss Ts lthy; val (_, Jconsts, Jconst_defs, mk_Jconsts) = @{split_list 4} Jbnf_consts; val (_, Jsetss, Jbds_Ds, _, _, _) = @{split_list 6} Jconsts; val (Jmap_defs, Jset_defss, Jbd_defs, _, Jrel_defs, Jpred_defs) = @{split_list 6} Jconst_defs; val (mk_Jmaps_Ds, mk_Jt_Ds, _, mk_Jrels_Ds, mk_Jpreds_Ds, _, _) = @{split_list 7} mk_Jconsts; val Jrel_unabs_defs = map (fn def => mk_unabs_def m (HOLogic.mk_obj_eq def)) Jrel_defs; val Jpred_unabs_defs = map (fn def => mk_unabs_def m (HOLogic.mk_obj_eq def)) Jpred_defs; val Jset_defs = flat Jset_defss; fun mk_Jmaps As Bs = map (fn mk => mk deads As Bs) mk_Jmaps_Ds; fun mk_Jsetss As = map2 (fn mk => fn Jsets => map (mk deads As) Jsets) mk_Jt_Ds Jsetss; val Jbds = map2 (fn mk => mk deads passiveAs) mk_Jt_Ds Jbds_Ds; fun mk_Jrels As Bs = map (fn mk => mk deads As Bs) mk_Jrels_Ds; fun mk_Jpreds As = map (fn mk => mk deads As) mk_Jpreds_Ds; val Jmaps = mk_Jmaps passiveAs passiveBs; val (Jsetss_by_range, Jsetss_by_bnf) = `transpose (mk_Jsetss passiveAs); val timer = time (timer "bnf constants for the new datatypes"); val ((((((((((((((((((((ys, ys'), (nat, nat')), (Jzs, Jzs')), Jz's), Jzs_copy), Jz's_copy), dtor_set_induct_phiss), Jphis), Jpsi1s), Jpsi2s), activeJphis), fs), fs_copy), gs), us), (Jys, Jys')), (Jys_copy, Jys'_copy)), (ys_copy, ys'_copy)), Kss), names_lthy) = lthy |> mk_Frees' "y" passiveAs ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "n") HOLogic.natT ||>> mk_Frees' "z" Ts ||>> mk_Frees "y" Ts' ||>> mk_Frees "z'" Ts ||>> mk_Frees "y'" Ts' ||>> mk_Freess "P" (map (fn A => map (mk_pred2T A) Ts) passiveAs) ||>> mk_Frees "R" JphiTs ||>> mk_Frees "R" Jpsi1Ts ||>> mk_Frees "Q" Jpsi2Ts ||>> mk_Frees "JR" activeJphiTs ||>> mk_Frees "f" fTs ||>> mk_Frees "f" fTs ||>> mk_Frees "g" gTs ||>> mk_Frees "u" uTs ||>> mk_Frees' "b" Ts' ||>> mk_Frees' "b" Ts' ||>> mk_Frees' "y" passiveAs ||>> mk_Freess "K" (map (fn AT => map (fn T => T --> AT) Ts) ATs); val fs_Jmaps = map (fn m => Term.list_comb (m, fs)) Jmaps; val fs_copy_Jmaps = map (fn m => Term.list_comb (m, fs_copy)) Jmaps; val gs_Jmaps = map (fn m => Term.list_comb (m, gs)) (mk_Jmaps passiveBs passiveCs); val fgs_Jmaps = map (fn m => Term.list_comb (m, map2 (curry HOLogic.mk_comp) gs fs)) (mk_Jmaps passiveAs passiveCs); val (dtor_Jmap_thms, Jmap_thms) = let fun mk_goal fs_Jmap map dtor dtor' = mk_Trueprop_eq (HOLogic.mk_comp (dtor', fs_Jmap), HOLogic.mk_comp (Term.list_comb (map, fs @ fs_Jmaps), dtor)); val goals = @{map 4} mk_goal fs_Jmaps map_FTFT's dtors dtor's; val maps = @{map 5} (fn goal => fn unfold => fn map_comp => fn map_cong0 => fn map_arg_cong => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jmap_defs THEN mk_map_tac ctxt m n map_arg_cong unfold map_comp map_cong0)) |> Thm.close_derivation \<^here>) goals dtor_unfold_thms map_comps map_cong0s map_arg_cong_thms; in map_split (fn thm => (thm RS @{thm comp_eq_dest}, thm)) maps end; val (dtor_Jmap_unique_thms, dtor_Jmap_unique_thm) = let fun mk_prem u map dtor dtor' = mk_Trueprop_eq (HOLogic.mk_comp (dtor', u), HOLogic.mk_comp (Term.list_comb (map, fs @ us), dtor)); val prems = @{map 4} mk_prem us map_FTFT's dtors dtor's; val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_eq) us fs_Jmaps)); val vars = fold (Variable.add_free_names lthy) (goal :: prems) []; in `split_conj_thm (Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, goal)) (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jmap_defs THEN mk_dtor_map_unique_tac ctxt dtor_unfold_unique_thm sym_map_comps) |> Thm.close_derivation \<^here>) end; val Jmap_comp0_thms = let val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (@{map 3} (fn fmap => fn gmap => fn fgmap => HOLogic.mk_eq (HOLogic.mk_comp (gmap, fmap), fgmap)) fs_Jmaps gs_Jmaps fgs_Jmaps)) val vars = Variable.add_free_names lthy goal []; in split_conj_thm (Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_map_comp0_tac ctxt Jmap_thms map_comp0s dtor_Jmap_unique_thm) |> Thm.close_derivation \<^here>) end; val timer = time (timer "map functions for the new codatatypes"); val Jset_minimal_thms = let fun mk_passive_prem set dtor x K = Logic.all x (HOLogic.mk_Trueprop (mk_leq (set $ (dtor $ x)) (K $ x))); fun mk_active_prem dtor x1 K1 set x2 K2 = fold_rev Logic.all [x1, x2] (Logic.mk_implies (mk_Trueprop_mem (x2, set $ (dtor $ x1)), HOLogic.mk_Trueprop (mk_leq (K2 $ x2) (K1 $ x1)))); val premss = map2 (fn j => fn Ks => @{map 4} mk_passive_prem (map (fn xs => nth xs (j - 1)) FTs_setss) dtors Jzs Ks @ flat (@{map 4} (fn sets => fn s => fn x1 => fn K1 => @{map 3} (mk_active_prem s x1 K1) (drop m sets) Jzs_copy Ks) FTs_setss dtors Jzs Ks)) ls Kss; val col_minimal_thms = let fun mk_conjunct j T i K x = mk_leq (mk_col Ts nat i j T $ x) (K $ x); fun mk_concl j T Ks = list_all_free Jzs (Library.foldr1 HOLogic.mk_conj (@{map 3} (mk_conjunct j T) ks Ks Jzs)); val concls = @{map 3} mk_concl ls passiveAs Kss; val goals = map2 (fn prems => fn concl => Logic.list_implies (prems, HOLogic.mk_Trueprop concl)) premss concls val ctss = map (fn phi => map (SOME o Thm.cterm_of lthy) [Term.absfree nat' phi, nat]) concls; in @{map 4} (fn goal => fn cts => fn col_0s => fn col_Sucs => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_col_minimal_tac ctxt m cts col_0s col_Sucs)) |> Thm.close_derivation \<^here>) goals ctss col_0ss' col_Sucss' end; fun mk_conjunct set K x = mk_leq (set $ x) (K $ x); fun mk_concl sets Ks = Library.foldr1 HOLogic.mk_conj (@{map 3} mk_conjunct sets Ks Jzs); val concls = map2 mk_concl Jsetss_by_range Kss; val goals = map2 (fn prems => fn concl => Logic.list_implies (prems, HOLogic.mk_Trueprop concl)) premss concls; in map2 (fn goal => fn col_minimal => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jset_defs THEN mk_Jset_minimal_tac ctxt n col_minimal)) |> Thm.close_derivation \<^here>) goals col_minimal_thms end; val (dtor_Jset_incl_thmss, dtor_set_Jset_incl_thmsss) = let fun mk_set_incl_Jset dtor x set Jset = HOLogic.mk_Trueprop (mk_leq (set $ (dtor $ x)) (Jset $ x)); fun mk_set_Jset_incl_Jset dtor x y set Jset1 Jset2 = Logic.mk_implies (mk_Trueprop_mem (x, set $ (dtor $ y)), HOLogic.mk_Trueprop (mk_leq (Jset1 $ x) (Jset2 $ y))); val set_incl_Jset_goalss = @{map 4} (fn dtor => fn x => fn sets => fn Jsets => map2 (mk_set_incl_Jset dtor x) (take m sets) Jsets) dtors Jzs FTs_setss Jsetss_by_bnf; (*x(k) : F(i)set(m+k) (dtor(i) y(i)) ==> J(k)set(j) x(k) <= J(i)set(j) y(i)*) val set_Jset_incl_Jset_goalsss = @{map 4} (fn dtori => fn yi => fn sets => fn Jsetsi => @{map 3} (fn xk => fn set => fn Jsetsk => map2 (mk_set_Jset_incl_Jset dtori xk yi set) Jsetsk Jsetsi) Jzs_copy (drop m sets) Jsetss_by_bnf) dtors Jzs FTs_setss Jsetss_by_bnf; in (map2 (fn goals => fn rec_Sucs => map2 (fn goal => fn rec_Suc => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jset_defs THEN mk_set_incl_Jset_tac ctxt rec_Suc)) |> Thm.close_derivation \<^here>) goals rec_Sucs) set_incl_Jset_goalss col_Sucss, map2 (fn goalss => fn rec_Sucs => map2 (fn k => fn goals => map2 (fn goal => fn rec_Suc => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jset_defs THEN mk_set_Jset_incl_Jset_tac ctxt n rec_Suc k)) |> Thm.close_derivation \<^here>) goals rec_Sucs) ks goalss) set_Jset_incl_Jset_goalsss col_Sucss) end; val set_incl_Jset_thmss' = transpose dtor_Jset_incl_thmss; val set_Jset_incl_Jset_thmsss' = transpose (map transpose dtor_set_Jset_incl_thmsss); val set_Jset_thmss = map (map (fn thm => thm RS @{thm set_mp})) dtor_Jset_incl_thmss; val set_Jset_Jset_thmsss = map (map (map (fn thm => thm RS @{thm set_mp}))) dtor_set_Jset_incl_thmsss; val set_Jset_thmss' = transpose set_Jset_thmss; val set_Jset_Jset_thmsss' = transpose (map transpose set_Jset_Jset_thmsss); val dtor_Jset_induct_thms = let val incls = maps (map (fn thm => thm RS @{thm subset_Collect_iff})) dtor_Jset_incl_thmss @ @{thms subset_Collect_iff[OF subset_refl]}; val cTs = map (SOME o Thm.ctyp_of lthy) params'; fun mk_induct_tinst phis jsets y y' = @{map 4} (fn phi => fn jset => fn Jz => fn Jz' => SOME (Thm.cterm_of lthy (Term.absfree Jz' (HOLogic.mk_Collect (fst y', snd y', HOLogic.mk_conj (HOLogic.mk_mem (y, jset $ Jz), phi $ y $ Jz)))))) phis jsets Jzs Jzs'; in @{map 6} (fn set_minimal => fn set_set_inclss => fn jsets => fn y => fn y' => fn phis => ((set_minimal |> Thm.instantiate' cTs (mk_induct_tinst phis jsets y y') |> unfold_thms lthy incls) OF (replicate n ballI @ maps (map (fn thm => thm RS @{thm subset_CollectI})) set_set_inclss)) |> singleton (Proof_Context.export names_lthy lthy) |> rule_by_tactic lthy (ALLGOALS (TRY o etac lthy asm_rl))) Jset_minimal_thms set_Jset_incl_Jset_thmsss' Jsetss_by_range ys ys' dtor_set_induct_phiss end; val (dtor_Jset_thmss', dtor_Jset_thmss) = let fun mk_simp_goal relate pas_set act_sets sets dtor z set = relate (set $ z, mk_union (pas_set $ (dtor $ z), Library.foldl1 mk_union (map2 (fn X => mk_UNION (X $ (dtor $ z))) act_sets sets))); fun mk_goals eq = map2 (fn i => fn sets => @{map 4} (fn Fsets => mk_simp_goal eq (nth Fsets (i - 1)) (drop m Fsets) sets) FTs_setss dtors Jzs sets) ls Jsetss_by_range; val le_goals = map (HOLogic.mk_Trueprop o Library.foldr1 HOLogic.mk_conj) (mk_goals (uncurry mk_leq)); val set_le_thmss = map split_conj_thm (@{map 4} (fn goal => fn Jset_minimal => fn set_Jsets => fn set_Jset_Jsetss => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_set_le_tac ctxt n Jset_minimal set_Jsets set_Jset_Jsetss)) |> Thm.close_derivation \<^here>) le_goals Jset_minimal_thms set_Jset_thmss' set_Jset_Jset_thmsss'); val ge_goalss = map (map HOLogic.mk_Trueprop) (mk_goals (uncurry mk_leq o swap)); val set_ge_thmss = @{map 3} (@{map 3} (fn goal => fn set_incl_Jset => fn set_Jset_incl_Jsets => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_set_ge_tac ctxt n set_incl_Jset set_Jset_incl_Jsets)) |> Thm.close_derivation \<^here>)) ge_goalss set_incl_Jset_thmss' set_Jset_incl_Jset_thmsss' in map2 (map2 (fn le => fn ge => equalityI OF [le, ge])) set_le_thmss set_ge_thmss |> `transpose end; val timer = time (timer "set functions for the new codatatypes"); val colss = map2 (fn j => fn T => map (fn i => mk_col Ts nat i j T) ks) ls passiveAs; val colss' = map2 (fn j => fn T => map (fn i => mk_col Ts' nat i j T) ks) ls passiveBs; val col_natural_thmss = let fun mk_col_natural f map z col col' = HOLogic.mk_eq (mk_image f $ (col $ z), col' $ (map $ z)); fun mk_goal f cols cols' = list_all_free Jzs (Library.foldr1 HOLogic.mk_conj (@{map 4} (mk_col_natural f) fs_Jmaps Jzs cols cols')); val goals = @{map 3} mk_goal fs colss colss'; val ctss = map (fn phi => map (SOME o Thm.cterm_of lthy) [Term.absfree nat' phi, nat]) goals; val thms = @{map 4} (fn goal => fn cts => fn rec_0s => fn rec_Sucs => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] (HOLogic.mk_Trueprop goal) (fn {context = ctxt, prems = _} => mk_col_natural_tac ctxt cts rec_0s rec_Sucs dtor_Jmap_thms set_mapss)) |> Thm.close_derivation \<^here>) goals ctss col_0ss' col_Sucss'; in map (split_conj_thm o mk_specN n) thms end; val col_bd_thmss = let fun mk_col_bd z col bd = mk_ordLeq (mk_card_of (col $ z)) bd; fun mk_goal bds cols = list_all_free Jzs (Library.foldr1 HOLogic.mk_conj (@{map 3} mk_col_bd Jzs cols bds)); val goals = map (mk_goal Jbds) colss; val ctss = map (fn phi => map (SOME o Thm.cterm_of lthy) [Term.absfree nat' phi, nat]) (map (mk_goal (replicate n sbd)) colss); val thms = @{map 5} (fn j => fn goal => fn cts => fn rec_0s => fn rec_Sucs => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] (HOLogic.mk_Trueprop goal) (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jbd_defs THEN mk_col_bd_tac ctxt m j cts rec_0s rec_Sucs sbd_Card_order sbd_Cinfinite set_sbdss)) |> Thm.close_derivation \<^here>) ls goals ctss col_0ss' col_Sucss'; in map (split_conj_thm o mk_specN n) thms end; val map_cong0_thms = let val cTs = map (SOME o Thm.ctyp_of lthy o Term.typ_subst_atomic (passiveAs ~~ passiveBs) o TFree) coinduct_params; fun mk_prem z set f g y y' = mk_Ball (set $ z) (Term.absfree y' (HOLogic.mk_eq (f $ y, g $ y))); fun mk_prems sets z = Library.foldr1 HOLogic.mk_conj (@{map 5} (mk_prem z) sets fs fs_copy ys ys') fun mk_map_cong0 sets z fmap gmap = HOLogic.mk_imp (mk_prems sets z, HOLogic.mk_eq (fmap $ z, gmap $ z)); fun mk_coind_body sets (x, T) z fmap gmap y y_copy = HOLogic.mk_conj (HOLogic.mk_mem (z, HOLogic.mk_Collect (x, T, mk_prems sets z)), HOLogic.mk_conj (HOLogic.mk_eq (y, fmap $ z), HOLogic.mk_eq (y_copy, gmap $ z))) fun mk_cphi sets (z' as (x, T)) z fmap gmap y' y y'_copy y_copy = HOLogic.mk_exists (x, T, mk_coind_body sets z' z fmap gmap y y_copy) |> Term.absfree y'_copy |> Term.absfree y' |> Thm.cterm_of lthy; val cphis = @{map 9} mk_cphi Jsetss_by_bnf Jzs' Jzs fs_Jmaps fs_copy_Jmaps Jys' Jys Jys'_copy Jys_copy; val coinduct = Thm.instantiate' cTs (map SOME cphis) dtor_coinduct_thm; val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (@{map 4} mk_map_cong0 Jsetss_by_bnf Jzs fs_Jmaps fs_copy_Jmaps)); val vars = Variable.add_free_names lthy goal []; val thm = Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_mcong_tac ctxt m (rtac ctxt coinduct) map_comps dtor_Jmap_thms map_cong0s set_mapss set_Jset_thmss set_Jset_Jset_thmsss in_rels) |> Thm.close_derivation \<^here>; in split_conj_thm thm end; val in_Jrels = map (fn def => trans OF [def, @{thm OO_Grp_alt}] RS @{thm predicate2_eqD}) Jrel_unabs_defs; val Jrels = mk_Jrels passiveAs passiveBs; val Jpreds = mk_Jpreds passiveAs; val Jrelphis = map (fn rel => Term.list_comb (rel, Jphis)) Jrels; val relphis = map (fn rel => Term.list_comb (rel, Jphis @ Jrelphis)) rels; val Jrelpsi1s = map (fn rel => Term.list_comb (rel, Jpsi1s)) (mk_Jrels passiveAs passiveCs); val Jrelpsi2s = map (fn rel => Term.list_comb (rel, Jpsi2s)) (mk_Jrels passiveCs passiveBs); val Jrelpsi12s = map (fn rel => Term.list_comb (rel, map2 (curry mk_rel_compp) Jpsi1s Jpsi2s)) Jrels; val dtor_Jrel_thms = let fun mk_goal Jz Jz' dtor dtor' Jrelphi relphi = mk_Trueprop_eq (Jrelphi $ Jz $ Jz', relphi $ (dtor $ Jz) $ (dtor' $ Jz')); val goals = @{map 6} mk_goal Jzs Jz's dtors dtor's Jrelphis relphis; in @{map 12} (fn i => fn goal => fn in_rel => fn map_comp0 => fn map_cong0 => fn dtor_map => fn dtor_sets => fn dtor_inject => fn dtor_ctor => fn set_map0s => fn dtor_set_incls => fn dtor_set_set_inclss => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_dtor_rel_tac ctxt in_Jrels i in_rel map_comp0 map_cong0 dtor_map dtor_sets dtor_inject dtor_ctor set_map0s dtor_set_incls dtor_set_set_inclss)) |> Thm.close_derivation \<^here>) ks goals in_rels map_comps map_cong0s dtor_Jmap_thms dtor_Jset_thmss' dtor_inject_thms dtor_ctor_thms set_mapss dtor_Jset_incl_thmss dtor_set_Jset_incl_thmsss end; val passiveABs = map2 (curry HOLogic.mk_prodT) passiveAs passiveBs; val zip_ranTs = passiveABs @ prodTsTs'; val allJphis = Jphis @ activeJphis; val zipFTs = mk_FTs zip_ranTs; val zipTs = @{map 3} (fn T => fn T' => fn FT => T --> T' --> FT) Ts Ts' zipFTs; val zip_zTs = mk_Ts passiveABs; val (((zips, (abs, abs')), (zip_zs, zip_zs')), _) = names_lthy |> mk_Frees "zip" zipTs ||>> mk_Frees' "ab" passiveABs ||>> mk_Frees' "z" zip_zTs; val Iphi_sets = map2 (fn phi => fn T => HOLogic.Collect_const T $ HOLogic.mk_case_prod phi) allJphis zip_ranTs; val in_phis = map2 (mk_in Iphi_sets) (mk_setss zip_ranTs) zipFTs; val fstABs = map fst_const passiveABs; val all_fsts = fstABs @ fstsTsTs'; val map_all_fsts = map2 (fn Ds => fn bnf => Term.list_comb (mk_map_of_bnf Ds zip_ranTs (passiveAs @ Ts) bnf, all_fsts)) Dss bnfs; val Jmap_fsts = map2 (fn map => fn T => if m = 0 then HOLogic.id_const T else Term.list_comb (map, fstABs)) (mk_Jmaps passiveABs passiveAs) Ts; val sndABs = map snd_const passiveABs; val all_snds = sndABs @ sndsTsTs'; val map_all_snds = map2 (fn Ds => fn bnf => Term.list_comb (mk_map_of_bnf Ds zip_ranTs (passiveBs @ Ts') bnf, all_snds)) Dss bnfs; val Jmap_snds = map2 (fn map => fn T => if m = 0 then HOLogic.id_const T else Term.list_comb (map, sndABs)) (mk_Jmaps passiveABs passiveBs) Ts; val zip_unfolds = map (mk_unfold zip_zTs (map HOLogic.mk_case_prod zips)) ks; val zip_setss = mk_Jsetss passiveABs |> transpose; fun Jrel_coinduct_tac {context = ctxt, prems = CIHs} = let fun mk_helper_prem phi in_phi zip x y map map' dtor dtor' = let val zipxy = zip $ x $ y; in HOLogic.mk_Trueprop (list_all_free [x, y] (HOLogic.mk_imp (phi $ x $ y, HOLogic.mk_conj (HOLogic.mk_mem (zipxy, in_phi), HOLogic.mk_conj (HOLogic.mk_eq (map $ zipxy, dtor $ x), HOLogic.mk_eq (map' $ zipxy, dtor' $ y)))))) end; val helper_prems = @{map 9} mk_helper_prem activeJphis in_phis zips Jzs Jz's map_all_fsts map_all_snds dtors dtor's; fun mk_helper_coind_phi fst phi x alt y map zip_unfold = list_exists_free [if fst then y else x] (HOLogic.mk_conj (phi $ x $ y, HOLogic.mk_eq (alt, map $ (zip_unfold $ HOLogic.mk_prod (x, y))))) val coind1_phis = @{map 6} (mk_helper_coind_phi true) activeJphis Jzs Jzs_copy Jz's Jmap_fsts zip_unfolds; val coind2_phis = @{map 6} (mk_helper_coind_phi false) activeJphis Jzs Jz's_copy Jz's Jmap_snds zip_unfolds; fun mk_cts zs z's phis = @{map 3} (fn z => fn z' => fn phi => SOME (Thm.cterm_of lthy (fold_rev (Term.absfree o Term.dest_Free) [z', z] phi))) zs z's phis @ map (SOME o Thm.cterm_of lthy) (splice z's zs); val cts1 = mk_cts Jzs Jzs_copy coind1_phis; val cts2 = mk_cts Jz's Jz's_copy coind2_phis; fun mk_helper_coind_concl z alt coind_phi = HOLogic.mk_imp (coind_phi, HOLogic.mk_eq (alt, z)); val helper_coind1_concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (@{map 3} mk_helper_coind_concl Jzs Jzs_copy coind1_phis)); val helper_coind2_concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (@{map 3} mk_helper_coind_concl Jz's Jz's_copy coind2_phis)); fun mk_helper_coind_thms fst concl cts = let val vars = fold (Variable.add_free_names lthy) (concl :: helper_prems) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (helper_prems, concl)) (fn {context = ctxt, prems = _} => mk_rel_coinduct_coind_tac ctxt fst m (infer_instantiate' ctxt cts dtor_coinduct_thm) ks map_comps map_cong0s map_arg_cong_thms set_mapss dtor_unfold_thms dtor_Jmap_thms in_rels) |> Thm.close_derivation \<^here> |> split_conj_thm end; val helper_coind1_thms = mk_helper_coind_thms true helper_coind1_concl cts1; val helper_coind2_thms = mk_helper_coind_thms false helper_coind2_concl cts2; fun mk_helper_ind_phi phi ab fst snd z active_phi x y zip_unfold = list_all_free [x, y] (HOLogic.mk_imp (HOLogic.mk_conj (active_phi $ x $ y, HOLogic.mk_eq (z, zip_unfold $ HOLogic.mk_prod (x, y))), phi $ (fst $ ab) $ (snd $ ab))); val helper_ind_phiss = @{map 4} (fn Jphi => fn ab => fn fst => fn snd => @{map 5} (mk_helper_ind_phi Jphi ab fst snd) zip_zs activeJphis Jzs Jz's zip_unfolds) Jphis abs fstABs sndABs; val ctss = map2 (fn ab' => fn phis => map2 (fn z' => fn phi => SOME (Thm.cterm_of lthy (Term.absfree ab' (Term.absfree z' phi)))) zip_zs' phis @ map (SOME o Thm.cterm_of lthy) zip_zs) abs' helper_ind_phiss; fun mk_helper_ind_concl ab' z ind_phi set = mk_Ball (set $ z) (Term.absfree ab' ind_phi); val mk_helper_ind_concls = @{map 3} (fn ab' => fn ind_phis => fn zip_sets => @{map 3} (mk_helper_ind_concl ab') zip_zs ind_phis zip_sets) abs' helper_ind_phiss zip_setss |> map (HOLogic.mk_Trueprop o Library.foldr1 HOLogic.mk_conj); val helper_ind_thmss = if m = 0 then replicate n [] else @{map 4} (fn concl => fn j => fn set_induct => fn cts => fold (Variable.add_free_names lthy) (concl :: helper_prems) [] |> (fn vars => Goal.prove_sorry lthy vars [] (Logic.list_implies (helper_prems, concl)) (fn {context = ctxt, prems = _} => mk_rel_coinduct_ind_tac ctxt m ks dtor_unfold_thms set_mapss j (infer_instantiate' ctxt cts set_induct))) |> Thm.close_derivation \<^here> |> split_conj_thm) mk_helper_ind_concls ls dtor_Jset_induct_thms ctss |> transpose; in mk_rel_coinduct_tac ctxt CIHs in_rels in_Jrels helper_ind_thmss helper_coind1_thms helper_coind2_thms end; val Jrel_coinduct_thm = mk_xtor_rel_co_induct_thm Greatest_FP rels activeJphis Jrels Jphis Jzs Jz's dtors dtor's Jrel_coinduct_tac lthy; val le_Jrel_OO_thm = let fun mk_le_Jrel_OO Jrelpsi1 Jrelpsi2 Jrelpsi12 = mk_leq (mk_rel_compp (Jrelpsi1, Jrelpsi2)) Jrelpsi12; val goals = @{map 3} mk_le_Jrel_OO Jrelpsi1s Jrelpsi2s Jrelpsi12s; val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj goals); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_le_rel_OO_tac ctxt Jrel_coinduct_thm dtor_Jrel_thms le_rel_OOs) |> Thm.close_derivation \<^here> end; val timer = time (timer "helpers for BNF properties"); fun close_wit I wit = (I, fold_rev Term.absfree (map (nth ys') I) wit); val all_unitTs = replicate live HOLogic.unitT; val unitTs = replicate n HOLogic.unitT; val unit_funs = replicate n (Term.absdummy HOLogic.unitT HOLogic.unit); fun mk_map_args I = map (fn i => if member (op =) I i then Term.absdummy HOLogic.unitT (nth ys i) else mk_undefined (HOLogic.unitT --> nth passiveAs i)) (0 upto (m - 1)); fun mk_nat_wit Ds bnf (I, wit) () = let val passiveI = filter (fn i => i < m) I; val map_args = mk_map_args passiveI; in Term.absdummy HOLogic.unitT (Term.list_comb (mk_map_of_bnf Ds all_unitTs (passiveAs @ unitTs) bnf, map_args @ unit_funs) $ wit) end; fun mk_dummy_wit Ds bnf I = let val map_args = mk_map_args I; in Term.absdummy HOLogic.unitT (Term.list_comb (mk_map_of_bnf Ds all_unitTs (passiveAs @ unitTs) bnf, map_args @ unit_funs) $ mk_undefined (mk_T_of_bnf Ds all_unitTs bnf)) end; val nat_witss = map2 (fn Ds => fn bnf => mk_wits_of_bnf (replicate (nwits_of_bnf bnf) Ds) (replicate (nwits_of_bnf bnf) (replicate live HOLogic.unitT)) bnf |> map (fn (I, wit) => (I, Lazy.lazy (mk_nat_wit Ds bnf (I, Term.list_comb (wit, map (K HOLogic.unit) I)))))) Dss bnfs; val nat_wit_thmss = map2 (curry op ~~) nat_witss (map wit_thmss_of_bnf bnfs) val Iss = map (map fst) nat_witss; fun filter_wits (I, wit) = let val J = filter (fn i => i < m) I; in (J, (length J < length I, wit)) end; val wit_treess = map_index (fn (i, Is) => map_index (finish Iss m [i+m] (i+m)) Is) Iss |> map (minimize_wits o map filter_wits o minimize_wits o flat); val coind_wit_argsss = map (map (tree_to_coind_wits nat_wit_thmss o snd o snd) o filter (fst o snd)) wit_treess; val nonredundant_coind_wit_argsss = fold (fn i => fn argsss => nth_map (i - 1) (filter_out (fn xs => exists (fn ys => let val xs' = (map (fst o fst) xs, snd (fst (hd xs))); val ys' = (map (fst o fst) ys, snd (fst (hd ys))); in eq_pair (subset (op =)) (eq_set (op =)) (xs', ys') andalso not (fst xs' = fst ys') end) (flat argsss))) argsss) ks coind_wit_argsss; fun prepare_args args = let val I = snd (fst (hd args)); val (dummys, args') = map_split (fn i => (case find_first (fn arg => fst (fst arg) = i - 1) args of SOME (_, ((_, wit), thms)) => (NONE, (Lazy.force wit, thms)) | NONE => (SOME (i - 1), (mk_dummy_wit (nth Dss (i - 1)) (nth bnfs (i - 1)) I, [])))) ks; in ((I, dummys), apsnd flat (split_list args')) end; fun mk_coind_wits ((I, dummys), (args, thms)) = ((I, dummys), (map (fn i => mk_unfold Ts args i $ HOLogic.unit) ks, thms)); val coind_witss = maps (map (mk_coind_wits o prepare_args)) nonredundant_coind_wit_argsss; val witss = map2 (fn Ds => fn bnf => mk_wits_of_bnf (replicate (nwits_of_bnf bnf) Ds) (replicate (nwits_of_bnf bnf) (passiveAs @ Ts)) bnf) Dss bnfs; val ctor_witss = map (map (uncurry close_wit o tree_to_ctor_wit ys ctors witss o snd o snd) o filter_out (fst o snd)) wit_treess; fun mk_coind_wit_thms ((I, dummys), (wits, wit_thms)) = let fun mk_goal sets y y_copy y'_copy j = let fun mk_conjunct set z dummy wit = mk_Ball (set $ z) (Term.absfree y'_copy (if dummy = NONE orelse member (op =) I (j - 1) then HOLogic.mk_imp (HOLogic.mk_eq (z, wit), if member (op =) I (j - 1) then HOLogic.mk_eq (y_copy, y) else \<^term>\False\) else \<^term>\True\)); in HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (@{map 4} mk_conjunct sets Jzs dummys wits)) end; val goals = @{map 5} mk_goal Jsetss_by_range ys ys_copy ys'_copy ls; in map2 (fn goal => fn induct => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_coind_wit_tac ctxt induct dtor_unfold_thms (flat set_mapss) wit_thms)) |> Thm.close_derivation \<^here>) goals dtor_Jset_induct_thms |> map split_conj_thm |> transpose |> map (map_filter (try (fn thm => thm RS bspec RS mp))) |> curry op ~~ (map_index Library.I (map (close_wit I) wits)) |> filter (fn (_, thms) => length thms = m) end; val coind_wit_thms = maps mk_coind_wit_thms coind_witss; val (wit_thmss, all_witss) = fold (fn ((i, wit), thms) => fn witss => nth_map i (fn (thms', wits) => (thms @ thms', wit :: wits)) witss) coind_wit_thms (map (pair []) ctor_witss) |> map (apsnd (map snd o minimize_wits)) |> split_list; val timer = time (timer "witnesses"); val map_id0_tacs = map2 (fn thm => fn thm' => fn ctxt => mk_map_id0_tac ctxt Jmap_thms thm thm') dtor_unfold_unique_thms unfold_dtor_thms; val map_comp0_tacs = map (fn thm => fn ctxt => rtac ctxt (thm RS sym) 1) Jmap_comp0_thms; val map_cong0_tacs = map (fn thm => fn ctxt => mk_map_cong0_tac ctxt m thm) map_cong0_thms; val set_map0_tacss = map (map (fn col => fn ctxt => unfold_thms_tac ctxt Jset_defs THEN mk_set_map0_tac ctxt col)) (transpose col_natural_thmss); val Jbd_card_orders = map (fn def => Local_Defs.fold lthy [def] sbd_card_order) Jbd_defs; val Jbd_Cinfinites = map (fn def => Local_Defs.fold lthy [def] sbd_Cinfinite) Jbd_defs; val bd_co_tacs = map (fn thm => fn ctxt => rtac ctxt thm 1) Jbd_card_orders; val bd_cinf_tacs = map (fn thm => fn ctxt => rtac ctxt (thm RS conjunct1) 1) Jbd_Cinfinites; val set_bd_tacss = map2 (fn Cinf => map (fn col => fn ctxt => unfold_thms_tac ctxt Jset_defs THEN mk_set_bd_tac ctxt Cinf col)) Jbd_Cinfinites (transpose col_bd_thmss); val le_rel_OO_tacs = map (fn i => fn ctxt => rtac ctxt (le_Jrel_OO_thm RS mk_conjunctN n i) 1) ks; val rel_OO_Grp_tacs = map (fn def => fn ctxt => rtac ctxt def 1) Jrel_unabs_defs; val pred_set_tacs = map (fn def => fn ctxt => rtac ctxt def 1) Jpred_unabs_defs; val tacss = @{map 10} zip_axioms map_id0_tacs map_comp0_tacs map_cong0_tacs set_map0_tacss bd_co_tacs bd_cinf_tacs set_bd_tacss le_rel_OO_tacs rel_OO_Grp_tacs pred_set_tacs; fun wit_tac thms ctxt = mk_wit_tac ctxt n dtor_ctor_thms (flat dtor_Jset_thmss) (maps wit_thms_of_bnf bnfs) thms; val (Jbnfs, lthy) = @{fold_map 7} (fn tacs => fn map_b => fn rel_b => fn pred_b => fn set_bs => fn wit_thms => fn consts => bnf_def Hardly_Inline (user_policy Note_Some) false I tacs (wit_tac wit_thms) (SOME deads) map_b rel_b pred_b set_bs consts) tacss map_bs rel_bs pred_bs set_bss wit_thmss (((((((replicate n Binding.empty ~~ Ts) ~~ Jmaps) ~~ Jsetss_by_bnf) ~~ Jbds) ~~ all_witss) ~~ map SOME Jrels) ~~ map SOME Jpreds) lthy; val timer = time (timer "registered new codatatypes as BNFs"); val ls' = if m = 1 then [0] else ls; val Jbnf_common_notes = map2 (fn i => fn thm => (mk_dtor_set_inductN i, [thm])) ls' dtor_Jset_induct_thms |> map (fn (thmN, thms) => ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])])); val Jbnf_notes = [(dtor_mapN, map single dtor_Jmap_thms), (dtor_map_uniqueN, map single dtor_Jmap_unique_thms), (dtor_relN, map single dtor_Jrel_thms), (dtor_set_inclN, dtor_Jset_incl_thmss), (dtor_set_set_inclN, map flat dtor_set_Jset_incl_thmsss)] @ map2 (fn i => fn thms => (mk_dtor_setN i, map single thms)) ls' dtor_Jset_thmss |> maps (fn (thmN, thmss) => map2 (fn b => fn thms => ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])])) bs thmss) in (timer, Jbnfs, (Jmap_thms, dtor_Jmap_thms), dtor_Jmap_unique_thm, dtor_Jset_thmss', dtor_Jrel_thms, Jrel_coinduct_thm, Jbnf_common_notes @ Jbnf_notes, dtor_Jset_induct_thms, lthy) end; val ((Jphis, activephis), _) = lthy |> mk_Frees "R" JphiTs ||>> mk_Frees "S" activephiTs; val dtor_unfold_o_Jmap_thms = mk_xtor_co_iter_o_map_thms Greatest_FP false m dtor_unfold_unique_thm dtor_Jmap_o_thms (map (mk_pointfree2 lthy) dtor_unfold_thms) sym_map_comps map_cong0s; val rels = map2 (fn Ds => mk_rel_of_bnf Ds allAs allBs') Dss bnfs; val Jrels = if m = 0 then map HOLogic.eq_const Ts else map (mk_rel_of_bnf deads passiveAs passiveBs) Jbnfs; val dtor_unfold_transfer_thms = mk_xtor_co_iter_transfer_thms Greatest_FP rels activephis activephis Jrels Jphis (mk_unfolds passiveAs activeAs) (mk_unfolds passiveBs activeBs) (fn {context = ctxt, prems = _} => mk_unfold_transfer_tac ctxt m Jrel_coinduct_thm (map map_transfer_of_bnf bnfs) dtor_unfold_thms) lthy; val timer = time (timer "relator coinduction"); fun mk_Ts As = map (typ_subst_atomic (passiveAs ~~ As)) Ts; val export = map (Morphism.term (Local_Theory.target_morphism lthy)) val ((corecs, (dtor_corec_thms, dtor_corec_unique_thm, dtor_corec_o_Jmap_thms, dtor_corec_transfer_thms)), lthy) = lthy |> derive_xtor_co_recs Greatest_FP external_bs mk_Ts (Dss, resDs) bnfs (export dtors) (export unfolds) dtor_unfold_unique_thm dtor_unfold_thms dtor_unfold_transfer_thms dtor_Jmap_thms dtor_Jrel_thms (replicate n NONE); val timer = time (timer "recursor"); val common_notes = [(dtor_coinductN, [dtor_coinduct_thm]), (dtor_rel_coinductN, [Jrel_coinduct_thm])] |> map (fn (thmN, thms) => ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])])); val notes = [(ctor_dtorN, ctor_dtor_thms), (ctor_exhaustN, ctor_exhaust_thms), (ctor_injectN, ctor_inject_thms), (dtor_ctorN, dtor_ctor_thms), (dtor_exhaustN, dtor_exhaust_thms), (dtor_injectN, dtor_inject_thms), (dtor_unfoldN, dtor_unfold_thms), (dtor_unfold_o_mapN, dtor_unfold_o_Jmap_thms), (dtor_unfold_transferN, dtor_unfold_transfer_thms), (dtor_unfold_uniqueN, dtor_unfold_unique_thms)] |> map (apsnd (map single)) |> maps (fn (thmN, thmss) => map2 (fn b => fn thms => ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])])) bs thmss); val lthy' = lthy |> internals ? snd o Local_Theory.notes (common_notes @ notes @ Jbnf_notes); val fp_res = {Ts = Ts, bnfs = Jbnfs, pre_bnfs = bnfs, absT_infos = absT_infos, ctors = ctors, dtors = dtors, xtor_un_folds = unfolds, xtor_co_recs = export corecs, xtor_co_induct = dtor_coinduct_thm, dtor_ctors = dtor_ctor_thms, ctor_dtors = ctor_dtor_thms, ctor_injects = ctor_inject_thms, dtor_injects = dtor_inject_thms, xtor_maps = dtor_Jmap_thms, xtor_map_unique = dtor_Jmap_unique_thm, xtor_setss = dtor_Jset_thmss', xtor_rels = dtor_Jrel_thms, xtor_un_fold_thms = dtor_unfold_thms, xtor_co_rec_thms = dtor_corec_thms, xtor_un_fold_unique = dtor_unfold_unique_thm, xtor_co_rec_unique = dtor_corec_unique_thm, xtor_un_fold_o_maps = dtor_unfold_o_Jmap_thms, xtor_co_rec_o_maps = dtor_corec_o_Jmap_thms, xtor_un_fold_transfers = dtor_unfold_transfer_thms, xtor_co_rec_transfers = dtor_corec_transfer_thms, xtor_rel_co_induct = Jrel_coinduct_thm, dtor_set_inducts = dtor_Jset_induct_thms}; in timer; (fp_res, lthy') end; val _ = Outer_Syntax.local_theory \<^command_keyword>\codatatype\ "define coinductive datatypes" (parse_co_datatype_cmd Greatest_FP construct_gfp); end; diff --git a/src/HOL/Tools/BNF/bnf_gfp_grec.ML b/src/HOL/Tools/BNF/bnf_gfp_grec.ML --- a/src/HOL/Tools/BNF/bnf_gfp_grec.ML +++ b/src/HOL/Tools/BNF/bnf_gfp_grec.ML @@ -1,3229 +1,3229 @@ (* Title: HOL/Tools/BNF/bnf_gfp_grec.ML Author: Jasmin Blanchette, Inria, LORIA, MPII Author: Aymeric Bouzy, Ecole polytechnique Author: Dmitriy Traytel, ETH Zürich Copyright 2015, 2016 Generalized corecursor construction. *) signature BNF_GFP_GREC = sig val Tsubst: typ -> typ -> typ -> typ val substT: typ -> typ -> term -> term val freeze_types: Proof.context -> (indexname * sort) list -> typ list -> typ list val dummify_atomic_types: term -> term val define_const: bool -> binding -> int -> string -> term -> local_theory -> (term * thm) * local_theory type buffer = {Oper: term, VLeaf: term, CLeaf: term, ctr_wrapper: term, friends: (typ * term) Symtab.table} val map_buffer: (term -> term) -> buffer -> buffer val specialize_buffer_types: buffer -> buffer type dtor_coinduct_info = {dtor_coinduct: thm, cong_def: thm, cong_locale: thm, cong_base: thm, cong_refl: thm, cong_sym: thm, cong_trans: thm, cong_alg_intros: thm list} type corec_info = {fp_b: binding, version: int, fpT: typ, Y: typ, Z: typ, friend_names: string list, sig_fp_sugars: BNF_FP_Def_Sugar.fp_sugar list, ssig_fp_sugar: BNF_FP_Def_Sugar.fp_sugar, Lam: term, proto_sctr: term, flat: term, eval_core: term, eval: term, algLam: term, corecUU: term, dtor_transfer: thm, Lam_transfer: thm, Lam_pointful_natural: thm, proto_sctr_transfer: thm, flat_simps: thm list, eval_core_simps: thm list, eval_thm: thm, eval_simps: thm list, all_algLam_algs: thm list, algLam_thm: thm, dtor_algLam: thm, corecUU_thm: thm, corecUU_unique: thm, corecUU_transfer: thm, buffer: buffer, all_dead_k_bnfs: BNF_Def.bnf list, Retr: term, equivp_Retr: thm, Retr_coinduct: thm, dtor_coinduct_info: dtor_coinduct_info} type friend_info = {algrho: term, dtor_algrho: thm, algLam_algrho: thm} val not_codatatype: Proof.context -> typ -> 'a val mk_fp_binding: binding -> string -> binding val bnf_kill_all_but: int -> BNF_Def.bnf -> local_theory -> BNF_Def.bnf * local_theory val print_corec_infos: Proof.context -> unit val has_no_corec_info: Proof.context -> string -> bool val corec_info_of: typ -> local_theory -> corec_info * local_theory val maybe_corec_info_of: Proof.context -> typ -> corec_info option val corec_infos_of: Proof.context -> string -> corec_info list val corec_infos_of_generic: Context.generic -> Symtab.key -> corec_info list val prepare_friend_corec: string -> typ -> local_theory -> (corec_info * binding * int * typ * typ * typ * typ * typ * BNF_Def.bnf * BNF_Def.bnf * BNF_FP_Def_Sugar.fp_sugar * BNF_FP_Def_Sugar.fp_sugar * buffer) * local_theory val register_friend_corec: string -> binding -> int -> typ -> typ -> typ -> BNF_Def.bnf -> BNF_FP_Def_Sugar.fp_sugar -> BNF_FP_Def_Sugar.fp_sugar -> term -> term -> thm -> corec_info -> local_theory -> friend_info * local_theory end; structure BNF_GFP_Grec : BNF_GFP_GREC = struct open Ctr_Sugar open BNF_Util open BNF_Def open BNF_Comp open BNF_FP_Util open BNF_LFP open BNF_FP_Def_Sugar open BNF_LFP_Rec_Sugar open BNF_GFP_Grec_Tactics val algLamN = "algLam"; val algLam_algLamN = "algLam_algLam"; val algLam_algrhoN = "algLam_algrho"; val algrhoN = "algrho"; val CLeafN = "CLeaf"; val congN = "congclp"; val cong_alg_introsN = "cong_alg_intros"; val cong_localeN = "cong_locale"; val corecUUN = "corecUU"; val corecUU_transferN = "corecUU_transfer"; val corecUU_uniqueN = "corecUU_unique"; val cutSsigN = "cutSsig"; val dtor_algLamN = "dtor_algLam"; val dtor_algrhoN = "dtor_algrho"; val dtor_coinductN = "dtor_coinduct"; val dtor_transferN = "dtor_transfer"; val embLN = "embL"; val embLLN = "embLL"; val embLRN = "embLR"; val embL_pointful_naturalN = "embL_pointful_natural"; val embL_transferN = "embL_transfer"; val equivp_RetrN = "equivp_Retr"; val evalN = "eval"; val eval_coreN = "eval_core"; val eval_core_pointful_naturalN = "eval_core_pointful_natural"; val eval_core_transferN = "eval_core_transfer"; val eval_flatN = "eval_flat"; val eval_simpsN = "eval_simps"; val flatN = "flat"; val flat_pointful_naturalN = "flat_pointful_natural"; val flat_transferN = "flat_transfer"; val k_as_ssig_naturalN = "k_as_ssig_natural"; val k_as_ssig_transferN = "k_as_ssig_transfer"; val LamN = "Lam"; val Lam_transferN = "Lam_transfer"; val Lam_pointful_naturalN = "Lam_pointful_natural"; val OperN = "Oper"; val proto_sctrN = "proto_sctr"; val proto_sctr_pointful_naturalN = "proto_sctr_pointful_natural"; val proto_sctr_transferN = "proto_sctr_transfer"; val rho_transferN = "rho_transfer"; val Retr_coinductN = "Retr_coinduct"; val sctrN = "sctr"; val sctr_transferN = "sctr_transfer"; val sctr_pointful_naturalN = "sctr_pointful_natural"; val sigN = "sig"; val SigN = "Sig"; val Sig_pointful_naturalN = "Sig_pointful_natural"; val corecUN = "corecU"; val corecU_ctorN = "corecU_ctor"; val corecU_uniqueN = "corecU_unique"; val unsigN = "unsig"; val VLeafN = "VLeaf"; val s_prefix = "s"; (* transforms "sig" into "ssig" *) fun not_codatatype ctxt T = error ("Not a codatatype: " ^ Syntax.string_of_typ ctxt T); fun mutual_codatatype () = error ("Mutually corecursive codatatypes are not supported (try " ^ quote (#1 \<^command_keyword>\primcorec\) ^ " instead of " ^ quote (#1 \<^command_keyword>\corec\) ^ ")"); fun noncorecursive_codatatype () = error ("Noncorecursive codatatypes are not supported (try " ^ quote (#1 \<^command_keyword>\definition\) ^ " instead of " ^ quote (#1 \<^command_keyword>\corec\) ^ ")"); fun singleton_codatatype ctxt = error ("Singleton corecursive codatatypes are not supported (use " ^ quote (Syntax.string_of_typ ctxt \<^typ>\unit\) ^ " instead)"); fun merge_lists eq old1 old2 = (old1 |> subtract eq old2) @ old2; fun add_type_namesT (Type (s, Ts)) = insert (op =) s #> fold add_type_namesT Ts | add_type_namesT _ = I; fun Tsubst Y T = Term.typ_subst_atomic [(Y, T)]; fun substT Y T = Term.subst_atomic_types [(Y, T)]; fun freeze_types ctxt except_tvars Ts = let val As = fold Term.add_tvarsT Ts [] |> subtract (op =) except_tvars; val (Bs, _) = ctxt |> mk_TFrees' (map snd As); in map (Term.typ_subst_TVars (map fst As ~~ Bs)) Ts end; fun typ_unify_disjointly thy (T, T') = if T = T' then T else let val tvars = Term.add_tvar_namesT T []; val tvars' = Term.add_tvar_namesT T' []; val maxidx' = maxidx_of_typ T'; val T = T |> exists (member (op =) tvars') tvars ? Logic.incr_tvar (maxidx' + 1); val maxidx = Integer.max (maxidx_of_typ T) maxidx'; val (tyenv, _) = Sign.typ_unify thy (T, T') (Vartab.empty, maxidx); in Envir.subst_type tyenv T end; val dummify_atomic_types = Term.map_types (Term.map_atyps (K Term.dummyT)); fun mk_internal internal ctxt f = if internal andalso not (Config.get ctxt bnf_internals) then f else I fun mk_fp_binding fp_b pre = Binding.map_name (K pre) fp_b |> Binding.qualify true (Binding.name_of fp_b); fun mk_version_binding version = Binding.qualify false ("v" ^ string_of_int version); fun mk_version_fp_binding internal ctxt = mk_internal internal ctxt Binding.concealed ooo (mk_fp_binding oo mk_version_binding); (*FIXME: get rid of ugly names when typedef and primrec respect qualification*) fun mk_version_binding_ugly version = Binding.suffix_name ("_v" ^ string_of_int version); fun mk_version_fp_binding_ugly internal ctxt version fp_b pre = Binding.prefix_name (pre ^ "_") fp_b |> mk_version_binding_ugly version |> mk_internal internal ctxt Binding.concealed; fun mk_mapN ctxt live_AsBs TA bnf = let val TB = Term.typ_subst_atomic live_AsBs TA in enforce_type ctxt (snd o strip_typeN (length live_AsBs)) (TA --> TB) (map_of_bnf bnf) end; fun mk_relN ctxt live_AsBs TA bnf = let val TB = Term.typ_subst_atomic live_AsBs TA in enforce_type ctxt (snd o strip_typeN (length live_AsBs)) (mk_pred2T TA TB) (rel_of_bnf bnf) end; fun mk_map1 ctxt Y Z = mk_mapN ctxt [(Y, Z)]; fun mk_rel1 ctxt Y Z = mk_relN ctxt [(Y, Z)]; fun define_const internal fp_b version name rhs lthy = let val b = mk_version_fp_binding internal lthy version fp_b name; val ((free, (_, def_free)), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> Local_Theory.define ((b, NoSyn), ((Thm.def_binding b |> Binding.concealed, []), rhs)) ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val const = Morphism.term phi free; val const' = enforce_type lthy I (fastype_of free) const; in ((const', Morphism.thm phi def_free), lthy) end; fun define_single_primrec b eqs lthy = let val (([free], [def_free], [simps_free]), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> Local_Theory.map_background_naming (mk_internal true lthy Name_Space.concealed) (*TODO check*) |> primrec false [] [(b, NONE, NoSyn)] (map (fn eq => ((Binding.empty_atts, eq), [], [])) eqs) ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val const = Morphism.term phi free; val const' = enforce_type lthy I (fastype_of free) const; in ((const', Morphism.thm phi def_free, map (Morphism.thm phi) simps_free), lthy) end; type buffer = {Oper: term, VLeaf: term, CLeaf: term, ctr_wrapper: term, friends: (typ * term) Symtab.table}; fun map_buffer f {Oper, VLeaf, CLeaf, ctr_wrapper, friends} = {Oper = f Oper, VLeaf = f VLeaf, CLeaf = f CLeaf, ctr_wrapper = f ctr_wrapper, friends = Symtab.map (K (apsnd f)) friends}; fun morph_buffer phi = map_buffer (Morphism.term phi); fun specialize_buffer_types {Oper, VLeaf, CLeaf, ctr_wrapper, friends} = let val ssig_T as Type (_, Ts) = body_type (fastype_of VLeaf); val Y = List.last Ts; val ssigifyT = substT Y ssig_T; in {Oper = Oper, VLeaf = VLeaf, CLeaf = CLeaf, ctr_wrapper = ssigifyT ctr_wrapper, friends = Symtab.map (K (apsnd ssigifyT)) friends} end; type dtor_coinduct_info = {dtor_coinduct: thm, cong_def: thm, cong_locale: thm, cong_base: thm, cong_refl: thm, cong_sym: thm, cong_trans: thm, cong_alg_intros: thm list}; fun map_dtor_coinduct_info f {dtor_coinduct, cong_def, cong_locale, cong_base, cong_refl, cong_sym, cong_trans, cong_alg_intros} = {dtor_coinduct = f dtor_coinduct, cong_def = f cong_def, cong_locale = f cong_locale, cong_base = f cong_base, cong_refl = f cong_refl, cong_sym = f cong_sym, cong_trans = f cong_trans, cong_alg_intros = map f cong_alg_intros}; fun morph_dtor_coinduct_info phi = map_dtor_coinduct_info (Morphism.thm phi); type corec_ad = {fpT: typ, friend_names: string list}; fun morph_corec_ad phi {fpT, friend_names} = {fpT = Morphism.typ phi fpT, friend_names = friend_names}; type corec_info = {fp_b: binding, version: int, fpT: typ, Y: typ, Z: typ, friend_names: string list, sig_fp_sugars: fp_sugar list, ssig_fp_sugar: fp_sugar, Lam: term, proto_sctr: term, flat: term, eval_core: term, eval: term, algLam: term, corecUU: term, dtor_transfer: thm, Lam_transfer: thm, Lam_pointful_natural: thm, proto_sctr_transfer: thm, flat_simps: thm list, eval_core_simps: thm list, eval_thm: thm, eval_simps: thm list, all_algLam_algs: thm list, algLam_thm: thm, dtor_algLam: thm, corecUU_thm: thm, corecUU_unique: thm, corecUU_transfer: thm, buffer: buffer, all_dead_k_bnfs: bnf list, Retr: term, equivp_Retr: thm, Retr_coinduct: thm, dtor_coinduct_info: dtor_coinduct_info}; fun morph_corec_info phi ({fp_b, version, fpT, Y, Z, friend_names, sig_fp_sugars, ssig_fp_sugar, Lam, proto_sctr, flat, eval_core, eval, algLam, corecUU, dtor_transfer, Lam_transfer, Lam_pointful_natural, proto_sctr_transfer, flat_simps, eval_core_simps, eval_thm, eval_simps, all_algLam_algs, algLam_thm, dtor_algLam, corecUU_thm, corecUU_unique, corecUU_transfer, buffer, all_dead_k_bnfs, Retr, equivp_Retr, Retr_coinduct, dtor_coinduct_info} : corec_info) = {fp_b = fp_b, version = version, fpT = Morphism.typ phi fpT, Y = Morphism.typ phi Y, Z = Morphism.typ phi Z, friend_names = friend_names, sig_fp_sugars = sig_fp_sugars (*no morph*), ssig_fp_sugar = ssig_fp_sugar (*no morph*), Lam = Morphism.term phi Lam, proto_sctr = Morphism.term phi proto_sctr, flat = Morphism.term phi flat, eval_core = Morphism.term phi eval_core, eval = Morphism.term phi eval, algLam = Morphism.term phi algLam, corecUU = Morphism.term phi corecUU, dtor_transfer = dtor_transfer, Lam_transfer = Morphism.thm phi Lam_transfer, Lam_pointful_natural = Morphism.thm phi Lam_pointful_natural, proto_sctr_transfer = Morphism.thm phi proto_sctr_transfer, flat_simps = map (Morphism.thm phi) flat_simps, eval_core_simps = map (Morphism.thm phi) eval_core_simps, eval_thm = Morphism.thm phi eval_thm, eval_simps = map (Morphism.thm phi) eval_simps, all_algLam_algs = map (Morphism.thm phi) all_algLam_algs, algLam_thm = Morphism.thm phi algLam_thm, dtor_algLam = Morphism.thm phi dtor_algLam, corecUU_thm = Morphism.thm phi corecUU_thm, corecUU_unique = Morphism.thm phi corecUU_unique, corecUU_transfer = Morphism.thm phi corecUU_transfer, buffer = morph_buffer phi buffer, all_dead_k_bnfs = map (morph_bnf phi) all_dead_k_bnfs, Retr = Morphism.term phi Retr, equivp_Retr = Morphism.thm phi equivp_Retr, Retr_coinduct = Morphism.thm phi Retr_coinduct, dtor_coinduct_info = morph_dtor_coinduct_info phi dtor_coinduct_info}; datatype ('a, 'b) expr = Ad of 'a * (local_theory -> 'b * local_theory) | Info of 'b; fun is_Ad (Ad _) = true | is_Ad _ = false; fun is_Info (Info _) = true | is_Info _ = false; type corec_info_expr = (corec_ad, corec_info) expr; fun morph_corec_info_expr phi (Ad (ad, f)) = Ad (morph_corec_ad phi ad, f) | morph_corec_info_expr phi (Info info) = Info (morph_corec_info phi info); val transfer_corec_info_expr = morph_corec_info_expr o Morphism.transfer_morphism; type corec_data = int Symtab.table * corec_info_expr list Symtab.table list; structure Data = Generic_Data ( type T = corec_data; val empty = (Symtab.empty, [Symtab.empty]); val extend = I; fun merge ((version_tab1, info_tabs1), (version_tab2, info_tabs2)) : T = (Symtab.join (K Int.max) (version_tab1, version_tab2), info_tabs1 @ info_tabs2); ); fun corec_ad_of_expr (Ad (ad, _)) = ad | corec_ad_of_expr (Info {fpT, friend_names, ...}) = {fpT = fpT, friend_names = friend_names}; fun corec_info_exprs_of_generic context fpT_name = let val thy = Context.theory_of context; val info_tabs = snd (Data.get context); in maps (fn info_tab => these (Symtab.lookup info_tab fpT_name)) info_tabs |> map (transfer_corec_info_expr thy) end; val corec_info_exprs_of = corec_info_exprs_of_generic o Context.Proof; val keep_corec_infos = map_filter (fn Ad _ => NONE | Info info => SOME info); val corec_infos_of_generic = keep_corec_infos oo corec_info_exprs_of_generic; val corec_infos_of = keep_corec_infos oo corec_info_exprs_of; fun str_of_corec_ad ctxt {fpT, friend_names} = "[" ^ Syntax.string_of_typ ctxt fpT ^ "; " ^ commas friend_names ^ "]"; fun str_of_corec_info ctxt {fpT, version, friend_names, ...} = "{" ^ Syntax.string_of_typ ctxt fpT ^ "; " ^ commas friend_names ^ "; v" ^ string_of_int version ^ "}"; fun str_of_corec_info_expr ctxt (Ad (ad, _)) = str_of_corec_ad ctxt ad | str_of_corec_info_expr ctxt (Info info) = str_of_corec_info ctxt info; fun print_corec_infos ctxt = Symtab.fold (fn (fpT_name, exprs) => fn () => writeln (fpT_name ^ ":\n" ^ cat_lines (map (prefix " " o str_of_corec_info_expr ctxt) exprs))) (the_single (snd (Data.get (Context.Proof ctxt)))) (); val has_no_corec_info = null oo corec_info_exprs_of; fun get_name_next_version_of fpT_name ctxt = let val (version_tab, info_tabs) = Data.get (Context.Theory (Proof_Context.theory_of ctxt)); val fp_base = Long_Name.base_name fpT_name; val fp_b = Binding.name fp_base; val version_tab' = Symtab.map_default (fp_base, ~1) (Integer.add 1) version_tab; val SOME version = Symtab.lookup version_tab' fp_base; val ctxt' = ctxt |> Local_Theory.background_theory (Context.theory_map (Data.put (version_tab', info_tabs))); in ((fp_b, version), ctxt') end; type friend_info = {algrho: term, dtor_algrho: thm, algLam_algrho: thm}; fun morph_friend_info phi ({algrho, dtor_algrho, algLam_algrho} : friend_info) = {algrho = Morphism.term phi algrho, dtor_algrho = Morphism.thm phi dtor_algrho, algLam_algrho = Morphism.thm phi algLam_algrho}; fun checked_fp_sugar_of ctxt fpT_name = let val fp_sugar as {X, fp_res = {Ts = fpTs, ...}, fp_ctr_sugar = {ctrXs_Tss, ...}, ...} = (case fp_sugar_of ctxt fpT_name of SOME (fp_sugar as {fp = Greatest_FP, ...}) => fp_sugar | _ => not_codatatype ctxt (Type (fpT_name, [] (*yuck*)))); val _ = if length fpTs > 1 then mutual_codatatype () else if not (exists (exists (Term.exists_subtype (curry (op =) X))) ctrXs_Tss) then noncorecursive_codatatype () else if ctrXs_Tss = [[X]] then singleton_codatatype ctxt else (); in fp_sugar end; fun bnf_kill_all_but nn bnf lthy = ((empty_comp_cache, empty_unfolds), lthy) |> kill_bnf I (live_of_bnf bnf - nn) bnf ||> snd; fun bnf_with_deads_and_lives dead_Es live_As Y fpT T lthy = let val qsoty = quote o Syntax.string_of_typ lthy; val unfreeze_fp = Tsubst Y fpT; fun flatten_tyargs Ass = map dest_TFree live_As |> filter (fn T => exists (fn Ts => member (op =) Ts T) Ass); val ((bnf, _), (_, lthy)) = bnf_of_typ false Do_Inline I flatten_tyargs [Term.dest_TFree Y] (map Term.dest_TFree dead_Es) T ((empty_comp_cache, empty_unfolds), lthy) handle BAD_DEAD (Y, Y_backdrop) => (case Y_backdrop of Type (bad_tc, _) => let val T = qsoty (unfreeze_fp Y); val T_backdrop = qsoty (unfreeze_fp Y_backdrop); fun register_hint () = "\nUse the " ^ quote (#1 \<^command_keyword>\bnf\) ^ " command to register " ^ quote bad_tc ^ " as a bounded natural functor to allow nested (co)recursion through \ \it"; in if is_some (bnf_of lthy bad_tc) orelse is_some (fp_sugar_of lthy bad_tc) then error ("Inadmissible occurrence of type " ^ T ^ " in type expression " ^ T_backdrop) else error ("Unsupported occurrence of type " ^ T ^ " via type constructor " ^ quote bad_tc ^ " in type expression " ^ T_backdrop ^ register_hint ()) end); val phi = Morphism.term_morphism "BNF" (Raw_Simplifier.rewrite_term (Proof_Context.theory_of lthy) @{thms BNF_Composition.id_bnf_def} []) $> Morphism.thm_morphism "BNF" (unfold_thms lthy @{thms BNF_Composition.id_bnf_def}); in (morph_bnf phi bnf, lthy) end; fun define_sig_type fp_b version fp_alives Es Y rhsT lthy = let val T_b = mk_version_fp_binding_ugly true lthy version fp_b sigN; val ctr_b = mk_version_fp_binding false lthy version fp_b SigN; val sel_b = mk_version_fp_binding true lthy version fp_b unsigN; - val lthy = Local_Theory.open_target lthy |> snd; + val lthy = Local_Theory.open_target lthy; val T_name = Local_Theory.full_name lthy T_b; val tyargs = map2 (fn alive => fn T => (if alive then SOME Binding.empty else NONE, (T, Type.sort_of_atyp T))) (fp_alives @ [true]) (Es @ [Y]); val ctr_specs = [(((Binding.empty, ctr_b), [(sel_b, rhsT)]), NoSyn)]; val spec = (((((tyargs, T_b), NoSyn), ctr_specs), (Binding.empty, Binding.empty, Binding.empty)), []); val plugins = Plugin_Name.make_filter lthy (K (curry (op =) transfer_plugin)); val discs_sels = true; val lthy = lthy |> Local_Theory.map_background_naming (mk_internal true lthy Name_Space.concealed) (*TODO check*) |> with_typedef_threshold ~1 (co_datatypes Least_FP construct_lfp ((plugins, discs_sels), [spec])) |> Local_Theory.close_target; val SOME fp_sugar = fp_sugar_of lthy T_name; in (fp_sugar, lthy) end; fun define_ssig_type fp_b version fp_alives Es Y fpT lthy = let val sig_T_b = mk_version_fp_binding_ugly true lthy version fp_b sigN; val T_b = Binding.prefix_name s_prefix sig_T_b; val Oper_b = mk_version_fp_binding false lthy version fp_b OperN; val VLeaf_b = mk_version_fp_binding false lthy version fp_b VLeafN; val CLeaf_b = mk_version_fp_binding false lthy version fp_b CLeafN; - val lthy = Local_Theory.open_target lthy |> snd; + val lthy = Local_Theory.open_target lthy; val sig_T_name = Local_Theory.full_name lthy sig_T_b; val T_name = Long_Name.map_base_name (prefix s_prefix) sig_T_name; val As = Es @ [Y]; val ssig_sig_T = Type (sig_T_name, Es @ [Type (T_name, As)]); val tyargs = map2 (fn alive => fn T => (if alive then SOME Binding.empty else NONE, (T, Type.sort_of_atyp T))) (fp_alives @ [true]) (Es @ [Y]); val ctr_specs = [(((Binding.empty, Oper_b), [(Binding.empty, ssig_sig_T)]), NoSyn), (((Binding.empty, VLeaf_b), [(Binding.empty, Y)]), NoSyn), (((Binding.empty, CLeaf_b), [(Binding.empty, fpT)]), NoSyn)]; val spec = (((((tyargs, T_b), NoSyn), ctr_specs), (Binding.empty, Binding.empty, Binding.empty)), []); val plugins = Plugin_Name.make_filter lthy (K (curry (op =) transfer_plugin)); val discs_sels = false; val lthy = lthy |> Local_Theory.map_background_naming (mk_internal true lthy Name_Space.concealed) (*TODO check*) |> with_typedef_threshold ~1 (co_datatypes Least_FP construct_lfp ((plugins, discs_sels), [spec])) |> Local_Theory.close_target; val SOME fp_sugar = fp_sugar_of lthy T_name; in (fp_sugar, lthy) end; fun embed_Sig ctxt Sig inl_or_r t = Library.foldl1 HOLogic.mk_comp [Sig, inl_or_r, dummify_atomic_types t] |> Syntax.check_term ctxt; fun mk_ctr_wrapper_friends ctxt friend_name friend_T old_sig_T k_T Sig old_buffer = let val embed_Sig_inl = embed_Sig ctxt Sig (Inl_const old_sig_T k_T); val ctr_wrapper = embed_Sig_inl (#ctr_wrapper old_buffer); val friends = Symtab.map (K (apsnd embed_Sig_inl)) (#friends old_buffer) |> Symtab.update_new (friend_name, (friend_T, HOLogic.mk_comp (Sig, Inr_const old_sig_T k_T))); in (ctr_wrapper, friends) end; fun pre_type_of_ctor Y ctor = let val (fp_preT, fpT) = dest_funT (fastype_of ctor); in typ_subst_nonatomic [(fpT, Y)] fp_preT end; fun mk_k_as_ssig Z old_sig_T k_T ssig_T Sig dead_sig_map Oper VLeaf = let val inr' = Inr_const old_sig_T k_T; val dead_sig_map' = substT Z ssig_T dead_sig_map; in Library.foldl1 HOLogic.mk_comp [Oper, dead_sig_map' $ VLeaf, Sig, inr'] end; fun define_embL name fp_b version Y Z fpT old_sig_T old_ssig_T other_summand_T ssig_T Inl_or_r_const dead_old_sig_map Sig old_Oper old_VLeaf old_CLeaf Oper VLeaf CLeaf lthy = let val embL_b = mk_version_fp_binding true lthy version fp_b name; val old_ssig_old_sig_T = Tsubst Y old_ssig_T old_sig_T; val ssig_old_sig_T = Tsubst Y ssig_T old_sig_T; val ssig_other_summand_T = Tsubst Y ssig_T other_summand_T; val sigx = Var (("s", 0), old_ssig_old_sig_T); val x = Var (("x", 0), Y); val j = Var (("j", 0), fpT); val embL = Free (Binding.name_of embL_b, old_ssig_T --> ssig_T); val dead_old_sig_map' = Term.subst_atomic_types [(Y, old_ssig_T), (Z, ssig_T)] dead_old_sig_map; val Sig' = substT Y ssig_T Sig; val inl' = Inl_or_r_const ssig_old_sig_T ssig_other_summand_T; val Oper_eq = mk_Trueprop_eq (embL $ (old_Oper $ sigx), Oper $ (Sig' $ (inl' $ (dead_old_sig_map' $ embL $ sigx)))) |> Logic.all sigx; val VLeaf_eq = mk_Trueprop_eq (embL $ (old_VLeaf $ x), VLeaf $ x) |> Logic.all x; val CLeaf_eq = mk_Trueprop_eq (embL $ (old_CLeaf $ j), CLeaf $ j) |> Logic.all j; in define_single_primrec embL_b [Oper_eq, VLeaf_eq, CLeaf_eq] lthy end; fun define_Lam_base fp_b version Y Z preT ssig_T dead_pre_map Sig unsig dead_sig_map Oper VLeaf lthy = let val YpreT = HOLogic.mk_prodT (Y, preT); val snd' = snd_const YpreT; val dead_pre_map' = substT Z ssig_T dead_pre_map; val Sig' = substT Y ssig_T Sig; val unsig' = substT Y ssig_T unsig; val dead_sig_map' = Term.subst_atomic_types [(Y, YpreT), (Z, ssig_T)] dead_sig_map; val rhs = HOLogic.mk_comp (unsig', dead_sig_map' $ Library.foldl1 HOLogic.mk_comp [Oper, Sig', dead_pre_map' $ VLeaf, snd']); in define_const true fp_b version LamN rhs lthy end; fun define_Lam_step_or_merge fp_b version Y preT unsig left_case right_case lthy = let val YpreT = HOLogic.mk_prodT (Y, preT); val unsig' = substT Y YpreT unsig; val rhs = HOLogic.mk_comp (mk_case_sum (left_case, right_case), unsig'); in define_const true fp_b version LamN rhs lthy end; fun define_Lam_step fp_b version Y Z preT old_ssig_T ssig_T dead_pre_map unsig rho embL old_Lam lthy = let val dead_pre_map' = Term.subst_atomic_types [(Y, old_ssig_T), (Z, ssig_T)] dead_pre_map; val left_case = HOLogic.mk_comp (dead_pre_map' $ embL, old_Lam); in define_Lam_step_or_merge fp_b version Y preT unsig left_case rho lthy end; fun define_Lam_merge fp_b version Y Z preT old1_ssig_T old2_ssig_T ssig_T dead_pre_map unsig embLL embLR old1_Lam old2_Lam lthy = let val dead_pre_map' = Term.subst_atomic_types [(Y, old1_ssig_T), (Z, ssig_T)] dead_pre_map; val dead_pre_map'' = Term.subst_atomic_types [(Y, old2_ssig_T), (Z, ssig_T)] dead_pre_map; val left_case = HOLogic.mk_comp (dead_pre_map' $ embLL, old1_Lam); val right_case = HOLogic.mk_comp (dead_pre_map'' $ embLR, old2_Lam); in define_Lam_step_or_merge fp_b version Y preT unsig left_case right_case lthy end; fun define_proto_sctr_step_or_merge fp_b version old_sig_T right_T Sig old_proto_sctr = let val rhs = Library.foldl1 HOLogic.mk_comp [Sig, Inl_const old_sig_T right_T, old_proto_sctr]; in define_const true fp_b version proto_sctrN rhs end; fun define_flat fp_b version Y Z fpT sig_T ssig_T Oper VLeaf CLeaf dead_sig_map lthy = let val flat_b = mk_version_fp_binding true lthy version fp_b flatN; val ssig_sig_T = Tsubst Y ssig_T sig_T; val ssig_ssig_sig_T = Tsubst Y ssig_T ssig_sig_T; val ssig_ssig_T = Tsubst Y ssig_T ssig_T; val sigx = Var (("s", 0), ssig_ssig_sig_T); val x = Var (("x", 0), ssig_T); val j = Var (("j", 0), fpT); val flat = Free (Binding.name_of flat_b, ssig_ssig_T --> ssig_T); val Oper' = substT Y ssig_T Oper; val VLeaf' = substT Y ssig_T VLeaf; val CLeaf' = substT Y ssig_T CLeaf; val dead_sig_map' = Term.subst_atomic_types [(Y, ssig_ssig_T), (Z, ssig_T)] dead_sig_map; val Oper_eq = mk_Trueprop_eq (flat $ (Oper' $ sigx), Oper $ (dead_sig_map' $ flat $ sigx)) |> Logic.all sigx; val VLeaf_eq = mk_Trueprop_eq (flat $ (VLeaf' $ x), x) |> Logic.all x; val CLeaf_eq = mk_Trueprop_eq (flat $ (CLeaf' $ j), CLeaf $ j) |> Logic.all j; in define_single_primrec flat_b [Oper_eq, VLeaf_eq, CLeaf_eq] lthy end; fun define_eval_core fp_b version Y Z preT fpT sig_T ssig_T dtor Oper VLeaf CLeaf dead_pre_map dead_sig_map dead_ssig_map flat Lam lthy = let val eval_core_b = mk_version_fp_binding true lthy version fp_b eval_coreN; val YpreT = HOLogic.mk_prodT (Y, preT); val Ypre_ssig_T = Tsubst Y YpreT ssig_T; val Ypre_ssig_sig_T = Tsubst Y Ypre_ssig_T sig_T; val ssig_preT = Tsubst Y ssig_T preT; val ssig_YpreT = Tsubst Y ssig_T YpreT; val ssig_ssig_T = Tsubst Y ssig_T ssig_T; val sigx = Var (("s", 0), Ypre_ssig_sig_T); val x = Var (("x", 0), YpreT); val j = Var (("j", 0), fpT); val eval_core = Free (Binding.name_of eval_core_b, Ypre_ssig_T --> ssig_preT); val Oper' = substT Y YpreT Oper; val VLeaf' = substT Y YpreT VLeaf; val CLeaf' = substT Y YpreT CLeaf; val dead_pre_map' = Term.subst_atomic_types [(Y, ssig_ssig_T), (Z, ssig_T)] dead_pre_map; val dead_pre_map'' = substT Z ssig_T dead_pre_map; val dead_pre_map''' = Term.subst_atomic_types [(Y, fpT), (Z, ssig_T)] dead_pre_map; val dead_sig_map' = Term.subst_atomic_types [(Y, Ypre_ssig_T), (Z, ssig_YpreT)] dead_sig_map; val dead_ssig_map' = Term.subst_atomic_types [(Y, YpreT), (Z, Y)] dead_ssig_map; val Lam' = substT Y ssig_T Lam; val fst' = fst_const YpreT; val snd' = snd_const YpreT; val Oper_eq = mk_Trueprop_eq (eval_core $ (Oper' $ sigx), dead_pre_map' $ flat $ (Lam' $ (dead_sig_map' $ (Abs (Name.uu, Ypre_ssig_T, HOLogic.mk_prod (dead_ssig_map' $ fst' $ Bound 0, eval_core $ Bound 0))) $ sigx))) |> Logic.all sigx; val VLeaf_eq = mk_Trueprop_eq (eval_core $ (VLeaf' $ x), dead_pre_map'' $ VLeaf $ (snd' $ x)) |> Logic.all x; val CLeaf_eq = mk_Trueprop_eq (eval_core $ (CLeaf' $ j), dead_pre_map''' $ CLeaf $ (dtor $ j)) |> Logic.all j; in define_single_primrec eval_core_b [Oper_eq, VLeaf_eq, CLeaf_eq] lthy end; fun define_eval fp_b version Y Z preT fpT ssig_T dtor dtor_unfold dead_ssig_map eval_core lthy = let val fp_preT = Tsubst Y fpT preT; val fppreT = HOLogic.mk_prodT (fpT, fp_preT); val fp_ssig_T = Tsubst Y fpT ssig_T; val dtor_unfold' = substT Z fp_ssig_T dtor_unfold; val dead_ssig_map' = Term.subst_atomic_types [(Y, fpT), (Z, fppreT)] dead_ssig_map; val eval_core' = substT Y fpT eval_core; val id' = HOLogic.id_const fpT; val rhs = dtor_unfold' $ HOLogic.mk_comp (eval_core', dead_ssig_map' $ mk_convol (id', dtor)); in define_const true fp_b version evalN rhs lthy end; fun define_cutSsig fp_b version Y Z preT ssig_T dead_pre_map VLeaf dead_ssig_map flat eval_core lthy = let val ssig_preT = Tsubst Y ssig_T preT; val ssig_ssig_T = Tsubst Y ssig_T ssig_T; val ssig_ssig_preT = HOLogic.mk_prodT (ssig_T, ssig_preT); val h = Var (("h", 0), Y --> ssig_preT); val dead_pre_map' = Term.subst_atomic_types [(Y, ssig_ssig_T), (Z, ssig_T)] dead_pre_map; val dead_ssig_map' = substT Z ssig_ssig_preT dead_ssig_map; val eval_core' = substT Y ssig_T eval_core; val rhs = Library.foldl1 HOLogic.mk_comp [dead_pre_map' $ flat, eval_core', dead_ssig_map' $ mk_convol (VLeaf, h)] |> Term.lambda h; in define_const true fp_b version cutSsigN rhs lthy end; fun define_algLam fp_b version Y Z fpT ssig_T Oper VLeaf dead_sig_map eval lthy = let val fp_ssig_T = Tsubst Y fpT ssig_T; val Oper' = substT Y fpT Oper; val VLeaf' = substT Y fpT VLeaf; val dead_sig_map' = Term.subst_atomic_types [(Y, fpT), (Z, fp_ssig_T)] dead_sig_map; val rhs = Library.foldl1 HOLogic.mk_comp [eval, Oper', dead_sig_map' $ VLeaf']; in define_const true fp_b version algLamN rhs lthy end; fun define_corecU fp_b version Y Z preT ssig_T dtor_unfold VLeaf cutSsig lthy = let val ssig_preT = Tsubst Y ssig_T preT; val h = Var (("h", 0), Y --> ssig_preT); val dtor_unfold' = substT Z ssig_T dtor_unfold; val rhs = HOLogic.mk_comp (dtor_unfold' $ (cutSsig $ h), VLeaf) |> Term.lambda h; in define_const true fp_b version corecUN rhs lthy end; fun define_corecUU fp_b version Y Z preT ssig_T dead_pre_map dead_ssig_map flat eval_core sctr corecU lthy = let val ssig_preT = Tsubst Y ssig_T preT; val ssig_ssig_T = Tsubst Y ssig_T ssig_T val ssig_ssig_preT = HOLogic.mk_prodT (ssig_T, ssig_preT); val ssig_pre_ssig_T = Tsubst Y ssig_preT ssig_T; val h = Var (("h", 0), Y --> ssig_pre_ssig_T); val dead_pre_map' = Term.subst_atomic_types [(Y, ssig_ssig_T), (Z, ssig_T)] dead_pre_map; val eval_core' = substT Y ssig_T eval_core; val dead_ssig_map' = Term.subst_atomic_types [(Y, ssig_preT), (Z, ssig_ssig_preT)] dead_ssig_map; val id' = HOLogic.id_const ssig_preT; val rhs = corecU $ Library.foldl1 HOLogic.mk_comp [dead_pre_map' $ flat, eval_core', dead_ssig_map' $ mk_convol (sctr, id'), h] |> Term.lambda h; in define_const true fp_b version corecUUN rhs lthy end; fun derive_sig_transfer maybe_swap ctxt live_AsBs pre_rel sig_rel Rs R const pre_rel_def preT_rel_eqs transfer_thm = let val RRpre_rel = list_comb (pre_rel, Rs) $ R; val RRsig_rel = list_comb (sig_rel, Rs) $ R; val constB = Term.subst_atomic_types live_AsBs const; val goal = uncurry mk_rel_fun (maybe_swap (RRpre_rel, RRsig_rel)) $ const $ constB |> HOLogic.mk_Trueprop; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_sig_transfer_tac ctxt pre_rel_def preT_rel_eqs transfer_thm)) |> Thm.close_derivation \<^here> end; fun derive_transfer_by_transfer_prover ctxt live_AsBs Rs R const const_defs rel_eqs transfers = let val constB = Term.subst_atomic_types live_AsBs const; val goal = mk_parametricity_goal ctxt (Rs @ [R]) const constB; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_transfer_by_transfer_prover_tac ctxt (const_defs @ map (fn thm => thm RS sym) rel_eqs) rel_eqs transfers)) |> Thm.close_derivation \<^here> end; fun derive_dtor_transfer ctxt live_EsFs Y Z pre_rel fp_rel Rs dtor dtor_rel_thm = let val Type (\<^type_name>\fun\, [fpT, Type (\<^type_name>\fun\, [fpTB, \<^typ>\bool\])]) = snd (strip_typeN (length live_EsFs) (fastype_of fp_rel)); val pre_rel' = Term.subst_atomic_types [(Y, fpT), (Z, fpTB)] pre_rel; val Rpre_rel = list_comb (pre_rel', Rs); val Rfp_rel = list_comb (fp_rel, Rs); val dtorB = Term.subst_atomic_types live_EsFs dtor; val goal = HOLogic.mk_Trueprop (mk_rel_fun Rfp_rel (Rpre_rel $ Rfp_rel) $ dtor $ dtorB); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_dtor_transfer_tac ctxt dtor_rel_thm)) |> Thm.close_derivation \<^here> end; fun derive_Lam_or_eval_core_transfer ctxt live_AsBs Y Z preT ssig_T Rs R pre_rel sig_or_ssig_rel ssig_rel const const_def rel_eqs transfers = let val YpreT = HOLogic.mk_prodT (Y, preT); val ZpreTB = typ_subst_atomic live_AsBs YpreT; val ssig_TB = typ_subst_atomic live_AsBs ssig_T; val pre_rel' = Term.subst_atomic_types [(Y, ssig_T), (Z, ssig_TB)] pre_rel; val sig_or_ssig_rel' = Term.subst_atomic_types [(Y, YpreT), (Z, ZpreTB)] sig_or_ssig_rel; val Rsig_or_ssig_rel' = list_comb (sig_or_ssig_rel', Rs); val RRpre_rel = list_comb (pre_rel, Rs) $ R; val RRssig_rel = list_comb (ssig_rel, Rs) $ R; val Rpre_rel' = list_comb (pre_rel', Rs); val constB = subst_atomic_types live_AsBs const; val goal = mk_rel_fun (Rsig_or_ssig_rel' $ mk_rel_prod R RRpre_rel) (Rpre_rel' $ RRssig_rel) $ const $ constB |> HOLogic.mk_Trueprop; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_transfer_by_transfer_prover_tac ctxt [const_def] rel_eqs transfers)) |> Thm.close_derivation \<^here> end; fun derive_proto_sctr_transfer_step_or_merge ctxt Y Z R dead_pre_rel dead_sig_rel proto_sctr proto_sctr_def fp_k_T_rel_eqs transfers = let val proto_sctrZ = substT Y Z proto_sctr; val goal = mk_rel_fun (dead_pre_rel $ R) (dead_sig_rel $ R) $ proto_sctr $ proto_sctrZ |> HOLogic.mk_Trueprop; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_transfer_by_transfer_prover_tac ctxt [proto_sctr_def] fp_k_T_rel_eqs transfers)) |> Thm.close_derivation \<^here> end; fun derive_sctr_transfer ctxt live_AsBs Y Z ssig_T Rs R pre_rel ssig_rel sctr sctr_def fp_k_T_rel_eqs transfers = let val ssig_TB = typ_subst_atomic live_AsBs ssig_T; val pre_rel' = Term.subst_atomic_types [(Y, ssig_T), (Z, ssig_TB)] pre_rel; val Rpre_rel' = list_comb (pre_rel', Rs); val RRssig_rel = list_comb (ssig_rel, Rs) $ R; val sctrB = subst_atomic_types live_AsBs sctr; val goal = HOLogic.mk_Trueprop (mk_rel_fun (Rpre_rel' $ RRssig_rel) RRssig_rel $ sctr $ sctrB); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_transfer_by_transfer_prover_tac ctxt [sctr_def] fp_k_T_rel_eqs transfers)) |> Thm.close_derivation \<^here> end; fun derive_corecUU_transfer ctxt live_AsBs Y Z Rs R preT ssig_T pre_rel fp_rel ssig_rel corecUU cutSsig_def corecU_def corecUU_def fp_k_T_rel_eqs transfers = let val ssig_preT = Tsubst Y ssig_T preT; val ssig_TB = typ_subst_atomic live_AsBs ssig_T; val ssig_preTB = typ_subst_atomic live_AsBs ssig_preT; val pre_rel' = Term.subst_atomic_types [(Y, ssig_T), (Z, ssig_TB)] pre_rel; val ssig_rel' = Term.subst_atomic_types [(Y, ssig_preT), (Z, ssig_preTB)] ssig_rel; val Rpre_rel' = list_comb (pre_rel', Rs); val Rfp_rel = list_comb (fp_rel, Rs); val RRssig_rel = list_comb (ssig_rel, Rs) $ R; val Rssig_rel' = list_comb (ssig_rel', Rs); val corecUUB = subst_atomic_types live_AsBs corecUU; val goal = mk_rel_fun (mk_rel_fun R (Rssig_rel' $ (Rpre_rel' $ RRssig_rel))) (mk_rel_fun R Rfp_rel) $ corecUU $ corecUUB |> HOLogic.mk_Trueprop; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_transfer_by_transfer_prover_tac ctxt [cutSsig_def, corecU_def, corecUU_def] fp_k_T_rel_eqs transfers)) |> Thm.close_derivation \<^here> end; fun mk_natural_goal ctxt simple_T_mapfs fs t u = let fun build_simple (T, _) = (case AList.lookup (op =) simple_T_mapfs T of SOME mapf => mapf | NONE => the (find_first (fn f => domain_type (fastype_of f) = T) fs)); val simple_Ts = map fst simple_T_mapfs; val t_map = build_map ctxt simple_Ts [] build_simple (apply2 (range_type o fastype_of) (t, u)); val u_map = build_map ctxt simple_Ts [] build_simple (apply2 (domain_type o fastype_of) (t, u)); in mk_Trueprop_eq (HOLogic.mk_comp (u, u_map), HOLogic.mk_comp (t_map, t)) end; fun derive_natural_by_unfolding ctxt live_AsBs preT pre_map fs f const map_thms = let val ffpre_map = list_comb (pre_map, fs) $ f; val constB = subst_atomic_types live_AsBs const; val goal = mk_natural_goal ctxt [(preT, ffpre_map)] (fs @ [f]) const constB; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_natural_by_unfolding_tac ctxt map_thms)) |> Thm.close_derivation \<^here> end; fun derive_natural_from_transfer ctxt live_AsBs simple_T_mapfs fs f const transfer bnfs subst_bnfs = let val m = length live_AsBs; val constB = Term.subst_atomic_types live_AsBs const; val goal = mk_natural_goal ctxt simple_T_mapfs (fs @ [f]) const constB; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_natural_from_transfer_tac ctxt m (replicate m true) transfer [] (map rel_Grp_of_bnf bnfs) (map rel_Grp_of_bnf subst_bnfs))) |> Thm.close_derivation \<^here> end; fun derive_natural_from_transfer_with_pre_type ctxt live_AsBs Y Z preT ssig_T pre_map ssig_map fs f = let val ssig_TB = typ_subst_atomic live_AsBs ssig_T; val preT' = Term.typ_subst_atomic [(Y, ssig_T), (Z, ssig_TB)] preT; val ffpre_map = list_comb (pre_map, fs) $ f; val pre_map' = Term.subst_atomic_types [(Y, ssig_T), (Z, ssig_TB)] pre_map; val fpre_map' = list_comb (pre_map', fs); val ffssig_map = list_comb (ssig_map, fs) $ f; val preT_mapfs = [(preT, ffpre_map), (preT', fpre_map' $ ffssig_map)]; in derive_natural_from_transfer ctxt live_AsBs preT_mapfs fs f end; fun derive_Lam_Inl_Inr ctxt Y Z preT old_sig_T old_ssig_T k_T ssig_T dead_pre_map Sig embL old_Lam Lam rho unsig_thm Lam_def = let val YpreT = HOLogic.mk_prodT (Y, preT); val Ypre_old_sig_T = Tsubst Y YpreT old_sig_T; val Ypre_k_T = Tsubst Y YpreT k_T; val inl' = Inl_const Ypre_old_sig_T Ypre_k_T; val inr' = Inr_const Ypre_old_sig_T Ypre_k_T; val dead_pre_map' = Term.subst_atomic_types [(Y, old_ssig_T), (Z, ssig_T)] dead_pre_map; val Sig' = substT Y YpreT Sig; val Lam_o_Sig = HOLogic.mk_comp (Lam, Sig'); val inl_goal = mk_Trueprop_eq (HOLogic.mk_comp (Lam_o_Sig, inl'), HOLogic.mk_comp (dead_pre_map' $ embL, old_Lam)); val inr_goal = mk_Trueprop_eq (HOLogic.mk_comp (Lam_o_Sig, inr'), rho); val goals = [inl_goal, inr_goal]; val goal = Logic.mk_conjunction_balanced goals; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_Lam_Inl_Inr_tac ctxt unsig_thm Lam_def)) |> Conjunction.elim_balanced (length goals) |> map (Thm.close_derivation \<^here>) end; fun derive_flat_VLeaf ctxt Y Z ssig_T x VLeaf dead_ssig_map flat ssig_induct fp_map_id sig_map_cong sig_map_ident sig_map_comp ssig_map_thms flat_simps = let val x' = substT Y ssig_T x; val dead_ssig_map' = substT Z ssig_T dead_ssig_map; val goal = mk_Trueprop_eq (flat $ (dead_ssig_map' $ VLeaf $ x'), x'); val ssig_induct' = infer_instantiate' ctxt [NONE, SOME (Thm.cterm_of ctxt x')] ssig_induct; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_flat_VLeaf_or_flat_tac ctxt ssig_induct' sig_map_cong (fp_map_id :: sig_map_ident :: sig_map_comp :: ssig_map_thms @ flat_simps @ @{thms o_apply id_apply id_def[symmetric]}))) |> Thm.close_derivation \<^here> end; fun derive_flat_flat ctxt Y Z ssig_T x dead_ssig_map flat ssig_induct fp_map_id sig_map_cong sig_map_comp ssig_map_thms flat_simps = let val ssig_ssig_T = Tsubst Y ssig_T ssig_T; val ssig_ssig_ssig_T = Tsubst Y ssig_T ssig_ssig_T; val x' = substT Y ssig_ssig_ssig_T x; val dead_ssig_map' = Term.subst_atomic_types [(Y, ssig_ssig_T), (Z, ssig_T)] dead_ssig_map; val flat' = substT Y ssig_T flat; val goal = mk_Trueprop_eq (flat $ (dead_ssig_map' $ flat $ x'), flat $ (flat' $ x')); val ssig_induct' = infer_instantiate' ctxt [NONE, SOME (Thm.cterm_of ctxt x')] ssig_induct; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_flat_VLeaf_or_flat_tac ctxt ssig_induct' sig_map_cong (o_apply :: fp_map_id :: sig_map_comp :: ssig_map_thms @ flat_simps))) |> Thm.close_derivation \<^here> end; fun derive_eval_core_flat ctxt Y Z preT ssig_T dead_pre_map dead_ssig_map flat eval_core x ssig_induct dead_pre_map_id dead_pre_map_comp0 dead_pre_map_comp fp_map_id sig_map_comp sig_map_cong ssig_map_thms ssig_map_comp flat_simps flat_pointful_natural flat_flat Lam_pointful_natural eval_core_simps = let val YpreT = HOLogic.mk_prodT (Y, preT); val ssig_ssig_T = Tsubst Y ssig_T ssig_T; val Ypre_ssig_T = Tsubst Y YpreT ssig_T; val Ypre_ssig_ssig_T = Tsubst Y YpreT ssig_ssig_T; val ssig_YpreT = Tsubst Y ssig_T YpreT; val dead_pre_map' = Term.subst_atomic_types [(Y, ssig_ssig_T), (Z, ssig_T)] dead_pre_map; val dead_ssig_map' = Term.subst_atomic_types [(Y, Ypre_ssig_T), (Z, ssig_YpreT)] dead_ssig_map; val dead_ssig_map'' = Term.subst_atomic_types [(Y, YpreT), (Z, Y)] dead_ssig_map; val flat' = substT Y YpreT flat; val eval_core' = substT Y ssig_T eval_core; val x' = substT Y Ypre_ssig_ssig_T x; val fst' = fst_const YpreT; val goal = mk_Trueprop_eq (eval_core $ (flat' $ x'), dead_pre_map' $ flat $ (eval_core' $ (dead_ssig_map' $ mk_convol (dead_ssig_map'' $ fst', eval_core) $ x'))); val ssig_induct' = infer_instantiate' ctxt [NONE, SOME (Thm.cterm_of ctxt x')] ssig_induct; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_eval_core_flat_tac ctxt ssig_induct' dead_pre_map_id dead_pre_map_comp0 dead_pre_map_comp fp_map_id sig_map_comp sig_map_cong ssig_map_thms ssig_map_comp flat_simps flat_pointful_natural flat_flat Lam_pointful_natural eval_core_simps)) |> Thm.close_derivation \<^here> end; fun derive_eval_thm ctxt dtor_inject dtor_unfold_thm eval_def = (trans OF [iffD2 OF [dtor_inject, HOLogic.mk_obj_eq eval_def RS fun_cong], dtor_unfold_thm]) |> unfold_thms ctxt [o_apply, eval_def RS symmetric_thm]; fun derive_eval_flat ctxt Y Z fpT ssig_T dead_ssig_map flat eval x dead_pre_map_comp0 dtor_unfold_unique ssig_map_id ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_core_flat eval_thm = let val fp_ssig_T = Tsubst Y fpT ssig_T; val fp_ssig_ssig_T = Tsubst Y fp_ssig_T ssig_T; val dead_ssig_map' = Term.subst_atomic_types [(Y, fp_ssig_T), (Z, fpT)] dead_ssig_map; val flat' = substT Y fpT flat; val x' = substT Y fp_ssig_ssig_T x; val goal = mk_Trueprop_eq (eval $ (flat' $ x'), eval $ (dead_ssig_map' $ eval $ x')); val cond_eval_o_flat = infer_instantiate' ctxt [SOME (Thm.cterm_of ctxt (HOLogic.mk_comp (eval, flat')))] (trans OF [dtor_unfold_unique, dtor_unfold_unique RS sym] RS fun_cong) OF [ext, ext]; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_eval_flat_tac ctxt dead_pre_map_comp0 ssig_map_id ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_core_flat eval_thm cond_eval_o_flat)) |> Thm.close_derivation \<^here> end; fun derive_eval_Oper ctxt live Y Z fpT sig_T ssig_T dead_sig_map Oper eval algLam x sig_map_ident sig_map_comp0 sig_map_comp Oper_natural_pointful VLeaf_natural flat_simps eval_flat algLam_def = let val fp_ssig_T = Tsubst Y fpT ssig_T; val fp_ssig_sig_T = Tsubst Y fp_ssig_T sig_T; val dead_sig_map' = Term.subst_atomic_types [(Y, fp_ssig_T), (Z, fpT)] dead_sig_map; val Oper' = substT Y fpT Oper; val x' = substT Y fp_ssig_sig_T x; val goal = mk_Trueprop_eq (eval $ (Oper' $ x'), algLam $ (dead_sig_map' $ eval $ x')); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_eval_Oper_tac ctxt live sig_map_ident sig_map_comp0 sig_map_comp Oper_natural_pointful VLeaf_natural flat_simps eval_flat algLam_def)) |> Thm.close_derivation \<^here> end; fun derive_eval_V_or_CLeaf ctxt Y fpT V_or_CLeaf eval x dead_pre_map_id dead_pre_map_comp fp_map_id dtor_unfold_unique V_or_CLeaf_map_thm eval_core_simps eval_thm = let val V_or_CLeaf' = substT Y fpT V_or_CLeaf; val x' = substT Y fpT x; val goal = mk_Trueprop_eq (eval $ (V_or_CLeaf' $ x'), x'); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_eval_V_or_CLeaf_tac ctxt dead_pre_map_id dead_pre_map_comp fp_map_id dtor_unfold_unique V_or_CLeaf_map_thm eval_core_simps eval_thm)) |> Thm.close_derivation \<^here> end; fun derive_extdd_mor ctxt Y Z preT fpT ssig_T dead_pre_map dtor extdd cutSsig f g dead_pre_map_comp0 dead_pre_map_comp VLeaf_map_thm ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_thm eval_flat eval_VLeaf cutSsig_def = let val ssig_preT = Tsubst Y ssig_T preT; val dead_pre_map' = Term.subst_atomic_types [(Y, ssig_T), (Z, fpT)] dead_pre_map; val f' = substT Z fpT f; val g' = substT Z ssig_preT g; val extdd_f = extdd $ f'; val prem = mk_Trueprop_eq (HOLogic.mk_comp (dead_pre_map' $ extdd_f, g'), HOLogic.mk_comp (dtor, f')); val goal = mk_Trueprop_eq (HOLogic.mk_comp (dead_pre_map' $ extdd_f, cutSsig $ g'), HOLogic.mk_comp (dtor, extdd_f)); in fold (Variable.add_free_names ctxt) [prem, goal] [] |> (fn vars => Goal.prove_sorry ctxt vars [prem] goal (fn {context = ctxt, prems = [prem]} => mk_extdd_mor_tac ctxt dead_pre_map_comp0 dead_pre_map_comp VLeaf_map_thm ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_thm eval_flat eval_VLeaf cutSsig_def prem)) |> Thm.close_derivation \<^here> end; fun derive_mor_cutSsig_flat ctxt Y Z preT fpT ssig_T dead_pre_map dead_ssig_map dtor flat eval_core eval cutSsig f g dead_pre_map_comp0 dead_pre_map_comp dead_pre_map_cong dtor_unfold_unique dead_ssig_map_comp0 ssig_map_comp flat_simps flat_pointful_natural eval_core_pointful_natural flat_flat flat_VLeaf eval_core_flat cutSsig_def cutSsig_def_pointful_natural eval_thm = let val ssig_preT = Tsubst Y ssig_T preT; val substYZ = Term.subst_atomic_types [(Y, ssig_T), (Z, fpT)]; val dead_pre_map' = substYZ dead_pre_map; val dead_ssig_map' = substYZ dead_ssig_map; val f' = substYZ f; val g' = substT Z ssig_preT g; val cutSsig_g = cutSsig $ g'; val id' = HOLogic.id_const ssig_T; val convol' = mk_convol (id', cutSsig_g); val dead_ssig_map'' = Term.subst_atomic_types [(Y, ssig_T), (Z, range_type (fastype_of convol'))] dead_ssig_map; val eval_core' = substT Y ssig_T eval_core; val eval_core_o_map = HOLogic.mk_comp (eval_core', dead_ssig_map'' $ convol'); val prem = mk_Trueprop_eq (HOLogic.mk_comp (dead_pre_map' $ f', cutSsig_g), HOLogic.mk_comp (dtor, f')); val goal = mk_Trueprop_eq (HOLogic.mk_comp (eval, dead_ssig_map' $ f'), HOLogic.mk_comp (f', flat)); in fold (Variable.add_free_names ctxt) [prem, goal] [] |> (fn vars => Goal.prove_sorry ctxt vars [prem] goal (fn {context = ctxt, prems = [prem]} => mk_mor_cutSsig_flat_tac ctxt eval_core_o_map dead_pre_map_comp0 dead_pre_map_comp dead_pre_map_cong dtor_unfold_unique dead_ssig_map_comp0 ssig_map_comp flat_simps flat_pointful_natural eval_core_pointful_natural flat_flat flat_VLeaf eval_core_flat cutSsig_def cutSsig_def_pointful_natural eval_thm prem)) |> Thm.close_derivation \<^here> end; fun derive_extdd_o_VLeaf ctxt Y Z preT fpT ssig_T dead_pre_map dtor VLeaf extdd f g dead_pre_map_comp0 dead_pre_map_comp dtor_inject ssig_map_thms eval_core_simps eval_thm eval_VLeaf = let val ssig_preT = Tsubst Y ssig_T preT; val substYZ = Term.subst_atomic_types [(Y, ssig_T), (Z, fpT)]; val dead_pre_map' = substYZ dead_pre_map; val f' = substT Z fpT f; val g' = substT Z ssig_preT g; val extdd_f = extdd $ f'; val prem = mk_Trueprop_eq (HOLogic.mk_comp (dead_pre_map' $ extdd_f, g'), HOLogic.mk_comp (dtor, f')); val goal = mk_Trueprop_eq (HOLogic.mk_comp (extdd_f, VLeaf), f'); in fold (Variable.add_free_names ctxt) [prem, goal] [] |> (fn vars => Goal.prove_sorry ctxt vars [prem] goal (fn {context = ctxt, prems = [prem]} => mk_extdd_o_VLeaf_tac ctxt dead_pre_map_comp0 dead_pre_map_comp dtor_inject ssig_map_thms eval_core_simps eval_thm eval_VLeaf prem)) |> Thm.close_derivation \<^here> end; fun derive_corecU_pointfree ctxt Y Z preT fpT ssig_T dead_pre_map dtor extdd corecU g dead_pre_map_comp dtor_unfold_thm ssig_map_thms dead_ssig_map_comp0 flat_simps flat_VLeaf eval_core_simps cutSsig_def mor_cutSsig_flat corecU_def = let val ssig_preT = Tsubst Y ssig_T preT; val substYZ = Term.subst_atomic_types [(Y, ssig_T), (Z, fpT)]; val dead_pre_map' = substYZ dead_pre_map; val g' = substT Z ssig_preT g; val corecU_g = corecU $ g'; val goal = mk_Trueprop_eq (HOLogic.mk_comp (dead_pre_map' $ (extdd $ corecU_g), g'), HOLogic.mk_comp (dtor, corecU_g)); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_corecU_pointfree_tac ctxt dead_pre_map_comp dtor_unfold_thm ssig_map_thms dead_ssig_map_comp0 flat_simps flat_VLeaf eval_core_simps cutSsig_def mor_cutSsig_flat corecU_def)) |> Thm.close_derivation \<^here> end; fun derive_corecU_ctor_unique ctxt Y Z preT fpT ssig_T dead_pre_map ctor dtor VLeaf extdd corecU f g dead_pre_map_comp ctor_dtor dtor_unfold_thm dtor_unfold_unique ssig_map_thms dead_ssig_map_comp0 flat_simps flat_VLeaf eval_core_simps extdd_mor extdd_o_VLeaf cutSsig_def mor_cutSsig_flat corecU_def = let val corecU_pointfree = derive_corecU_pointfree ctxt Y Z preT fpT ssig_T dead_pre_map dtor extdd corecU g dead_pre_map_comp dtor_unfold_thm ssig_map_thms dead_ssig_map_comp0 flat_simps flat_VLeaf eval_core_simps cutSsig_def mor_cutSsig_flat corecU_def; val corecU_thm = corecU_pointfree RS @{thm comp_eq_dest}; val corecU_ctor = let val arg_cong' = infer_instantiate' ctxt [NONE, NONE, SOME (Thm.cterm_of ctxt ctor)] arg_cong; in unfold_thms ctxt [ctor_dtor] (corecU_thm RS arg_cong') end; val corecU_unique = let val substYZ = Term.subst_atomic_types [(Y, ssig_T), (Z, fpT)]; val f' = substYZ f; val abs_f_o_VLeaf = Term.lambda f' (HOLogic.mk_comp (f', VLeaf)); val inject_refine' = infer_instantiate' ctxt [SOME (Thm.cterm_of ctxt abs_f_o_VLeaf), SOME (Thm.cterm_of ctxt extdd)] @{thm inject_refine}; in unfold_thms ctxt @{thms atomize_imp} (((inject_refine' OF [extdd_o_VLeaf, extdd_o_VLeaf] RS iffD1) OF [asm_rl, corecU_pointfree]) OF [asm_rl, trans OF [dtor_unfold_unique, dtor_unfold_unique RS sym] OF [extdd_mor, corecU_pointfree RS extdd_mor]]) RS @{thm obj_distinct_prems} end; in (corecU_ctor, corecU_unique) end; fun derive_dtor_algLam ctxt Y Z preT fpT sig_T ssig_T dead_pre_map dtor dead_sig_map Lam eval algLam x pre_map_comp dead_pre_map_id dead_pre_map_comp0 dead_pre_map_comp sig_map_comp Oper_pointful_natural ssig_map_thms dead_ssig_map_comp0 Lam_pointful_natural eval_core_simps eval_thm eval_flat eval_VLeaf algLam_def = let val fp_preT = Tsubst Y fpT preT; val fppreT = HOLogic.mk_prodT (fpT, fp_preT); val fp_sig_T = Tsubst Y fpT sig_T; val fp_ssig_T = Tsubst Y fpT ssig_T; val id' = HOLogic.id_const fpT; val convol' = mk_convol (id', dtor); val dead_pre_map' = Term.subst_atomic_types [(Y, fp_ssig_T), (Z, fpT)] dead_pre_map; val dead_sig_map' = Term.subst_atomic_types [(Y, fpT), (Z, fppreT)] dead_sig_map; val Lam' = substT Y fpT Lam; val x' = substT Y fp_sig_T x; val goal = mk_Trueprop_eq (dtor $ (algLam $ x'), dead_pre_map' $ eval $ (Lam' $ (dead_sig_map' $ convol' $ x'))); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_dtor_algLam_tac ctxt pre_map_comp dead_pre_map_id dead_pre_map_comp0 dead_pre_map_comp sig_map_comp Oper_pointful_natural ssig_map_thms dead_ssig_map_comp0 Lam_pointful_natural eval_core_simps eval_thm eval_flat eval_VLeaf algLam_def)) |> Thm.close_derivation \<^here> end; fun derive_algLam_base ctxt Y Z preT fpT dead_pre_map ctor dtor algLam proto_sctr dead_pre_map_id dead_pre_map_comp ctor_dtor dtor_ctor dtor_unfold_unique unsig_thm Sig_pointful_natural ssig_map_thms Lam_def flat_simps eval_core_simps eval_thm algLam_def = let val fp_preT = Tsubst Y fpT preT; val proto_sctr' = substT Y fpT proto_sctr; val dead_pre_map' = Term.subst_atomic_types [(Y, fpT), (Z, fp_preT)] dead_pre_map; val dead_pre_map_dtor = dead_pre_map' $ dtor; val goal = mk_Trueprop_eq (HOLogic.mk_comp (algLam, proto_sctr'), ctor); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_algLam_base_tac ctxt dead_pre_map_dtor dead_pre_map_id dead_pre_map_comp ctor_dtor dtor_ctor dtor_unfold_unique unsig_thm Sig_pointful_natural ssig_map_thms Lam_def flat_simps eval_core_simps eval_thm algLam_def)) |> Thm.close_derivation \<^here> end; fun derive_flat_embL ctxt Y Z old_ssig_T ssig_T dead_old_ssig_map embL old_flat flat x old_ssig_induct fp_map_id Sig_pointful_natural old_sig_map_comp old_sig_map_cong old_ssig_map_thms old_flat_simps flat_simps embL_simps = let val old_ssig_old_ssig_T = Tsubst Y old_ssig_T old_ssig_T; val dead_old_ssig_map' = Term.subst_atomic_types [(Y, old_ssig_T), (Z, ssig_T)] dead_old_ssig_map; val embL' = substT Y ssig_T embL; val x' = substT Y old_ssig_old_ssig_T x; val goal = mk_Trueprop_eq (flat $ (embL' $ (dead_old_ssig_map' $ embL $ x')), embL $ (old_flat $ x')); val old_ssig_induct' = infer_instantiate' ctxt [NONE, SOME (Thm.cterm_of ctxt x')] old_ssig_induct; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_flat_embL_tac ctxt old_ssig_induct' fp_map_id Sig_pointful_natural old_sig_map_comp old_sig_map_cong old_ssig_map_thms old_flat_simps flat_simps embL_simps)) |> Thm.close_derivation \<^here> end; fun derive_eval_core_embL ctxt Y Z preT old_ssig_T ssig_T dead_pre_map embL old_eval_core eval_core x old_ssig_induct dead_pre_map_comp0 dead_pre_map_comp Sig_pointful_natural unsig_thm old_sig_map_comp old_sig_map_cong old_Lam_pointful_natural Lam_def flat_embL embL_simps embL_pointful_natural old_eval_core_simps eval_core_simps = let val YpreT = HOLogic.mk_prodT (Y, preT); val Ypre_old_ssig_T = Tsubst Y YpreT old_ssig_T; val dead_pre_map' = Term.subst_atomic_types [(Y, old_ssig_T), (Z, ssig_T)] dead_pre_map; val embL' = substT Y YpreT embL; val x' = substT Y Ypre_old_ssig_T x; val goal = mk_Trueprop_eq (eval_core $ (embL' $ x'), dead_pre_map' $ embL $ (old_eval_core $ x')); val old_ssig_induct' = infer_instantiate' ctxt [NONE, SOME (Thm.cterm_of ctxt x')] old_ssig_induct; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_eval_core_embL_tac ctxt old_ssig_induct' dead_pre_map_comp0 dead_pre_map_comp Sig_pointful_natural unsig_thm old_sig_map_comp old_sig_map_cong old_Lam_pointful_natural Lam_def flat_embL old_eval_core_simps eval_core_simps embL_simps embL_pointful_natural)) |> Thm.close_derivation \<^here> end; fun derive_eval_embL ctxt Y fpT embL old_eval eval dead_pre_map_comp0 dtor_unfold_unique embL_pointful_natural eval_core_embL old_eval_thm eval_thm = let val embL' = substT Y fpT embL; val goal = mk_Trueprop_eq (HOLogic.mk_comp (eval, embL'), old_eval); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_eval_embL_tac ctxt dead_pre_map_comp0 dtor_unfold_unique embL_pointful_natural eval_core_embL old_eval_thm eval_thm)) |> Thm.close_derivation \<^here> end; fun derive_algLam_algLam ctxt Inx_const Y fpT Sig old_algLam algLam dead_pre_map_comp dtor_inject unsig_thm sig_map_thm Lam_def eval_embL old_dtor_algLam dtor_algLam = let val Sig' = substT Y fpT Sig; val (left_T, right_T) = dest_sumT (domain_type (fastype_of Sig')); val inx' = Inx_const left_T right_T; val goal = mk_Trueprop_eq (Library.foldl1 HOLogic.mk_comp [algLam, Sig', inx'], old_algLam); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_algLam_algLam_tac ctxt dead_pre_map_comp dtor_inject unsig_thm sig_map_thm Lam_def eval_embL old_dtor_algLam dtor_algLam)) |> Thm.close_derivation \<^here> end; fun derive_eval_core_k_as_ssig ctxt Y preT k_T rho eval_core k_as_ssig x pre_map_comp dead_pre_map_id sig_map_comp ssig_map_thms Lam_natural_pointful Lam_Inr flat_VLeaf eval_core_simps = let val YpreT = HOLogic.mk_prodT (Y, preT); val Ypre_k_T = Tsubst Y YpreT k_T; val k_as_ssig' = substT Y YpreT k_as_ssig; val x' = substT Y Ypre_k_T x; val goal = mk_Trueprop_eq (eval_core $ (k_as_ssig' $ x'), rho $ x'); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_eval_core_k_as_ssig_tac ctxt pre_map_comp dead_pre_map_id sig_map_comp ssig_map_thms Lam_natural_pointful Lam_Inr flat_VLeaf eval_core_simps)) |> Thm.close_derivation \<^here> end; fun derive_algLam_algrho ctxt Y fpT Sig algLam algrho algLam_def algrho_def = let val Sig' = substT Y fpT Sig; val (left_T, right_T) = dest_sumT (domain_type (fastype_of Sig')); val inr' = Inr_const left_T right_T; val goal = mk_Trueprop_eq (Library.foldl1 HOLogic.mk_comp [algLam, Sig', inr'], algrho); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_algLam_algrho_tac ctxt algLam_def algrho_def)) |> Thm.close_derivation \<^here> end; fun derive_dtor_algrho ctxt Y Z preT fpT k_T ssig_T dead_pre_map dead_k_map dtor rho eval algrho x eval_thm k_as_ssig_natural_pointful eval_core_k_as_ssig algrho_def = let val YpreT = HOLogic.mk_prodT (Y, preT); val fppreT = Tsubst Y fpT YpreT; val fp_k_T = Tsubst Y fpT k_T; val fp_ssig_T = Tsubst Y fpT ssig_T; val id' = HOLogic.id_const fpT; val convol' = mk_convol (id', dtor); val dead_pre_map' = Term.subst_atomic_types [(Y, fp_ssig_T), (Z, fpT)] dead_pre_map; val dead_k_map' = Term.subst_atomic_types [(Y, fpT), (Z, fppreT)] dead_k_map; val rho' = substT Y fpT rho; val x' = substT Y fp_k_T x; val goal = mk_Trueprop_eq (dtor $ (algrho $ x'), dead_pre_map' $ eval $ (rho' $ (dead_k_map' $ convol' $ x'))); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_dtor_algrho_tac ctxt eval_thm k_as_ssig_natural_pointful eval_core_k_as_ssig algrho_def)) |> Thm.close_derivation \<^here> end; fun derive_algLam_step_or_merge ctxt Y fpT ctor proto_sctr algLam proto_sctr_def old_algLam_pointful algLam_algLam = let val proto_sctr' = substT Y fpT proto_sctr; val goal = mk_Trueprop_eq (HOLogic.mk_comp (algLam, proto_sctr'), ctor); val algLam_algLam_pointful = mk_pointful ctxt algLam_algLam; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_algLam_step_tac ctxt proto_sctr_def old_algLam_pointful algLam_algLam_pointful)) |> Thm.close_derivation \<^here> end; fun derive_eval_sctr ctxt Y Z fpT ssig_T dead_pre_map ctor eval sctr proto_sctr_pointful_natural eval_Oper algLam_thm sctr_def = let val fp_ssig_T = Tsubst Y fpT ssig_T; val dead_pre_map' = Term.subst_atomic_types [(Y, fp_ssig_T), (Z, fpT)] dead_pre_map; val sctr' = substT Y fpT sctr; val goal = mk_Trueprop_eq (HOLogic.mk_comp (eval, sctr'), HOLogic.mk_comp (ctor, dead_pre_map' $ eval)); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_eval_sctr_tac ctxt proto_sctr_pointful_natural eval_Oper algLam_thm sctr_def)) |> Thm.close_derivation \<^here> end; fun derive_corecUU_pointfree_unique ctxt Y Z preT fpT ssig_T dead_pre_map ctor dead_ssig_map eval corecUU f g dead_pre_map_comp0 dead_pre_map_comp dtor_ctor dtor_inject ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_thm eval_flat corecU_ctor corecU_unique sctr_pointful_natural eval_sctr_pointful corecUU_def = let val ssig_preT = Tsubst Y ssig_T preT; val ssig_pre_ssig_T = Tsubst Y ssig_preT ssig_T; val fp_ssig_T = Tsubst Y fpT ssig_T; val dead_pre_map' = Term.subst_atomic_types [(Y, fp_ssig_T), (Z, fpT)] dead_pre_map; val dead_pre_map'' = Term.subst_atomic_types [(Y, ssig_T), (Z, fp_ssig_T)] dead_pre_map; val dead_ssig_map' = Term.subst_atomic_types [(Y, ssig_preT), (Z, fpT)] dead_ssig_map; val dead_ssig_map'' = substT Z fpT dead_ssig_map; val f' = substT Z ssig_pre_ssig_T f; val g' = substT Z fpT g; val corecUU_f = corecUU $ f'; fun mk_eq fpf = mk_Trueprop_eq (fpf, Library.foldl1 HOLogic.mk_comp [eval, dead_ssig_map' $ Library.foldl1 HOLogic.mk_comp [ctor, dead_pre_map' $ eval, dead_pre_map'' $ (dead_ssig_map'' $ fpf)], f']); val corecUU_pointfree = let val goal = mk_eq corecUU_f; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_corecUU_pointfree_tac ctxt dead_pre_map_comp0 dead_pre_map_comp dtor_ctor dtor_inject ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_thm eval_flat corecU_ctor sctr_pointful_natural eval_sctr_pointful corecUU_def)) |> Thm.close_derivation \<^here> end; val corecUU_unique = let val prem = mk_eq g'; val goal = mk_Trueprop_eq (g', corecUU_f); in fold (Variable.add_free_names ctxt) [prem, goal] [] |> (fn vars => Goal.prove_sorry ctxt vars [prem] goal (fn {context = ctxt, prems = [prem]} => mk_corecUU_unique_tac ctxt dead_pre_map_comp0 dead_pre_map_comp dtor_ctor ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_thm eval_flat corecU_unique sctr_pointful_natural eval_sctr_pointful corecUU_def prem)) |> Thm.close_derivation \<^here> end; in (corecUU_pointfree, corecUU_unique) end; fun define_flat_etc fp_b version live_AsBs Y Z preT fpT sig_T ssig_T Oper VLeaf CLeaf pre_rel dead_pre_map dtor dtor_unfold dead_sig_map ssig_rel dead_ssig_map Lam Rs R pre_map_transfer fp_k_T_rel_eqs sig_map_transfer ssig_map_transfer Lam_transfer dtor_transfer lthy = let val (flat_data as (flat, flat_def, _), lthy) = lthy |> define_flat fp_b version Y Z fpT sig_T ssig_T Oper VLeaf CLeaf dead_sig_map; val (eval_core_data as (eval_core, eval_core_def, _), lthy) = lthy |> define_eval_core fp_b version Y Z preT fpT sig_T ssig_T dtor Oper VLeaf CLeaf dead_pre_map dead_sig_map dead_ssig_map flat Lam; val ((eval_data as (eval, _), cutSsig_data as (cutSsig, _)), lthy) = lthy |> define_eval fp_b version Y Z preT fpT ssig_T dtor dtor_unfold dead_ssig_map eval_core ||>> define_cutSsig fp_b version Y Z preT ssig_T dead_pre_map VLeaf dead_ssig_map flat eval_core; val ((algLam_data, unfold_data), lthy) = lthy |> define_algLam fp_b version Y Z fpT ssig_T Oper VLeaf dead_sig_map eval ||>> define_corecU fp_b version Y Z preT ssig_T dtor_unfold VLeaf cutSsig; val flat_transfer = derive_transfer_by_transfer_prover lthy live_AsBs Rs R flat [flat_def] [] [sig_map_transfer]; val eval_core_transfer = derive_Lam_or_eval_core_transfer lthy live_AsBs Y Z preT ssig_T Rs R pre_rel ssig_rel ssig_rel eval_core eval_core_def fp_k_T_rel_eqs [pre_map_transfer, sig_map_transfer, ssig_map_transfer, flat_transfer, Lam_transfer, dtor_transfer]; in (((((((flat_data, flat_transfer), (eval_core_data, eval_core_transfer)), eval_data), cutSsig_data), algLam_data), unfold_data), lthy) end; fun derive_Sig_natural_etc ctxt live live_AsBs Y Z preT fpT k_or_fpT sig_T ssig_T pre_map dead_pre_map ctor dtor Sig dead_sig_map Oper VLeaf CLeaf ssig_map dead_ssig_map Lam flat eval_core eval cutSsig algLam corecU x fs f g ctor_dtor dtor_inject dtor_unfold_thm dtor_unfold_unique sig_map_thm ssig_induct ssig_map_thms Oper_map_thm VLeaf_map_thm CLeaf_map_thm Lam_transfer flat_simps flat_transfer eval_core_simps eval_core_transfer eval_def cutSsig_def algLam_def corecU_def live_pre_bnf pre_bnf dead_pre_bnf fp_bnf sig_bnf ssig_bnf dead_ssig_bnf = let val SOME prod_bnf = bnf_of ctxt \<^type_name>\prod\; val f' = substT Z fpT f; val dead_ssig_map' = substT Z fpT dead_ssig_map; val extdd = Term.lambda f' (HOLogic.mk_comp (eval, dead_ssig_map' $ f')); val live_pre_map_def = map_def_of_bnf live_pre_bnf; val pre_map_comp = map_comp_of_bnf pre_bnf; val dead_pre_map_id = map_id_of_bnf dead_pre_bnf; val dead_pre_map_comp0 = map_comp0_of_bnf dead_pre_bnf; val dead_pre_map_comp = map_comp_of_bnf dead_pre_bnf; val dead_pre_map_cong = map_cong_of_bnf dead_pre_bnf; val fp_map_id = map_id_of_bnf fp_bnf; val sig_map_ident = map_ident_of_bnf sig_bnf; val sig_map_comp0 = map_comp0_of_bnf sig_bnf; val sig_map_comp = map_comp_of_bnf sig_bnf; val sig_map_cong = map_cong_of_bnf sig_bnf; val ssig_map_id = map_id_of_bnf ssig_bnf; val ssig_map_comp = map_comp_of_bnf ssig_bnf; val dead_ssig_map_comp0 = map_comp0_of_bnf dead_ssig_bnf; val k_preT_map_id0s = map map_id0_of_bnf (map_filter (bnf_of ctxt) (fold add_type_namesT [preT, k_or_fpT] [])); val Sig_natural = derive_natural_by_unfolding ctxt live_AsBs preT pre_map fs f Sig ([sig_map_thm, live_pre_map_def, @{thm BNF_Composition.id_bnf_def}] @ k_preT_map_id0s); val Oper_natural = derive_natural_by_unfolding ctxt live_AsBs preT pre_map fs f Oper [Oper_map_thm]; val VLeaf_natural = derive_natural_by_unfolding ctxt live_AsBs preT pre_map fs f VLeaf [VLeaf_map_thm]; val Lam_natural = derive_natural_from_transfer_with_pre_type ctxt live_AsBs Y Z preT ssig_T pre_map ssig_map fs f Lam Lam_transfer [prod_bnf, pre_bnf, sig_bnf, ssig_bnf] []; val flat_natural = derive_natural_from_transfer ctxt live_AsBs [] fs f flat flat_transfer [ssig_bnf] []; val eval_core_natural = derive_natural_from_transfer_with_pre_type ctxt live_AsBs Y Z preT ssig_T pre_map ssig_map fs f eval_core eval_core_transfer [prod_bnf, pre_bnf, ssig_bnf] []; val Sig_pointful_natural = mk_pointful ctxt Sig_natural RS sym; val Oper_natural_pointful = mk_pointful ctxt Oper_natural; val Oper_pointful_natural = Oper_natural_pointful RS sym; val flat_pointful_natural = mk_pointful ctxt flat_natural RS sym; val Lam_natural_pointful = mk_pointful ctxt Lam_natural; val Lam_pointful_natural = Lam_natural_pointful RS sym; val eval_core_pointful_natural = mk_pointful ctxt eval_core_natural RS sym; val cutSsig_def_pointful_natural = mk_pointful ctxt (HOLogic.mk_obj_eq cutSsig_def) RS sym; val flat_VLeaf = derive_flat_VLeaf ctxt Y Z ssig_T x VLeaf dead_ssig_map flat ssig_induct fp_map_id sig_map_cong sig_map_ident sig_map_comp ssig_map_thms flat_simps; val flat_flat = derive_flat_flat ctxt Y Z ssig_T x dead_ssig_map flat ssig_induct fp_map_id sig_map_cong sig_map_comp ssig_map_thms flat_simps; val eval_core_flat = derive_eval_core_flat ctxt Y Z preT ssig_T dead_pre_map dead_ssig_map flat eval_core x ssig_induct dead_pre_map_id dead_pre_map_comp0 dead_pre_map_comp fp_map_id sig_map_comp sig_map_cong ssig_map_thms ssig_map_comp flat_simps flat_pointful_natural flat_flat Lam_pointful_natural eval_core_simps; val eval_thm = derive_eval_thm ctxt dtor_inject dtor_unfold_thm eval_def; val eval_flat = derive_eval_flat ctxt Y Z fpT ssig_T dead_ssig_map flat eval x dead_pre_map_comp0 dtor_unfold_unique ssig_map_id ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_core_flat eval_thm; val eval_Oper = derive_eval_Oper ctxt live Y Z fpT sig_T ssig_T dead_sig_map Oper eval algLam x sig_map_ident sig_map_comp0 sig_map_comp Oper_natural_pointful VLeaf_natural flat_simps eval_flat algLam_def; val eval_VLeaf = derive_eval_V_or_CLeaf ctxt Y fpT VLeaf eval x dead_pre_map_id dead_pre_map_comp fp_map_id dtor_unfold_unique VLeaf_map_thm eval_core_simps eval_thm; val eval_CLeaf = derive_eval_V_or_CLeaf ctxt Y fpT CLeaf eval x dead_pre_map_id dead_pre_map_comp fp_map_id dtor_unfold_unique CLeaf_map_thm eval_core_simps eval_thm; val extdd_mor = derive_extdd_mor ctxt Y Z preT fpT ssig_T dead_pre_map dtor extdd cutSsig f g dead_pre_map_comp0 dead_pre_map_comp VLeaf_map_thm ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_thm eval_flat eval_VLeaf cutSsig_def; val mor_cutSsig_flat = derive_mor_cutSsig_flat ctxt Y Z preT fpT ssig_T dead_pre_map dead_ssig_map dtor flat eval_core eval cutSsig f g dead_pre_map_comp0 dead_pre_map_comp dead_pre_map_cong dtor_unfold_unique dead_ssig_map_comp0 ssig_map_comp flat_simps flat_pointful_natural eval_core_pointful_natural flat_flat flat_VLeaf eval_core_flat cutSsig_def cutSsig_def_pointful_natural eval_thm; val extdd_o_VLeaf = derive_extdd_o_VLeaf ctxt Y Z preT fpT ssig_T dead_pre_map dtor VLeaf extdd f g dead_pre_map_comp0 dead_pre_map_comp dtor_inject ssig_map_thms eval_core_simps eval_thm eval_VLeaf; val (corecU_ctor, corecU_unique) = derive_corecU_ctor_unique ctxt Y Z preT fpT ssig_T dead_pre_map ctor dtor VLeaf extdd corecU f g dead_pre_map_comp ctor_dtor dtor_unfold_thm dtor_unfold_unique ssig_map_thms dead_ssig_map_comp0 flat_simps flat_VLeaf eval_core_simps extdd_mor extdd_o_VLeaf cutSsig_def mor_cutSsig_flat corecU_def; val dtor_algLam = derive_dtor_algLam ctxt Y Z preT fpT sig_T ssig_T dead_pre_map dtor dead_sig_map Lam eval algLam x pre_map_comp dead_pre_map_id dead_pre_map_comp0 dead_pre_map_comp sig_map_comp Oper_pointful_natural ssig_map_thms dead_ssig_map_comp0 Lam_pointful_natural eval_core_simps eval_thm eval_flat eval_VLeaf algLam_def; in (Sig_pointful_natural, flat_pointful_natural, Lam_natural_pointful, Lam_pointful_natural, flat_VLeaf, eval_core_pointful_natural, eval_thm, eval_flat, [eval_Oper, eval_VLeaf, eval_CLeaf], corecU_ctor, corecU_unique, dtor_algLam) end; fun derive_embL_natural_etc ctxt Inx_const old_ssig_bnf ssig_bnf Y Z preT fpT old_ssig_T ssig_T dead_pre_map Sig dead_old_ssig_map embL old_algLam algLam old_flat flat old_eval_core eval_core old_eval eval x f old_ssig_induct dead_pre_map_comp0 dead_pre_map_comp fp_map_id dtor_inject dtor_unfold_unique Sig_pointful_natural unsig_thm sig_map_thm old_sig_map_comp old_sig_map_cong old_ssig_map_thms old_Lam_pointful_natural Lam_def old_flat_simps flat_simps embL_simps embL_transfer old_eval_core_simps eval_core_simps old_eval_thm eval_thm old_dtor_algLam dtor_algLam old_algLam_thm = let val embL_natural = derive_natural_from_transfer ctxt [(Y, Z)] [] [] f embL embL_transfer [old_ssig_bnf, ssig_bnf] []; val embL_pointful_natural = mk_pointful ctxt embL_natural RS sym; val old_algLam_pointful = mk_pointful ctxt old_algLam_thm; val flat_embL = derive_flat_embL ctxt Y Z old_ssig_T ssig_T dead_old_ssig_map embL old_flat flat x old_ssig_induct fp_map_id Sig_pointful_natural old_sig_map_comp old_sig_map_cong old_ssig_map_thms old_flat_simps flat_simps embL_simps; val eval_core_embL = derive_eval_core_embL ctxt Y Z preT old_ssig_T ssig_T dead_pre_map embL old_eval_core eval_core x old_ssig_induct dead_pre_map_comp0 dead_pre_map_comp Sig_pointful_natural unsig_thm old_sig_map_comp old_sig_map_cong old_Lam_pointful_natural Lam_def flat_embL embL_simps embL_pointful_natural old_eval_core_simps eval_core_simps; val eval_embL = derive_eval_embL ctxt Y fpT embL old_eval eval dead_pre_map_comp0 dtor_unfold_unique embL_pointful_natural eval_core_embL old_eval_thm eval_thm; val algLam_algLam = derive_algLam_algLam ctxt Inx_const Y fpT Sig old_algLam algLam dead_pre_map_comp dtor_inject unsig_thm sig_map_thm Lam_def eval_embL old_dtor_algLam dtor_algLam; in (embL_pointful_natural, old_algLam_pointful, eval_embL, algLam_algLam) end; fun define_corecUU_etc fp_b version live_AsBs Y Z preT fpT ssig_T pre_map dead_pre_map pre_rel fp_rel ctor Oper ssig_map dead_ssig_map ssig_rel proto_sctr flat eval_core eval corecU fs f g Rs R pre_map_transfer fp_k_T_rel_eqs dtor_unfold_transfer dtor_transfer ssig_map_transfer proto_sctr_transfer proto_sctr_pointful_natural flat_transfer flat_pointful_natural eval_core_transfer eval_core_pointful_natural eval_thm eval_flat eval_Oper algLam_thm cutSsig_def corecU_def corecU_ctor corecU_unique pre_bnf dead_pre_bnf fp_res ssig_fp_sugar lthy = let val ssig_bnf = #fp_bnf ssig_fp_sugar; val dead_pre_map_comp0 = map_comp0_of_bnf dead_pre_bnf; val dead_pre_map_comp = map_comp_of_bnf dead_pre_bnf; val [dtor_ctor] = #dtor_ctors fp_res; val [dtor_inject] = #dtor_injects fp_res; val ssig_map_comp = map_comp_of_bnf ssig_bnf; val sctr_rhs = HOLogic.mk_comp (Oper, substT Y ssig_T proto_sctr); val ((sctr, sctr_def), lthy) = lthy |> define_const true fp_b version sctrN sctr_rhs; val (corecUU_data as (corecUU, corecUU_def), lthy) = lthy |> define_corecUU fp_b version Y Z preT ssig_T dead_pre_map dead_ssig_map flat eval_core sctr corecU; val eval_sctr = derive_eval_sctr lthy Y Z fpT ssig_T dead_pre_map ctor eval sctr proto_sctr_pointful_natural eval_Oper algLam_thm sctr_def; val sctr_transfer = derive_sctr_transfer lthy live_AsBs Y Z ssig_T Rs R pre_rel ssig_rel sctr sctr_def fp_k_T_rel_eqs [proto_sctr_transfer]; val sctr_natural = derive_natural_from_transfer_with_pre_type lthy live_AsBs Y Z preT ssig_T pre_map ssig_map fs f sctr sctr_transfer [pre_bnf, ssig_bnf] []; val sctr_pointful_natural = mk_pointful lthy sctr_natural RS sym; val eval_sctr_pointful = mk_pointful lthy eval_sctr RS sym; val (corecUU_pointfree, corecUU_unique) = derive_corecUU_pointfree_unique lthy Y Z preT fpT ssig_T dead_pre_map ctor dead_ssig_map eval corecUU f g dead_pre_map_comp0 dead_pre_map_comp dtor_ctor dtor_inject ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_thm eval_flat corecU_ctor corecU_unique sctr_pointful_natural eval_sctr_pointful corecUU_def; val corecUU_thm = mk_pointful lthy corecUU_pointfree; val corecUU_transfer = derive_corecUU_transfer lthy live_AsBs Y Z Rs R preT ssig_T pre_rel fp_rel ssig_rel corecUU cutSsig_def corecU_def corecUU_def fp_k_T_rel_eqs [pre_map_transfer, dtor_unfold_transfer, dtor_transfer, ssig_map_transfer, flat_transfer, eval_core_transfer, sctr_transfer, @{thm convol_transfer} (*FIXME: needed?*)]; in ((corecUU_data, corecUU_thm, corecUU_unique, corecUU_transfer, eval_sctr, sctr_transfer, sctr_pointful_natural), lthy) end; fun mk_equivp T = Const (\<^const_name>\equivp\, mk_predT [mk_pred2T T T]); fun derive_equivp_Retr ctxt fpT Retr R dead_pre_rel_refl_thm dead_pre_rel_flip_thm dead_pre_rel_mono_thm dead_pre_rel_compp_thm = let val prem = HOLogic.mk_Trueprop (mk_equivp fpT $ R); val goal = Logic.mk_implies (prem, HOLogic.mk_Trueprop (mk_equivp fpT $ (betapply (Retr, R)))); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_equivp_Retr_tac ctxt dead_pre_rel_refl_thm dead_pre_rel_flip_thm dead_pre_rel_mono_thm dead_pre_rel_compp_thm)) |> Thm.close_derivation \<^here> end; fun derive_Retr_coinduct ctxt fpT Retr R dtor_rel_coinduct_thm rel_eq_thm = let val goal = HOLogic.mk_Trueprop (list_all_free [R] (HOLogic.mk_imp (mk_leq R (Retr $ R), mk_leq R (HOLogic.eq_const fpT)))); in Goal.prove_sorry ctxt [] [] goal (fn {context = ctxt, prems = _} => mk_Retr_coinduct_tac ctxt dtor_rel_coinduct_thm rel_eq_thm) |> Thm.close_derivation \<^here> end; fun derive_Retr fp_sugar fpT dead_pre_bnf ctor dtor names_lthy lthy = let val (R, _) = names_lthy |> yield_singleton (mk_Frees "R") (mk_pred2T fpT fpT); val pre_fpT = pre_type_of_ctor fpT ctor; val fp_pre_rel = mk_rel1 lthy fpT fpT pre_fpT dead_pre_bnf; val Retr = Abs ("R", fastype_of R, Abs ("a", fpT, Abs ("b", fpT, list_comb (fp_pre_rel, [Bound 2, dtor $ Bound 1, dtor $ Bound 0])))); val equivp_Retr = derive_equivp_Retr lthy fpT Retr R (rel_refl_of_bnf dead_pre_bnf) (rel_flip_of_bnf dead_pre_bnf) (rel_mono_of_bnf dead_pre_bnf) (rel_OO_of_bnf dead_pre_bnf); val Retr_coinduct = derive_Retr_coinduct lthy fpT Retr R (fp_sugar |> #fp_res |> #xtor_rel_co_induct) (fp_sugar |> #fp_bnf |> rel_eq_of_bnf); in (Retr, equivp_Retr, Retr_coinduct) end; fun mk_gen_cong fpT eval_domT = let val fp_relT = mk_pred2T fpT fpT in Const (\<^const_name>\cong.gen_cong\, [mk_predT [fp_relT, eval_domT, eval_domT], eval_domT --> fpT, fp_relT] ---> fp_relT) end; fun mk_cong_locale rel eval Retr = Const (\<^const_name>\cong\, mk_predT (map fastype_of [rel, eval, Retr])); fun derive_cong_locale ctxt rel eval Retr0 tac = let val Retr = enforce_type ctxt domain_type (domain_type (fastype_of rel)) Retr0; val goal = HOLogic.mk_Trueprop (list_comb (mk_cong_locale rel eval Retr, [rel, eval, Retr])); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => tac ctxt)) |> Thm.close_derivation \<^here> end; fun derive_cong_general fp_b version fpT dead_ssig_bnf dead_pre_bnf eval Retr equivp_Retr Retr_coinduct eval_thm eval_core_transfer lthy = let val eval_domT = domain_type (fastype_of eval); fun cong_locale_tac ctxt = mk_cong_locale_tac ctxt (rel_mono_of_bnf dead_pre_bnf) (rel_map_of_bnf dead_pre_bnf) equivp_Retr (rel_mono_of_bnf dead_ssig_bnf) (rel_map_of_bnf dead_ssig_bnf) eval_thm eval_core_transfer; val rel = mk_rel1 lthy fpT fpT eval_domT dead_ssig_bnf; val cong_rhs = list_comb (mk_gen_cong fpT eval_domT, [rel, eval]); val ((_, cong_def), lthy) = lthy |> define_const false fp_b version congN cong_rhs; val cong_locale = derive_cong_locale lthy rel eval Retr cong_locale_tac; val fold_cong_def = Local_Defs.fold lthy [cong_def]; fun instance_of_gen thm = fold_cong_def (thm OF [cong_locale]); val cong_base = instance_of_gen @{thm cong.imp_gen_cong}; val cong_refl = instance_of_gen @{thm cong.gen_cong_reflp}; val cong_sym = instance_of_gen @{thm cong.gen_cong_symp}; val cong_trans = instance_of_gen @{thm cong.gen_cong_transp}; fun mk_cong_rho thm = thm RS instance_of_gen @{thm cong.gen_cong_rho}; val dtor_coinduct = @{thm predicate2I_obj} RS (Retr_coinduct RS instance_of_gen @{thm cong.coinduction} RS @{thm predicate2D_obj}); in (cong_def, cong_locale, cong_base, cong_refl, cong_sym, cong_trans, dtor_coinduct, mk_cong_rho, lthy) end; fun derive_cong_base fp_b version fpT dead_ssig_bnf ssig_fp_bnf_sugar dead_pre_bnf eval eval_thm eval_core_transfer eval_VLeaf eval_sctr sctr_transfer Retr equivp_Retr Retr_coinduct lthy = let val (cong_def, cong_locale, cong_base, cong_refl, cong_sym, cong_trans, dtor_coinduct, mk_cong_rho, lthy) = derive_cong_general fp_b version fpT dead_ssig_bnf dead_pre_bnf eval Retr equivp_Retr Retr_coinduct eval_thm eval_core_transfer lthy; val dead_pre_map_id0 = map_id0_of_bnf dead_pre_bnf; val dead_pre_map_comp0 = map_comp0_of_bnf dead_pre_bnf; val dead_pre_map_cong0 = map_cong0_of_bnf dead_pre_bnf; val dead_pre_map_cong0' = @{thm box_equals[OF _ o_apply[symmetric] id_apply[symmetric]]} RS dead_pre_map_cong0 RS ext; val dead_pre_rel_map = rel_map_of_bnf dead_pre_bnf; val ctor_alt_thm = eval_VLeaf RS (@{thm eq_comp_compI} OF [eval_sctr, trans OF [dead_pre_map_comp0 RS sym, trans OF [dead_pre_map_cong0', dead_pre_map_id0]]]); val cong_ctor_intro = mk_cong_rho ctor_alt_thm |> unfold_thms lthy [o_apply] |> (fn thm => sctr_transfer RS rel_funD RS thm) |> unfold_thms lthy (id_apply :: dead_pre_rel_map @ #rel_injects ssig_fp_bnf_sugar); in ({cong_def = cong_def, cong_locale = cong_locale, cong_base = cong_base, cong_refl = cong_refl, cong_sym = cong_sym, cong_trans = cong_trans, dtor_coinduct = dtor_coinduct, cong_alg_intros = [cong_ctor_intro]}, lthy) end; fun update_cong_alg_intros ctxt cong_def cong_locale old_cong_def old_cong_locale emb = let fun instance_of_gen thm = Local_Defs.fold ctxt [cong_def] (thm OF [cong_locale]); fun instance_of_old_gen thm = Local_Defs.fold ctxt [old_cong_def] (thm OF [old_cong_locale]); val emb_idem = @{thm ord_le_eq_trans} OF [emb, instance_of_gen @{thm cong.gen_cong_idem}]; fun mk_rel_mono bnf = instance_of_old_gen @{thm cong.leq_gen_cong} RS rel_mono_of_bnf bnf RS @{thm predicate2D}; fun mk_intro bnf thm = mk_rel_mono bnf RS (@{thm predicate2D} OF [emb_idem, thm]); in map2 mk_intro end fun derive_cong_step fp_b version fpT dead_ssig_bnf dead_pre_bnf eval eval_thm eval_core_transfer old_dtor_coinduct_info algrho_def k_as_ssig_transfer Retr equivp_Retr Retr_coinduct eval_embL embL_transfer old_all_dead_k_bnfs lthy = let val old_cong_def = #cong_def old_dtor_coinduct_info; val old_cong_locale = #cong_locale old_dtor_coinduct_info; val old_cong_alg_intros = #cong_alg_intros old_dtor_coinduct_info; val (cong_def, cong_locale, cong_base, cong_refl, cong_sym, cong_trans, dtor_coinduct, mk_cong_rho, lthy) = derive_cong_general fp_b version fpT dead_ssig_bnf dead_pre_bnf eval Retr equivp_Retr Retr_coinduct eval_thm eval_core_transfer lthy; val cong_alg_intro = k_as_ssig_transfer RS rel_funD RS mk_cong_rho (HOLogic.mk_obj_eq algrho_def); val gen_cong_emb = (@{thm gen_cong_emb} OF [old_cong_locale, cong_locale, eval_embL, embL_transfer]) |> Local_Defs.fold lthy [old_cong_def, cong_def]; val cong_alg_intros = update_cong_alg_intros lthy cong_def cong_locale old_cong_def old_cong_locale gen_cong_emb old_all_dead_k_bnfs old_cong_alg_intros; in ({cong_def = cong_def, cong_locale = cong_locale, cong_base = cong_base, cong_refl = cong_refl, cong_sym = cong_sym, cong_trans = cong_trans, dtor_coinduct = dtor_coinduct, cong_alg_intros = cong_alg_intro :: cong_alg_intros}, lthy) end; fun derive_cong_merge fp_b version fpT old1_friend_names old2_friend_names dead_ssig_bnf dead_pre_bnf eval eval_thm eval_core_transfer old1_dtor_coinduct_info old2_dtor_coinduct_info Retr equivp_Retr Retr_coinduct eval_embLL embLL_transfer eval_embLR embLR_transfer old1_all_dead_k_bnfs old2_all_dead_k_bnfs lthy = let val old1_cong_def = #cong_def old1_dtor_coinduct_info; val old1_cong_locale = #cong_locale old1_dtor_coinduct_info; val old1_cong_alg_intros = #cong_alg_intros old1_dtor_coinduct_info; val old2_cong_def = #cong_def old2_dtor_coinduct_info; val old2_cong_locale = #cong_locale old2_dtor_coinduct_info; val old2_cong_alg_intros = #cong_alg_intros old2_dtor_coinduct_info; val (cong_def, cong_locale, cong_base, cong_refl, cong_sym, cong_trans, dtor_coinduct, _, lthy) = derive_cong_general fp_b version fpT dead_ssig_bnf dead_pre_bnf eval Retr equivp_Retr Retr_coinduct eval_thm eval_core_transfer lthy; val emb1 = (@{thm gen_cong_emb} OF [old1_cong_locale, cong_locale, eval_embLL, embLL_transfer]) |> Local_Defs.fold lthy [old1_cong_def, cong_def]; val emb2 = (@{thm gen_cong_emb} OF [old2_cong_locale, cong_locale, eval_embLR, embLR_transfer]) |> Local_Defs.fold lthy [old2_cong_def, cong_def]; val cong_alg_intros1 = update_cong_alg_intros lthy cong_def cong_locale old1_cong_def old1_cong_locale emb1 old1_all_dead_k_bnfs old1_cong_alg_intros; val cong_alg_intros2 = update_cong_alg_intros lthy cong_def cong_locale old2_cong_def old2_cong_locale emb2 old2_all_dead_k_bnfs old2_cong_alg_intros; val (cong_algrho_intros1, cong_ctor_intro1) = split_last cong_alg_intros1; val (cong_algrho_intros2, _) = split_last cong_alg_intros2; val (old1_all_rho_k_bnfs, old1_Sig_bnf) = split_last old1_all_dead_k_bnfs; val (old2_all_rho_k_bnfs, _) = split_last old2_all_dead_k_bnfs; val (friend_names, (cong_algrho_intros, all_rho_k_bnfs)) = merge_lists (op = o apply2 fst) (old1_friend_names ~~ (cong_algrho_intros1 ~~ old1_all_rho_k_bnfs)) (old2_friend_names ~~ (cong_algrho_intros2 ~~ old2_all_rho_k_bnfs)) |> split_list ||> split_list; in (({cong_def = cong_def, cong_locale = cong_locale, cong_base = cong_base, cong_refl = cong_refl, cong_sym = cong_sym, cong_trans = cong_trans, dtor_coinduct = dtor_coinduct, cong_alg_intros = cong_algrho_intros @ [cong_ctor_intro1]}, all_rho_k_bnfs @ [old1_Sig_bnf], friend_names), lthy) end; fun derive_corecUU_base fpT_name lthy = let val fp_sugar as {T = Type (_, fpT_args0), pre_bnf, fp_bnf, fp_res, ...} = checked_fp_sugar_of lthy fpT_name; val fpT_Ss = map Type.sort_of_atyp fpT_args0; val fp_alives = liveness_of_fp_bnf (length fpT_args0) fp_bnf; val (((Es, Fs0), [Y, Z]), names_lthy) = lthy |> mk_TFrees' fpT_Ss ||>> mk_TFrees' fpT_Ss ||>> mk_TFrees 2; val Fs = @{map 3} (fn alive => fn E => fn F => if alive then F else E) fp_alives Es Fs0; val As = Es @ [Y]; val Bs = Es @ [Z]; val live_EsFs = filter (op <>) (Es ~~ Fs); val live_AsBs = live_EsFs @ [(Y, Z)]; val fTs = map (op -->) live_EsFs; val RTs = map (uncurry mk_pred2T) live_EsFs; val live = length live_EsFs; val ((((((x, fs), f), g), Rs), R), names_lthy) = names_lthy |> yield_singleton (mk_Frees "x") Y ||>> mk_Frees "f" fTs ||>> yield_singleton (mk_Frees "f") (Y --> Z) ||>> yield_singleton (mk_Frees "g") (Y --> Z) ||>> mk_Frees "R" RTs ||>> yield_singleton (mk_Frees "R") (mk_pred2T Y Z); val ctor = mk_ctor Es (the_single (#ctors fp_res)); val dtor = mk_dtor Es (the_single (#dtors fp_res)); val fpT = Type (fpT_name, Es); val preT = pre_type_of_ctor Y ctor; val ((fp_b, version), lthy) = lthy |> get_name_next_version_of fpT_name; val ((sig_fp_sugar, ssig_fp_sugar), lthy) = lthy |> define_sig_type fp_b version fp_alives Es Y preT ||>> define_ssig_type fp_b version fp_alives Es Y fpT; val sig_bnf = #fp_bnf sig_fp_sugar; val ssig_bnf = #fp_bnf ssig_fp_sugar; val (((dead_pre_bnf, dead_sig_bnf), dead_ssig_bnf), lthy) = lthy |> bnf_kill_all_but 1 pre_bnf ||>> bnf_kill_all_but 1 sig_bnf ||>> bnf_kill_all_but 1 ssig_bnf; val sig_fp_ctr_sugar = #fp_ctr_sugar sig_fp_sugar; val ssig_fp_ctr_sugar = #fp_ctr_sugar ssig_fp_sugar; val sig_fp_bnf_sugar = #fp_bnf_sugar sig_fp_sugar; val ssig_fp_bnf_sugar = #fp_bnf_sugar ssig_fp_sugar; val ssig_fp_induct_sugar = the (#fp_co_induct_sugar ssig_fp_sugar); val sig_ctr_sugar = #ctr_sugar sig_fp_ctr_sugar; val ssig_ctr_sugar = #ctr_sugar ssig_fp_ctr_sugar; val sig_T_name = fst (dest_Type (#T sig_fp_sugar)); val ssig_T_name = fst (dest_Type (#T ssig_fp_sugar)); val sig_T = Type (sig_T_name, As); val ssig_T = Type (ssig_T_name, As); val pre_map = mk_mapN lthy live_AsBs preT pre_bnf; val pre_rel = mk_relN lthy live_AsBs preT pre_bnf; val dead_pre_map = mk_map1 lthy Y Z preT dead_pre_bnf; val fp_rel = mk_relN lthy live_EsFs fpT fp_bnf; val dtor_unfold = mk_co_rec (Proof_Context.theory_of lthy) Greatest_FP [Z] fpT (the_single (#xtor_un_folds fp_res)); val Sig = mk_ctr As (the_single (#ctrs sig_ctr_sugar)); val unsig = mk_disc_or_sel As (the_single (the_single (#selss sig_ctr_sugar))); val sig_rel = mk_relN lthy live_AsBs sig_T sig_bnf; val dead_sig_map = mk_map 1 As Bs (map_of_bnf dead_sig_bnf); val [Oper, VLeaf, CLeaf] = map (mk_ctr As) (#ctrs ssig_ctr_sugar); val ssig_map = mk_mapN lthy live_AsBs ssig_T ssig_bnf; val ssig_rel = mk_relN lthy live_AsBs ssig_T ssig_bnf; val dead_ssig_map = mk_map 1 As Bs (map_of_bnf dead_ssig_bnf); val ((Lam, Lam_def), lthy) = lthy |> define_Lam_base fp_b version Y Z preT ssig_T dead_pre_map Sig unsig dead_sig_map Oper VLeaf; val proto_sctr = Sig; val pre_map_transfer = map_transfer_of_bnf pre_bnf; val pre_rel_def = rel_def_of_bnf pre_bnf; val dead_pre_map_id = map_id_of_bnf dead_pre_bnf; val dead_pre_map_comp = map_comp_of_bnf dead_pre_bnf; val fp_rel_eq = rel_eq_of_bnf fp_bnf; val [ctor_dtor] = #ctor_dtors fp_res; val [dtor_ctor] = #dtor_ctors fp_res; val [dtor_inject] = #dtor_injects fp_res; val [dtor_unfold_thm] = #xtor_un_fold_thms fp_res; val dtor_unfold_unique = #xtor_un_fold_unique fp_res; val [dtor_unfold_transfer] = #xtor_un_fold_transfers fp_res; val [dtor_rel_thm] = #xtor_rels fp_res; val unsig_thm = the_single (the_single (#sel_thmss sig_ctr_sugar)); val [sig_map_thm] = #map_thms sig_fp_bnf_sugar; val [Oper_map_thm, VLeaf_map_thm, CLeaf_map_thm] = #map_thms ssig_fp_bnf_sugar; val sig_map_transfer = map_transfer_of_bnf sig_bnf; val ssig_map_thms = #map_thms ssig_fp_bnf_sugar; val ssig_map_transfer = map_transfer_of_bnf ssig_bnf; val ssig_induct = the_single (#co_inducts ssig_fp_induct_sugar); val dtor_transfer = derive_dtor_transfer lthy live_EsFs Y Z pre_rel fp_rel Rs dtor dtor_rel_thm; val preT_rel_eqs = map rel_eq_of_bnf (map_filter (bnf_of lthy) (add_type_namesT preT [])); val Sig_transfer = derive_sig_transfer I lthy live_AsBs pre_rel sig_rel Rs R Sig pre_rel_def preT_rel_eqs (the_single (#ctr_transfers sig_fp_ctr_sugar)); val proto_sctr_transfer = Sig_transfer; val unsig_transfer = derive_sig_transfer swap lthy live_AsBs pre_rel sig_rel Rs R unsig pre_rel_def preT_rel_eqs (the_single (#sel_transfers sig_fp_ctr_sugar)); val Lam_transfer = derive_Lam_or_eval_core_transfer lthy live_AsBs Y Z preT ssig_T Rs R pre_rel sig_rel ssig_rel Lam Lam_def [] [pre_map_transfer, sig_map_transfer, Sig_transfer, unsig_transfer]; val ((((((((flat, _, flat_simps), flat_transfer), ((eval_core, _, eval_core_simps), eval_core_transfer)), (eval, eval_def)), (cutSsig, cutSsig_def)), (algLam, algLam_def)), (corecU, corecU_def)), lthy) = lthy |> define_flat_etc fp_b version live_AsBs Y Z preT fpT sig_T ssig_T Oper VLeaf CLeaf pre_rel dead_pre_map dtor dtor_unfold dead_sig_map ssig_rel dead_ssig_map Lam Rs R pre_map_transfer [fp_rel_eq] sig_map_transfer ssig_map_transfer Lam_transfer dtor_transfer; val (Sig_pointful_natural, flat_pointful_natural, _, Lam_pointful_natural, _, eval_core_pointful_natural, eval_thm, eval_flat, eval_simps as [eval_Oper, eval_VLeaf, _], corecU_ctor, corecU_unique, dtor_algLam) = derive_Sig_natural_etc lthy live live_AsBs Y Z preT fpT fpT sig_T ssig_T pre_map dead_pre_map ctor dtor Sig dead_sig_map Oper VLeaf CLeaf ssig_map dead_ssig_map Lam flat eval_core eval cutSsig algLam corecU x fs f g ctor_dtor dtor_inject dtor_unfold_thm dtor_unfold_unique sig_map_thm ssig_induct ssig_map_thms Oper_map_thm VLeaf_map_thm CLeaf_map_thm Lam_transfer flat_simps flat_transfer eval_core_simps eval_core_transfer eval_def cutSsig_def algLam_def corecU_def pre_bnf pre_bnf dead_pre_bnf fp_bnf sig_bnf ssig_bnf dead_ssig_bnf; val proto_sctr_pointful_natural = Sig_pointful_natural; val algLam_thm = derive_algLam_base lthy Y Z preT fpT dead_pre_map ctor dtor algLam proto_sctr dead_pre_map_id dead_pre_map_comp ctor_dtor dtor_ctor dtor_unfold_unique unsig_thm Sig_pointful_natural ssig_map_thms Lam_def flat_simps eval_core_simps eval_thm algLam_def; val (((corecUU, _), corecUU_thm, corecUU_unique, corecUU_transfer, eval_sctr, sctr_transfer, sctr_pointful_natural), lthy) = lthy |> define_corecUU_etc fp_b version live_AsBs Y Z preT fpT ssig_T pre_map dead_pre_map pre_rel fp_rel ctor Oper ssig_map dead_ssig_map ssig_rel proto_sctr flat eval_core eval corecU fs f g Rs R pre_map_transfer [] dtor_unfold_transfer dtor_transfer ssig_map_transfer proto_sctr_transfer proto_sctr_pointful_natural flat_transfer flat_pointful_natural eval_core_transfer eval_core_pointful_natural eval_thm eval_flat eval_Oper algLam_thm cutSsig_def corecU_def corecU_ctor corecU_unique pre_bnf dead_pre_bnf fp_res ssig_fp_sugar; val (Retr, equivp_Retr, Retr_coinduct) = lthy |> derive_Retr fp_sugar fpT dead_pre_bnf ctor dtor names_lthy; val (dtor_coinduct_info, lthy) = lthy |> derive_cong_base fp_b version fpT dead_ssig_bnf ssig_fp_bnf_sugar dead_pre_bnf eval eval_thm eval_core_transfer eval_VLeaf eval_sctr sctr_transfer Retr equivp_Retr Retr_coinduct; val buffer = {Oper = Oper, VLeaf = VLeaf, CLeaf = CLeaf, ctr_wrapper = Sig, friends = Symtab.empty}; val notes = [(corecUU_transferN, [corecUU_transfer])] @ (if Config.get lthy bnf_internals then [(algLamN, [algLam_thm]), (cong_alg_introsN, #cong_alg_intros dtor_coinduct_info), (cong_localeN, [#cong_locale dtor_coinduct_info]), (corecU_ctorN, [corecU_ctor]), (corecU_uniqueN, [corecU_unique]), (corecUUN, [corecUU_thm]), (corecUU_uniqueN, [corecUU_unique]), (dtor_algLamN, [dtor_algLam]), (dtor_coinductN, [#dtor_coinduct dtor_coinduct_info]), (dtor_transferN, [dtor_transfer]), (equivp_RetrN, [equivp_Retr]), (evalN, [eval_thm]), (eval_core_pointful_naturalN, [eval_core_pointful_natural]), (eval_core_transferN, [eval_core_transfer]), (eval_flatN, [eval_flat]), (eval_simpsN, eval_simps), (flat_pointful_naturalN, [flat_pointful_natural]), (flat_transferN, [flat_transfer]), (Lam_pointful_naturalN, [Lam_pointful_natural]), (Lam_transferN, [Lam_transfer]), (proto_sctr_pointful_naturalN, [proto_sctr_pointful_natural]), (proto_sctr_transferN, [proto_sctr_transfer]), (Retr_coinductN, [Retr_coinduct]), (sctr_pointful_naturalN, [sctr_pointful_natural]), (sctr_transferN, [sctr_transfer]), (Sig_pointful_naturalN, [Sig_pointful_natural])] else []) |> map (fn (thmN, thms) => ((mk_version_fp_binding true lthy version fp_b thmN, []), [(thms, [])])); in ({fp_b = fp_b, version = version, fpT = fpT, Y = Y, Z = Z, friend_names = [], sig_fp_sugars = [sig_fp_sugar], ssig_fp_sugar = ssig_fp_sugar, Lam = Lam, proto_sctr = proto_sctr, flat = flat, eval_core = eval_core, eval = eval, algLam = algLam, corecUU = corecUU, dtor_transfer = dtor_transfer, Lam_transfer = Lam_transfer, Lam_pointful_natural = Lam_pointful_natural, proto_sctr_transfer = proto_sctr_transfer, flat_simps = flat_simps, eval_core_simps = eval_core_simps, eval_thm = eval_thm, eval_simps = eval_simps, all_algLam_algs = [algLam_thm], algLam_thm = algLam_thm, dtor_algLam = dtor_algLam, corecUU_thm = corecUU_thm, corecUU_unique = corecUU_unique, corecUU_transfer = corecUU_transfer, buffer = buffer, all_dead_k_bnfs = [dead_pre_bnf], Retr = Retr, equivp_Retr = equivp_Retr, Retr_coinduct = Retr_coinduct, dtor_coinduct_info = dtor_coinduct_info} |> morph_corec_info (Local_Theory.target_morphism lthy), lthy |> Local_Theory.notes notes |> snd) end; fun derive_corecUU_step (fpT as Type (fpT_name, res_Ds)) ({friend_names = old_friend_names, sig_fp_sugars = old_sig_fp_sugars as old_sig_fp_sugar :: _, ssig_fp_sugar = old_ssig_fp_sugar, Lam = old_Lam0, proto_sctr = old_proto_sctr0, flat = old_flat0, eval_core = old_eval_core0, eval = old_eval0, algLam = old_algLam0, dtor_transfer, Lam_transfer = old_Lam_transfer, Lam_pointful_natural = old_Lam_pointful_natural, proto_sctr_transfer = old_proto_sctr_transfer, flat_simps = old_flat_simps, eval_core_simps = old_eval_core_simps, eval_thm = old_eval_thm, all_algLam_algs = old_all_algLam_algs, algLam_thm = old_algLam_thm, dtor_algLam = old_dtor_algLam, buffer = old_buffer, all_dead_k_bnfs = old_all_dead_k_bnfs, Retr = old_Retr0, equivp_Retr, Retr_coinduct, dtor_coinduct_info = old_dtor_coinduct_info, ...} : corec_info) friend_name friend_T fp_b version Y Z k_T dead_k_bnf sig_fp_sugar ssig_fp_sugar rho rho_transfer lthy = let val {pre_bnf = live_pre_bnf, fp_bnf = live_fp_bnf, fp_res, ...} = checked_fp_sugar_of lthy fpT_name; val names_lthy = lthy |> fold Variable.declare_typ [Y, Z]; (* FIXME *) val live_EsFs = []; val live_AsBs = live_EsFs @ [(Y, Z)]; val live = length live_EsFs; val ((((x, f), g), R), _) = names_lthy |> yield_singleton (mk_Frees "x") Y ||>> yield_singleton (mk_Frees "f") (Y --> Z) ||>> yield_singleton (mk_Frees "g") (Y --> Z) ||>> yield_singleton (mk_Frees "R") (mk_pred2T Y Z); (* FIXME *) val fs = []; val Rs = []; val ctor = mk_ctor res_Ds (the_single (#ctors fp_res)); val dtor = mk_dtor res_Ds (the_single (#dtors fp_res)); val friend_names = friend_name :: old_friend_names; val old_sig_bnf = #fp_bnf old_sig_fp_sugar; val old_ssig_bnf = #fp_bnf old_ssig_fp_sugar; val sig_bnf = #fp_bnf sig_fp_sugar; val ssig_bnf = #fp_bnf ssig_fp_sugar; val ((((((dead_pre_bnf, dead_fp_bnf), dead_old_sig_bnf), dead_old_ssig_bnf), dead_sig_bnf), dead_ssig_bnf), lthy) = lthy |> bnf_kill_all_but 1 live_pre_bnf ||>> bnf_kill_all_but 0 live_fp_bnf ||>> bnf_kill_all_but 1 old_sig_bnf ||>> bnf_kill_all_but 1 old_ssig_bnf ||>> bnf_kill_all_but 1 sig_bnf ||>> bnf_kill_all_but 1 ssig_bnf; (* FIXME *) val pre_bnf = dead_pre_bnf; val fp_bnf = dead_fp_bnf; val old_ssig_fp_ctr_sugar = #fp_ctr_sugar old_ssig_fp_sugar; val sig_fp_ctr_sugar = #fp_ctr_sugar sig_fp_sugar; val ssig_fp_ctr_sugar = #fp_ctr_sugar ssig_fp_sugar; val sig_fp_bnf_sugar = #fp_bnf_sugar sig_fp_sugar; val old_ssig_fp_bnf_sugar = #fp_bnf_sugar old_ssig_fp_sugar; val ssig_fp_bnf_sugar = #fp_bnf_sugar ssig_fp_sugar; val old_ssig_fp_induct_sugar = the (#fp_co_induct_sugar old_ssig_fp_sugar); val ssig_fp_induct_sugar = the (#fp_co_induct_sugar ssig_fp_sugar); val old_ssig_ctr_sugar = #ctr_sugar old_ssig_fp_ctr_sugar; val sig_ctr_sugar = #ctr_sugar sig_fp_ctr_sugar; val ssig_ctr_sugar = #ctr_sugar ssig_fp_ctr_sugar; val old_sig_T_name = fst (dest_Type (#T old_sig_fp_sugar)); val old_ssig_T_name = fst (dest_Type (#T old_ssig_fp_sugar)); val sig_T_name = fst (dest_Type (#T sig_fp_sugar)); val ssig_T_name = fst (dest_Type (#T ssig_fp_sugar)); val res_As = res_Ds @ [Y]; val res_Bs = res_Ds @ [Z]; val preT = pre_type_of_ctor Y ctor; val YpreT = HOLogic.mk_prodT (Y, preT); val old_sig_T = Type (old_sig_T_name, res_As); val old_ssig_T = Type (old_ssig_T_name, res_As); val sig_T = Type (sig_T_name, res_As); val ssig_T = Type (ssig_T_name, res_As); val old_Lam_domT = Tsubst Y YpreT old_sig_T; val old_eval_core_domT = Tsubst Y YpreT old_ssig_T; val pre_map = mk_mapN lthy live_AsBs preT pre_bnf; val pre_rel = mk_relN lthy live_AsBs preT pre_bnf; val dead_pre_map = mk_map1 lthy Y Z preT dead_pre_bnf; val dead_pre_rel = mk_rel1 lthy Y Z preT dead_pre_bnf; val fp_rel = mk_relN lthy live_EsFs fpT fp_bnf; val dtor_unfold = mk_co_rec (Proof_Context.theory_of lthy) Greatest_FP [Z] fpT (the_single (#xtor_un_folds fp_res)); val dead_k_map = mk_map1 lthy Y Z k_T dead_k_bnf; val Sig = mk_ctr res_As (the_single (#ctrs sig_ctr_sugar)); val unsig = mk_disc_or_sel res_As (the_single (the_single (#selss sig_ctr_sugar))); val sig_rel = mk_relN lthy live_AsBs sig_T sig_bnf; val dead_old_sig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_old_sig_bnf); val dead_sig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_sig_bnf); val dead_sig_rel = mk_rel 1 res_As res_Bs (rel_of_bnf dead_sig_bnf); val [old_Oper, old_VLeaf, old_CLeaf] = map (mk_ctr res_As) (#ctrs old_ssig_ctr_sugar); val [Oper, VLeaf, CLeaf] = map (mk_ctr res_As) (#ctrs ssig_ctr_sugar); val dead_old_ssig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_old_ssig_bnf); val ssig_map = mk_mapN lthy live_AsBs ssig_T ssig_bnf; val ssig_rel = mk_relN lthy live_AsBs ssig_T ssig_bnf; val dead_ssig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_ssig_bnf); val old_Lam = enforce_type lthy domain_type old_Lam_domT old_Lam0; val old_proto_sctr = enforce_type lthy domain_type preT old_proto_sctr0; val old_flat = enforce_type lthy range_type old_ssig_T old_flat0; val old_eval_core = enforce_type lthy domain_type old_eval_core_domT old_eval_core0; val old_eval = enforce_type lthy range_type fpT old_eval0; val old_algLam = enforce_type lthy range_type fpT old_algLam0; val ((embL, embL_def, embL_simps), lthy) = lthy |> define_embL embLN fp_b version Y Z fpT old_sig_T old_ssig_T k_T ssig_T Inl_const dead_old_sig_map Sig old_Oper old_VLeaf old_CLeaf Oper VLeaf CLeaf; val ((Lam, Lam_def), lthy) = lthy |> define_Lam_step fp_b version Y Z preT old_ssig_T ssig_T dead_pre_map unsig rho embL old_Lam; val ((proto_sctr, proto_sctr_def), lthy) = lthy |> define_proto_sctr_step_or_merge fp_b version old_sig_T k_T Sig old_proto_sctr; val pre_map_comp = map_comp_of_bnf pre_bnf; val pre_map_transfer = map_transfer_of_bnf pre_bnf; val dead_pre_map_id = map_id_of_bnf dead_pre_bnf; val dead_pre_map_comp0 = map_comp0_of_bnf dead_pre_bnf; val dead_pre_map_comp = map_comp_of_bnf dead_pre_bnf; val fp_map_id = map_id_of_bnf fp_bnf; val [ctor_dtor] = #ctor_dtors fp_res; val [dtor_inject] = #dtor_injects fp_res; val [dtor_unfold_thm] = #xtor_un_fold_thms fp_res; val dtor_unfold_unique = #xtor_un_fold_unique fp_res; val [dtor_unfold_transfer] = #xtor_un_fold_transfers fp_res; val fp_k_T_rel_eqs = map rel_eq_of_bnf (map_filter (bnf_of lthy) (fold add_type_namesT [fpT, k_T] [])); val unsig_thm = the_single (the_single (#sel_thmss sig_ctr_sugar)); val [sig_map_thm] = #map_thms sig_fp_bnf_sugar; val old_sig_map_comp = map_comp_of_bnf old_sig_bnf; val old_sig_map_cong = map_cong_of_bnf old_sig_bnf; val old_ssig_map_thms = #map_thms old_ssig_fp_bnf_sugar; val [Oper_map_thm, VLeaf_map_thm, CLeaf_map_thm] = #map_thms ssig_fp_bnf_sugar; val old_sig_map_transfer = map_transfer_of_bnf old_sig_bnf; val sig_map_comp = map_comp_of_bnf sig_bnf; val sig_map_transfer = map_transfer_of_bnf sig_bnf; val ssig_map_thms = #map_thms ssig_fp_bnf_sugar; val ssig_map_transfer = map_transfer_of_bnf ssig_bnf; val old_ssig_induct = the_single (#co_inducts old_ssig_fp_induct_sugar); val ssig_induct = the_single (#co_inducts ssig_fp_induct_sugar); val proto_sctr_transfer = derive_proto_sctr_transfer_step_or_merge lthy Y Z R dead_pre_rel dead_sig_rel proto_sctr proto_sctr_def fp_k_T_rel_eqs [old_proto_sctr_transfer]; val embL_transfer = derive_transfer_by_transfer_prover lthy live_AsBs Rs R embL [embL_def] fp_k_T_rel_eqs [old_sig_map_transfer]; val Lam_transfer = derive_Lam_or_eval_core_transfer lthy live_AsBs Y Z preT ssig_T Rs R pre_rel sig_rel ssig_rel Lam Lam_def fp_k_T_rel_eqs [pre_map_transfer, old_Lam_transfer, embL_transfer, rho_transfer]; val ((((((((flat, _, flat_simps), flat_transfer), ((eval_core, _, eval_core_simps), eval_core_transfer)), (eval, eval_def)), (cutSsig, cutSsig_def)), (algLam, algLam_def)), (corecU, corecU_def)), lthy) = lthy |> define_flat_etc fp_b version live_AsBs Y Z preT fpT sig_T ssig_T Oper VLeaf CLeaf pre_rel dead_pre_map dtor dtor_unfold dead_sig_map ssig_rel dead_ssig_map Lam Rs R pre_map_transfer fp_k_T_rel_eqs sig_map_transfer ssig_map_transfer Lam_transfer dtor_transfer; val (Sig_pointful_natural, flat_pointful_natural, Lam_natural_pointful, Lam_pointful_natural, flat_VLeaf, eval_core_pointful_natural, eval_thm, eval_flat, eval_simps as [eval_Oper, _, _], corecU_ctor, corecU_unique, dtor_algLam) = derive_Sig_natural_etc lthy live live_AsBs Y Z preT fpT k_T sig_T ssig_T pre_map dead_pre_map ctor dtor Sig dead_sig_map Oper VLeaf CLeaf ssig_map dead_ssig_map Lam flat eval_core eval cutSsig algLam corecU x fs f g ctor_dtor dtor_inject dtor_unfold_thm dtor_unfold_unique sig_map_thm ssig_induct ssig_map_thms Oper_map_thm VLeaf_map_thm CLeaf_map_thm Lam_transfer flat_simps flat_transfer eval_core_simps eval_core_transfer eval_def cutSsig_def algLam_def corecU_def live_pre_bnf pre_bnf dead_pre_bnf fp_bnf sig_bnf ssig_bnf dead_ssig_bnf; val proto_sctr_natural = derive_natural_from_transfer_with_pre_type lthy live_AsBs Y Z preT ssig_T pre_map ssig_map fs f proto_sctr proto_sctr_transfer [pre_bnf, sig_bnf] []; val proto_sctr_pointful_natural = mk_pointful lthy proto_sctr_natural RS sym; val (embL_pointful_natural, old_algLam_pointful, eval_embL, algLam_algLam) = derive_embL_natural_etc lthy Inl_const old_ssig_bnf ssig_bnf Y Z preT fpT old_ssig_T ssig_T dead_pre_map Sig dead_old_ssig_map embL old_algLam algLam old_flat flat old_eval_core eval_core old_eval eval x f old_ssig_induct dead_pre_map_comp0 dead_pre_map_comp fp_map_id dtor_inject dtor_unfold_unique Sig_pointful_natural unsig_thm sig_map_thm old_sig_map_comp old_sig_map_cong old_ssig_map_thms old_Lam_pointful_natural Lam_def old_flat_simps flat_simps embL_simps embL_transfer old_eval_core_simps eval_core_simps old_eval_thm eval_thm old_dtor_algLam dtor_algLam old_algLam_thm; val algLam_thm = derive_algLam_step_or_merge lthy Y fpT ctor proto_sctr algLam proto_sctr_def old_algLam_pointful algLam_algLam; val k_as_ssig = mk_k_as_ssig Z old_sig_T k_T ssig_T Sig dead_sig_map Oper VLeaf; val k_as_ssig' = substT Y fpT k_as_ssig; val algrho_rhs = HOLogic.mk_comp (eval, k_as_ssig'); val ((algrho, algrho_def), lthy) = lthy |> define_const true fp_b version algrhoN algrho_rhs; val k_as_ssig_transfer = derive_transfer_by_transfer_prover lthy live_AsBs Rs R k_as_ssig [] fp_k_T_rel_eqs [sig_map_transfer]; val k_as_ssig_natural = derive_natural_from_transfer lthy [(Y, Z)] [] [] f k_as_ssig k_as_ssig_transfer [ssig_bnf] [dead_k_bnf]; val k_as_ssig_natural_pointful = mk_pointful lthy k_as_ssig_natural; val [_, Lam_Inr] = derive_Lam_Inl_Inr lthy Y Z preT old_sig_T old_ssig_T k_T ssig_T dead_pre_map Sig embL old_Lam Lam rho unsig_thm Lam_def; val eval_core_k_as_ssig = derive_eval_core_k_as_ssig lthy Y preT k_T rho eval_core k_as_ssig x pre_map_comp dead_pre_map_id sig_map_comp ssig_map_thms Lam_natural_pointful Lam_Inr flat_VLeaf eval_core_simps; val algLam_algrho = derive_algLam_algrho lthy Y fpT Sig algLam algrho algLam_def algrho_def; val dtor_algrho = derive_dtor_algrho lthy Y Z preT fpT k_T ssig_T dead_pre_map dead_k_map dtor rho eval algrho x eval_thm k_as_ssig_natural_pointful eval_core_k_as_ssig algrho_def; val all_algLam_algs = algLam_algLam :: algLam_algrho :: old_all_algLam_algs; val (((corecUU, _), corecUU_thm, corecUU_unique, corecUU_transfer, _, sctr_transfer, sctr_pointful_natural), lthy) = lthy |> define_corecUU_etc fp_b version live_AsBs Y Z preT fpT ssig_T pre_map dead_pre_map pre_rel fp_rel ctor Oper ssig_map dead_ssig_map ssig_rel proto_sctr flat eval_core eval corecU fs f g Rs R pre_map_transfer fp_k_T_rel_eqs dtor_unfold_transfer dtor_transfer ssig_map_transfer proto_sctr_transfer proto_sctr_pointful_natural flat_transfer flat_pointful_natural eval_core_transfer eval_core_pointful_natural eval_thm eval_flat eval_Oper algLam_thm cutSsig_def corecU_def corecU_ctor corecU_unique pre_bnf dead_pre_bnf fp_res ssig_fp_sugar; val (ctr_wrapper, friends) = mk_ctr_wrapper_friends lthy friend_name friend_T old_sig_T k_T Sig old_buffer; val Retr = enforce_type lthy (domain_type o domain_type) fpT old_Retr0; val (dtor_coinduct_info, lthy) = lthy |> derive_cong_step fp_b version fpT dead_ssig_bnf dead_pre_bnf eval eval_thm eval_core_transfer old_dtor_coinduct_info algrho_def k_as_ssig_transfer Retr equivp_Retr Retr_coinduct eval_embL embL_transfer old_all_dead_k_bnfs; val buffer = {Oper = Oper, VLeaf = VLeaf, CLeaf = CLeaf, ctr_wrapper = ctr_wrapper, friends = friends}; val notes = [(corecUU_transferN, [corecUU_transfer])] @ (if Config.get lthy bnf_internals then [(algLamN, [algLam_thm]), (algLam_algLamN, [algLam_algLam]), (algLam_algrhoN, [algLam_algrho]), (cong_alg_introsN, #cong_alg_intros dtor_coinduct_info), (cong_localeN, [#cong_locale dtor_coinduct_info]), (corecU_ctorN, [corecU_ctor]), (corecU_uniqueN, [corecU_unique]), (corecUUN, [corecUU_thm]), (corecUU_uniqueN, [corecUU_unique]), (dtor_algLamN, [dtor_algLam]), (dtor_algrhoN, [dtor_algrho]), (dtor_coinductN, [#dtor_coinduct dtor_coinduct_info]), (embL_pointful_naturalN, [embL_pointful_natural]), (embL_transferN, [embL_transfer]), (evalN, [eval_thm]), (eval_core_pointful_naturalN, [eval_core_pointful_natural]), (eval_core_transferN, [eval_core_transfer]), (eval_flatN, [eval_flat]), (eval_simpsN, eval_simps), (flat_pointful_naturalN, [flat_pointful_natural]), (flat_transferN, [flat_transfer]), (k_as_ssig_naturalN, [k_as_ssig_natural]), (k_as_ssig_transferN, [k_as_ssig_transfer]), (Lam_pointful_naturalN, [Lam_pointful_natural]), (Lam_transferN, [Lam_transfer]), (proto_sctr_pointful_naturalN, [proto_sctr_pointful_natural]), (proto_sctr_transferN, [proto_sctr_transfer]), (rho_transferN, [rho_transfer]), (sctr_pointful_naturalN, [sctr_pointful_natural]), (sctr_transferN, [sctr_transfer]), (Sig_pointful_naturalN, [Sig_pointful_natural])] else []) |> map (fn (thmN, thms) => ((mk_version_fp_binding true lthy version fp_b thmN, []), [(thms, [])])); val phi = Local_Theory.target_morphism lthy; in (({fp_b = fp_b, version = version, fpT = fpT, Y = Y, Z = Z, friend_names = friend_names, sig_fp_sugars = sig_fp_sugar :: old_sig_fp_sugars, ssig_fp_sugar = ssig_fp_sugar, Lam = Lam, proto_sctr = proto_sctr, flat = flat, eval_core = eval_core, eval = eval, algLam = algLam, corecUU = corecUU, dtor_transfer = dtor_transfer, Lam_transfer = Lam_transfer, Lam_pointful_natural = Lam_pointful_natural, proto_sctr_transfer = proto_sctr_transfer, flat_simps = flat_simps, eval_core_simps = eval_core_simps, eval_thm = eval_thm, eval_simps = eval_simps, all_algLam_algs = all_algLam_algs, algLam_thm = algLam_thm, dtor_algLam = dtor_algLam, corecUU_thm = corecUU_thm, corecUU_unique = corecUU_unique, corecUU_transfer = corecUU_transfer, buffer = buffer, all_dead_k_bnfs = dead_k_bnf :: old_all_dead_k_bnfs, Retr = Retr, equivp_Retr = equivp_Retr, Retr_coinduct = Retr_coinduct, dtor_coinduct_info = dtor_coinduct_info} |> morph_corec_info phi, ({algrho = algrho, dtor_algrho = dtor_algrho, algLam_algrho = algLam_algrho} |> morph_friend_info phi)), lthy |> Local_Theory.notes notes |> snd) end; fun derive_corecUU_merge (fpT as Type (fpT_name, res_Ds)) ({friend_names = old1_friend_names, sig_fp_sugars = old1_sig_fp_sugars as old1_sig_fp_sugar :: _, ssig_fp_sugar = old1_ssig_fp_sugar, Lam = old1_Lam0, proto_sctr = old1_proto_sctr0, flat = old1_flat0, eval_core = old1_eval_core0, eval = old1_eval0, algLam = old1_algLam0, dtor_transfer, Lam_transfer = old1_Lam_transfer, Lam_pointful_natural = old1_Lam_pointful_natural, proto_sctr_transfer = old1_proto_sctr_transfer, flat_simps = old1_flat_simps, eval_core_simps = old1_eval_core_simps, eval_thm = old1_eval_thm, all_algLam_algs = old1_all_algLam_algs, algLam_thm = old1_algLam_thm, dtor_algLam = old1_dtor_algLam, buffer = old1_buffer, all_dead_k_bnfs = old1_all_dead_k_bnfs, Retr = old1_Retr0, equivp_Retr, Retr_coinduct, dtor_coinduct_info = old1_dtor_coinduct_info, ...} : corec_info) ({friend_names = old2_friend_names, sig_fp_sugars = old2_sig_fp_sugars as old2_sig_fp_sugar :: _, ssig_fp_sugar = old2_ssig_fp_sugar, Lam = old2_Lam0, flat = old2_flat0, eval_core = old2_eval_core0, eval = old2_eval0, algLam = old2_algLam0, Lam_transfer = old2_Lam_transfer, Lam_pointful_natural = old2_Lam_pointful_natural, flat_simps = old2_flat_simps, eval_core_simps = old2_eval_core_simps, eval_thm = old2_eval_thm, all_algLam_algs = old2_all_algLam_algs, algLam_thm = old2_algLam_thm, dtor_algLam = old2_dtor_algLam, buffer = old2_buffer, all_dead_k_bnfs = old2_all_dead_k_bnfs, dtor_coinduct_info = old2_dtor_coinduct_info, ...} : corec_info) lthy = let val {T = Type (_, fpT_args0), pre_bnf = live_pre_bnf, fp_bnf = live_fp_bnf, fp_res, ...} = checked_fp_sugar_of lthy fpT_name; val fpT_Ss = map Type.sort_of_atyp fpT_args0; val live_fp_alives = liveness_of_fp_bnf (length fpT_args0) live_fp_bnf; val ((Ds, [Y, Z]), names_lthy) = lthy |> mk_TFrees' fpT_Ss ||>> mk_TFrees 2; (* FIXME *) val live_EsFs = []; val live_AsBs = live_EsFs @ [(Y, Z)]; val live = length live_EsFs; val ((((x, f), g), R), _) = names_lthy |> yield_singleton (mk_Frees "x") Y ||>> yield_singleton (mk_Frees "f") (Y --> Z) ||>> yield_singleton (mk_Frees "g") (Y --> Z) ||>> yield_singleton (mk_Frees "R") (mk_pred2T Y Z); (* FIXME *) val fs = []; val Rs = []; val ctor = mk_ctor res_Ds (the_single (#ctors fp_res)); val dtor = mk_dtor res_Ds (the_single (#dtors fp_res)); val old1_sig_T_name = fst (dest_Type (#T old1_sig_fp_sugar)); val old2_sig_T_name = fst (dest_Type (#T old2_sig_fp_sugar)); val old1_ssig_T_name = fst (dest_Type (#T old1_ssig_fp_sugar)); val old2_ssig_T_name = fst (dest_Type (#T old2_ssig_fp_sugar)); val fp_alives = map (K false) live_fp_alives; val As = Ds @ [Y]; val res_As = res_Ds @ [Y]; val res_Bs = res_Ds @ [Z]; val preT = pre_type_of_ctor Y ctor; val YpreT = HOLogic.mk_prodT (Y, preT); val fpT0 = Type (fpT_name, Ds); val old1_sig_T0 = Type (old1_sig_T_name, As); val old2_sig_T0 = Type (old2_sig_T_name, As); val old1_sig_T = Type (old1_sig_T_name, res_As); val old2_sig_T = Type (old2_sig_T_name, res_As); val old1_ssig_T = Type (old1_ssig_T_name, res_As); val old2_ssig_T = Type (old2_ssig_T_name, res_As); val old1_Lam_domT = Tsubst Y YpreT old1_sig_T; val old2_Lam_domT = Tsubst Y YpreT old2_sig_T; val old1_eval_core_domT = Tsubst Y YpreT old1_ssig_T; val old2_eval_core_domT = Tsubst Y YpreT old2_ssig_T; val ((fp_b, version), lthy) = lthy |> get_name_next_version_of fpT_name; val ((sig_fp_sugar, ssig_fp_sugar), lthy) = lthy |> define_sig_type fp_b version fp_alives Ds Y (mk_sumT (old1_sig_T0, old2_sig_T0)) ||>> define_ssig_type fp_b version fp_alives Ds Y fpT0; val sig_T_name = fst (dest_Type (#T sig_fp_sugar)); val ssig_T_name = fst (dest_Type (#T ssig_fp_sugar)); val old1_sig_bnf = #fp_bnf old1_sig_fp_sugar; val old2_sig_bnf = #fp_bnf old2_sig_fp_sugar; val old1_ssig_bnf = #fp_bnf old1_ssig_fp_sugar; val old2_ssig_bnf = #fp_bnf old2_ssig_fp_sugar; val sig_bnf = #fp_bnf sig_fp_sugar; val ssig_bnf = #fp_bnf ssig_fp_sugar; val ((((((((dead_pre_bnf, dead_fp_bnf), dead_old1_sig_bnf), dead_old2_sig_bnf), dead_old1_ssig_bnf), dead_old2_ssig_bnf), dead_sig_bnf), dead_ssig_bnf), lthy) = lthy |> bnf_kill_all_but 1 live_pre_bnf ||>> bnf_kill_all_but 0 live_fp_bnf ||>> bnf_kill_all_but 1 old1_sig_bnf ||>> bnf_kill_all_but 1 old2_sig_bnf ||>> bnf_kill_all_but 1 old1_ssig_bnf ||>> bnf_kill_all_but 1 old2_ssig_bnf ||>> bnf_kill_all_but 1 sig_bnf ||>> bnf_kill_all_but 1 ssig_bnf; (* FIXME *) val pre_bnf = dead_pre_bnf; val fp_bnf = dead_fp_bnf; val old1_ssig_fp_ctr_sugar = #fp_ctr_sugar old1_ssig_fp_sugar; val old2_ssig_fp_ctr_sugar = #fp_ctr_sugar old2_ssig_fp_sugar; val sig_fp_ctr_sugar = #fp_ctr_sugar sig_fp_sugar; val ssig_fp_ctr_sugar = #fp_ctr_sugar ssig_fp_sugar; val sig_fp_bnf_sugar = #fp_bnf_sugar sig_fp_sugar; val old1_ssig_fp_bnf_sugar = #fp_bnf_sugar old1_ssig_fp_sugar; val old2_ssig_fp_bnf_sugar = #fp_bnf_sugar old2_ssig_fp_sugar; val ssig_fp_bnf_sugar = #fp_bnf_sugar ssig_fp_sugar; val old1_ssig_fp_induct_sugar = the (#fp_co_induct_sugar old1_ssig_fp_sugar); val old2_ssig_fp_induct_sugar = the (#fp_co_induct_sugar old2_ssig_fp_sugar); val ssig_fp_induct_sugar = the (#fp_co_induct_sugar ssig_fp_sugar); val old1_ssig_ctr_sugar = #ctr_sugar old1_ssig_fp_ctr_sugar; val old2_ssig_ctr_sugar = #ctr_sugar old2_ssig_fp_ctr_sugar; val sig_ctr_sugar = #ctr_sugar sig_fp_ctr_sugar; val ssig_ctr_sugar = #ctr_sugar ssig_fp_ctr_sugar; val sig_T = Type (sig_T_name, res_As); val ssig_T = Type (ssig_T_name, res_As); val pre_map = mk_mapN lthy live_AsBs preT pre_bnf; val pre_rel = mk_relN lthy live_AsBs preT pre_bnf; val dead_pre_map = mk_map1 lthy Y Z preT dead_pre_bnf; val dead_pre_rel = mk_rel1 lthy Y Z preT dead_pre_bnf; val fp_rel = mk_relN lthy live_EsFs fpT fp_bnf; val dtor_unfold = mk_co_rec (Proof_Context.theory_of lthy) Greatest_FP [Z] fpT (the_single (#xtor_un_folds fp_res)); val Sig = mk_ctr res_As (the_single (#ctrs sig_ctr_sugar)); val unsig = mk_disc_or_sel res_As (the_single (the_single (#selss sig_ctr_sugar))); val sig_rel = mk_relN lthy live_AsBs sig_T sig_bnf; val dead_old1_sig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_old1_sig_bnf); val dead_old2_sig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_old2_sig_bnf); val dead_sig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_sig_bnf); val dead_sig_rel = mk_rel 1 res_As res_Bs (rel_of_bnf dead_sig_bnf); val [old1_Oper, old1_VLeaf, old1_CLeaf] = map (mk_ctr res_As) (#ctrs old1_ssig_ctr_sugar); val [old2_Oper, old2_VLeaf, old2_CLeaf] = map (mk_ctr res_As) (#ctrs old2_ssig_ctr_sugar); val [Oper, VLeaf, CLeaf] = map (mk_ctr res_As) (#ctrs ssig_ctr_sugar); val old1_ssig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_old1_ssig_bnf); val old2_ssig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_old2_ssig_bnf); val ssig_map = mk_mapN lthy live_AsBs ssig_T ssig_bnf; val ssig_rel = mk_relN lthy live_AsBs ssig_T ssig_bnf; val dead_ssig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_ssig_bnf); val old1_Lam = enforce_type lthy domain_type old1_Lam_domT old1_Lam0; val old2_Lam = enforce_type lthy domain_type old2_Lam_domT old2_Lam0; val old1_proto_sctr = enforce_type lthy domain_type preT old1_proto_sctr0; val old1_flat = enforce_type lthy range_type old1_ssig_T old1_flat0; val old2_flat = enforce_type lthy range_type old2_ssig_T old2_flat0; val old1_eval_core = enforce_type lthy domain_type old1_eval_core_domT old1_eval_core0; val old2_eval_core = enforce_type lthy domain_type old2_eval_core_domT old2_eval_core0; val old1_eval = enforce_type lthy range_type fpT old1_eval0; val old2_eval = enforce_type lthy range_type fpT old2_eval0; val old1_algLam = enforce_type lthy range_type fpT old1_algLam0; val old2_algLam = enforce_type lthy range_type fpT old2_algLam0; val ((embLL, embLL_def, embLL_simps), lthy) = lthy |> define_embL embLLN fp_b version Y Z fpT old1_sig_T old1_ssig_T old2_sig_T ssig_T Inl_const dead_old1_sig_map Sig old1_Oper old1_VLeaf old1_CLeaf Oper VLeaf CLeaf; val ((embLR, embLR_def, embLR_simps), lthy) = lthy |> define_embL embLRN fp_b version Y Z fpT old2_sig_T old2_ssig_T old1_sig_T ssig_T (fn T => fn U => Inr_const U T) dead_old2_sig_map Sig old2_Oper old2_VLeaf old2_CLeaf Oper VLeaf CLeaf; val ((Lam, Lam_def), lthy) = lthy |> define_Lam_merge fp_b version Y Z preT old1_ssig_T old2_ssig_T ssig_T dead_pre_map unsig embLL embLR old1_Lam old2_Lam; val ((proto_sctr, proto_sctr_def), lthy) = lthy |> define_proto_sctr_step_or_merge fp_b version old1_sig_T old2_sig_T Sig old1_proto_sctr; val pre_map_transfer = map_transfer_of_bnf pre_bnf; val dead_pre_map_comp0 = map_comp0_of_bnf dead_pre_bnf; val dead_pre_map_comp = map_comp_of_bnf dead_pre_bnf; val fp_map_id = map_id_of_bnf fp_bnf; val fp_rel_eq = rel_eq_of_bnf fp_bnf; val [ctor_dtor] = #ctor_dtors fp_res; val [dtor_inject] = #dtor_injects fp_res; val [dtor_unfold_thm] = #xtor_un_fold_thms fp_res; val dtor_unfold_unique = #xtor_un_fold_unique fp_res; val [dtor_unfold_transfer] = #xtor_un_fold_transfers fp_res; val unsig_thm = the_single (the_single (#sel_thmss sig_ctr_sugar)); val [sig_map_thm] = #map_thms sig_fp_bnf_sugar; val old1_sig_map_comp = map_comp_of_bnf old1_sig_bnf; val old2_sig_map_comp = map_comp_of_bnf old2_sig_bnf; val old1_sig_map_cong = map_cong_of_bnf old1_sig_bnf; val old2_sig_map_cong = map_cong_of_bnf old2_sig_bnf; val old1_ssig_map_thms = #map_thms old1_ssig_fp_bnf_sugar; val old2_ssig_map_thms = #map_thms old2_ssig_fp_bnf_sugar; val [Oper_map_thm, VLeaf_map_thm, CLeaf_map_thm] = #map_thms ssig_fp_bnf_sugar; val old1_sig_map_transfer = map_transfer_of_bnf old1_sig_bnf; val old2_sig_map_transfer = map_transfer_of_bnf old2_sig_bnf; val sig_map_transfer = map_transfer_of_bnf sig_bnf; val ssig_map_thms = #map_thms ssig_fp_bnf_sugar; val ssig_map_transfer = map_transfer_of_bnf ssig_bnf; val old1_ssig_induct = the_single (#co_inducts old1_ssig_fp_induct_sugar); val old2_ssig_induct = the_single (#co_inducts old2_ssig_fp_induct_sugar); val ssig_induct = the_single (#co_inducts ssig_fp_induct_sugar); val proto_sctr_transfer = derive_proto_sctr_transfer_step_or_merge lthy Y Z R dead_pre_rel dead_sig_rel proto_sctr proto_sctr_def [] [old1_proto_sctr_transfer]; val embLL_transfer = derive_transfer_by_transfer_prover lthy live_AsBs Rs R embLL [embLL_def] [] [old1_sig_map_transfer]; val embLR_transfer = derive_transfer_by_transfer_prover lthy live_AsBs Rs R embLR [embLR_def] [] [old2_sig_map_transfer]; val Lam_transfer = derive_Lam_or_eval_core_transfer lthy live_AsBs Y Z preT ssig_T Rs R pre_rel sig_rel ssig_rel Lam Lam_def [] [pre_map_transfer, old1_Lam_transfer, old2_Lam_transfer, embLL_transfer, embLR_transfer]; val ((((((((flat, _, flat_simps), flat_transfer), ((eval_core, _, eval_core_simps), eval_core_transfer)), (eval, eval_def)), (cutSsig, cutSsig_def)), (algLam, algLam_def)), (corecU, corecU_def)), lthy) = lthy |> define_flat_etc fp_b version live_AsBs Y Z preT fpT sig_T ssig_T Oper VLeaf CLeaf pre_rel dead_pre_map dtor dtor_unfold dead_sig_map ssig_rel dead_ssig_map Lam Rs R pre_map_transfer [fp_rel_eq] sig_map_transfer ssig_map_transfer Lam_transfer dtor_transfer; val (Sig_pointful_natural, flat_pointful_natural, _, Lam_pointful_natural, _, eval_core_pointful_natural, eval_thm, eval_flat, eval_simps as [eval_Oper, _, _], corecU_ctor, corecU_unique, dtor_algLam) = derive_Sig_natural_etc lthy live live_AsBs Y Z preT fpT fpT sig_T ssig_T pre_map dead_pre_map ctor dtor Sig dead_sig_map Oper VLeaf CLeaf ssig_map dead_ssig_map Lam flat eval_core eval cutSsig algLam corecU x fs f g ctor_dtor dtor_inject dtor_unfold_thm dtor_unfold_unique sig_map_thm ssig_induct ssig_map_thms Oper_map_thm VLeaf_map_thm CLeaf_map_thm Lam_transfer flat_simps flat_transfer eval_core_simps eval_core_transfer eval_def cutSsig_def algLam_def corecU_def live_pre_bnf pre_bnf dead_pre_bnf fp_bnf sig_bnf ssig_bnf dead_ssig_bnf; val proto_sctr_natural = derive_natural_from_transfer_with_pre_type lthy live_AsBs Y Z preT ssig_T pre_map ssig_map fs f proto_sctr proto_sctr_transfer [pre_bnf, sig_bnf] []; val proto_sctr_pointful_natural = mk_pointful lthy proto_sctr_natural RS sym; val (embLL_pointful_natural, old1_algLam_pointful, eval_embLL, algLam_algLamL) = derive_embL_natural_etc lthy Inl_const old1_ssig_bnf ssig_bnf Y Z preT fpT old1_ssig_T ssig_T dead_pre_map Sig old1_ssig_map embLL old1_algLam algLam old1_flat flat old1_eval_core eval_core old1_eval eval x f old1_ssig_induct dead_pre_map_comp0 dead_pre_map_comp fp_map_id dtor_inject dtor_unfold_unique Sig_pointful_natural unsig_thm sig_map_thm old1_sig_map_comp old1_sig_map_cong old1_ssig_map_thms old1_Lam_pointful_natural Lam_def old1_flat_simps flat_simps embLL_simps embLL_transfer old1_eval_core_simps eval_core_simps old1_eval_thm eval_thm old1_dtor_algLam dtor_algLam old1_algLam_thm; val (embLR_pointful_natural, _, eval_embLR, algLam_algLamR) = derive_embL_natural_etc lthy Inr_const old2_ssig_bnf ssig_bnf Y Z preT fpT old2_ssig_T ssig_T dead_pre_map Sig old2_ssig_map embLR old2_algLam algLam old2_flat flat old2_eval_core eval_core old2_eval eval x f old2_ssig_induct dead_pre_map_comp0 dead_pre_map_comp fp_map_id dtor_inject dtor_unfold_unique Sig_pointful_natural unsig_thm sig_map_thm old2_sig_map_comp old2_sig_map_cong old2_ssig_map_thms old2_Lam_pointful_natural Lam_def old2_flat_simps flat_simps embLR_simps embLR_transfer old2_eval_core_simps eval_core_simps old2_eval_thm eval_thm old2_dtor_algLam dtor_algLam old2_algLam_thm; val algLam_thm = derive_algLam_step_or_merge lthy Y fpT ctor proto_sctr algLam proto_sctr_def old1_algLam_pointful algLam_algLamL; val all_algLam_algs = algLam_algLamL :: algLam_algLamR :: merge_lists (Thm.eq_thm_prop o apply2 zero_var_indexes) old1_all_algLam_algs old2_all_algLam_algs; val (((corecUU, _), corecUU_thm, corecUU_unique, corecUU_transfer, _, sctr_transfer, sctr_pointful_natural), lthy) = lthy |> define_corecUU_etc fp_b version live_AsBs Y Z preT fpT ssig_T pre_map dead_pre_map pre_rel fp_rel ctor Oper ssig_map dead_ssig_map ssig_rel proto_sctr flat eval_core eval corecU fs f g Rs R pre_map_transfer [] dtor_unfold_transfer dtor_transfer ssig_map_transfer proto_sctr_transfer proto_sctr_pointful_natural flat_transfer flat_pointful_natural eval_core_transfer eval_core_pointful_natural eval_thm eval_flat eval_Oper algLam_thm cutSsig_def corecU_def corecU_ctor corecU_unique pre_bnf dead_pre_bnf fp_res ssig_fp_sugar; val Retr = enforce_type lthy (domain_type o domain_type) fpT old1_Retr0; val embed_Sig_inl = embed_Sig lthy Sig (Inl_const old1_sig_T old2_sig_T); val embed_Sig_inr = embed_Sig lthy Sig (Inr_const old1_sig_T old2_sig_T); val ctr_wrapper = embed_Sig_inl (#ctr_wrapper old1_buffer); val friends = Symtab.merge (K true) (Symtab.map (K (apsnd embed_Sig_inl)) (#friends old1_buffer), Symtab.map (K (apsnd embed_Sig_inr)) (#friends old2_buffer)); val old_fp_sugars = merge_lists (op = o apply2 (fst o dest_Type o #T)) old1_sig_fp_sugars old2_sig_fp_sugars; val ((dtor_coinduct_info, all_dead_k_bnfs, friend_names), lthy) = lthy |> derive_cong_merge fp_b version fpT old1_friend_names old2_friend_names dead_ssig_bnf dead_pre_bnf eval eval_thm eval_core_transfer old1_dtor_coinduct_info old2_dtor_coinduct_info Retr equivp_Retr Retr_coinduct eval_embLL embLL_transfer eval_embLR embLR_transfer old1_all_dead_k_bnfs old2_all_dead_k_bnfs; val buffer = {Oper = Oper, VLeaf = VLeaf, CLeaf = CLeaf, ctr_wrapper = ctr_wrapper, friends = friends}; val notes = [(corecUU_transferN, [corecUU_transfer])] @ (if Config.get lthy bnf_internals then [(algLamN, [algLam_thm]), (algLam_algLamN, [algLam_algLamL, algLam_algLamR]), (cong_alg_introsN, #cong_alg_intros dtor_coinduct_info), (cong_localeN, [#cong_locale dtor_coinduct_info]), (corecU_ctorN, [corecU_ctor]), (corecU_uniqueN, [corecU_unique]), (corecUUN, [corecUU_thm]), (corecUU_uniqueN, [corecUU_unique]), (dtor_algLamN, [dtor_algLam]), (dtor_coinductN, [#dtor_coinduct dtor_coinduct_info]), (eval_core_pointful_naturalN, [eval_core_pointful_natural]), (eval_core_transferN, [eval_core_transfer]), (embL_pointful_naturalN, [embLL_pointful_natural, embLR_pointful_natural]), (embL_transferN, [embLL_transfer, embLR_transfer]), (evalN, [eval_thm]), (eval_flatN, [eval_flat]), (eval_simpsN, eval_simps), (flat_pointful_naturalN, [flat_pointful_natural]), (flat_transferN, [flat_transfer]), (Lam_pointful_naturalN, [Lam_pointful_natural]), (Lam_transferN, [Lam_transfer]), (proto_sctr_pointful_naturalN, [proto_sctr_pointful_natural]), (proto_sctr_transferN, [proto_sctr_transfer]), (sctr_pointful_naturalN, [sctr_pointful_natural]), (sctr_transferN, [sctr_transfer]), (Sig_pointful_naturalN, [Sig_pointful_natural])] else []) |> map (fn (thmN, thms) => ((mk_version_fp_binding true lthy version fp_b thmN, []), [(thms, [])])); in ({fp_b = fp_b, version = version, fpT = fpT, Y = Y, Z = Z, friend_names = friend_names, sig_fp_sugars = sig_fp_sugar :: old_fp_sugars, ssig_fp_sugar = ssig_fp_sugar, Lam = Lam, proto_sctr = proto_sctr, flat = flat, eval_core = eval_core, eval = eval, algLam = algLam, corecUU = corecUU, dtor_transfer = dtor_transfer, Lam_transfer = Lam_transfer, Lam_pointful_natural = Lam_pointful_natural, proto_sctr_transfer = proto_sctr_transfer, flat_simps = flat_simps, eval_core_simps = eval_core_simps, eval_thm = eval_thm, eval_simps = eval_simps, all_algLam_algs = all_algLam_algs, algLam_thm = algLam_thm, dtor_algLam = dtor_algLam, corecUU_thm = corecUU_thm, corecUU_unique = corecUU_unique, corecUU_transfer = corecUU_transfer, buffer = buffer, all_dead_k_bnfs = all_dead_k_bnfs, Retr = Retr, equivp_Retr = equivp_Retr, Retr_coinduct = Retr_coinduct, dtor_coinduct_info = dtor_coinduct_info} |> morph_corec_info (Local_Theory.target_morphism lthy), lthy |> Local_Theory.notes notes |> snd) end; fun set_corec_info_exprs fpT_name f = Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => let val exprs = f phi in Data.map (apsnd (fn [info_tab] => [Symtab.map_default (fpT_name, exprs) (K exprs) info_tab])) end); fun subsume_corec_info_ad ctxt {fpT = fpT1, friend_names = friend_names1} {fpT = fpT2, friend_names = friend_names2} = Sign.typ_instance (Proof_Context.theory_of ctxt) (fpT1, fpT2) andalso subset (op =) (friend_names1, friend_names2); fun subsume_corec_info_expr ctxt expr1 expr2 = subsume_corec_info_ad ctxt (corec_ad_of_expr expr1) (corec_ad_of_expr expr2); fun instantiate_corec_info thy res_T (info as {fpT, ...}) = let val As_rho = tvar_subst thy [fpT] [res_T]; val substAT = Term.typ_subst_TVars As_rho; val substA = Term.subst_TVars As_rho; val phi = Morphism.typ_morphism "BNF" substAT $> Morphism.term_morphism "BNF" substA; in morph_corec_info phi info end; fun instantiate_corec_info_expr thy res_T (Ad ({friend_names, ...}, f)) = Ad ({fpT = res_T, friend_names = friend_names}, f #>> instantiate_corec_info thy res_T) | instantiate_corec_info_expr thy res_T (Info info) = Info (instantiate_corec_info thy res_T info); fun ensure_Info expr = corec_info_of_expr expr #>> Info and ensure_Info_if_Info old_expr (expr, lthy) = if is_Info old_expr then ensure_Info expr lthy else (expr, lthy) and merge_corec_info_exprs old_exprs expr1 expr2 lthy = if subsume_corec_info_expr lthy expr2 expr1 then if subsume_corec_info_expr lthy expr1 expr2 andalso is_Ad expr1 then (expr2, lthy) else ensure_Info_if_Info expr2 (expr1, lthy) else if subsume_corec_info_expr lthy expr1 expr2 then ensure_Info_if_Info expr1 (expr2, lthy) else let val thy = Proof_Context.theory_of lthy; val {fpT = fpT1, friend_names = friend_names1} = corec_ad_of_expr expr1; val {fpT = fpT2, friend_names = friend_names2} = corec_ad_of_expr expr2; val fpT0 = typ_unify_disjointly thy (fpT1, fpT2); val fpT = singleton (freeze_types lthy []) fpT0; val friend_names = merge_lists (op =) friend_names1 friend_names2; val expr = Ad ({fpT = fpT, friend_names = friend_names}, corec_info_of_expr expr1 ##>> corec_info_of_expr expr2 #-> uncurry (derive_corecUU_merge fpT)); val old_same_type_exprs = if old_exprs then [] |> Sign.typ_instance thy (fpT1, fpT0) ? cons expr1 |> Sign.typ_instance thy (fpT2, fpT0) ? cons expr2 else []; in (expr, lthy) |> fold ensure_Info_if_Info old_same_type_exprs end and insert_corec_info_expr expr exprs lthy = let val thy = Proof_Context.theory_of lthy; val {fpT = new_fpT, ...} = corec_ad_of_expr expr; val is_Tinst = curry (Sign.typ_instance thy); fun is_Tequiv T U = is_Tinst T U andalso is_Tinst U T; val (((equiv_exprs, sub_exprs), sup_exprs), incomp_exprs) = exprs |> List.partition ((fn {fpT, ...} => is_Tequiv fpT new_fpT) o corec_ad_of_expr) ||>> List.partition ((fn {fpT, ...} => is_Tinst fpT new_fpT) o corec_ad_of_expr) ||>> List.partition ((fn {fpT, ...} => is_Tinst new_fpT fpT) o corec_ad_of_expr); fun add_instantiated_incomp_expr expr exprs = let val {fpT, ...} = corec_ad_of_expr expr in (case try (typ_unify_disjointly thy) (fpT, new_fpT) of SOME new_T => let val subsumes = (fn {fpT, ...} => is_Tinst new_T fpT) o corec_ad_of_expr in if exists (exists subsumes) [exprs, sub_exprs] then exprs else instantiate_corec_info_expr thy new_T expr :: exprs end | NONE => exprs) end; val unincomp_exprs = fold add_instantiated_incomp_expr incomp_exprs []; val ((merged_sub_exprs, merged_unincomp_exprs), lthy) = lthy |> fold_map (merge_corec_info_exprs true expr) sub_exprs ||>> fold_map (merge_corec_info_exprs false expr) unincomp_exprs; val (merged_equiv_expr, lthy) = (expr, lthy) |> fold (uncurry o merge_corec_info_exprs true) equiv_exprs; in (merged_unincomp_exprs @ merged_sub_exprs @ merged_equiv_expr :: sup_exprs @ incomp_exprs |> sort (rev_order o int_ord o apply2 (length o #friend_names o corec_ad_of_expr)), lthy) end and register_corec_info (info as {fpT = Type (fpT_name, _), ...}) lthy = let val (exprs, lthy) = insert_corec_info_expr (Info info) (corec_info_exprs_of lthy fpT_name) lthy; in lthy |> set_corec_info_exprs fpT_name (fn phi => map (morph_corec_info_expr phi) exprs) end and corec_info_of_expr (Ad (_, f)) lthy = f lthy | corec_info_of_expr (Info info) lthy = (info, lthy); fun nonempty_corec_info_exprs_of fpT_name lthy = (case corec_info_exprs_of lthy fpT_name of [] => derive_corecUU_base fpT_name lthy |> (fn (info, lthy) => ([Info info], lthy |> set_corec_info_exprs fpT_name (fn phi => [Info (morph_corec_info phi info)]))) | exprs => (exprs, lthy)); fun corec_info_of res_T lthy = (case res_T of Type (fpT_name, _) => let val (exprs, lthy) = nonempty_corec_info_exprs_of fpT_name lthy; val thy = Proof_Context.theory_of lthy; val expr = (case find_first ((fn {fpT, ...} => Sign.typ_instance thy (res_T, fpT)) o corec_ad_of_expr) exprs of SOME expr => expr | NONE => error ("Invalid type: " ^ Syntax.string_of_typ lthy res_T)); val (info, lthy) = corec_info_of_expr expr lthy; in (instantiate_corec_info thy res_T info, lthy |> is_Ad expr ? register_corec_info info) end | _ => not_codatatype lthy res_T); fun maybe_corec_info_of ctxt res_T = (case res_T of Type (fpT_name, _) => let val thy = Proof_Context.theory_of ctxt; val infos = corec_infos_of ctxt fpT_name; in find_first (fn {fpT, ...} => Sign.typ_instance thy (res_T, fpT)) infos |> Option.map (instantiate_corec_info thy res_T) end | _ => not_codatatype ctxt res_T); fun prepare_friend_corec friend_name friend_T lthy = let val (arg_Ts, res_T) = strip_type friend_T; val Type (fpT_name, res_Ds) = (case res_T of T as Type _ => T | T => error (not_codatatype lthy T)); val _ = not (null arg_Ts) orelse error "Function with no argument cannot be registered as friend"; val {T = Type (fpT_name, fpT_args0), pre_bnf, fp_bnf = live_fp_bnf, fp_res, ...} = checked_fp_sugar_of lthy fpT_name; val num_fp_tyargs = length fpT_args0; val fpT_Ss = map Type.sort_of_atyp fpT_args0; val live_fp_alives = liveness_of_fp_bnf num_fp_tyargs live_fp_bnf; val (old_info as {friend_names = old_friend_names, sig_fp_sugars = old_sig_fp_sugar :: _, buffer = old_buffer, ...}, lthy) = corec_info_of res_T lthy; val old_sig_T_name = fst (dest_Type (#T old_sig_fp_sugar)); val old_sig_alives = liveness_of_fp_bnf (num_fp_tyargs + 1) (#fp_bnf old_sig_fp_sugar); (* FIXME: later *) val fp_alives = fst (split_last old_sig_alives); val fp_alives = map (K false) live_fp_alives; val _ = not (member (op =) old_friend_names friend_name) orelse error ("Function " ^ quote (Syntax.string_of_term lthy (Const (friend_name, friend_T))) ^ " already registered as friend"); val lthy = lthy |> Variable.declare_typ friend_T; val ((Ds, [Y, Z]), _) = lthy |> mk_TFrees' fpT_Ss ||>> mk_TFrees 2; (* FIXME *) val dead_Ds = Ds; val live_As = [Y]; val ctor = mk_ctor res_Ds (the_single (#ctors fp_res)); val fpT0 = Type (fpT_name, Ds); val k_Ts0 = map (typ_subst_nonatomic (res_Ds ~~ Ds) o typ_subst_nonatomic [(res_T, Y)]) arg_Ts; val k_T0 = mk_tupleT_balanced k_Ts0; val As = Ds @ [Y]; val res_As = res_Ds @ [Y]; val k_As = fold Term.add_tfreesT k_Ts0 []; val _ = (case filter_out (member (op =) As o TFree) k_As of [] => () | k_A :: _ => error ("Cannot have type variable " ^ quote (Syntax.string_of_typ lthy (TFree k_A)) ^ " in the argument types when it does not occur as an immediate argument of the result \ \type constructor")); val substDT = Term.typ_subst_atomic (Ds ~~ res_Ds); val old_sig_T0 = Type (old_sig_T_name, As); val ((fp_b, version), lthy) = lthy |> get_name_next_version_of fpT_name; val (((dead_k_bnf, sig_fp_sugar), ssig_fp_sugar), lthy) = lthy |> bnf_with_deads_and_lives dead_Ds live_As Y fpT0 k_T0 ||>> define_sig_type fp_b version fp_alives Ds Y (mk_sumT (old_sig_T0, k_T0)) ||>> define_ssig_type fp_b version fp_alives Ds Y fpT0; val _ = live_of_bnf dead_k_bnf = 1 orelse error "Impossible type for friend (the result codatatype must occur live in the arguments)"; val (dead_pre_bnf, lthy) = lthy |> bnf_kill_all_but 1 pre_bnf; val sig_fp_ctr_sugar = #fp_ctr_sugar sig_fp_sugar; val ssig_fp_ctr_sugar = #fp_ctr_sugar ssig_fp_sugar; val sig_ctr_sugar = #ctr_sugar sig_fp_ctr_sugar; val ssig_ctr_sugar = #ctr_sugar ssig_fp_ctr_sugar; val ssig_T_name = fst (dest_Type (#T ssig_fp_sugar)); val preT = pre_type_of_ctor Y ctor; val old_sig_T = substDT old_sig_T0; val k_T = substDT k_T0; val ssig_T = Type (ssig_T_name, res_As); val Sig = mk_ctr res_As (the_single (#ctrs sig_ctr_sugar)); val [Oper, VLeaf, CLeaf] = map (mk_ctr res_As) (#ctrs ssig_ctr_sugar); val (ctr_wrapper, friends) = mk_ctr_wrapper_friends lthy friend_name friend_T old_sig_T k_T Sig old_buffer; val buffer = {Oper = Oper, VLeaf = VLeaf, CLeaf = CLeaf, ctr_wrapper = ctr_wrapper, friends = friends}; in ((old_info, fp_b, version, Y, Z, preT, k_T, ssig_T, dead_pre_bnf, dead_k_bnf, sig_fp_sugar, ssig_fp_sugar, buffer), lthy) end; fun register_friend_corec key fp_b version Y Z k_T dead_k_bnf sig_fp_sugar ssig_fp_sugar friend_const rho rho_transfer old_info lthy = let val friend_T = fastype_of friend_const; val res_T = body_type friend_T; in derive_corecUU_step res_T old_info key friend_T fp_b version Y Z k_T dead_k_bnf sig_fp_sugar ssig_fp_sugar rho rho_transfer lthy |> (fn ((info, friend_info), lthy) => (friend_info, register_corec_info info lthy)) end; fun merge_corec_info_exprss exprs1 exprs2 lthy = let fun all_friend_names_of exprs = fold (union (op =)) (map (#friend_names o corec_ad_of_expr) exprs) []; val friend_names1 = all_friend_names_of exprs1; val friend_names2 = all_friend_names_of exprs2; in if subset (op =) (friend_names2, friend_names1) then if subset (op =) (friend_names1, friend_names2) andalso length (filter is_Info exprs2) > length (filter is_Info exprs1) then (exprs2, lthy) else (exprs1, lthy) else if subset (op =) (friend_names1, friend_names2) then (exprs2, lthy) else fold_rev (uncurry o insert_corec_info_expr) exprs2 (exprs1, lthy) end; fun merge_corec_info_tabs info_tab1 info_tab2 lthy = let val fpT_names = union (op =) (Symtab.keys info_tab1) (Symtab.keys info_tab2); fun add_infos_of fpT_name (info_tab, lthy) = (case Symtab.lookup info_tab1 fpT_name of NONE => (Symtab.update_new (fpT_name, the (Symtab.lookup info_tab2 fpT_name)) info_tab, lthy) | SOME exprs1 => (case Symtab.lookup info_tab2 fpT_name of NONE => (Symtab.update_new (fpT_name, exprs1) info_tab, lthy) | SOME exprs2 => let val (exprs, lthy) = merge_corec_info_exprss exprs1 exprs2 lthy in (Symtab.update_new (fpT_name, exprs) info_tab, lthy) end)); in fold add_infos_of fpT_names (Symtab.empty, lthy) end; fun consolidate lthy = (case snd (Data.get (Context.Proof lthy)) of [_] => raise Same.SAME | info_tab :: info_tabs => let val (info_tab', lthy) = fold_rev (uncurry o merge_corec_info_tabs) info_tabs (info_tab, lthy); in Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => Data.map (apsnd (fn _ => [Symtab.map (K (map (morph_corec_info_expr phi))) info_tab']))) lthy end); fun consolidate_global thy = SOME (Named_Target.theory_map consolidate thy) handle Same.SAME => NONE; val _ = Theory.setup (Theory.at_begin consolidate_global); end; diff --git a/src/HOL/Tools/BNF/bnf_gfp_grec_sugar.ML b/src/HOL/Tools/BNF/bnf_gfp_grec_sugar.ML --- a/src/HOL/Tools/BNF/bnf_gfp_grec_sugar.ML +++ b/src/HOL/Tools/BNF/bnf_gfp_grec_sugar.ML @@ -1,2391 +1,2391 @@ (* Title: HOL/Tools/BNF/bnf_gfp_grec_sugar.ML Author: Aymeric Bouzy, Ecole polytechnique Author: Jasmin Blanchette, Inria, LORIA, MPII Author: Dmitriy Traytel, ETH Zürich Copyright 2015, 2016 Generalized corecursor sugar ("corec" and friends). *) signature BNF_GFP_GREC_SUGAR = sig datatype corec_option = Plugins_Option of Proof.context -> Plugin_Name.filter | Friend_Option | Transfer_Option val parse_corec_equation: Proof.context -> term list -> term -> term list * term val explore_corec_equation: Proof.context -> bool -> bool -> string -> term -> BNF_GFP_Grec_Sugar_Util.s_parse_info -> typ -> term list * term -> term list * term val build_corecUU_arg_and_goals: bool -> term -> term list * term -> local_theory -> (((thm list * thm list * thm list) * term list) * term) * local_theory val derive_eq_corecUU: Proof.context -> BNF_GFP_Grec.corec_info -> term -> term -> thm -> thm val derive_unique: Proof.context -> morphism -> term -> BNF_GFP_Grec.corec_info -> string -> thm -> thm val corec_cmd: bool -> corec_option list -> (binding * string option * mixfix) list * string -> local_theory -> local_theory val corecursive_cmd: bool -> corec_option list -> (binding * string option * mixfix) list * string -> local_theory -> Proof.state val friend_of_corec_cmd: (string * string option) * string -> local_theory -> Proof.state val coinduction_upto_cmd: string * string -> local_theory -> local_theory end; structure BNF_GFP_Grec_Sugar : BNF_GFP_GREC_SUGAR = struct open Ctr_Sugar open BNF_Util open BNF_Tactics open BNF_Def open BNF_Comp open BNF_FP_Util open BNF_FP_Def_Sugar open BNF_FP_N2M_Sugar open BNF_GFP_Util open BNF_GFP_Rec_Sugar open BNF_FP_Rec_Sugar_Transfer open BNF_GFP_Grec open BNF_GFP_Grec_Sugar_Util open BNF_GFP_Grec_Sugar_Tactics val cong_N = "cong_"; val baseN = "base"; val reflN = "refl"; val symN = "sym"; val transN = "trans"; val cong_introsN = prefix cong_N "intros"; val codeN = "code"; val coinductN = "coinduct"; val coinduct_uptoN = "coinduct_upto"; val corecN = "corec"; val ctrN = "ctr"; val discN = "disc"; val disc_iffN = "disc_iff"; val eq_algrhoN = "eq_algrho"; val eq_corecUUN = "eq_corecUU"; val friendN = "friend"; val inner_elimN = "inner_elim"; val inner_inductN = "inner_induct"; val inner_simpN = "inner_simp"; val rhoN = "rho"; val selN = "sel"; val uniqueN = "unique"; val inner_fp_suffix = "_inner_fp"; val nitpicksimp_attrs = @{attributes [nitpick_simp]}; val simp_attrs = @{attributes [simp]}; val unfold_id_thms1 = map (fn thm => thm RS eq_reflection) @{thms id_bnf_o o_id_bnf id_apply o_apply} @ @{thms fst_def[abs_def, symmetric] snd_def[abs_def, symmetric]}; fun unfold_id_bnf_etc lthy = let val thy = Proof_Context.theory_of lthy in Raw_Simplifier.rewrite_term thy unfold_id_thms1 [] #> Raw_Simplifier.rewrite_term thy @{thms BNF_Composition.id_bnf_def} [] end; datatype corec_option = Plugins_Option of Proof.context -> Plugin_Name.filter | Friend_Option | Transfer_Option; val corec_option_parser = Parse.group (K "option") (Plugin_Name.parse_filter >> Plugins_Option || Parse.reserved "friend" >> K Friend_Option || Parse.reserved "transfer" >> K Transfer_Option); type codatatype_extra = {case_dtor: thm, case_trivial: thm, abs_rep_transfers: thm list}; fun morph_codatatype_extra phi ({case_dtor, case_trivial, abs_rep_transfers} : codatatype_extra) = {case_dtor = Morphism.thm phi case_dtor, case_trivial = Morphism.thm phi case_trivial, abs_rep_transfers = map (Morphism.thm phi) abs_rep_transfers}; val transfer_codatatype_extra = morph_codatatype_extra o Morphism.transfer_morphism; type coinduct_extra = {coinduct: thm, coinduct_attrs: Token.src list, cong_intro_pairs: (string * thm) list}; fun morph_coinduct_extra phi ({coinduct, coinduct_attrs, cong_intro_pairs} : coinduct_extra) = {coinduct = Morphism.thm phi coinduct, coinduct_attrs = coinduct_attrs, cong_intro_pairs = map (apsnd (Morphism.thm phi)) cong_intro_pairs}; val transfer_coinduct_extra = morph_coinduct_extra o Morphism.transfer_morphism; type friend_extra = {eq_algrhos: thm list, algrho_eqs: thm list}; val empty_friend_extra = {eq_algrhos = [], algrho_eqs = []}; fun merge_friend_extras ({eq_algrhos = eq_algrhos1, algrho_eqs = algrho_eqs1}, {eq_algrhos = eq_algrhos2, algrho_eqs = algrho_eqs2}) = {eq_algrhos = union Thm.eq_thm_prop eq_algrhos1 eq_algrhos2, algrho_eqs = union Thm.eq_thm_prop algrho_eqs1 algrho_eqs2}; type corec_sugar_data = codatatype_extra Symtab.table * coinduct_extra Symtab.table * friend_extra Symtab.table; structure Data = Generic_Data ( type T = corec_sugar_data; val empty = (Symtab.empty, Symtab.empty, Symtab.empty); val extend = I; fun merge data : T = (Symtab.merge (K true) (apply2 #1 data), Symtab.merge (K true) (apply2 #2 data), Symtab.join (K merge_friend_extras) (apply2 #3 data)); ); fun register_codatatype_extra fpT_name extra = Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => Data.map (@{apply 3(1)} (Symtab.update (fpT_name, morph_codatatype_extra phi extra)))); fun codatatype_extra_of ctxt = Symtab.lookup (#1 (Data.get (Context.Proof ctxt))) #> Option.map (transfer_codatatype_extra (Proof_Context.theory_of ctxt)); fun all_codatatype_extras_of ctxt = Symtab.dest (#1 (Data.get (Context.Proof ctxt))); fun register_coinduct_extra fpT_name extra = Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => Data.map (@{apply 3(2)} (Symtab.update (fpT_name, morph_coinduct_extra phi extra)))); fun coinduct_extra_of ctxt = Symtab.lookup (#2 (Data.get (Context.Proof ctxt))) #> Option.map (transfer_coinduct_extra (Proof_Context.theory_of ctxt)); fun register_friend_extra fun_name eq_algrho algrho_eq = Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => Data.map (@{apply 3(3)} (Symtab.map_default (fun_name, empty_friend_extra) (fn {eq_algrhos, algrho_eqs} => {eq_algrhos = Morphism.thm phi eq_algrho :: eq_algrhos, algrho_eqs = Morphism.thm phi algrho_eq :: algrho_eqs})))); fun all_friend_extras_of ctxt = Symtab.dest (#3 (Data.get (Context.Proof ctxt))); fun coinduct_extras_of_generic context = corec_infos_of_generic context #> map (#corecUU #> dest_Const #> fst #> Symtab.lookup (#2 (Data.get context)) #> the #> transfer_coinduct_extra (Context.theory_of context)); fun get_coinduct_uptos fpT_name context = coinduct_extras_of_generic context fpT_name |> map #coinduct; fun get_cong_all_intros fpT_name context = coinduct_extras_of_generic context fpT_name |> maps (#cong_intro_pairs #> map snd); fun get_cong_intros fpT_name name context = coinduct_extras_of_generic context fpT_name |> map_filter (#cong_intro_pairs #> (fn ps => AList.lookup (op =) ps name)); fun ctr_names_of_fp_name lthy fpT_name = fpT_name |> fp_sugar_of lthy |> the |> #fp_ctr_sugar |> #ctr_sugar |> #ctrs |> map (Long_Name.base_name o name_of_ctr); fun register_coinduct_dynamic_base fpT_name lthy = let val fp_b = Binding.name (Long_Name.base_name fpT_name) in lthy |> fold Local_Theory.add_thms_dynamic ((mk_fp_binding fp_b coinduct_uptoN, get_coinduct_uptos fpT_name) :: map (fn N => let val N = cong_N ^ N in (mk_fp_binding fp_b N, get_cong_intros fpT_name N) end) ([baseN, reflN, symN, transN] @ ctr_names_of_fp_name lthy fpT_name)) |> Local_Theory.add_thms_dynamic (mk_fp_binding fp_b cong_introsN, get_cong_all_intros fpT_name) end; fun register_coinduct_dynamic_friend fpT_name friend_name = let val fp_b = Binding.name (Long_Name.base_name fpT_name); val friend_base_name = cong_N ^ Long_Name.base_name friend_name; in Local_Theory.add_thms_dynamic (mk_fp_binding fp_b friend_base_name, get_cong_intros fpT_name friend_base_name) end; fun derive_case_dtor ctxt fpT_name = let val thy = Proof_Context.theory_of ctxt; val SOME ({fp_res_index, fp_res = {dtors = dtors0, dtor_ctors, ...}, absT_info = {rep = rep0, abs_inverse, ...}, fp_ctr_sugar = {ctr_defs, ctr_sugar = {casex, exhaust, case_thms, ...}, ...}, ...}) = fp_sugar_of ctxt fpT_name; val (f_Ts, Type (_, [fpT as Type (_, As), _])) = strip_fun_type (fastype_of casex); val x_Tss = map binder_types f_Ts; val (((u, fs), xss), _) = ctxt |> yield_singleton (mk_Frees "y") fpT ||>> mk_Frees "f" f_Ts ||>> mk_Freess "x" x_Tss; val dtor0 = nth dtors0 fp_res_index; val dtor = mk_dtor As dtor0; val u' = dtor $ u; val absT = fastype_of u'; val rep = mk_rep absT rep0; val goal = mk_Trueprop_eq (list_comb (casex, fs) $ u, mk_case_absumprod absT rep fs xss xss $ u') |> Raw_Simplifier.rewrite_term thy @{thms comp_def[THEN eq_reflection]} []; val vars = map (fst o dest_Free) (u :: fs); val dtor_ctor = nth dtor_ctors fp_res_index; in Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_case_dtor_tac ctxt u abs_inverse dtor_ctor ctr_defs exhaust case_thms) |> Thm.close_derivation \<^here> end; fun derive_case_trivial ctxt fpT_name = let val SOME {casex, exhaust, case_thms, ...} = ctr_sugar_of ctxt fpT_name; val Type (_, As0) = domain_type (body_fun_type (fastype_of casex)); val (As, _) = ctxt |> mk_TFrees' (map Type.sort_of_atyp As0); val fpT = Type (fpT_name, As); val (var_name, ()) = singleton (Variable.variant_frees ctxt []) ("x", ()); val var = Free (var_name, fpT); val goal = mk_Trueprop_eq (expand_to_ctr_term ctxt fpT var, var); val exhaust' = infer_instantiate' ctxt [SOME (Thm.cterm_of ctxt var)] exhaust; in Goal.prove_sorry ctxt [var_name] [] goal (fn {context = ctxt, prems = _} => HEADGOAL (rtac ctxt exhaust') THEN ALLGOALS (hyp_subst_tac ctxt) THEN unfold_thms_tac ctxt case_thms THEN ALLGOALS (rtac ctxt refl)) |> Thm.close_derivation \<^here> end; fun mk_abs_rep_transfers ctxt fpT_name = [mk_abs_transfer ctxt fpT_name, mk_rep_transfer ctxt fpT_name] handle Fail _ => []; fun ensure_codatatype_extra fpT_name ctxt = (case codatatype_extra_of ctxt fpT_name of NONE => let val abs_rep_transfers = mk_abs_rep_transfers ctxt fpT_name in ctxt |> register_codatatype_extra fpT_name {case_dtor = derive_case_dtor ctxt fpT_name, case_trivial = derive_case_trivial ctxt fpT_name, abs_rep_transfers = abs_rep_transfers} |> set_transfer_rule_attrs abs_rep_transfers end | SOME {abs_rep_transfers, ...} => ctxt |> set_transfer_rule_attrs abs_rep_transfers); fun setup_base fpT_name = register_coinduct_dynamic_base fpT_name #> ensure_codatatype_extra fpT_name; fun is_set ctxt (const_name, T) = (case T of Type (\<^type_name>\fun\, [Type (fpT_name, _), Type (\<^type_name>\set\, [_])]) => (case bnf_of ctxt fpT_name of SOME bnf => exists (fn Const (s, _) => s = const_name | _ => false) (sets_of_bnf bnf) | NONE => false) | _ => false); fun case_eq_if_thms_of_term ctxt t = let val ctr_sugars = map_filter (ctr_sugar_of_case ctxt o fst) (Term.add_consts t []) in maps #case_eq_ifs ctr_sugars end; fun all_algrho_eqs_of ctxt = maps (#algrho_eqs o snd) (all_friend_extras_of ctxt); fun derive_code ctxt inner_fp_simps goal {sig_fp_sugars, ssig_fp_sugar, eval, eval_simps, all_algLam_algs, corecUU_thm, ...} fun_t fun_def = let val fun_T = fastype_of fun_t; val (arg_Ts, Type (fpT_name, _)) = strip_type fun_T; val num_args = length arg_Ts; val SOME {pre_bnf, fp_bnf, absT_info, fp_nesting_bnfs, live_nesting_bnfs, fp_ctr_sugar, ...} = fp_sugar_of ctxt fpT_name; val SOME {case_trivial, ...} = codatatype_extra_of ctxt fpT_name; val ctr_sugar = #ctr_sugar fp_ctr_sugar; val pre_map_def = map_def_of_bnf pre_bnf; val abs_inverse = #abs_inverse absT_info; val ctr_defs = #ctr_defs fp_ctr_sugar; val case_eq_ifs = #case_eq_ifs ctr_sugar @ case_eq_if_thms_of_term ctxt goal; val all_sig_map_thms = maps (#map_thms o #fp_bnf_sugar) sig_fp_sugars; val fp_map_ident = map_ident_of_bnf fp_bnf; val fpsig_nesting_bnfs = fp_nesting_bnfs @ maps #live_nesting_bnfs sig_fp_sugars; val fpsig_nesting_T_names = map (fst o dest_Type o T_of_bnf) fpsig_nesting_bnfs; val fpsig_nesting_fp_sugars = map_filter (fp_sugar_of ctxt) fpsig_nesting_T_names; val fpsig_nesting_fp_bnf_sugars = map #fp_bnf_sugar fpsig_nesting_fp_sugars; val ssig_fp_bnf_sugar = #fp_bnf_sugar ssig_fp_sugar; val ssig_bnf = #fp_bnf ssig_fp_sugar; val ssig_map = map_of_bnf ssig_bnf; val fpsig_nesting_maps = map map_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_ident0s = map map_ident0_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_comps = map map_comp_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_thms = maps #map_thms fpsig_nesting_fp_bnf_sugars; val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs; val ssig_map_thms = #map_thms ssig_fp_bnf_sugar; val all_algLam_alg_pointfuls = map (mk_pointful ctxt) all_algLam_algs; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_code_tac ctxt num_args fpsig_nesting_maps ssig_map eval pre_map_def abs_inverse fpsig_nesting_map_ident0s fpsig_nesting_map_comps fpsig_nesting_map_thms live_nesting_map_ident0s fp_map_ident case_trivial ctr_defs case_eq_ifs corecUU_thm all_sig_map_thms ssig_map_thms all_algLam_alg_pointfuls (all_algrho_eqs_of ctxt) eval_simps inner_fp_simps fun_def)) |> Thm.close_derivation \<^here> end; fun derive_unique ctxt phi code_goal {sig_fp_sugars, ssig_fp_sugar, eval, eval_simps, all_algLam_algs, corecUU_unique, ...} fpT_name eq_corecUU = let val SOME {pre_bnf, fp_bnf, absT_info, fp_nesting_bnfs, live_nesting_bnfs, fp_ctr_sugar, ...} = fp_sugar_of ctxt fpT_name; val SOME {case_trivial, ...} = codatatype_extra_of ctxt fpT_name; val ctr_sugar = #ctr_sugar fp_ctr_sugar; val pre_map_def = map_def_of_bnf pre_bnf; val abs_inverse = #abs_inverse absT_info; val ctr_defs = #ctr_defs fp_ctr_sugar; val case_eq_ifs = #case_eq_ifs ctr_sugar @ case_eq_if_thms_of_term ctxt code_goal; val all_sig_map_thms = maps (#map_thms o #fp_bnf_sugar) sig_fp_sugars; val fp_map_ident = map_ident_of_bnf fp_bnf; val fpsig_nesting_bnfs = fp_nesting_bnfs @ maps #live_nesting_bnfs sig_fp_sugars; val fpsig_nesting_T_names = map (fst o dest_Type o T_of_bnf) fpsig_nesting_bnfs; val fpsig_nesting_fp_sugars = map_filter (fp_sugar_of ctxt) fpsig_nesting_T_names; val fpsig_nesting_fp_bnf_sugars = map #fp_bnf_sugar fpsig_nesting_fp_sugars; val ssig_fp_bnf_sugar = #fp_bnf_sugar ssig_fp_sugar; val ssig_bnf = #fp_bnf ssig_fp_sugar; val ssig_map = map_of_bnf ssig_bnf; val fpsig_nesting_maps = map map_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_ident0s = map map_ident0_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_comps = map map_comp_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_thms = maps #map_thms fpsig_nesting_fp_bnf_sugars; val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs; val ssig_map_thms = #map_thms ssig_fp_bnf_sugar; val all_algLam_alg_pointfuls = map (mk_pointful ctxt) all_algLam_algs; val \<^const>\Trueprop\ $ (Const (\<^const_name>\HOL.eq\, _) $ lhs $ rhs) = code_goal; val (fun_t, args) = strip_comb lhs; val closed_rhs = fold_rev lambda args rhs; val fun_T = fastype_of fun_t; val num_args = num_binder_types fun_T; val f = Free (singleton (Variable.variant_frees ctxt []) ("f", fun_T)); val is_self_call = curry (op aconv) fun_t; val has_self_call = exists_subterm is_self_call; fun fify args (t $ u) = fify (u :: args) t $ fify [] u | fify _ (Abs (s, T, t)) = Abs (s, T, fify [] t) | fify args t = if t = fun_t andalso not (exists has_self_call args) then f else t; val goal = Logic.mk_implies (mk_Trueprop_eq (f, fify [] closed_rhs), mk_Trueprop_eq (f, fun_t)) |> Morphism.term phi; in Goal.prove_sorry ctxt [fst (dest_Free f)] [] goal (fn {context = ctxt, prems = _} => mk_unique_tac ctxt num_args fpsig_nesting_maps ssig_map eval pre_map_def abs_inverse fpsig_nesting_map_ident0s fpsig_nesting_map_comps fpsig_nesting_map_thms live_nesting_map_ident0s fp_map_ident case_trivial ctr_defs case_eq_ifs all_sig_map_thms ssig_map_thms all_algLam_alg_pointfuls (all_algrho_eqs_of ctxt) eval_simps corecUU_unique eq_corecUU) |> Thm.close_derivation \<^here> end; fun derive_last_disc ctxt fcT_name = let val SOME {T = fcT, discs, exhaust, disc_thmss, ...} = ctr_sugar_of ctxt fcT_name; val (u, _) = ctxt |> yield_singleton (mk_Frees "x") fcT; val udiscs = map (rapp u) discs; val (not_udiscs, last_udisc) = split_last udiscs |>> map HOLogic.mk_not; val goal = mk_Trueprop_eq (last_udisc, foldr1 HOLogic.mk_conj not_udiscs); in Goal.prove_sorry ctxt [fst (dest_Free u)] [] goal (fn {context = ctxt, prems = _} => mk_last_disc_tac ctxt u exhaust (flat disc_thmss)) |> Thm.close_derivation \<^here> end; fun derive_eq_algrho ctxt {sig_fp_sugars, ssig_fp_sugar, eval, eval_simps, all_algLam_algs, corecUU_unique, ...} ({algrho = algrho0, dtor_algrho, ...} : friend_info) fun_t k_T code_goal const_transfers rho_def eq_corecUU = let val fun_T = fastype_of fun_t; val (arg_Ts, Type (fpT_name, Ts)) = strip_type fun_T; val num_args = length arg_Ts; val SOME {fp_res_index, fp_res, pre_bnf, fp_bnf, absT_info, fp_nesting_bnfs, live_nesting_bnfs, fp_ctr_sugar, ...} = fp_sugar_of ctxt fpT_name; val SOME {case_dtor, ...} = codatatype_extra_of ctxt fpT_name; val fp_nesting_Ts = map T_of_bnf fp_nesting_bnfs; fun is_nullary_disc_def (\<^const>\Trueprop\ $ (Const (\<^const_name>\HOL.eq\, _) $ _ $ (Const (\<^const_name>\HOL.eq\, _) $ _ $ _))) = true | is_nullary_disc_def (Const (\<^const_name>\Pure.eq\, _) $ _ $ (Const (\<^const_name>\HOL.eq\, _) $ _ $ _)) = true | is_nullary_disc_def _ = false; val dtor_ctor = nth (#dtor_ctors fp_res) fp_res_index; val ctor_iff_dtor = #ctor_iff_dtor fp_ctr_sugar; val ctr_sugar = #ctr_sugar fp_ctr_sugar; val pre_map_def = map_def_of_bnf pre_bnf; val abs_inverse = #abs_inverse absT_info; val ctr_defs = #ctr_defs fp_ctr_sugar; val nullary_disc_defs = filter (is_nullary_disc_def o Thm.prop_of) (#disc_defs ctr_sugar); val disc_sel_eq_cases = #disc_eq_cases ctr_sugar @ #sel_defs ctr_sugar; val case_eq_ifs = #case_eq_ifs ctr_sugar @ case_eq_if_thms_of_term ctxt code_goal; val all_sig_map_thms = maps (#map_thms o #fp_bnf_sugar) sig_fp_sugars; fun add_tnameT (Type (s, Ts)) = insert (op =) s #> fold add_tnameT Ts | add_tnameT _ = I; fun map_disc_sels'_of s = (case fp_sugar_of ctxt s of SOME {fp_bnf_sugar = {map_disc_iffs, map_selss, ...}, ...} => let val map_selss' = if length map_selss <= 1 then map_selss else map (map (unfold_thms ctxt (no_refl [derive_last_disc ctxt s]))) map_selss; in map_disc_iffs @ flat map_selss' end | NONE => []); fun mk_const_pointful_natural const_transfer = SOME (mk_pointful_natural_from_transfer ctxt const_transfer) handle UNNATURAL () => NONE; val const_pointful_natural_opts = map mk_const_pointful_natural const_transfers; val const_pointful_naturals = map_filter I const_pointful_natural_opts; val fp_nesting_k_T_names = fold add_tnameT (k_T :: fp_nesting_Ts) []; val fp_nesting_k_map_disc_sels' = maps map_disc_sels'_of fp_nesting_k_T_names; val fp_map_ident = map_ident_of_bnf fp_bnf; val fpsig_nesting_bnfs = fp_nesting_bnfs @ maps #live_nesting_bnfs sig_fp_sugars; val fpsig_nesting_T_names = map (fst o dest_Type o T_of_bnf) fpsig_nesting_bnfs; val fpsig_nesting_fp_sugars = map_filter (fp_sugar_of ctxt) fpsig_nesting_T_names; val fpsig_nesting_fp_bnf_sugars = map #fp_bnf_sugar fpsig_nesting_fp_sugars; val ssig_fp_bnf_sugar = #fp_bnf_sugar ssig_fp_sugar; val ssig_bnf = #fp_bnf ssig_fp_sugar; val ssig_map = map_of_bnf ssig_bnf; val fpsig_nesting_maps = map map_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_ident0s = map map_ident0_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_comps = map map_comp_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_thms = maps #map_thms fpsig_nesting_fp_bnf_sugars; val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs; val ssig_map_thms = #map_thms ssig_fp_bnf_sugar; val all_algLam_alg_pointfuls = map (mk_pointful ctxt) all_algLam_algs; val ctor = nth (#ctors fp_res) fp_res_index; val abs = #abs absT_info; val rep = #rep absT_info; val algrho = mk_ctr Ts algrho0; val goal = mk_Trueprop_eq (fun_t, abs_curried_balanced arg_Ts algrho); fun const_of_transfer thm = (case Thm.prop_of thm of \<^const>\Trueprop\ $ (_ $ cst $ _) => cst); val eq_algrho = Goal.prove (*no sorry*) ctxt [] [] goal (fn {context = ctxt, prems = _} => mk_eq_algrho_tac ctxt fpsig_nesting_maps abs rep ctor ssig_map eval pre_map_def abs_inverse fpsig_nesting_map_ident0s fpsig_nesting_map_comps fpsig_nesting_map_thms live_nesting_map_ident0s fp_map_ident dtor_ctor ctor_iff_dtor ctr_defs nullary_disc_defs disc_sel_eq_cases case_dtor case_eq_ifs const_pointful_naturals fp_nesting_k_map_disc_sels' rho_def dtor_algrho corecUU_unique eq_corecUU all_sig_map_thms ssig_map_thms all_algLam_alg_pointfuls (all_algrho_eqs_of ctxt) eval_simps) |> Thm.close_derivation \<^here> handle e as ERROR _ => (case filter (is_none o snd) (const_transfers ~~ const_pointful_natural_opts) of [] => Exn.reraise e | thm_nones => error ("Failed to state naturality property for " ^ commas (map (Syntax.string_of_term ctxt o const_of_transfer o fst) thm_nones))); val algrho_eq = eq_algrho RS (mk_curry_uncurryN_balanced ctxt num_args RS iffD2) RS sym; in (eq_algrho, algrho_eq) end; fun prime_rho_transfer_goal ctxt fpT_name rho_def goal = let val thy = Proof_Context.theory_of ctxt; val SOME {pre_bnf, ...} = fp_sugar_of ctxt fpT_name; val SOME {abs_rep_transfers, ...} = codatatype_extra_of ctxt fpT_name; val simps = rel_def_of_bnf pre_bnf :: rho_transfer_simps; val fold_rho = unfold_thms ctxt [rho_def RS @{thm symmetric}]; fun derive_unprimed rho_transfer' = Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt simps THEN HEADGOAL (rtac ctxt rho_transfer'))) |> Thm.close_derivation \<^here>; val goal' = Raw_Simplifier.rewrite_term thy simps [] goal; in if null abs_rep_transfers then (goal', derive_unprimed #> fold_rho) else (goal, fold_rho) end; fun derive_rho_transfer_folded ctxt fpT_name const_transfers rho_def goal = let val SOME {pre_bnf, ...} = fp_sugar_of ctxt fpT_name; val SOME {abs_rep_transfers, ...} = codatatype_extra_of ctxt fpT_name; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_rho_transfer_tac ctxt (null abs_rep_transfers) (rel_def_of_bnf pre_bnf) const_transfers)) |> unfold_thms ctxt [rho_def RS @{thm symmetric}] |> Thm.close_derivation \<^here> end; fun mk_cong_intro_ctr_or_friend_goal ctxt fpT Rcong alg = let val xy_Ts = binder_types (fastype_of alg); val ((xs, ys), _) = ctxt |> mk_Frees "x" xy_Ts ||>> mk_Frees "y" xy_Ts; fun mk_prem xy_T x y = build_rel [] ctxt [fpT] [] (fn (T, _) => if T = fpT then Rcong else HOLogic.eq_const T) (xy_T, xy_T) $ x $ y; val prems = @{map 3} mk_prem xy_Ts xs ys; val concl = Rcong $ list_comb (alg, xs) $ list_comb (alg, ys); in Logic.list_implies (map HOLogic.mk_Trueprop prems, HOLogic.mk_Trueprop concl) end; fun derive_cong_ctr_intros ctxt cong_ctor_intro = let val \<^const>\Pure.imp\ $ _ $ (\<^const>\Trueprop\ $ ((Rcong as _ $ _) $ _ $ (ctor $ _))) = Thm.prop_of cong_ctor_intro; val fpT as Type (fpT_name, fp_argTs) = range_type (fastype_of ctor); val SOME {pre_bnf, absT_info = {abs_inverse, ...}, fp_ctr_sugar = {ctr_defs, ctr_sugar = {ctrs = ctrs0, ...}, ...}, ...} = fp_sugar_of ctxt fpT_name; val ctrs = map (mk_ctr fp_argTs) ctrs0; val pre_rel_def = rel_def_of_bnf pre_bnf; fun prove ctr_def goal = Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_cong_intro_ctr_or_friend_tac ctxt ctr_def [pre_rel_def, abs_inverse] cong_ctor_intro)) |> Thm.close_derivation \<^here>; val goals = map (mk_cong_intro_ctr_or_friend_goal ctxt fpT Rcong) ctrs; in map2 prove ctr_defs goals end; fun derive_cong_friend_intro ctxt cong_algrho_intro = let val \<^const>\Pure.imp\ $ _ $ (\<^const>\Trueprop\ $ ((Rcong as _ $ _) $ _ $ ((algrho as Const (algrho_name, _)) $ _))) = Thm.prop_of cong_algrho_intro; val fpT as Type (_, fp_argTs) = range_type (fastype_of algrho); fun has_algrho (\<^const>\Trueprop\ $ (Const (\<^const_name>\HOL.eq\, _) $ _ $ rhs)) = fst (dest_Const (head_of (strip_abs_body rhs))) = algrho_name; val eq_algrho :: _ = maps (filter (has_algrho o Thm.prop_of) o #eq_algrhos o snd) (all_friend_extras_of ctxt); val \<^const>\Trueprop\ $ (Const (\<^const_name>\HOL.eq\, _) $ friend0 $ _) = Thm.prop_of eq_algrho; val friend = mk_ctr fp_argTs friend0; val goal = mk_cong_intro_ctr_or_friend_goal ctxt fpT Rcong friend; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_cong_intro_ctr_or_friend_tac ctxt eq_algrho [] cong_algrho_intro)) |> Thm.close_derivation \<^here> end; fun derive_cong_intros lthy ctr_names friend_names ({cong_base, cong_refl, cong_sym, cong_trans, cong_alg_intros, ...} : dtor_coinduct_info) = let val cong_ctor_intro :: cong_algrho_intros = rev cong_alg_intros; val names = map (prefix cong_N) ([baseN, reflN, symN, transN] @ ctr_names @ friend_names); val thms = [cong_base, cong_refl, cong_sym, cong_trans] @ derive_cong_ctr_intros lthy cong_ctor_intro @ map (derive_cong_friend_intro lthy) cong_algrho_intros; in names ~~ thms end; fun derive_coinduct ctxt (fpT as Type (fpT_name, fpT_args)) dtor_coinduct = let val thy = Proof_Context.theory_of ctxt; val \<^const>\Pure.imp\ $ (\<^const>\Trueprop\ $ (_ $ Abs (_, _, _ $ Abs (_, _, \<^const>\implies\ $ _ $ (_ $ (cong0 $ _) $ _ $ _))))) $ _ = Thm.prop_of dtor_coinduct; val SOME {X as TVar ((X_s, _), _), fp_res = {dtor_ctors, ...}, pre_bnf, absT_info = {abs_inverse, ...}, live_nesting_bnfs, fp_ctr_sugar = {ctrXs_Tss, ctr_defs, ctr_sugar = ctr_sugar0 as {T = Type (_, T0_args), ctrs = ctrs0, discs = discs0, ...}, ...}, ...} = fp_sugar_of ctxt fpT_name; val n = length ctrXs_Tss; val ms = map length ctrXs_Tss; val X' = TVar ((X_s, maxidx_of_typ fpT + 1), \<^sort>\type\); val As_rho = tvar_subst thy T0_args fpT_args; val substXAT = Term.typ_subst_TVars As_rho o Tsubst X X'; val substXA = Term.subst_TVars As_rho o substT X X'; val phi = Morphism.typ_morphism "BNF" substXAT $> Morphism.term_morphism "BNF" substXA; fun mk_applied_cong arg = enforce_type ctxt domain_type (fastype_of arg) cong0 $ arg; val thm = derive_coinduct_thms_for_types ctxt false mk_applied_cong [pre_bnf] dtor_coinduct dtor_ctors live_nesting_bnfs [fpT] [substXAT X] [map (map substXAT) ctrXs_Tss] [n] [abs_inverse] [abs_inverse] I [ctr_defs] [morph_ctr_sugar phi ctr_sugar0] |> map snd |> the_single; val (attrs, _) = mk_coinduct_attrs [fpT] [ctrs0] [discs0] [ms]; in (thm, attrs) end; type explore_parameters = {bound_Us: typ list, bound_Ts: typ list, U: typ, T: typ}; fun update_UT {bound_Us, bound_Ts, ...} U T = {bound_Us = bound_Us, bound_Ts = bound_Ts, U = U, T = T}; fun explore_nested lthy explore {bound_Us, bound_Ts, U, T} t = let fun build_simple (T, U) = if T = U then \<^term>\%y. y\ else Bound 0 |> explore {bound_Us = T :: bound_Us, bound_Ts = T :: bound_Ts, U = U, T = T} |> (fn t => Abs (Name.uu, T, t)); in betapply (build_map lthy [] [] build_simple (T, U), t) end; fun add_boundvar t = betapply (incr_boundvars 1 t, Bound 0); fun explore_fun (arg_U :: arg_Us) explore {bound_Us, bound_Ts, U, T} t = let val arg_name = the_default Name.uu (try (fn (Abs (s, _, _)) => s) t) in add_boundvar t |> explore_fun arg_Us explore {bound_Us = arg_U :: bound_Us, bound_Ts = domain_type T :: bound_Ts, U = range_type U, T = range_type T} |> (fn t => Abs (arg_name, arg_U, t)) end | explore_fun [] explore params t = explore params t; fun massage_fun explore (params as {T, U, ...}) = if can dest_funT T then explore_fun [domain_type U] explore params else explore params; fun massage_star massages explore = let fun after_massage massages' t params t' = if t aconv t' then massage_any massages' params t else massage_any massages params t' and massage_any [] params t = explore params t | massage_any (massage :: massages') params t = massage (after_massage massages' t) params t; in massage_any massages end; fun massage_let explore params t = (case strip_comb t of (Const (\<^const_name>\Let\, _), [_, _]) => unfold_lets_splits t | _ => t) |> explore params; fun check_corec_equation ctxt fun_frees (lhs, rhs) = let val (fun_t, arg_ts) = strip_comb lhs; fun check_fun_name () = null fun_frees orelse member (op aconv) fun_frees fun_t orelse ill_formed_equation_head ctxt [] fun_t; fun check_no_other_frees () = (case Term.add_frees rhs [] |> map Free |> subtract (op =) (fun_frees @ arg_ts) |> find_first (not o Variable.is_fixed ctxt o fst o dest_Free) of NONE => () | SOME t => extra_variable_in_rhs ctxt [] t); in check_duplicate_variables_in_lhs ctxt [] arg_ts; check_fun_name (); check_all_fun_arg_frees ctxt [] (filter_out is_Var arg_ts); check_no_other_frees () end; fun parse_corec_equation ctxt fun_frees eq = let val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop (drop_all eq)) handle TERM _ => ill_formed_equation_lhs_rhs ctxt [eq]; val _ = check_corec_equation ctxt fun_frees (lhs, rhs); val (fun_t, arg_ts) = strip_comb lhs; val (arg_Ts, _) = strip_type (fastype_of fun_t); val added_Ts = drop (length arg_ts) arg_Ts; val free_names = mk_names (length added_Ts) "x" ~~ added_Ts; val free_args = Variable.variant_frees ctxt [rhs, lhs] free_names |> map Free; in (arg_ts @ free_args, list_comb (rhs, free_args)) end; fun morph_views phi (code, ctrs, discs, disc_iffs, sels) = (Morphism.term phi code, map (Morphism.term phi) ctrs, map (Morphism.term phi) discs, map (Morphism.term phi) disc_iffs, map (Morphism.term phi) sels); fun generate_views ctxt eq fun_t (lhs_free_args, rhs) = let val lhs = list_comb (fun_t, lhs_free_args); val T as Type (T_name, Ts) = fastype_of rhs; val SOME {fp_ctr_sugar = {ctr_sugar = {ctrs = ctrs0, discs = discs0, selss = selss0, ...}, ...}, ...} = fp_sugar_of ctxt T_name; val ctrs = map (mk_ctr Ts) ctrs0; val discs = map (mk_disc_or_sel Ts) discs0; val selss = map (map (mk_disc_or_sel Ts)) selss0; val code_view = drop_all eq; fun can_case_expand t = not (can (dest_ctr ctxt T_name) t); fun generate_raw_views conds t raw_views = let fun analyse (ctr :: ctrs) (disc :: discs) ctr' = if ctr = ctr' then (conds, disc, ctr) else analyse ctrs discs ctr'; in (analyse ctrs discs (fst (strip_comb t))) :: raw_views end; fun generate_disc_views raw_views = if length discs = 1 then ([], []) else let fun collect_condss_disc condss [] _ = condss | collect_condss_disc condss ((conds, disc', _) :: raw_views) disc = collect_condss_disc (condss |> disc = disc' ? cons conds) raw_views disc; val grouped_disc_views = discs |> map (collect_condss_disc [] raw_views) |> curry (op ~~) (map (fn disc => disc $ lhs) discs); fun mk_disc_iff_props props [] = props | mk_disc_iff_props _ ((lhs, \<^const>\HOL.True\) :: _) = [lhs] | mk_disc_iff_props props ((lhs, rhs) :: views) = mk_disc_iff_props ((HOLogic.mk_eq (lhs, rhs)) :: props) views; in (grouped_disc_views |> map swap, grouped_disc_views |> map (apsnd (s_dnf #> mk_conjs)) |> mk_disc_iff_props [] |> map (fn eq => ([[]], eq))) end; fun generate_ctr_views raw_views = let fun collect_condss_ctr condss [] _ = condss | collect_condss_ctr condss ((conds, _, ctr') :: raw_views) ctr = collect_condss_ctr (condss |> ctr = ctr' ? cons conds) raw_views ctr; fun mk_ctr_eq ctr_sels ctr = let fun extract_arg n sel _(*bound_Ts*) fun_t arg_ts = if ctr = fun_t then nth arg_ts n else let val t = list_comb (fun_t, arg_ts) in if can_case_expand t then sel $ t else Term.dummy_pattern (range_type (fastype_of sel)) end; in ctr_sels |> map_index (uncurry extract_arg) |> map (fn extract => massage_corec_code_rhs ctxt extract [] rhs) |> curry list_comb ctr |> curry HOLogic.mk_eq lhs end; fun remove_condss_if_alone [(_, concl)] = [([[]], concl)] | remove_condss_if_alone views = views; in ctrs |> `(map (collect_condss_ctr [] raw_views)) ||> map2 mk_ctr_eq selss |> op ~~ |> filter_out (null o fst) |> remove_condss_if_alone end; fun generate_sel_views raw_views only_one_ctr = let fun mk_sel_positions sel = let fun get_sel_position _ [] = NONE | get_sel_position i (sel' :: sels) = if sel = sel' then SOME i else get_sel_position (i + 1) sels; in ctrs ~~ map (get_sel_position 0) selss |> map_filter (fn (ctr, pos_opt) => if is_some pos_opt then SOME (ctr, the pos_opt) else NONE) end; fun collect_sel_condss0 condss [] _ = condss | collect_sel_condss0 condss ((conds, _, ctr) :: raw_views) sel_positions = let val condss' = condss |> is_some (AList.lookup (op =) sel_positions ctr) ? cons conds in collect_sel_condss0 condss' raw_views sel_positions end; val collect_sel_condss = if only_one_ctr then K [[]] else collect_sel_condss0 [] raw_views; fun mk_sel_rhs sel_positions sel = let val sel_T = range_type (fastype_of sel); fun extract_sel_value _(*bound_Ts*) fun_t arg_ts = (case AList.lookup (op =) sel_positions fun_t of SOME n => nth arg_ts n | NONE => let val t = list_comb (fun_t, arg_ts) in if can_case_expand t then sel $ t else Term.dummy_pattern sel_T end); in massage_corec_code_rhs ctxt extract_sel_value [] rhs end; val ordered_sels = distinct (op =) (flat selss); val sel_positionss = map mk_sel_positions ordered_sels; val sel_rhss = map2 mk_sel_rhs sel_positionss ordered_sels; val sel_lhss = map (rapp lhs o mk_disc_or_sel Ts) ordered_sels; val sel_condss = map collect_sel_condss sel_positionss; fun is_undefined (Const (\<^const_name>\undefined\, _)) = true | is_undefined _ = false; in sel_condss ~~ (sel_lhss ~~ sel_rhss) |> filter_out (is_undefined o snd o snd) |> map (apsnd HOLogic.mk_eq) end; fun mk_atomic_prop fun_args (condss, concl) = (Logic.list_all (map dest_Free fun_args, abstract_over_list fun_args (Logic.list_implies (map HOLogic.mk_Trueprop (s_dnf condss), HOLogic.mk_Trueprop concl)))); val raw_views = rhs |> massage_let_if_case ctxt (K false) (fn _(*bound_Ts*) => fn t => t |> can_case_expand t ? expand_to_ctr_term ctxt T) (K (K ())) (K I) [] |> (fn expanded_rhs => fold_rev_let_if_case ctxt generate_raw_views [] expanded_rhs []) |> rev; val (disc_views, disc_iff_views) = generate_disc_views raw_views; val ctr_views = generate_ctr_views raw_views; val sel_views = generate_sel_views raw_views (length ctr_views = 1); val mk_props = filter_out (null o fst) #> map (mk_atomic_prop lhs_free_args); in (code_view, mk_props ctr_views, mk_props disc_views, mk_props disc_iff_views, mk_props sel_views) end; fun find_all_associated_types [] _ = [] | find_all_associated_types ((Type (_, Ts1), Type (_, Ts2)) :: TTs) T = find_all_associated_types ((Ts1 ~~ Ts2) @ TTs) T | find_all_associated_types ((T1, T2) :: TTs) T = find_all_associated_types TTs T |> T1 = T ? cons T2; fun as_member_of tab = try dest_Const #> Option.mapPartial (fst #> Symtab.lookup tab); fun extract_rho_from_equation ({ctr_guards, inner_buffer = {Oper, VLeaf, CLeaf, ctr_wrapper, friends}, ...}, {pattern_ctrs, discs, sels, it, mk_case}) b version Y preT ssig_T friend_tm (lhs_args, rhs) lthy = let val thy = Proof_Context.theory_of lthy; val res_T = fastype_of rhs; val YpreT = HOLogic.mk_prodT (Y, preT); fun fpT_to new_T T = if T = res_T then new_T else (case T of Type (s, Ts) => Type (s, map (fpT_to new_T) Ts) | _ => T); fun build_params bound_Us bound_Ts T = {bound_Us = bound_Us, bound_Ts = bound_Ts, U = T, T = T}; fun typ_before explore {bound_Us, bound_Ts, ...} t = explore (build_params bound_Us bound_Ts (fastype_of1 (bound_Ts, t))) t; val is_self_call = curry (op aconv) friend_tm; val has_self_call = Term.exists_subterm is_self_call; fun has_res_T bound_Ts t = fastype_of1 (bound_Ts, t) = res_T; fun contains_res_T (Type (s, Ts)) = s = fst (dest_Type res_T) orelse exists contains_res_T Ts | contains_res_T _ = false; val is_lhs_arg = member (op =) lhs_args; fun is_constant t = not (Term.exists_subterm is_lhs_arg t orelse has_self_call t orelse loose_bvar (t, 0)); fun is_nested_type T = T <> res_T andalso T <> YpreT andalso T <> ssig_T; val is_valid_case_argumentT = not o member (op =) [Y, ssig_T]; fun is_same_type_constr (Type (s, _)) (Type (s', _)) = (s = s') | is_same_type_constr _ _ = false; exception NO_ENCAPSULATION of unit; val parametric_consts = Unsynchronized.ref []; (* We are assuming that set functions are marked with "[transfer_rule]" (cf. the "transfer" plugin). Otherwise, the "eq_algrho" tactic might fail. *) fun is_special_parametric_const (x as (s, _)) = s = \<^const_name>\id\ orelse is_set lthy x; fun add_parametric_const s general_T T U = let fun tupleT_of_funT T = let val (Ts, T) = strip_type T in mk_tupleT_balanced (Ts @ [T]) end; fun funT_of_tupleT n = dest_tupleT_balanced (n + 1) #> split_last #> op --->; val m = num_binder_types general_T; val param1_T = Type_Infer.paramify_vars general_T; val param2_T = Type_Infer.paramify_vars param1_T; val deadfixed_T = build_map lthy [] [] (mk_undefined o op -->) (apply2 tupleT_of_funT (param1_T, param2_T)) |> singleton (Type_Infer_Context.infer_types lthy) |> singleton (Type_Infer.fixate lthy false) |> type_of |> dest_funT |-> generalize_types 1 |> funT_of_tupleT m; val j = maxidx_of_typ deadfixed_T + 1; fun varifyT (Type (s, Ts)) = Type (s, map varifyT Ts) | varifyT (TFree (s, T)) = TVar ((s, j), T) | varifyT T = T; val dedvarified_T = varifyT deadfixed_T; val new_vars = Sign.typ_match thy (dedvarified_T, T) Vartab.empty |> Vartab.dest |> filter (curry (op =) j o snd o fst) |> Vartab.make; val deadinstantiated_T = map_atyps (Type.devar new_vars) dedvarified_T; val final_T = if Sign.typ_instance thy (U, deadinstantiated_T) then deadfixed_T else general_T; in parametric_consts := insert (op =) (s, final_T) (!parametric_consts) end; fun encapsulate (params as {U, T, ...}) t = if U = T then t else if T = Y then VLeaf $ t else if T = res_T then CLeaf $ t else if T = YpreT then it $ t else if is_nested_type T andalso is_same_type_constr T U then explore_nested lthy encapsulate params t else raise NO_ENCAPSULATION (); fun build_function_after_encapsulation fun_t fun_t' (params as {bound_Us, ...}) arg_ts arg_ts' = let val arg_Us' = fst (strip_typeN (length arg_ts) (fastype_of1 (bound_Us, fun_t'))); fun the_or_error arg NONE = error ("Illegal argument " ^ quote (Syntax.string_of_term lthy arg) ^ " to " ^ quote (Syntax.string_of_term lthy fun_t)) | the_or_error _ (SOME arg') = arg'; in arg_ts' |> `(map (curry fastype_of1 bound_Us)) |>> map2 (update_UT params) arg_Us' |> op ~~ |> map (try (uncurry encapsulate)) |> map2 the_or_error arg_ts |> curry list_comb fun_t' end; fun rebuild_function_after_exploration old_fn new_fn explore params arg_ts = arg_ts |> map (typ_before explore params) |> build_function_after_encapsulation old_fn new_fn params arg_ts; fun update_case Us U casex = let val Type (T_name, _) = domain_type (snd (strip_fun_type (fastype_of casex))); val SOME {fp_ctr_sugar = {ctr_sugar = {T = Type (_, Ts), casex, ...}, ...}, ...} = fp_sugar_of lthy T_name; val T = body_type (fastype_of casex); in Term.subst_atomic_types ((T :: Ts) ~~ (U :: Us)) casex end; fun deduce_according_type default_T [] = default_T | deduce_according_type default_T Ts = (case distinct (op =) Ts of U :: [] => U | _ => fpT_to ssig_T default_T); fun massage_if explore_cond explore (params as {bound_Us, bound_Ts, ...}) t = (case strip_comb t of (const as Const (\<^const_name>\If\, _), obj :: (branches as [_, _])) => (case List.partition Term.is_dummy_pattern (map (explore params) branches) of (dummy_branch' :: _, []) => dummy_branch' | (_, [branch']) => branch' | (_, branches') => let val brancheUs = map (curry fastype_of1 bound_Us) branches'; val U = deduce_according_type (fastype_of1 (bound_Ts, hd branches)) brancheUs; val const_obj' = (If_const U, obj) ||> explore_cond (update_UT params \<^typ>\bool\ \<^typ>\bool\) |> op $; in build_function_after_encapsulation (const $ obj) const_obj' params branches branches' end) | _ => explore params t); fun massage_map explore (params as {bound_Us, bound_Ts, T = Type (T_name, Ts), ...}) (t as func $ mapped_arg) = if is_self_call (head_of func) then explore params t else (case try (dest_map lthy T_name) func of SOME (map_tm, fs) => let val n = length fs; val mapped_arg' = mapped_arg |> `(curry fastype_of1 bound_Ts) |>> build_params bound_Us bound_Ts |-> explore; in (case fastype_of1 (bound_Us, mapped_arg') of Type (U_name, Us0) => if U_name = T_name then let val Us = map (fpT_to ssig_T) Us0; val temporary_map = map_tm |> mk_map n Us Ts; val map_fn_Ts = fastype_of #> strip_fun_type #> fst; val binder_Uss = map_fn_Ts temporary_map |> map (map (fpT_to ssig_T) o binder_types); val fun_paramss = map_fn_Ts (head_of func) |> map (build_params bound_Us bound_Ts); val fs' = fs |> @{map 4} explore_fun binder_Uss (replicate n explore) fun_paramss; val SOME bnf = bnf_of lthy T_name; val Type (_, bnf_Ts) = T_of_bnf bnf; val typ_alist = lives_of_bnf bnf ~~ map (curry fastype_of1 bound_Us #> range_type) fs'; val Us' = map2 the_default Us (map (AList.lookup (op =) typ_alist) bnf_Ts); val map_tm' = map_tm |> mk_map n Us Us'; in build_function_after_encapsulation func (list_comb (map_tm', fs')) params [mapped_arg] [mapped_arg'] end else explore params t | _ => explore params t) end | NONE => explore params t) | massage_map explore params t = explore params t; fun massage_comp explore (params as {bound_Us, ...}) t = (case strip_comb t of (Const (\<^const_name>\comp\, _), f1 :: f2 :: args) => let val args' = map (typ_before explore params) args; val f2' = typ_before (explore_fun (map (curry fastype_of1 bound_Us) args') explore) params f2; val f1' = typ_before (explore_fun [range_type (fastype_of1 (bound_Us, f2'))] explore) params f1; in betapply (f1', list_comb (f2', args')) end | _ => explore params t); fun massage_ctr explore (params as {T = T as Type (s, Ts), bound_Us, ...}) t = if T <> res_T then (case try (dest_ctr lthy s) t of SOME (ctr, args) => let val args' = map (typ_before explore params) args; val SOME {T = Type (_, ctr_Ts), ...} = ctr_sugar_of lthy s; val temp_ctr = mk_ctr ctr_Ts ctr; val argUs = map (curry fastype_of1 bound_Us) args'; val typ_alist = binder_types (fastype_of temp_ctr) ~~ argUs; val Us = ctr_Ts |> map (find_all_associated_types typ_alist) |> map2 deduce_according_type Ts; val ctr' = mk_ctr Us ctr; in build_function_after_encapsulation ctr ctr' params args args' end | NONE => explore params t) else explore params t | massage_ctr explore params t = explore params t; fun const_of [] _ = NONE | const_of ((sel as Const (s1, _)) :: r) (const as Const (s2, _)) = if s1 = s2 then SOME sel else const_of r const | const_of _ _ = NONE; fun massage_disc explore (params as {T, bound_Us, bound_Ts, ...}) t = (case (strip_comb t, T = \<^typ>\bool\) of ((fun_t, arg :: []), true) => let val arg_T = fastype_of1 (bound_Ts, arg) in if arg_T <> res_T then (case arg_T |> try (fst o dest_Type) |> Option.mapPartial (ctr_sugar_of lthy) of SOME {discs, T = Type (_, Ts), ...} => (case const_of discs fun_t of SOME disc => let val arg' = arg |> typ_before explore params; val Type (_, Us) = fastype_of1 (bound_Us, arg'); val disc' = disc |> Term.subst_TVars (map (fst o dest_TVar) Ts ~~ Us); in disc' $ arg' end | NONE => explore params t) | NONE => explore params t) else explore params t end | _ => explore params t); fun massage_sel explore (params as {bound_Us, bound_Ts, ...}) t = let val (fun_t, args) = strip_comb t in if args = [] then explore params t else let val T = fastype_of1 (bound_Ts, hd args) in (case (Option.mapPartial (ctr_sugar_of lthy) (try (fst o dest_Type) T), T <> res_T) of (SOME {selss, T = Type (_, Ts), ...}, true) => (case const_of (flat selss) fun_t of SOME sel => let val args' = args |> map (typ_before explore params); val Type (_, Us) = fastype_of1 (bound_Us, hd args'); val sel' = sel |> Term.subst_TVars (map (fst o dest_TVar) Ts ~~ Us); in build_function_after_encapsulation sel sel' params args args' end | NONE => explore params t) | _ => explore params t) end end; fun massage_equality explore (params as {bound_Us, bound_Ts, ...}) (t as Const (\<^const_name>\HOL.eq\, _) $ t1 $ t2) = let val check_is_VLeaf = not o (Term.exists_subterm (fn t => t aconv CLeaf orelse t aconv Oper)); fun try_pattern_matching (fun_t, arg_ts) t = (case as_member_of pattern_ctrs fun_t of SOME (disc, sels) => let val t' = typ_before explore params t in if fastype_of1 (bound_Us, t') = YpreT then let val arg_ts' = map (typ_before explore params) arg_ts; val sels_t' = map (fn sel => betapply (sel, t')) sels; val Ts = map (curry fastype_of1 bound_Us) arg_ts'; val Us = map (curry fastype_of1 bound_Us) sels_t'; val arg_ts' = map2 encapsulate (map2 (update_UT params) Us Ts) arg_ts'; in if forall check_is_VLeaf arg_ts' then SOME (Library.foldl1 HOLogic.mk_conj (betapply (disc, t') :: (map HOLogic.mk_eq (arg_ts' ~~ sels_t')))) else NONE end else NONE end | NONE => NONE); in (case try_pattern_matching (strip_comb t1) t2 of SOME cond => cond | NONE => (case try_pattern_matching (strip_comb t2) t1 of SOME cond => cond | NONE => let val T = fastype_of1 (bound_Ts, t1); val params' = build_params bound_Us bound_Ts T; val t1' = explore params' t1; val t2' = explore params' t2; in if fastype_of1 (bound_Us, t1') = T andalso fastype_of1 (bound_Us, t2') = T then HOLogic.mk_eq (t1', t2') else error ("Unsupported condition: " ^ quote (Syntax.string_of_term lthy t)) end)) end | massage_equality explore params t = explore params t; fun infer_types (TVar _) (TVar _) = [] | infer_types (U as TVar _) T = [(U, T)] | infer_types (Type (s', Us)) (Type (s, Ts)) = if s' = s then flat (map2 infer_types Us Ts) else [] | infer_types _ _ = []; fun group_by_fst associations [] = associations | group_by_fst associations ((a, b) :: r) = group_by_fst (add_association a b associations) r and add_association a b [] = [(a, [b])] | add_association a b ((c, d) :: r) = if a = c then (c, b :: d) :: r else (c, d) :: (add_association a b r); fun new_TVar known_TVars = Name.invent_list (map (fst o fst o dest_TVar) known_TVars) "x" 1 |> (fn [s] => TVar ((s, 0), [])); fun instantiate_type inferred_types = Term.typ_subst_TVars (map (apfst (fst o dest_TVar)) inferred_types); fun chose_unknown_TVar (T as TVar _) = SOME T | chose_unknown_TVar (Type (_, Ts)) = fold (curry merge_options) (map chose_unknown_TVar Ts) NONE | chose_unknown_TVar _ = NONE; (* The function under definition might not be defined yet when this is queried. *) fun maybe_const_type ctxt (s, T) = Sign.const_type (Proof_Context.theory_of ctxt) s |> the_default T; fun massage_const polymorphic explore (params as {bound_Us, ...}) t = let val (fun_t, arg_ts) = strip_comb t in (case fun_t of Const (fun_x as (s, fun_T)) => let val general_T = if polymorphic then maybe_const_type lthy fun_x else fun_T in if fun_t aconv friend_tm orelse contains_res_T (body_type general_T) orelse is_constant t then explore params t else let val inferred_types = infer_types general_T fun_T; fun prepare_skeleton [] _ = [] | prepare_skeleton ((T, U) :: inferred_types) As = let fun schematize_res_T U As = if U = res_T then let val A = new_TVar As in (A, A :: As) end else (case U of Type (s, Us) => fold_map schematize_res_T Us As |>> curry Type s | _ => (U, As)); val (U', As') = schematize_res_T U As; in (T, U') :: (prepare_skeleton inferred_types As') end; val inferred_types' = prepare_skeleton inferred_types (map fst inferred_types); val skeleton_T = instantiate_type inferred_types' general_T; fun explore_if_possible (exp_arg as (_, true)) _ = exp_arg | explore_if_possible (exp_arg as (arg, false)) arg_T = if exists (exists_subtype is_TVar) (binder_types arg_T) then exp_arg else (typ_before (explore_fun (binder_types arg_T) explore) params arg, true); fun collect_inferred_types [] _ = [] | collect_inferred_types ((arg, explored) :: exp_arg_ts) (arg_T :: arg_Ts) = (if explored then infer_types arg_T (fastype_of1 (bound_Us, arg)) else []) @ collect_inferred_types exp_arg_ts arg_Ts; fun propagate exp_arg_ts skeleton_T = let val arg_gen_Ts = binder_types skeleton_T; val exp_arg_ts = map2 explore_if_possible exp_arg_ts arg_gen_Ts; val inferred_types = collect_inferred_types exp_arg_ts arg_gen_Ts |> group_by_fst [] |> map (apsnd (deduce_according_type ssig_T)); in (exp_arg_ts, instantiate_type inferred_types skeleton_T) end; val remaining_to_be_explored = filter_out snd #> length; fun try_exploring_args exp_arg_ts skeleton_T = let val n = remaining_to_be_explored exp_arg_ts; val (exp_arg_ts', skeleton_T') = propagate exp_arg_ts skeleton_T; val n' = remaining_to_be_explored exp_arg_ts'; fun try_instantiating A T = try (try_exploring_args exp_arg_ts') (instantiate_type [(A, T)] skeleton_T'); in if n' = 0 then SOME (exp_arg_ts', skeleton_T') else if n = n' then if exists_subtype is_TVar skeleton_T' then let val SOME A = chose_unknown_TVar skeleton_T' in (case try_instantiating A ssig_T of SOME result => result | NONE => (case try_instantiating A YpreT of SOME result => result | NONE => (case try_instantiating A res_T of SOME result => result | NONE => NONE))) end else NONE else try_exploring_args exp_arg_ts' skeleton_T' end; in (case try_exploring_args (map (fn arg => (arg, false)) arg_ts) skeleton_T of SOME (exp_arg_ts, fun_U) => let val arg_ts' = map fst exp_arg_ts; val fun_t' = Const (s, fun_U); fun finish_off () = let val t' = build_function_after_encapsulation fun_t fun_t' params arg_ts arg_ts'; in if can type_of1 (bound_Us, t') then (if fun_T = fun_U orelse is_special_parametric_const (s, fun_T) then () else add_parametric_const s general_T fun_T fun_U; t') else explore params t end; in if polymorphic then finish_off () else (case try finish_off () of SOME t' => t' | NONE => explore params t) end | NONE => explore params t) end end | _ => explore params t) end; fun massage_rho explore = massage_star [massage_let, massage_if explore_cond, massage_case, massage_fun, massage_comp, massage_map, massage_ctr, massage_sel, massage_disc, massage_equality, massage_const false, massage_const true] explore and massage_case explore (params as {bound_Ts, bound_Us, ...}) t = (case strip_comb t of (casex as Const (case_x as (c, _)), args as _ :: _ :: _) => (case try strip_fun_type (maybe_const_type lthy case_x) of SOME (gen_branch_Ts, gen_body_fun_T) => let val gen_branch_ms = map num_binder_types gen_branch_Ts; val n = length gen_branch_ms; val (branches, obj_leftovers) = chop n args; in if n < length args then (case gen_body_fun_T of Type (_, [Type (T_name, _), _]) => if case_of lthy T_name = SOME (c, true) then let val brancheTs = binder_fun_types (fastype_of1 (bound_Ts, casex)); val obj_leftover_Ts = map (curry fastype_of1 bound_Ts) obj_leftovers; val obj_leftovers' = if is_constant (hd obj_leftovers) then obj_leftovers else (obj_leftover_Ts, obj_leftovers) |>> map (build_params bound_Us bound_Ts) |> op ~~ |> map (uncurry explore_inner); val obj_leftoverUs = obj_leftovers' |> map (curry fastype_of1 bound_Us); val _ = is_valid_case_argumentT (hd obj_leftoverUs) orelse error (quote (Syntax.string_of_term lthy (hd obj_leftovers)) ^ " is not a valid case argument"); val Us = obj_leftoverUs |> hd |> dest_Type |> snd; val branche_binderUss = (if hd obj_leftoverUs = YpreT then mk_case HOLogic.boolT else update_case Us HOLogic.boolT casex) |> fastype_of |> binder_fun_types |> map binder_types; val b_params = map (build_params bound_Us bound_Ts) brancheTs; val branches' = branches |> @{map 4} explore_fun branche_binderUss (replicate n explore) b_params; val brancheUs = map (curry fastype_of1 bound_Us) branches'; val U = deduce_according_type (body_type (hd brancheTs)) (map body_type brancheUs); val casex' = if hd obj_leftoverUs = YpreT then mk_case U else update_case Us U casex; in build_function_after_encapsulation casex casex' params (branches @ obj_leftovers) (branches' @ obj_leftovers') end else explore params t | _ => explore params t) else explore params t end | NONE => explore params t) | _ => explore params t) and explore_cond params t = if has_self_call t then unexpected_rec_call_in lthy [] t else explore_inner params t and explore_inner params t = massage_rho explore_inner_general params t and explore_inner_general (params as {bound_Us, bound_Ts, T, ...}) t = let val (fun_t, arg_ts) = strip_comb t in if is_constant t then t else (case (as_member_of discs fun_t, length arg_ts = 1 andalso has_res_T bound_Ts (the_single arg_ts)) of (SOME disc', true) => let val arg' = explore_inner params (the_single arg_ts); val arg_U = fastype_of1 (bound_Us, arg'); in if arg_U = res_T then fun_t $ arg' else if arg_U = YpreT then disc' $ arg' else error ("Discriminator " ^ quote (Syntax.string_of_term lthy fun_t) ^ " cannot be applied to non-variable " ^ quote (Syntax.string_of_term lthy (hd arg_ts))) end | _ => (case as_member_of sels fun_t of SOME sel' => let val arg_ts' = map (explore_inner params) arg_ts; val arg_U = fastype_of1 (bound_Us, hd arg_ts'); in if arg_U = res_T then build_function_after_encapsulation fun_t fun_t params arg_ts arg_ts' else if arg_U = YpreT then build_function_after_encapsulation fun_t sel' params arg_ts arg_ts' else error ("Selector " ^ quote (Syntax.string_of_term lthy fun_t) ^ " cannot be applied to non-variable " ^ quote (Syntax.string_of_term lthy (hd arg_ts))) end | NONE => (case as_member_of friends fun_t of SOME (_, friend') => rebuild_function_after_exploration fun_t friend' explore_inner params arg_ts |> curry (op $) Oper | NONE => (case as_member_of ctr_guards fun_t of SOME ctr_guard' => rebuild_function_after_exploration fun_t ctr_guard' explore_inner params arg_ts |> curry (op $) ctr_wrapper |> curry (op $) Oper | NONE => if is_Bound fun_t then rebuild_function_after_exploration fun_t fun_t explore_inner params arg_ts else if is_Free fun_t then let val fun_t' = map_types (fpT_to YpreT) fun_t in rebuild_function_after_exploration fun_t fun_t' explore_inner params arg_ts end else if T = res_T then error (quote (Syntax.string_of_term lthy fun_t) ^ " not polymorphic enough to be applied like this and no friend") else error (quote (Syntax.string_of_term lthy fun_t) ^ " not polymorphic enough to be applied like this"))))) end; fun explore_ctr params t = massage_rho explore_ctr_general params t and explore_ctr_general params t = let val (fun_t, arg_ts) = strip_comb t; val ctr_opt = as_member_of ctr_guards fun_t; in if is_some ctr_opt then rebuild_function_after_exploration fun_t (the ctr_opt) explore_inner params arg_ts else not_constructor_in_rhs lthy [] fun_t end; val rho_rhs = rhs |> explore_ctr (build_params [] [] (fastype_of rhs)) |> abs_tuple_balanced (map (map_types (fpT_to YpreT)) lhs_args) |> unfold_id_bnf_etc lthy; in lthy |> define_const false b version rhoN rho_rhs |>> pair (!parametric_consts, rho_rhs) end; fun mk_rho_parametricity_goal ctxt Y Z preT ssig_T dead_pre_rel dead_k_rel dead_ssig_rel rho_rhs = let val YpreT = HOLogic.mk_prodT (Y, preT); val ZpreT = Tsubst Y Z YpreT; val ssigZ_T = Tsubst Y Z ssig_T; val dead_pre_rel' = Term.subst_atomic_types [(Y, ssig_T), (Z, ssigZ_T)] dead_pre_rel; val dead_k_rel' = Term.subst_atomic_types [(Y, YpreT), (Z, ZpreT)] dead_k_rel; val (R, _) = ctxt |> yield_singleton (mk_Frees "R") (mk_pred2T Y Z); val rho_rel = mk_rel_fun (dead_k_rel' $ mk_rel_prod R (dead_pre_rel $ R)) (dead_pre_rel' $ (dead_ssig_rel $ R)); val rho_rhsZ = substT Y Z rho_rhs; in HOLogic.mk_Trueprop (rho_rel $ rho_rhs $ rho_rhsZ) end; fun extract_rho_return_transfer_goals fun_b version dead_pre_bnf dead_k_bnf Y Z preT k_T ssig_T ssig_fp_sugar friend_parse_info fun_t parsed_eq lthy = let fun mk_rel T bnf = let val ZT = Tsubst Y Z T; val rel_T = mk_predT [mk_pred2T Y Z, T, ZT]; in enforce_type lthy I rel_T (rel_of_bnf bnf) end; val ssig_bnf = #fp_bnf ssig_fp_sugar; val (dead_ssig_bnf, lthy') = bnf_kill_all_but 1 ssig_bnf lthy; val dead_pre_rel = mk_rel preT dead_pre_bnf; val dead_k_rel = mk_rel k_T dead_k_bnf; val dead_ssig_rel = mk_rel ssig_T dead_ssig_bnf; val (((parametric_consts, rho_rhs), rho_data), lthy'') = extract_rho_from_equation friend_parse_info fun_b version Y preT ssig_T fun_t parsed_eq lthy'; val const_transfer_goals = map (mk_const_transfer_goal lthy'') parametric_consts; val rho_transfer_goal = mk_rho_parametricity_goal lthy'' Y Z preT ssig_T dead_pre_rel dead_k_rel dead_ssig_rel rho_rhs; in ((rho_data, (const_transfer_goals, rho_transfer_goal)), lthy'') end; fun explore_corec_equation ctxt could_be_friend friend fun_name fun_free {outer_buffer, ctr_guards, inner_buffer} res_T (args, rhs) = let val is_self_call = curry (op aconv) fun_free; val has_self_call = Term.exists_subterm is_self_call; val outer_ssig_T = body_type (fastype_of (#Oper outer_buffer)); fun inner_fp_of (Free (s, _)) = Free (s ^ inner_fp_suffix, mk_tupleT_balanced (map fastype_of args) --> outer_ssig_T); fun build_params bound_Ts U T = {bound_Us = bound_Ts, bound_Ts = bound_Ts, U = U, T = T}; fun rebuild_function_after_exploration new_fn explore {bound_Ts, ...} arg_ts = let val binder_types_old_fn = map (curry fastype_of1 bound_Ts) arg_ts; val binder_types_new_fn = new_fn |> binder_types o (curry fastype_of1 bound_Ts) |> take (length binder_types_old_fn); val paramss = map2 (build_params bound_Ts) binder_types_new_fn binder_types_old_fn; in map2 explore paramss arg_ts |> curry list_comb new_fn end; fun massage_map_corec explore {bound_Ts, U, T, ...} t = let val explore' = explore ooo build_params in massage_nested_corec_call ctxt has_self_call explore' explore' bound_Ts U T t end; fun massage_comp explore params t = (case strip_comb t of (Const (\<^const_name>\comp\, _), f1 :: f2 :: args) => explore params (betapply (f1, (betapplys (f2, args)))) | _ => explore params t); fun massage_fun explore (params as {bound_Us, bound_Ts, U, T}) t = if can dest_funT T then let val arg_T = domain_type T; val arg_name = the_default Name.uu (try (fn (Abs (s, _, _)) => s) t); in add_boundvar t |> explore {bound_Us = arg_T :: bound_Us, bound_Ts = arg_T :: bound_Ts, U = range_type U, T = range_type T} |> (fn t => Abs (arg_name, arg_T, t)) end else explore params t fun massage_let_if_case_corec explore {bound_Ts, U, T, ...} t = massage_let_if_case ctxt has_self_call (fn bound_Ts => explore (build_params bound_Ts U T)) (K (unexpected_corec_call_in ctxt [t])) (K (unsupported_case_around_corec_call ctxt [t])) bound_Ts t; val massage_map_let_if_case = massage_star [massage_map_corec, massage_fun, massage_comp, massage_let_if_case_corec]; fun explore_arg _ t = if has_self_call t then error (quote (Syntax.string_of_term ctxt t) ^ " contains a nested corecursive call" ^ (if could_be_friend then " (try specifying \"(friend)\")" else "")) else t; fun explore_inner params t = massage_map_let_if_case explore_inner_general params t and explore_inner_general (params as {bound_Ts, T, ...}) t = if T = res_T then let val (f_t, arg_ts) = strip_comb t in if has_self_call t then (case as_member_of (#friends inner_buffer) f_t of SOME (_, friend') => rebuild_function_after_exploration friend' explore_inner params arg_ts |> curry (op $) (#Oper inner_buffer) | NONE => (case as_member_of ctr_guards f_t of SOME ctr_guard' => rebuild_function_after_exploration ctr_guard' explore_inner params arg_ts |> curry (op $) (#ctr_wrapper inner_buffer) |> curry (op $) (#Oper inner_buffer) | NONE => if is_self_call f_t then if friend andalso exists has_self_call arg_ts then (case Symtab.lookup (#friends inner_buffer) fun_name of SOME (_, friend') => rebuild_function_after_exploration friend' explore_inner params arg_ts |> curry (op $) (#Oper inner_buffer)) else let val arg_Ts = binder_types (fastype_of1 (bound_Ts, f_t)) in map2 explore_arg (map2 (update_UT params) arg_Ts arg_Ts) arg_ts |> mk_tuple1_balanced bound_Ts |> curry (op $) (#VLeaf inner_buffer) end else error (quote (Syntax.string_of_term ctxt f_t) ^ " not registered as friend"))) else #CLeaf inner_buffer $ t end else if has_self_call t then error (quote (Syntax.string_of_term ctxt t) ^ " contains a corecursive call but has type " ^ quote (Syntax.string_of_typ ctxt T)) else explore_nested ctxt explore_inner_general params t; fun explore_outer params t = massage_map_let_if_case explore_outer_general params t and explore_outer_general (params as {bound_Ts, T, ...}) t = if T = res_T then let val (f_t, arg_ts) = strip_comb t in (case as_member_of ctr_guards f_t of SOME ctr_guard' => rebuild_function_after_exploration ctr_guard' explore_inner params arg_ts |> curry (op $) (#VLeaf outer_buffer) | NONE => if not (has_self_call t) then t |> expand_to_ctr_term ctxt T |> massage_let_if_case_corec explore_outer_general params else (case as_member_of (#friends outer_buffer) f_t of SOME (_, friend') => rebuild_function_after_exploration friend' explore_outer params arg_ts |> curry (op $) (#Oper outer_buffer) | NONE => if is_self_call f_t then let val arg_Ts = binder_types (fastype_of1 (bound_Ts, f_t)) in map2 explore_arg (map2 (update_UT params) arg_Ts arg_Ts) arg_ts |> mk_tuple1_balanced bound_Ts |> curry (op $) (inner_fp_of f_t) end else error (quote (Syntax.string_of_term ctxt f_t) ^ " not registered as friend"))) end else if has_self_call t then error (quote (Syntax.string_of_term ctxt t) ^ " contains a corecursive call but has type " ^ quote (Syntax.string_of_typ ctxt T)) else explore_nested ctxt explore_outer_general params t; in (args, rhs |> explore_outer (build_params [] outer_ssig_T res_T) |> abs_tuple_balanced args) end; fun mk_corec_fun_def_rhs ctxt arg_Ts corecUU0 corecUU_arg = let val corecUU = enforce_type ctxt domain_type (fastype_of corecUU_arg) corecUU0 in abs_curried_balanced arg_Ts (corecUU $ unfold_id_bnf_etc ctxt corecUU_arg) end; fun get_options ctxt opts = let val plugins = get_first (fn Plugins_Option f => SOME (f ctxt) | _ => NONE) (rev opts) |> the_default Plugin_Name.default_filter; val friend = exists (can (fn Friend_Option => ())) opts; val transfer = exists (can (fn Transfer_Option => ())) opts; in (plugins, friend, transfer) end; fun add_function binding parsed_eq lthy = let fun pat_completeness_auto ctxt = Pat_Completeness.pat_completeness_tac ctxt 1 THEN auto_tac ctxt; val ({defname, pelims = [[pelim]], pinducts = [pinduct], psimps = [psimp], ...}, lthy') = Function.add_function [(Binding.concealed binding, NONE, NoSyn)] [(((Binding.concealed Binding.empty, []), parsed_eq), [], [])] Function_Common.default_config pat_completeness_auto lthy; in ((defname, (pelim, pinduct, psimp)), lthy') end; fun build_corecUU_arg_and_goals prove_termin (Free (fun_base_name, _)) (arg_ts, explored_rhs) lthy = let val inner_fp_name0 = fun_base_name ^ inner_fp_suffix; val inner_fp_free = Free (inner_fp_name0, fastype_of explored_rhs); in if Term.exists_subterm (curry (op aconv) inner_fp_free) explored_rhs then let val arg = mk_tuple_balanced arg_ts; val inner_fp_eq = mk_Trueprop_eq (betapply (inner_fp_free, arg), betapply (explored_rhs, arg)); val ((inner_fp_name, (pelim, pinduct, psimp)), lthy') = add_function (Binding.name inner_fp_name0) inner_fp_eq lthy; fun mk_triple elim induct simp = ([elim], [induct], [simp]); fun prepare_termin () = let val {goal, ...} = Proof.goal (Function.termination NONE lthy'); val termin_goal = goal |> Thm.concl_of |> Logic.unprotect |> Envir.beta_eta_contract; in (lthy', (mk_triple pelim pinduct psimp, [termin_goal])) end; val (lthy'', (inner_fp_triple, termin_goals)) = if prove_termin then (case try (Function.prove_termination NONE (Function_Common.termination_prover_tac true lthy')) lthy' of NONE => prepare_termin () | SOME ({elims = SOME [[elim]], inducts = SOME [induct], simps = SOME [simp], ...}, lthy'') => (lthy'', (mk_triple elim induct simp, []))) else prepare_termin (); val inner_fp_const = (Binding.name_of inner_fp_name, fastype_of explored_rhs) |>> Proof_Context.read_const {proper = true, strict = false} lthy' |> (fn (Const (s, _), T) => Const (s, T)); in (((inner_fp_triple, termin_goals), inner_fp_const), lthy'') end else (((([], [], []), []), explored_rhs), lthy) end; fun derive_eq_corecUU ctxt {sig_fp_sugars, ssig_fp_sugar, eval, corecUU, eval_simps, all_algLam_algs, corecUU_unique, ...} fun_t corecUU_arg fun_code = let val fun_T = fastype_of fun_t; val (arg_Ts, Type (fpT_name, _)) = strip_type fun_T; val num_args = length arg_Ts; val SOME {pre_bnf, fp_bnf, absT_info, fp_nesting_bnfs, live_nesting_bnfs, fp_ctr_sugar, ...} = fp_sugar_of ctxt fpT_name; val SOME {case_trivial, ...} = codatatype_extra_of ctxt fpT_name; val ctr_sugar = #ctr_sugar fp_ctr_sugar; val pre_map_def = map_def_of_bnf pre_bnf; val abs_inverse = #abs_inverse absT_info; val ctr_defs = #ctr_defs fp_ctr_sugar; val case_eq_ifs = #case_eq_ifs ctr_sugar @ case_eq_if_thms_of_term ctxt (Thm.prop_of fun_code); val all_sig_map_thms = maps (#map_thms o #fp_bnf_sugar) sig_fp_sugars; val fp_map_ident = map_ident_of_bnf fp_bnf; val fpsig_nesting_bnfs = fp_nesting_bnfs @ maps #live_nesting_bnfs sig_fp_sugars; val fpsig_nesting_T_names = map (fst o dest_Type o T_of_bnf) fpsig_nesting_bnfs; val fpsig_nesting_fp_sugars = map_filter (fp_sugar_of ctxt) fpsig_nesting_T_names; val fpsig_nesting_fp_bnf_sugars = map #fp_bnf_sugar fpsig_nesting_fp_sugars; val ssig_fp_bnf_sugar = #fp_bnf_sugar ssig_fp_sugar; val ssig_bnf = #fp_bnf ssig_fp_sugar; val ssig_map = map_of_bnf ssig_bnf; val fpsig_nesting_maps = map map_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_ident0s = map map_ident0_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_comps = map map_comp_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_thms = maps #map_thms fpsig_nesting_fp_bnf_sugars; val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs; val ssig_map_thms = #map_thms ssig_fp_bnf_sugar; val all_algLam_alg_pointfuls = map (mk_pointful ctxt) all_algLam_algs; val def_rhs = mk_corec_fun_def_rhs ctxt arg_Ts corecUU corecUU_arg; val goal = mk_Trueprop_eq (fun_t, def_rhs); in Goal.prove_sorry ctxt [] [] goal (fn {context = ctxt, prems = _} => mk_eq_corecUU_tac ctxt num_args fpsig_nesting_maps ssig_map eval pre_map_def abs_inverse fpsig_nesting_map_ident0s fpsig_nesting_map_comps fpsig_nesting_map_thms live_nesting_map_ident0s fp_map_ident case_trivial ctr_defs case_eq_ifs all_sig_map_thms ssig_map_thms all_algLam_alg_pointfuls (all_algrho_eqs_of ctxt) eval_simps corecUU_unique fun_code) |> Thm.close_derivation \<^here> end; fun derive_coinduct_cong_intros ({fpT = fpT0 as Type (fpT_name, _), friend_names = friend_names0, corecUU = Const (corecUU_name, _), dtor_coinduct_info as {dtor_coinduct, ...}, ...}) lthy = let val thy = Proof_Context.theory_of lthy; val phi = Proof_Context.export_morphism lthy (Local_Theory.target_of lthy); val fpT = Morphism.typ phi fpT0; val general_fpT = body_type (Sign.the_const_type thy corecUU_name); val most_general = Sign.typ_instance thy (general_fpT, fpT); in (case (most_general, coinduct_extra_of lthy corecUU_name) of (true, SOME extra) => ((false, extra), lthy) | _ => let val ctr_names = ctr_names_of_fp_name lthy fpT_name; val friend_names = friend_names0 |> map Long_Name.base_name |> rev; val cong_intro_pairs = derive_cong_intros lthy ctr_names friend_names dtor_coinduct_info; val (coinduct, coinduct_attrs) = derive_coinduct lthy fpT0 dtor_coinduct; val ((_, [coinduct]), lthy) = (* TODO check: only if most_general?*) Local_Theory.note ((Binding.empty, coinduct_attrs), [coinduct]) lthy; val extra = {coinduct = coinduct, coinduct_attrs = coinduct_attrs, cong_intro_pairs = cong_intro_pairs}; in ((most_general, extra), lthy |> most_general ? register_coinduct_extra corecUU_name extra) end) end; fun update_coinduct_cong_intross_dynamic fpT_name lthy = let val all_corec_infos = corec_infos_of lthy fpT_name in lthy |> fold_map (apfst snd oo derive_coinduct_cong_intros) all_corec_infos |> snd end; fun derive_and_update_coinduct_cong_intross [] = pair (false, []) | derive_and_update_coinduct_cong_intross (corec_infos as {fpT = Type (fpT_name, _), ...} :: _) = fold_map derive_coinduct_cong_intros corec_infos #>> split_list #> (fn ((changeds, extras), lthy) => if exists I changeds then ((true, extras), lthy |> update_coinduct_cong_intross_dynamic fpT_name) else ((false, extras), lthy)); fun prepare_corec_ursive_cmd int long_cmd opts (raw_fixes, raw_eq) lthy = let val _ = can the_single raw_fixes orelse error "Mutually corecursive functions not supported"; val (plugins, friend, transfer) = get_options lthy opts; val ([((b, fun_T), mx)], [(_, eq)]) = fst (Specification.read_multi_specs raw_fixes [((Binding.empty_atts, raw_eq), [], [])] lthy); val _ = check_top_sort lthy b fun_T; val (arg_Ts, res_T) = strip_type fun_T; val fpT_name = (case res_T of Type (s, _) => s | _ => not_codatatype lthy res_T); val fun_free = Free (Binding.name_of b, fun_T); val parsed_eq = parse_corec_equation lthy [fun_free] eq; val fun_name = Local_Theory.full_name lthy b; val fun_t = Const (fun_name, fun_T); (* FIXME: does this work with locales that fix variables? *) val no_base = has_no_corec_info lthy fpT_name; val lthy1 = lthy |> no_base ? setup_base fpT_name; fun extract_rho lthy' = let val lthy'' = lthy' |> Variable.declare_typ fun_T; val (prepared as (_, _, version, Y, Z, preT, k_T, ssig_T, dead_pre_bnf, dead_k_bnf, _, ssig_fp_sugar, buffer), lthy''') = prepare_friend_corec fun_name fun_T lthy''; val friend_parse_info = friend_parse_info_of lthy''' arg_Ts res_T buffer; val parsed_eq' = parsed_eq ||> subst_atomic [(fun_free, fun_t)]; in lthy''' |> extract_rho_return_transfer_goals b version dead_pre_bnf dead_k_bnf Y Z preT k_T ssig_T ssig_fp_sugar friend_parse_info fun_t parsed_eq' |>> pair prepared end; val ((prepareds, (rho_datas, transfer_goal_datas)), lthy2) = if friend then extract_rho lthy1 |>> (apfst single ##> (apfst single #> apsnd single)) else (([], ([], [])), lthy1); val ((buffer, corec_infos), lthy3) = if friend then ((#13 (the_single prepareds), []), lthy2) else corec_info_of res_T lthy2 ||> no_base ? update_coinduct_cong_intross_dynamic fpT_name |>> (fn info as {buffer, ...} => (buffer, [info])); val corec_parse_info = corec_parse_info_of lthy3 arg_Ts res_T buffer; val explored_eq = explore_corec_equation lthy3 true friend fun_name fun_free corec_parse_info res_T parsed_eq; val (((inner_fp_triple, termin_goals), corecUU_arg), lthy4) = build_corecUU_arg_and_goals (not long_cmd) fun_free explored_eq lthy3; fun def_fun (inner_fp_elims0, inner_fp_inducts0, inner_fp_simps0) const_transfers rho_transfers_foldeds lthy5 = let fun register_friend lthy' = let val [(old_corec_info, fp_b, version, Y, Z, _, k_T, _, _, dead_k_bnf, sig_fp_sugar, ssig_fp_sugar, _)] = prepareds; val [(rho, rho_def)] = rho_datas; val [(_, rho_transfer_goal)] = transfer_goal_datas; val Type (fpT_name, _) = res_T; val rho_transfer_folded = (case rho_transfers_foldeds of [] => derive_rho_transfer_folded lthy' fpT_name const_transfers rho_def rho_transfer_goal | [thm] => thm); in lthy' |> register_coinduct_dynamic_friend fpT_name fun_name |> register_friend_corec fun_name fp_b version Y Z k_T dead_k_bnf sig_fp_sugar ssig_fp_sugar fun_t rho rho_transfer_folded old_corec_info end; val (friend_infos, lthy6) = lthy5 |> (if friend then register_friend #>> single else pair []); val (corec_info as {corecUU = corecUU0, ...}, lthy7) = (case corec_infos of [] => corec_info_of res_T lthy6 | [info] => (info, lthy6)); val def_rhs = mk_corec_fun_def_rhs lthy7 arg_Ts corecUU0 corecUU_arg; val def = ((b, mx), ((Binding.concealed (Thm.def_binding b), []), def_rhs)); val ((fun_lhs0, (_, fun_def0)), (lthy9, lthy8')) = lthy7 - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> Local_Theory.define def |> tap (fn (def, lthy') => print_def_consts int [def] lthy') ||> `Local_Theory.close_target; val parsed_eq = parse_corec_equation lthy9 [fun_free] eq; val views0 = generate_views lthy9 eq fun_free parsed_eq; val lthy9' = lthy9 |> fold Variable.declare_typ (res_T :: arg_Ts); val phi = Proof_Context.export_morphism lthy8' lthy9'; val fun_lhs = Morphism.term phi fun_lhs0; val fun_def = Morphism.thm phi fun_def0; val inner_fp_elims = map (Morphism.thm phi) inner_fp_elims0; val inner_fp_inducts = map (Morphism.thm phi) inner_fp_inducts0; val inner_fp_simps = map (Morphism.thm phi) inner_fp_simps0; val (code_goal, _, _, _, _) = morph_views phi views0; fun derive_and_note_friend_extra_theorems lthy' = let val k_T = #7 (the_single prepareds); val rho_def = snd (the_single rho_datas); val (eq_algrho, algrho_eq) = derive_eq_algrho lthy' corec_info (the_single friend_infos) fun_lhs k_T code_goal const_transfers rho_def fun_def; val notes = (if Config.get lthy' bnf_internals then [(eq_algrhoN, [eq_algrho])] else []) |> map (fn (thmN, thms) => ((Binding.qualify true (Binding.name_of b) (Binding.qualify false friendN (Binding.name thmN)), []), [(thms, [])])); in lthy' |> register_friend_extra fun_name eq_algrho algrho_eq |> Local_Theory.notes notes |> snd end; val lthy10 = lthy9 |> friend ? derive_and_note_friend_extra_theorems; val code_thm = derive_code lthy10 inner_fp_simps code_goal corec_info fun_lhs fun_def; (* TODO: val ctr_thmss = map mk_thm (#2 views); val disc_thmss = map mk_thm (#3 views); val disc_iff_thmss = map mk_thm (#4 views); val sel_thmss = map mk_thm (#5 views); *) val uniques = if null inner_fp_simps then [derive_unique lthy10 phi (#1 views0) corec_info fpT_name fun_def] else []; (* TODO: val disc_iff_or_disc_thmss = map2 (fn [] => I | disc_iffs => K disc_iffs) disc_iff_thmss disc_thmss; val simp_thmss = map2 append disc_iff_or_disc_thmss sel_thmss; *) val ((_, [{cong_intro_pairs, coinduct, coinduct_attrs}]), lthy11) = lthy10 |> derive_and_update_coinduct_cong_intross [corec_info]; val cong_intros_pairs = AList.group (op =) cong_intro_pairs; val anonymous_notes = []; (* TODO: [(flat disc_iff_or_disc_thmss, simp_attrs)] |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])])); *) val notes = [(cong_introsN, maps snd cong_intros_pairs, []), (codeN, [code_thm], nitpicksimp_attrs), (coinductN, [coinduct], coinduct_attrs), (inner_inductN, inner_fp_inducts, []), (uniqueN, uniques, [])] @ map (fn (thmN, thms) => (thmN, thms, [])) cong_intros_pairs @ (if Config.get lthy11 bnf_internals then [(inner_elimN, inner_fp_elims, []), (inner_simpN, inner_fp_simps, [])] else []) (* TODO: (ctrN, ctr_thms, []), (discN, disc_thms, []), (disc_iffN, disc_iff_thms, []), (selN, sel_thms, simp_attrs), (simpsN, simp_thms, []), *) |> map (fn (thmN, thms, attrs) => ((Binding.qualify true (Binding.name_of b) (Binding.qualify false corecN (Binding.name thmN)), attrs), [(thms, [])])) |> filter_out (null o fst o hd o snd); in lthy11 (* TODO: |> Spec_Rules.add Spec_Rules.equational ([fun_lhs], flat sel_thmss) |> Spec_Rules.add Spec_Rules.equational ([fun_lhs], flat ctr_thmss) *) |> Spec_Rules.add Binding.empty Spec_Rules.equational [fun_lhs] [code_thm] |> plugins code_plugin ? Code.declare_default_eqns [(code_thm, true)] |> Local_Theory.notes (anonymous_notes @ notes) |> snd end; fun prove_transfer_goal ctxt goal = Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} => HEADGOAL (Transfer.transfer_prover_tac ctxt))) |> Thm.close_derivation \<^here>; fun maybe_prove_transfer_goal ctxt goal = (case try (prove_transfer_goal ctxt) goal of SOME thm => apfst (cons thm) | NONE => apsnd (cons goal)); val const_transfer_goals = fold (union (op aconv) o fst) transfer_goal_datas []; val (const_transfers, const_transfer_goals') = if long_cmd then ([], const_transfer_goals) else fold (maybe_prove_transfer_goal lthy4) const_transfer_goals ([], []); in ((def_fun, (([res_T], prepareds, rho_datas, map snd transfer_goal_datas), (inner_fp_triple, termin_goals), (const_transfers, const_transfer_goals'))), lthy4) end; fun corec_cmd int opts (raw_fixes, raw_eq) lthy = let val ((def_fun, (_, (inner_fp_triple, termin_goals), (const_transfers, const_transfer_goals))), lthy') = prepare_corec_ursive_cmd int false opts (raw_fixes, raw_eq) lthy; in if not (null termin_goals) then error ("Termination prover failed (try " ^ quote (#1 \<^command_keyword>\corecursive\) ^ " instead of " ^ quote (#1 \<^command_keyword>\corec\) ^ ")") else if not (null const_transfer_goals) then error ("Transfer prover failed (try " ^ quote (#1 \<^command_keyword>\corecursive\) ^ " instead of " ^ quote (#1 \<^command_keyword>\corec\) ^ ")") else def_fun inner_fp_triple const_transfers [] lthy' end; fun corecursive_cmd int opts (raw_fixes, raw_eq) lthy = let val ((def_fun, (([Type (fpT_name, _)], prepareds, rho_datas, rho_transfer_goals), (inner_fp_triple, termin_goals), (const_transfers, const_transfer_goals))), lthy') = prepare_corec_ursive_cmd int true opts (raw_fixes, raw_eq) lthy; val (rho_transfer_goals', unprime_rho_transfer_and_folds) = @{map 3} (fn (_, _, _, _, _, _, _, _, _, _, _, _, _) => fn (_, rho_def) => prime_rho_transfer_goal lthy' fpT_name rho_def) prepareds rho_datas rho_transfer_goals |> split_list; in Proof.theorem NONE (fn [termin_thms, const_transfers', rho_transfers'] => let val remove_domain_condition = full_simplify (put_simpset HOL_basic_ss lthy' addsimps (@{thm True_implies_equals} :: termin_thms)); in def_fun (@{apply 3} (map remove_domain_condition) inner_fp_triple) (const_transfers @ const_transfers') (map2 (fn f => f) unprime_rho_transfer_and_folds rho_transfers') end) (map (map (rpair [])) [termin_goals, const_transfer_goals, rho_transfer_goals']) lthy' end; fun friend_of_corec_cmd ((raw_fun_name, raw_fun_T_opt), raw_eq) lthy = let val Const (fun_name, _) = Proof_Context.read_const {proper = true, strict = false} lthy raw_fun_name; val fake_lthy = lthy |> (case raw_fun_T_opt of SOME raw_T => Proof_Context.add_const_constraint (fun_name, SOME (Syntax.read_typ lthy raw_T)) | NONE => I) handle TYPE (s, _, _) => error s; val fun_b = Binding.name (Long_Name.base_name fun_name); val code_goal = Syntax.read_prop fake_lthy raw_eq; val fun_T = (case code_goal of \<^const>\Trueprop\ $ (Const (\<^const_name>\HOL.eq\, _) $ t $ _) => fastype_of (head_of t) | _ => ill_formed_equation_lhs_rhs lthy [code_goal]); val fun_t = Const (fun_name, fun_T); val (arg_Ts, res_T as Type (fpT_name, _)) = strip_type fun_T; val no_base = has_no_corec_info lthy fpT_name; val lthy1 = lthy |> no_base ? setup_base fpT_name; val lthy2 = lthy1 |> Variable.declare_typ fun_T; val ((old_corec_info, fp_b, version, Y, Z, preT, k_T, ssig_T, dead_pre_bnf, dead_k_bnf, sig_fp_sugar, ssig_fp_sugar, buffer), lthy3) = prepare_friend_corec fun_name fun_T lthy2; val friend_parse_info = friend_parse_info_of lthy3 arg_Ts res_T buffer; val parsed_eq = parse_corec_equation lthy3 [] code_goal; val (((rho, rho_def), (const_transfer_goals, rho_transfer_goal)), lthy4) = extract_rho_return_transfer_goals fun_b version dead_pre_bnf dead_k_bnf Y Z preT k_T ssig_T ssig_fp_sugar friend_parse_info fun_t parsed_eq lthy3; fun register_friend_extra_and_note_thms code_goal code_thm const_transfers k_T friend_info lthy5 = let val (corec_info, lthy6) = corec_info_of res_T lthy5; val fun_free = Free (Binding.name_of fun_b, fun_T); fun freeze_fun (t as Const (s, T)) = if s = fun_name andalso T = fun_T then fun_free else t | freeze_fun t = t; val eq = Term.map_aterms freeze_fun code_goal; val parsed_eq = parse_corec_equation lthy6 [fun_free] eq; val corec_parse_info = corec_parse_info_of lthy6 arg_Ts res_T buffer; val explored_eq = explore_corec_equation lthy6 false false fun_name fun_free corec_parse_info res_T parsed_eq; val ((_, corecUU_arg), _) = build_corecUU_arg_and_goals false fun_free explored_eq lthy6; val eq_corecUU = derive_eq_corecUU lthy6 corec_info fun_t corecUU_arg code_thm; val (eq_algrho, algrho_eq) = derive_eq_algrho lthy6 corec_info friend_info fun_t k_T code_goal const_transfers rho_def eq_corecUU; val ((_, [{cong_intro_pairs, coinduct, coinduct_attrs}]), lthy7) = lthy6 |> register_friend_extra fun_name eq_algrho algrho_eq |> register_coinduct_dynamic_friend fpT_name fun_name |> derive_and_update_coinduct_cong_intross [corec_info]; val cong_intros_pairs = AList.group (op =) cong_intro_pairs; val unique = derive_unique lthy7 Morphism.identity code_goal corec_info fpT_name eq_corecUU; val notes = [(codeN, [code_thm], []), (coinductN, [coinduct], coinduct_attrs), (cong_introsN, maps snd cong_intros_pairs, []), (uniqueN, [unique], [])] @ map (fn (thmN, thms) => (thmN, thms, [])) cong_intros_pairs @ (if Config.get lthy7 bnf_internals then [(eq_algrhoN, [eq_algrho], []), (eq_corecUUN, [eq_corecUU], [])] else []) |> map (fn (thmN, thms, attrs) => ((Binding.qualify true (Binding.name_of fun_b) (Binding.qualify false friendN (Binding.name thmN)), attrs), [(thms, [])])); in lthy7 |> Local_Theory.notes notes |> snd end; val (rho_transfer_goal', unprime_rho_transfer_and_fold) = prime_rho_transfer_goal lthy4 fpT_name rho_def rho_transfer_goal; in lthy4 |> Proof.theorem NONE (fn [[code_thm], const_transfers, [rho_transfer']] => register_friend_corec fun_name fp_b version Y Z k_T dead_k_bnf sig_fp_sugar ssig_fp_sugar fun_t rho (unprime_rho_transfer_and_fold rho_transfer') old_corec_info #-> register_friend_extra_and_note_thms code_goal code_thm const_transfers k_T) (map (map (rpair [])) [[code_goal], const_transfer_goals, [rho_transfer_goal']]) |> Proof.refine_singleton (Method.primitive_text (K I)) end; fun coinduction_upto_cmd (base_name, raw_fpT) lthy = let val fpT as Type (fpT_name, _) = Syntax.read_typ lthy raw_fpT; val no_base = has_no_corec_info lthy fpT_name; val (corec_info as {version, ...}, lthy1) = lthy |> corec_info_of fpT; val lthy2 = lthy1 |> no_base ? setup_base fpT_name; val ((changed, [{cong_intro_pairs, coinduct, coinduct_attrs}]), lthy3) = lthy2 |> derive_and_update_coinduct_cong_intross [corec_info]; val lthy4 = lthy3 |> (changed orelse no_base) ? update_coinduct_cong_intross_dynamic fpT_name; val cong_intros_pairs = AList.group (op =) cong_intro_pairs; val notes = [(cong_introsN, maps snd cong_intros_pairs, []), (coinduct_uptoN, [coinduct], coinduct_attrs)] @ map (fn (thmN, thms) => (thmN, thms, [])) cong_intros_pairs |> map (fn (thmN, thms, attrs) => (((Binding.qualify true base_name (Binding.qualify false ("v" ^ string_of_int version) (Binding.name thmN))), attrs), [(thms, [])])); in lthy4 |> Local_Theory.notes notes |> snd end; fun consolidate lthy = let val corec_infoss = map (corec_infos_of lthy o fst) (all_codatatype_extras_of lthy); val (changeds, lthy') = lthy |> fold_map (apfst fst oo derive_and_update_coinduct_cong_intross) corec_infoss; in if exists I changeds then lthy' else raise Same.SAME end; fun consolidate_global thy = SOME (Named_Target.theory_map consolidate thy) handle Same.SAME => NONE; val _ = Outer_Syntax.local_theory \<^command_keyword>\corec\ "define nonprimitive corecursive functions" ((Scan.optional (\<^keyword>\(\ |-- Parse.!!! (Parse.list1 corec_option_parser) --| \<^keyword>\)\) []) -- (Parse.vars --| Parse.where_ -- Parse.prop) >> uncurry (corec_cmd true)); val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\corecursive\ "define nonprimitive corecursive functions" ((Scan.optional (\<^keyword>\(\ |-- Parse.!!! (Parse.list1 corec_option_parser) --| \<^keyword>\)\) []) -- (Parse.vars --| Parse.where_ -- Parse.prop) >> uncurry (corecursive_cmd true)); val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\friend_of_corec\ "register a function as a legal context for nonprimitive corecursion" (Parse.const -- Scan.option (\<^keyword>\::\ |-- Parse.typ) --| Parse.where_ -- Parse.prop >> friend_of_corec_cmd); val _ = Outer_Syntax.local_theory \<^command_keyword>\coinduction_upto\ "derive a coinduction up-to principle and a corresponding congruence closure" (Parse.name --| \<^keyword>\:\ -- Parse.typ >> coinduction_upto_cmd); val _ = Theory.setup (Theory.at_begin consolidate_global); end; diff --git a/src/HOL/Tools/BNF/bnf_lfp.ML b/src/HOL/Tools/BNF/bnf_lfp.ML --- a/src/HOL/Tools/BNF/bnf_lfp.ML +++ b/src/HOL/Tools/BNF/bnf_lfp.ML @@ -1,1873 +1,1873 @@ (* Title: HOL/Tools/BNF/bnf_lfp.ML Author: Dmitriy Traytel, TU Muenchen Author: Andrei Popescu, TU Muenchen Copyright 2012 Datatype construction. *) signature BNF_LFP = sig val construct_lfp: mixfix list -> binding list -> binding list -> binding list -> binding list list -> binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list -> BNF_Comp.absT_info list -> local_theory -> BNF_FP_Util.fp_result * local_theory end; structure BNF_LFP : BNF_LFP = struct open BNF_Def open BNF_Util open BNF_Tactics open BNF_Comp open BNF_FP_Util open BNF_FP_Def_Sugar open BNF_LFP_Util open BNF_LFP_Tactics (*all BNFs have the same lives*) fun construct_lfp mixfixes map_bs rel_bs pred_bs set_bss0 bs resBs (resDs, Dss) bnfs absT_infos lthy = let val time = time lthy; val timer = time (Timer.startRealTimer ()); val live = live_of_bnf (hd bnfs); val n = length bnfs; (*active*) val ks = 1 upto n; val m = live - n; (*passive, if 0 don't generate a new BNF*) val internals = Config.get lthy bnf_internals; val b_names = map Binding.name_of bs; val b_name = mk_common_name b_names; val b = Binding.name b_name; fun mk_internal_of_b name = Binding.prefix_name (name ^ "_") #> Binding.prefix true b_name #> Binding.concealed; fun mk_internal_b name = mk_internal_of_b name b; fun mk_internal_bs name = map (mk_internal_of_b name) bs; val external_bs = map2 (Binding.prefix false) b_names bs |> not internals ? map Binding.concealed; val deads = fold (union (op =)) Dss resDs; val names_lthy = fold Variable.declare_typ deads lthy; val passives = map fst (subtract (op = o apsnd TFree) deads resBs); (* tvars *) val (((((passiveAs, activeAs), passiveBs), activeBs), passiveCs), activeCs) = names_lthy |> variant_tfrees passives ||>> mk_TFrees n ||>> variant_tfrees passives ||>> mk_TFrees n ||>> variant_tfrees passives ||>> mk_TFrees n |> fst; val allAs = passiveAs @ activeAs; val allBs' = passiveBs @ activeBs; val Ass = replicate n allAs; val allBs = passiveAs @ activeBs; val Bss = replicate n allBs; val allCs = passiveAs @ activeCs; val allCs' = passiveBs @ activeCs; val Css' = replicate n allCs'; (* types *) val dead_poss = map (fn x => if member (op =) deads (TFree x) then SOME (TFree x) else NONE) resBs; fun mk_param NONE passive = (hd passive, tl passive) | mk_param (SOME a) passive = (a, passive); val mk_params = fold_map mk_param dead_poss #> fst; fun mk_FTs Ts = map2 (fn Ds => mk_T_of_bnf Ds Ts) Dss bnfs; val (params, params') = `(map Term.dest_TFree) (mk_params passiveAs); val FTsAs = mk_FTs allAs; val FTsBs = mk_FTs allBs; val FTsCs = mk_FTs allCs; val BTs = map HOLogic.mk_setT activeAs; val B'Ts = map HOLogic.mk_setT activeBs; val B''Ts = map HOLogic.mk_setT activeCs; val sTs = map2 (curry op -->) FTsAs activeAs; val s'Ts = map2 (curry op -->) FTsBs activeBs; val s''Ts = map2 (curry op -->) FTsCs activeCs; val fTs = map2 (curry op -->) activeAs activeBs; val inv_fTs = map2 (curry op -->) activeBs activeAs; val self_fTs = map2 (curry op -->) activeAs activeAs; val gTs = map2 (curry op -->) activeBs activeCs; val all_gTs = map2 (curry op -->) allBs allCs'; (* terms *) val mapsAsAs = @{map 4} mk_map_of_bnf Dss Ass Ass bnfs; val mapsAsBs = @{map 4} mk_map_of_bnf Dss Ass Bss bnfs; val mapsBsCs' = @{map 4} mk_map_of_bnf Dss Bss Css' bnfs; val mapsAsCs' = @{map 4} mk_map_of_bnf Dss Ass Css' bnfs; fun mk_setss Ts = @{map 3} mk_sets_of_bnf (map (replicate live) Dss) (map (replicate live) (replicate n Ts)) bnfs; val setssAs = mk_setss allAs; val bd0s = @{map 3} mk_bd_of_bnf Dss Ass bnfs; val bds = @{map 3} (fn bd0 => fn Ds => fn bnf => mk_csum bd0 (mk_card_of (HOLogic.mk_UNIV (mk_T_of_bnf Ds (replicate live (fst (dest_relT (fastype_of bd0)))) bnf)))) bd0s Dss bnfs; val witss = map wits_of_bnf bnfs; val ((((((((zs, zs'), Bs), ss), fs), self_fs), all_gs), (xFs, xFs')), _) = lthy |> mk_Frees' "z" activeAs ||>> mk_Frees "B" BTs ||>> mk_Frees "s" sTs ||>> mk_Frees "f" fTs ||>> mk_Frees "f" self_fTs ||>> mk_Frees "g" all_gTs ||>> mk_Frees' "x" FTsAs; val passive_UNIVs = map HOLogic.mk_UNIV passiveAs; val active_UNIVs = map HOLogic.mk_UNIV activeAs; val passive_ids = map HOLogic.id_const passiveAs; val active_ids = map HOLogic.id_const activeAs; (* thms *) val bd0_card_orders = map bd_card_order_of_bnf bnfs; val bd0_Card_orders = map bd_Card_order_of_bnf bnfs; val bd0_Cinfinites = map bd_Cinfinite_of_bnf bnfs; val set_bd0ss = map set_bd_of_bnf bnfs; val bd_Card_order = @{thm Card_order_csum}; val bd_Card_orders = replicate n bd_Card_order; val bd_Cinfinites = map (fn thm => thm RS @{thm Cinfinite_csum1}) bd0_Cinfinites; val bd_Cnotzeros = map (fn thm => thm RS @{thm Cinfinite_Cnotzero}) bd_Cinfinites; val bd_Cinfinite = hd bd_Cinfinites; val set_bdss = map2 (fn set_bd0s => fn bd0_Card_order => map (fn thm => ctrans OF [thm, bd0_Card_order RS @{thm ordLeq_csum1}]) set_bd0s) set_bd0ss bd0_Card_orders; val in_bds = map in_bd_of_bnf bnfs; val sym_map_comps = map (fn bnf => map_comp0_of_bnf bnf RS sym) bnfs; val map_comps = map map_comp_of_bnf bnfs; val map_cong0s = map map_cong0_of_bnf bnfs; val map_id0s = map map_id0_of_bnf bnfs; val map_ids = map map_id_of_bnf bnfs; val set_mapss = map set_map_of_bnf bnfs; val rel_mono_strong0s = map rel_mono_strong0_of_bnf bnfs; val le_rel_OOs = map le_rel_OO_of_bnf bnfs; val timer = time (timer "Extracted terms & thms"); (* nonemptiness check *) fun new_wit X (wit: nonemptiness_witness) = subset (op =) (#I wit, (0 upto m - 1) @ map snd X); val all = m upto m + n - 1; fun enrich X = map_filter (fn i => (case find_first (fn (_, i') => i = i') X of NONE => (case find_index (new_wit X) (nth witss (i - m)) of ~1 => NONE | j => SOME (j, i)) | SOME ji => SOME ji)) all; val reachable = fixpoint (op =) enrich []; val _ = (case subtract (op =) (map snd reachable) all of [] => () | i :: _ => raise EMPTY_DATATYPE (Binding.name_of (nth bs (i - m)))); val wit_thms = flat (map2 (fn bnf => fn (j, _) => nth (wit_thmss_of_bnf bnf) j) bnfs reachable); val timer = time (timer "Checked nonemptiness"); (* derived thms *) (*map g1 ... gm g(m+1) ... g(m+n) (map id ... id f(m+1) ... f(m+n) x) = map g1 ... gm (g(m+1) o f(m+1)) ... (g(m+n) o f(m+n)) x*) fun mk_map_comp_id x mapAsBs mapBsCs mapAsCs map_comp0 = let val lhs = Term.list_comb (mapBsCs, all_gs) $ (Term.list_comb (mapAsBs, passive_ids @ fs) $ x); val rhs = Term.list_comb (mapAsCs, take m all_gs @ map HOLogic.mk_comp (drop m all_gs ~~ fs)) $ x; val vars = fold (Variable.add_free_names lthy) [lhs, rhs] []; in Goal.prove_sorry lthy vars [] (mk_Trueprop_eq (lhs, rhs)) (fn {context = ctxt, prems = _} => mk_map_comp_id_tac ctxt map_comp0) |> Thm.close_derivation \<^here> end; val map_comp_id_thms = @{map 5} mk_map_comp_id xFs mapsAsBs mapsBsCs' mapsAsCs' map_comps; (*forall a : set(m+1) x. f(m+1) a = a; ...; forall a : set(m+n) x. f(m+n) a = a ==> map id ... id f(m+1) ... f(m+n) x = x*) fun mk_map_cong0L x mapAsAs sets map_cong0 map_id = let fun mk_prem set f z z' = HOLogic.mk_Trueprop (mk_Ball (set $ x) (Term.absfree z' (HOLogic.mk_eq (f $ z, z)))); val prems = @{map 4} mk_prem (drop m sets) self_fs zs zs'; val goal = mk_Trueprop_eq (Term.list_comb (mapAsAs, passive_ids @ self_fs) $ x, x); val vars = fold (Variable.add_free_names lthy) (goal :: prems) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, goal)) (fn {context = ctxt, prems = _} => mk_map_cong0L_tac ctxt m map_cong0 map_id) |> Thm.close_derivation \<^here> end; val map_cong0L_thms = @{map 5} mk_map_cong0L xFs mapsAsAs setssAs map_cong0s map_ids; val in_mono'_thms = map (fn bnf => in_mono_of_bnf bnf OF (replicate m subset_refl)) bnfs; val in_cong'_thms = map (fn bnf => in_cong_of_bnf bnf OF (replicate m refl)) bnfs; val timer = time (timer "Derived simple theorems"); (* algebra *) val alg_bind = mk_internal_b algN; val alg_def_bind = (Thm.def_binding alg_bind, []); (*forall i = 1 ... n: (\x \ Fi_in UNIV .. UNIV B1 ... Bn. si x \ Bi)*) val alg_spec = let val ins = @{map 3} mk_in (replicate n (passive_UNIVs @ Bs)) setssAs FTsAs; fun mk_alg_conjunct B s X x x' = mk_Ball X (Term.absfree x' (HOLogic.mk_mem (s $ x, B))); val rhs = Library.foldr1 HOLogic.mk_conj (@{map 5} mk_alg_conjunct Bs ss ins xFs xFs') in fold_rev (Term.absfree o Term.dest_Free) (Bs @ ss) rhs end; val ((alg_free, (_, alg_def_free)), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> Local_Theory.define ((alg_bind, NoSyn), (alg_def_bind, alg_spec)) ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val alg = fst (Term.dest_Const (Morphism.term phi alg_free)); val alg_def = mk_unabs_def (2 * n) (HOLogic.mk_obj_eq (Morphism.thm phi alg_def_free)); fun mk_alg Bs ss = let val args = Bs @ ss; val Ts = map fastype_of args; val algT = Library.foldr (op -->) (Ts, HOLogic.boolT); in Term.list_comb (Const (alg, algT), args) end; val ((((((((zs, zs'), Bs), B's), ss), s's), fs), (xFs, xFs')), _) = lthy |> mk_Frees' "z" activeAs ||>> mk_Frees "B" BTs ||>> mk_Frees "B'" B'Ts ||>> mk_Frees "s" sTs ||>> mk_Frees "s'" s'Ts ||>> mk_Frees "f" fTs ||>> mk_Frees' "x" FTsAs; val alg_set_thms = let val alg_prem = HOLogic.mk_Trueprop (mk_alg Bs ss); fun mk_prem x set B = HOLogic.mk_Trueprop (mk_leq (set $ x) B); fun mk_concl s x B = mk_Trueprop_mem (s $ x, B); val premss = map2 ((fn x => fn sets => map2 (mk_prem x) (drop m sets) Bs)) xFs setssAs; val concls = @{map 3} mk_concl ss xFs Bs; val goals = map2 (fn prems => fn concl => Logic.list_implies (alg_prem :: prems, concl)) premss concls; in map (fn goal => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_alg_set_tac ctxt alg_def)) |> Thm.close_derivation \<^here>) goals end; val timer = time (timer "Algebra definition & thms"); val alg_not_empty_thms = let val alg_prem = HOLogic.mk_Trueprop (mk_alg Bs ss); val concls = map (HOLogic.mk_Trueprop o mk_not_empty) Bs; val goals = map (fn concl => Logic.mk_implies (alg_prem, concl)) concls; in map2 (fn goal => fn alg_set => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_alg_not_empty_tac ctxt alg_set alg_set_thms wit_thms)) |> Thm.close_derivation \<^here>) goals alg_set_thms end; val timer = time (timer "Proved nonemptiness"); (* morphism *) val mor_bind = mk_internal_b morN; val mor_def_bind = (Thm.def_binding mor_bind, []); (*fbetw) forall i = 1 ... n: (\x \ Bi. f x \ B'i)*) (*mor) forall i = 1 ... n: (\x \ Fi_in UNIV ... UNIV B1 ... Bn. f (s1 x) = s1' (Fi_map id ... id f1 ... fn x))*) val mor_spec = let fun mk_fbetw f B1 B2 z z' = mk_Ball B1 (Term.absfree z' (HOLogic.mk_mem (f $ z, B2))); fun mk_mor sets mapAsBs f s s' T x x' = mk_Ball (mk_in (passive_UNIVs @ Bs) sets T) (Term.absfree x' (HOLogic.mk_eq (f $ (s $ x), s' $ (Term.list_comb (mapAsBs, passive_ids @ fs) $ x)))); val rhs = HOLogic.mk_conj (Library.foldr1 HOLogic.mk_conj (@{map 5} mk_fbetw fs Bs B's zs zs'), Library.foldr1 HOLogic.mk_conj (@{map 8} mk_mor setssAs mapsAsBs fs ss s's FTsAs xFs xFs')) in fold_rev (Term.absfree o Term.dest_Free) (Bs @ ss @ B's @ s's @ fs) rhs end; val ((mor_free, (_, mor_def_free)), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> Local_Theory.define ((mor_bind, NoSyn), (mor_def_bind, mor_spec)) ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val mor = fst (Term.dest_Const (Morphism.term phi mor_free)); val mor_def = mk_unabs_def (5 * n) (HOLogic.mk_obj_eq (Morphism.thm phi mor_def_free)); fun mk_mor Bs1 ss1 Bs2 ss2 fs = let val args = Bs1 @ ss1 @ Bs2 @ ss2 @ fs; val Ts = map fastype_of (Bs1 @ ss1 @ Bs2 @ ss2 @ fs); val morT = Library.foldr (op -->) (Ts, HOLogic.boolT); in Term.list_comb (Const (mor, morT), args) end; val (((((((((((Bs, Bs_copy), B's), B''s), ss), s's), s''s), fs), fs_copy), gs), xFs), _) = lthy |> mk_Frees "B" BTs ||>> mk_Frees "B" BTs ||>> mk_Frees "B'" B'Ts ||>> mk_Frees "B''" B''Ts ||>> mk_Frees "s" sTs ||>> mk_Frees "s'" s'Ts ||>> mk_Frees "s''" s''Ts ||>> mk_Frees "f" fTs ||>> mk_Frees "f" fTs ||>> mk_Frees "g" gTs ||>> mk_Frees "x" FTsAs; val morE_thms = let val prem = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs); fun mk_elim_prem sets x T = HOLogic.mk_Trueprop (HOLogic.mk_mem (x, mk_in (passive_UNIVs @ Bs) sets T)); fun mk_elim_goal sets mapAsBs f s s' x T = Logic.list_implies ([prem, mk_elim_prem sets x T], mk_Trueprop_eq (f $ (s $ x), s' $ Term.list_comb (mapAsBs, passive_ids @ fs @ [x]))); val elim_goals = @{map 7} mk_elim_goal setssAs mapsAsBs fs ss s's xFs FTsAs; fun prove goal = Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_mor_elim_tac ctxt mor_def)) |> Thm.close_derivation \<^here>; in map prove elim_goals end; val mor_incl_thm = let val prems = map2 (HOLogic.mk_Trueprop oo mk_leq) Bs Bs_copy; val concl = HOLogic.mk_Trueprop (mk_mor Bs ss Bs_copy ss active_ids); val vars = fold (Variable.add_free_names lthy) (concl :: prems) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, concl)) (fn {context = ctxt, prems = _} => mk_mor_incl_tac ctxt mor_def map_ids) |> Thm.close_derivation \<^here> end; val mor_comp_thm = let val prems = [HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs), HOLogic.mk_Trueprop (mk_mor B's s's B''s s''s gs)]; val concl = HOLogic.mk_Trueprop (mk_mor Bs ss B''s s''s (map2 (curry HOLogic.mk_comp) gs fs)); val vars = fold (Variable.add_free_names lthy) (concl :: prems) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, concl)) (fn {context = ctxt, prems = _} => mk_mor_comp_tac ctxt mor_def set_mapss map_comp_id_thms) |> Thm.close_derivation \<^here> end; val mor_cong_thm = let val prems = map HOLogic.mk_Trueprop (map2 (curry HOLogic.mk_eq) fs_copy fs @ [mk_mor Bs ss B's s's fs]) val concl = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs_copy); val vars = fold (Variable.add_free_names lthy) (concl :: prems) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, concl)) (fn {context = ctxt, prems = _} => (hyp_subst_tac ctxt THEN' assume_tac ctxt) 1) |> Thm.close_derivation \<^here> end; val mor_str_thm = let val maps = map2 (fn Ds => fn bnf => Term.list_comb (mk_map_of_bnf Ds (passiveAs @ FTsAs) allAs bnf, passive_ids @ ss)) Dss bnfs; val goal = HOLogic.mk_Trueprop (mk_mor (map HOLogic.mk_UNIV FTsAs) maps active_UNIVs ss ss); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_mor_str_tac ctxt ks mor_def) |> Thm.close_derivation \<^here> end; val mor_UNIV_thm = let fun mk_conjunct mapAsBs f s s' = HOLogic.mk_eq (HOLogic.mk_comp (f, s), HOLogic.mk_comp (s', Term.list_comb (mapAsBs, passive_ids @ fs))); val lhs = mk_mor active_UNIVs ss (map HOLogic.mk_UNIV activeBs) s's fs; val rhs = Library.foldr1 HOLogic.mk_conj (@{map 4} mk_conjunct mapsAsBs fs ss s's); val vars = fold (Variable.add_free_names lthy) [lhs, rhs] []; in Goal.prove_sorry lthy vars [] (mk_Trueprop_eq (lhs, rhs)) (fn {context = ctxt, prems = _} => mk_mor_UNIV_tac ctxt m morE_thms mor_def) |> Thm.close_derivation \<^here> end; val timer = time (timer "Morphism definition & thms"); (* bounds *) val sum_bd = Library.foldr1 (uncurry mk_csum) bds; val sum_bdT = fst (dest_relT (fastype_of sum_bd)); val (sum_bdT_params, sum_bdT_params') = `(map TFree) (Term.add_tfreesT sum_bdT []); val (lthy, sbd, sbd_Cinfinite, sbd_Card_order, set_sbdss, in_sbds) = if n = 1 then (lthy, sum_bd, bd_Cinfinite, bd_Card_order, set_bdss, in_bds) else let val sbdT_bind = mk_internal_b sum_bdTN; val ((sbdT_name, (sbdT_glob_info, sbdT_loc_info)), lthy) = typedef (sbdT_bind, sum_bdT_params', NoSyn) (HOLogic.mk_UNIV sum_bdT) NONE (fn ctxt => EVERY' [rtac ctxt exI, rtac ctxt UNIV_I] 1) lthy; val sbdT = Type (sbdT_name, sum_bdT_params); val Abs_sbdT = Const (#Abs_name sbdT_glob_info, sum_bdT --> sbdT); val sbd_bind = mk_internal_b sum_bdN; val sbd_def_bind = (Thm.def_binding sbd_bind, []); val sbd_spec = mk_dir_image sum_bd Abs_sbdT; val ((sbd_free, (_, sbd_def_free)), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> Local_Theory.define ((sbd_bind, NoSyn), (sbd_def_bind, sbd_spec)) ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val sbd_def = HOLogic.mk_obj_eq (Morphism.thm phi sbd_def_free); val sbd = Const (fst (Term.dest_Const (Morphism.term phi sbd_free)), mk_relT (`I sbdT)); val Abs_sbdT_inj = mk_Abs_inj_thm (#Abs_inject sbdT_loc_info); val sum_Cinfinite = mk_sum_Cinfinite bd_Cinfinites; val sum_Card_order = sum_Cinfinite RS conjunct2; val sbd_ordIso = @{thm ssubst_Pair_rhs} OF [@{thm dir_image} OF [Abs_sbdT_inj, sum_Card_order], sbd_def]; val sbd_Cinfinite = @{thm Cinfinite_cong} OF [sbd_ordIso, sum_Cinfinite]; val sbd_Card_order = sbd_Cinfinite RS conjunct2; fun mk_set_sbd i bd_Card_order bds = map (fn thm => @{thm ordLeq_ordIso_trans} OF [bd_Card_order RS mk_ordLeq_csum n i thm, sbd_ordIso]) bds; val set_sbdss = @{map 3} mk_set_sbd ks bd_Card_orders set_bdss; fun mk_in_bd_sum i Co Cnz bd = Cnz RS ((@{thm ordLeq_ordIso_trans} OF [Co RS mk_ordLeq_csum n i (Co RS @{thm ordLeq_refl}), sbd_ordIso]) RS (bd RS @{thm ordLeq_transitive[OF _ cexp_mono2_Cnotzero[OF _ Card_order_csum]]})); val in_sbds = @{map 4} mk_in_bd_sum ks bd_Card_orders bd_Cnotzeros in_bds; in (lthy, sbd, sbd_Cinfinite, sbd_Card_order, set_sbdss, in_sbds) end; val sbd_Cnotzero = sbd_Cinfinite RS @{thm Cinfinite_Cnotzero}; val suc_bd = mk_cardSuc sbd; val field_suc_bd = mk_Field suc_bd; val suc_bdT = fst (dest_relT (fastype_of suc_bd)); fun mk_Asuc_bd [] = mk_cexp ctwo suc_bd | mk_Asuc_bd As = mk_cexp (mk_csum (Library.foldr1 (uncurry mk_csum) (map mk_card_of As)) ctwo) suc_bd; val suc_bd_Card_order = sbd_Card_order RS @{thm cardSuc_Card_order}; val suc_bd_Cinfinite = sbd_Cinfinite RS @{thm Cinfinite_cardSuc}; val suc_bd_Cnotzero = suc_bd_Cinfinite RS @{thm Cinfinite_Cnotzero}; val suc_bd_worel = suc_bd_Card_order RS @{thm Card_order_wo_rel} val basis_Asuc = if m = 0 then @{thm ordLeq_refl[OF Card_order_ctwo]} else @{thm ordLeq_csum2[OF Card_order_ctwo]}; val Asuc_bd_Cinfinite = suc_bd_Cinfinite RS (basis_Asuc RS @{thm Cinfinite_cexp}); val suc_bd_Asuc_bd = @{thm ordLess_ordLeq_trans[OF ordLess_ctwo_cexp cexp_mono1]} OF [suc_bd_Card_order, basis_Asuc, suc_bd_Card_order]; val Asuc_bd = mk_Asuc_bd passive_UNIVs; val Asuc_bdT = fst (dest_relT (fastype_of Asuc_bd)); val II_BTs = replicate n (HOLogic.mk_setT Asuc_bdT); val II_sTs = map2 (fn Ds => fn bnf => mk_T_of_bnf Ds (passiveAs @ replicate n Asuc_bdT) bnf --> Asuc_bdT) Dss bnfs; val ((((((Bs, ss), idxs), Asi_name), (idx, idx')), (jdx, jdx')), _) = lthy |> mk_Frees "B" BTs ||>> mk_Frees "s" sTs ||>> mk_Frees "i" (replicate n suc_bdT) ||>> (fn ctxt => apfst the_single (mk_fresh_names ctxt 1 "Asi")) ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "i") suc_bdT ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "j") suc_bdT; val suc_bd_limit_thm = let val prem = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map (fn idx => HOLogic.mk_mem (idx, field_suc_bd)) idxs)); fun mk_conjunct idx = HOLogic.mk_conj (mk_not_eq idx jdx, HOLogic.mk_mem (HOLogic.mk_prod (idx, jdx), suc_bd)); val concl = HOLogic.mk_Trueprop (mk_Bex field_suc_bd (Term.absfree jdx' (Library.foldr1 HOLogic.mk_conj (map mk_conjunct idxs)))); val vars = fold (Variable.add_free_names lthy) [prem, concl] []; in Goal.prove_sorry lthy vars [] (Logic.list_implies ([prem], concl)) (fn {context = ctxt, prems = _} => mk_bd_limit_tac ctxt n suc_bd_Cinfinite) |> Thm.close_derivation \<^here> end; val timer = time (timer "Bounds"); (* minimal algebra *) fun mk_minG Asi i k = mk_UNION (mk_underS suc_bd $ i) (Term.absfree jdx' (mk_nthN n (Asi $ jdx) k)); fun mk_minH_component Asi i sets Ts s k = HOLogic.mk_binop \<^const_name>\sup\ (mk_minG Asi i k, mk_image s $ mk_in (passive_UNIVs @ map (mk_minG Asi i) ks) sets Ts); fun mk_min_algs ss = let val BTs = map (range_type o fastype_of) ss; val Ts = passiveAs @ BTs; val (Asi, Asi') = `Free (Asi_name, suc_bdT --> Library.foldr1 HOLogic.mk_prodT (map HOLogic.mk_setT BTs)); in mk_worec suc_bd (Term.absfree Asi' (Term.absfree idx' (HOLogic.mk_tuple (@{map 4} (mk_minH_component Asi idx) (mk_setss Ts) (mk_FTs Ts) ss ks)))) end; val (min_algs_thms, min_algs_mono_thms, card_of_min_algs_thm, least_min_algs_thm) = let val i_field = HOLogic.mk_mem (idx, field_suc_bd); val min_algs = mk_min_algs ss; val min_algss = map (fn k => mk_nthN n (min_algs $ idx) k) ks; val concl = HOLogic.mk_Trueprop (HOLogic.mk_eq (min_algs $ idx, HOLogic.mk_tuple (@{map 4} (mk_minH_component min_algs idx) setssAs FTsAs ss ks))); val goal = Logic.mk_implies (HOLogic.mk_Trueprop i_field, concl); val vars = Variable.add_free_names lthy goal []; val min_algs_thm = Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_min_algs_tac ctxt suc_bd_worel in_cong'_thms) |> Thm.close_derivation \<^here>; val min_algs_thms = map (fn k => min_algs_thm RS mk_nthI n k) ks; fun mk_mono_goal min_alg = HOLogic.mk_Trueprop (mk_relChain suc_bd (Term.absfree idx' min_alg)); val monos = map2 (fn goal => fn min_algs => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_min_algs_mono_tac ctxt min_algs)) |> Thm.close_derivation \<^here>) (map mk_mono_goal min_algss) min_algs_thms; fun mk_card_conjunct min_alg = mk_ordLeq (mk_card_of min_alg) Asuc_bd; val card_conjunction = Library.foldr1 HOLogic.mk_conj (map mk_card_conjunct min_algss); val card_cT = Thm.ctyp_of lthy suc_bdT; val card_ct = Thm.cterm_of lthy (Term.absfree idx' card_conjunction); val card_of = let val goal = HOLogic.mk_Trueprop (HOLogic.mk_imp (i_field, card_conjunction)); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_min_algs_card_of_tac ctxt card_cT card_ct m suc_bd_worel min_algs_thms in_sbds sbd_Card_order sbd_Cnotzero suc_bd_Card_order suc_bd_Cinfinite suc_bd_Cnotzero suc_bd_Asuc_bd Asuc_bd_Cinfinite) |> Thm.close_derivation \<^here> end; val least_prem = HOLogic.mk_Trueprop (mk_alg Bs ss); val least_conjunction = Library.foldr1 HOLogic.mk_conj (map2 mk_leq min_algss Bs); val least_cT = Thm.ctyp_of lthy suc_bdT; val least_ct = Thm.cterm_of lthy (Term.absfree idx' least_conjunction); val least = let val goal = Logic.mk_implies (least_prem, HOLogic.mk_Trueprop (HOLogic.mk_imp (i_field, least_conjunction))); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_min_algs_least_tac ctxt least_cT least_ct suc_bd_worel min_algs_thms alg_set_thms) |> Thm.close_derivation \<^here> end; in (min_algs_thms, monos, card_of, least) end; val timer = time (timer "min_algs definition & thms"); val min_alg_binds = mk_internal_bs min_algN; fun min_alg_bind i = nth min_alg_binds (i - 1); val min_alg_def_bind = rpair [] o Thm.def_binding o min_alg_bind; fun min_alg_spec i = let val rhs = mk_UNION (field_suc_bd) (Term.absfree idx' (mk_nthN n (mk_min_algs ss $ idx) i)); in fold_rev (Term.absfree o Term.dest_Free) ss rhs end; val ((min_alg_frees, (_, min_alg_def_frees)), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> fold_map (fn i => Local_Theory.define ((min_alg_bind i, NoSyn), (min_alg_def_bind i, min_alg_spec i))) ks |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val min_algs = map (fst o Term.dest_Const o Morphism.term phi) min_alg_frees; val min_alg_defs = map (fn def => mk_unabs_def n (HOLogic.mk_obj_eq (Morphism.thm phi def))) min_alg_def_frees; fun mk_min_alg ss i = let val T = HOLogic.mk_setT (range_type (fastype_of (nth ss (i - 1)))) val Ts = map fastype_of ss; val min_algT = Library.foldr (op -->) (Ts, T); in Term.list_comb (Const (nth min_algs (i - 1), min_algT), ss) end; val min_algs = map (mk_min_alg ss) ks; val ((Bs, ss), _) = lthy |> mk_Frees "B" BTs ||>> mk_Frees "s" sTs; val (alg_min_alg_thm, card_of_min_alg_thms, least_min_alg_thms, mor_incl_min_alg_thm) = let val alg_min_alg = let val goal = HOLogic.mk_Trueprop (mk_alg min_algs ss); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_alg_min_alg_tac ctxt m alg_def min_alg_defs suc_bd_limit_thm sbd_Cinfinite set_sbdss min_algs_thms min_algs_mono_thms) |> Thm.close_derivation \<^here> end; fun mk_card_of_thm min_alg def = let val goal = HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of min_alg) Asuc_bd); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_card_of_min_alg_tac ctxt def card_of_min_algs_thm suc_bd_Card_order suc_bd_Asuc_bd Asuc_bd_Cinfinite) |> Thm.close_derivation \<^here> end; fun mk_least_thm min_alg B def = let val prem = HOLogic.mk_Trueprop (mk_alg Bs ss); val goal = Logic.mk_implies (prem, HOLogic.mk_Trueprop (mk_leq min_alg B)); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_least_min_alg_tac ctxt def least_min_algs_thm) |> Thm.close_derivation \<^here> end; val leasts = @{map 3} mk_least_thm min_algs Bs min_alg_defs; val incl = let val prem = HOLogic.mk_Trueprop (mk_alg Bs ss); val goal = Logic.mk_implies (prem, HOLogic.mk_Trueprop (mk_mor min_algs ss Bs ss active_ids)); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => EVERY' (rtac ctxt mor_incl_thm :: map (etac ctxt) leasts) 1) |> Thm.close_derivation \<^here> end; in (alg_min_alg, map2 mk_card_of_thm min_algs min_alg_defs, leasts, incl) end; val timer = time (timer "Minimal algebra definition & thms"); val II_repT = HOLogic.mk_prodT (HOLogic.mk_tupleT II_BTs, HOLogic.mk_tupleT II_sTs); val IIT_bind = mk_internal_b IITN; val ((IIT_name, (IIT_glob_info, IIT_loc_info)), lthy) = typedef (IIT_bind, params, NoSyn) (HOLogic.mk_UNIV II_repT) NONE (fn ctxt => EVERY' [rtac ctxt exI, rtac ctxt UNIV_I] 1) lthy; val IIT = Type (IIT_name, params'); val Abs_IIT = Const (#Abs_name IIT_glob_info, II_repT --> IIT); val Rep_IIT = Const (#Rep_name IIT_glob_info, IIT --> II_repT); val Abs_IIT_inverse_thm = UNIV_I RS #Abs_inverse IIT_loc_info; val initT = IIT --> Asuc_bdT; val active_initTs = replicate n initT; val init_FTs = map2 (fn Ds => mk_T_of_bnf Ds (passiveAs @ active_initTs)) Dss bnfs; val init_fTs = map (fn T => initT --> T) activeAs; val ((((II_Bs, II_ss), (iidx, iidx')), init_xFs), _) = lthy |> mk_Frees "IIB" II_BTs ||>> mk_Frees "IIs" II_sTs ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "i") IIT ||>> mk_Frees "x" init_FTs; val II = HOLogic.mk_Collect (fst iidx', IIT, list_exists_free (II_Bs @ II_ss) (HOLogic.mk_conj (HOLogic.mk_eq (iidx, Abs_IIT $ (HOLogic.mk_prod (HOLogic.mk_tuple II_Bs, HOLogic.mk_tuple II_ss))), mk_alg II_Bs II_ss))); val select_Bs = map (mk_nthN n (HOLogic.mk_fst (Rep_IIT $ iidx))) ks; val select_ss = map (mk_nthN n (HOLogic.mk_snd (Rep_IIT $ iidx))) ks; val str_init_binds = mk_internal_bs str_initN; fun str_init_bind i = nth str_init_binds (i - 1); val str_init_def_bind = rpair [] o Thm.def_binding o str_init_bind; fun str_init_spec i = let val init_xF = nth init_xFs (i - 1) val select_s = nth select_ss (i - 1); val map = mk_map_of_bnf (nth Dss (i - 1)) (passiveAs @ active_initTs) (passiveAs @ replicate n Asuc_bdT) (nth bnfs (i - 1)); val map_args = passive_ids @ replicate n (mk_rapp iidx Asuc_bdT); val rhs = select_s $ (Term.list_comb (map, map_args) $ init_xF); in fold_rev (Term.absfree o Term.dest_Free) [init_xF, iidx] rhs end; val ((str_init_frees, (_, str_init_def_frees)), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> fold_map (fn i => Local_Theory.define ((str_init_bind i, NoSyn), (str_init_def_bind i, str_init_spec i))) ks |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val str_inits = map (Term.subst_atomic_types (map (`(Morphism.typ phi)) params') o Morphism.term phi) str_init_frees; val str_init_defs = map (fn def => mk_unabs_def 2 (HOLogic.mk_obj_eq (Morphism.thm phi def))) str_init_def_frees; val car_inits = map (mk_min_alg str_inits) ks; val (((((((((Bs, ss), Asuc_fs), (iidx, iidx')), init_xs), (init_xFs, init_xFs')), init_fs), init_fs_copy), init_phis), _) = lthy |> mk_Frees "B" BTs ||>> mk_Frees "s" sTs ||>> mk_Frees "f" (map (fn T => Asuc_bdT --> T) activeAs) ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "i") IIT ||>> mk_Frees "ix" active_initTs ||>> mk_Frees' "x" init_FTs ||>> mk_Frees "f" init_fTs ||>> mk_Frees "f" init_fTs ||>> mk_Frees "P" (replicate n (mk_pred1T initT)); val alg_init_thm = infer_instantiate' lthy (map (SOME o Thm.cterm_of lthy) str_inits) alg_min_alg_thm; val alg_select_thm = Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (mk_Ball II (Term.absfree iidx' (mk_alg select_Bs select_ss)))) (fn {context = ctxt, prems = _} => mk_alg_select_tac ctxt Abs_IIT_inverse_thm) |> Thm.close_derivation \<^here>; val mor_select_thm = let val i_prem = mk_Trueprop_mem (iidx, II); val mor_prem = HOLogic.mk_Trueprop (mk_mor select_Bs select_ss active_UNIVs ss Asuc_fs); val prems = [i_prem, mor_prem]; val concl = HOLogic.mk_Trueprop (mk_mor car_inits str_inits active_UNIVs ss (map (fn f => HOLogic.mk_comp (f, mk_rapp iidx Asuc_bdT)) Asuc_fs)); val vars = fold (Variable.add_free_names lthy) (concl :: prems) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, concl)) (fn {context = ctxt, prems = _} => mk_mor_select_tac ctxt mor_def mor_cong_thm mor_comp_thm mor_incl_min_alg_thm alg_def alg_select_thm alg_set_thms set_mapss str_init_defs) |> Thm.close_derivation \<^here> end; val init_unique_mor_thms = let val prems = map2 (HOLogic.mk_Trueprop oo curry HOLogic.mk_mem) init_xs car_inits val mor_prems = map HOLogic.mk_Trueprop [mk_mor car_inits str_inits Bs ss init_fs, mk_mor car_inits str_inits Bs ss init_fs_copy]; fun mk_fun_eq f g x = HOLogic.mk_eq (f $ x, g $ x); val unique = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (@{map 3} mk_fun_eq init_fs init_fs_copy init_xs)); val cts = map (Thm.cterm_of lthy) ss; val all_prems = prems @ mor_prems; val vars = fold (Variable.add_free_names lthy) (unique :: all_prems) []; val unique_mor = Goal.prove_sorry lthy vars [] (Logic.list_implies (all_prems, unique)) (fn {context = ctxt, prems = _} => mk_init_unique_mor_tac ctxt cts m alg_def alg_init_thm least_min_alg_thms in_mono'_thms alg_set_thms morE_thms map_cong0s) |> Thm.close_derivation \<^here>; in split_conj_thm unique_mor end; val init_setss = mk_setss (passiveAs @ active_initTs); val active_init_setss = map (drop m) init_setss; val init_ins = map2 (fn sets => mk_in (passive_UNIVs @ car_inits) sets) init_setss init_FTs; fun mk_closed phis = let fun mk_conjunct phi str_init init_sets init_in x x' = let val prem = Library.foldr1 HOLogic.mk_conj (map2 (fn set => mk_Ball (set $ x)) init_sets phis); val concl = phi $ (str_init $ x); in mk_Ball init_in (Term.absfree x' (HOLogic.mk_imp (prem, concl))) end; in Library.foldr1 HOLogic.mk_conj (@{map 6} mk_conjunct phis str_inits active_init_setss init_ins init_xFs init_xFs') end; val init_induct_thm = let val prem = HOLogic.mk_Trueprop (mk_closed init_phis); val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_Ball car_inits init_phis)); val vars = fold (Variable.add_free_names lthy) [concl, prem] []; in Goal.prove_sorry lthy vars [] (Logic.mk_implies (prem, concl)) (fn {context = ctxt, prems = _} => mk_init_induct_tac ctxt m alg_def alg_init_thm least_min_alg_thms alg_set_thms) |> Thm.close_derivation \<^here> end; val timer = time (timer "Initiality definition & thms"); val ((T_names, (T_glob_infos, T_loc_infos)), lthy) = lthy |> @{fold_map 3} (fn b => fn mx => fn car_init => typedef (b, params, mx) car_init NONE (fn ctxt => EVERY' [rtac ctxt iffD2, rtac ctxt @{thm ex_in_conv}, resolve_tac ctxt alg_not_empty_thms, rtac ctxt alg_init_thm] 1)) bs mixfixes car_inits |>> apsnd split_list o split_list; val Ts = map (fn name => Type (name, params')) T_names; fun mk_Ts passive = map (Term.typ_subst_atomic (passiveAs ~~ passive)) Ts; val Ts' = mk_Ts passiveBs; val Rep_Ts = map2 (fn info => fn T => Const (#Rep_name info, T --> initT)) T_glob_infos Ts; val Abs_Ts = map2 (fn info => fn T => Const (#Abs_name info, initT --> T)) T_glob_infos Ts; val type_defs = map #type_definition T_loc_infos; val Reps = map #Rep T_loc_infos; val Rep_inverses = map #Rep_inverse T_loc_infos; val Abs_inverses = map #Abs_inverse T_loc_infos; val timer = time (timer "THE TYPEDEFs & Rep/Abs thms"); val UNIVs = map HOLogic.mk_UNIV Ts; val FTs = mk_FTs (passiveAs @ Ts); val FTs' = mk_FTs (passiveBs @ Ts'); fun mk_set_Ts T = passiveAs @ replicate n (HOLogic.mk_setT T); val setFTss = map (mk_FTs o mk_set_Ts) passiveAs; val FTs_setss = mk_setss (passiveAs @ Ts); val FTs'_setss = mk_setss (passiveBs @ Ts'); val map_FT_inits = map2 (fn Ds => mk_map_of_bnf Ds (passiveAs @ Ts) (passiveAs @ active_initTs)) Dss bnfs; val fTs = map2 (curry op -->) Ts activeAs; val foldT = Library.foldr1 HOLogic.mk_prodT (map2 (curry op -->) Ts activeAs); val ((ss, (fold_f, fold_f')), _) = lthy |> mk_Frees "s" sTs ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "f") foldT; fun ctor_bind i = nth external_bs (i - 1) |> Binding.prefix_name (ctorN ^ "_"); val ctor_def_bind = rpair [] o Binding.concealed o Thm.def_binding o ctor_bind; fun ctor_spec abs str map_FT_init = Library.foldl1 HOLogic.mk_comp [abs, str, Term.list_comb (map_FT_init, map HOLogic.id_const passiveAs @ Rep_Ts)]; val ((ctor_frees, (_, ctor_def_frees)), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> @{fold_map 4} (fn i => fn abs => fn str => fn mapx => Local_Theory.define ((ctor_bind i, NoSyn), (ctor_def_bind i, ctor_spec abs str mapx))) ks Abs_Ts str_inits map_FT_inits |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; fun mk_ctors passive = map (Term.subst_atomic_types (map (Morphism.typ phi) params' ~~ (mk_params passive)) o Morphism.term phi) ctor_frees; val ctors = mk_ctors passiveAs; val ctor's = mk_ctors passiveBs; val ctor_defs = map (fn def => HOLogic.mk_obj_eq (Morphism.thm phi def)) ctor_def_frees; val (mor_Rep_thm, mor_Abs_thm) = let val defs = mor_def :: ctor_defs; val mor_Rep = Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (mk_mor UNIVs ctors car_inits str_inits Rep_Ts)) (fn {context = ctxt, prems = _} => mk_mor_Rep_tac ctxt m defs Reps Abs_inverses alg_min_alg_thm alg_set_thms set_mapss) |> Thm.close_derivation \<^here>; fun mk_ct initFT str abs = Term.absdummy initFT (abs $ (str $ Bound 0)) val cts = @{map 3} (Thm.cterm_of lthy ooo mk_ct) init_FTs str_inits Abs_Ts; val mor_Abs = Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (mk_mor car_inits str_inits UNIVs ctors Abs_Ts)) (fn {context = ctxt, prems = _} => mk_mor_Abs_tac ctxt cts defs Abs_inverses map_comp_id_thms map_cong0L_thms) |> Thm.close_derivation \<^here>; in (mor_Rep, mor_Abs) end; val timer = time (timer "ctor definitions & thms"); val fold_fun = Term.absfree fold_f' (mk_mor UNIVs ctors active_UNIVs ss (map (mk_nthN n fold_f) ks)); val foldx = HOLogic.choice_const foldT $ fold_fun; fun fold_bind i = nth external_bs (i - 1) |> Binding.prefix_name (ctor_foldN ^ "_"); val fold_def_bind = rpair [] o Binding.concealed o Thm.def_binding o fold_bind; fun fold_spec i = fold_rev (Term.absfree o Term.dest_Free) ss (mk_nthN n foldx i); val ((fold_frees, (_, fold_def_frees)), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> fold_map (fn i => Local_Theory.define ((fold_bind i, NoSyn), (fold_def_bind i, fold_spec i))) ks |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val folds = map (Morphism.term phi) fold_frees; val fold_names = map (fst o dest_Const) folds; fun mk_folds passives actives = @{map 3} (fn name => fn T => fn active => Const (name, Library.foldr (op -->) (map2 (curry op -->) (mk_FTs (passives @ actives)) actives, T --> active))) fold_names (mk_Ts passives) actives; fun mk_fold Ts ss i = Term.list_comb (Const (nth fold_names (i - 1), Library.foldr (op -->) (map fastype_of ss, nth Ts (i - 1) --> range_type (fastype_of (nth ss (i - 1))))), ss); val fold_defs = map (fn def => mk_unabs_def n (HOLogic.mk_obj_eq (Morphism.thm phi def))) fold_def_frees; (* algebra copies *) val ((((((Bs, B's), ss), s's), inv_fs), fs), _) = lthy |> mk_Frees "B" BTs ||>> mk_Frees "B'" B'Ts ||>> mk_Frees "s" sTs ||>> mk_Frees "s'" s'Ts ||>> mk_Frees "f" inv_fTs ||>> mk_Frees "f" fTs; val copy_thm = let val prems = HOLogic.mk_Trueprop (mk_alg Bs ss) :: @{map 3} (HOLogic.mk_Trueprop ooo mk_bij_betw) inv_fs B's Bs; val concl = HOLogic.mk_Trueprop (list_exists_free s's (HOLogic.mk_conj (mk_alg B's s's, mk_mor B's s's Bs ss inv_fs))); val vars = fold (Variable.add_free_names lthy) (concl :: prems) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, concl)) (fn {context = ctxt, prems = _} => mk_copy_tac ctxt m alg_def mor_def alg_set_thms set_mapss) |> Thm.close_derivation \<^here> end; val init_ex_mor_thm = let val goal = HOLogic.mk_Trueprop (list_exists_free fs (mk_mor UNIVs ctors active_UNIVs ss fs)); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_init_ex_mor_tac ctxt Abs_IIT_inverse_thm (alg_min_alg_thm RS copy_thm) card_of_min_alg_thms mor_Rep_thm mor_comp_thm mor_select_thm mor_incl_thm) |> Thm.close_derivation \<^here> end; val mor_fold_thm = let val mor_cong = mor_cong_thm OF (map (mk_nth_conv n) ks); val cT = Thm.ctyp_of lthy foldT; val ct = Thm.cterm_of lthy fold_fun val goal = HOLogic.mk_Trueprop (mk_mor UNIVs ctors active_UNIVs ss (map (mk_fold Ts ss) ks)); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_mor_fold_tac ctxt cT ct fold_defs init_ex_mor_thm mor_cong) |> Thm.close_derivation \<^here> end; val ctor_fold_thms = map (fn morE => rule_by_tactic lthy ((rtac lthy CollectI THEN' CONJ_WRAP' (K (rtac lthy @{thm subset_UNIV})) (1 upto m + n)) 1) (mor_fold_thm RS morE)) morE_thms; val (fold_unique_mor_thms, fold_unique_mor_thm) = let val prem = HOLogic.mk_Trueprop (mk_mor UNIVs ctors active_UNIVs ss fs); fun mk_fun_eq f i = HOLogic.mk_eq (f, mk_fold Ts ss i); val unique = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_fun_eq fs ks)); val vars = fold (Variable.add_free_names lthy) [prem, unique] []; val unique_mor = Goal.prove_sorry lthy vars [] (Logic.mk_implies (prem, unique)) (fn {context = ctxt, prems = _} => mk_fold_unique_mor_tac ctxt type_defs init_unique_mor_thms Reps mor_comp_thm mor_Abs_thm mor_fold_thm) |> Thm.close_derivation \<^here>; in `split_conj_thm unique_mor end; val (ctor_fold_unique_thms, ctor_fold_unique_thm) = `split_conj_thm (mk_conjIN n RS (mor_UNIV_thm RS iffD2 RS fold_unique_mor_thm)) val fold_ctor_thms = map (fn thm => (mor_incl_thm OF replicate n @{thm subset_UNIV}) RS thm RS sym) fold_unique_mor_thms; val ctor_o_fold_thms = let val mor = mor_comp_thm OF [mor_fold_thm, mor_str_thm]; in map2 (fn unique => fn fold_ctor => trans OF [mor RS unique, fold_ctor]) fold_unique_mor_thms fold_ctor_thms end; val timer = time (timer "fold definitions & thms"); val map_ctors = map2 (fn Ds => fn bnf => Term.list_comb (mk_map_of_bnf Ds (passiveAs @ FTs) (passiveAs @ Ts) bnf, map HOLogic.id_const passiveAs @ ctors)) Dss bnfs; fun dtor_bind i = nth external_bs (i - 1) |> Binding.prefix_name (dtorN ^ "_"); val dtor_def_bind = rpair [] o Binding.concealed o Thm.def_binding o dtor_bind; fun dtor_spec i = mk_fold Ts map_ctors i; val ((dtor_frees, (_, dtor_def_frees)), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> fold_map (fn i => Local_Theory.define ((dtor_bind i, NoSyn), (dtor_def_bind i, dtor_spec i))) ks |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; fun mk_dtors params = map (Term.subst_atomic_types (map (Morphism.typ phi) params' ~~ params) o Morphism.term phi) dtor_frees; val dtors = mk_dtors params'; val dtor_defs = map (fn def => HOLogic.mk_obj_eq (Morphism.thm phi def)) dtor_def_frees; val ctor_o_dtor_thms = map2 (Local_Defs.fold lthy o single) dtor_defs ctor_o_fold_thms; val dtor_o_ctor_thms = let fun mk_goal dtor ctor FT = mk_Trueprop_eq (HOLogic.mk_comp (dtor, ctor), HOLogic.id_const FT); val goals = @{map 3} mk_goal dtors ctors FTs; in @{map 5} (fn goal => fn dtor_def => fn foldx => fn map_comp_id => fn map_cong0L => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_dtor_o_ctor_tac ctxt dtor_def foldx map_comp_id map_cong0L ctor_o_fold_thms) |> Thm.close_derivation \<^here>) goals dtor_defs ctor_fold_thms map_comp_id_thms map_cong0L_thms end; val dtor_ctor_thms = map (fn thm => thm RS @{thm pointfree_idE}) dtor_o_ctor_thms; val ctor_dtor_thms = map (fn thm => thm RS @{thm pointfree_idE}) ctor_o_dtor_thms; val bij_dtor_thms = map2 (fn thm1 => fn thm2 => @{thm o_bij} OF [thm1, thm2]) ctor_o_dtor_thms dtor_o_ctor_thms; val inj_dtor_thms = map (fn thm => thm RS @{thm bij_is_inj}) bij_dtor_thms; val surj_dtor_thms = map (fn thm => thm RS @{thm bij_is_surj}) bij_dtor_thms; val dtor_nchotomy_thms = map (fn thm => thm RS @{thm surjD}) surj_dtor_thms; val dtor_inject_thms = map (fn thm => thm RS @{thm inj_eq}) inj_dtor_thms; val dtor_exhaust_thms = map (fn thm => thm RS exE) dtor_nchotomy_thms; val bij_ctor_thms = map2 (fn thm1 => fn thm2 => @{thm o_bij} OF [thm1, thm2]) dtor_o_ctor_thms ctor_o_dtor_thms; val inj_ctor_thms = map (fn thm => thm RS @{thm bij_is_inj}) bij_ctor_thms; val surj_ctor_thms = map (fn thm => thm RS @{thm bij_is_surj}) bij_ctor_thms; val ctor_nchotomy_thms = map (fn thm => thm RS @{thm surjD}) surj_ctor_thms; val ctor_inject_thms = map (fn thm => thm RS @{thm inj_eq}) inj_ctor_thms; val ctor_exhaust_thms = map (fn thm => thm RS exE) ctor_nchotomy_thms; val timer = time (timer "dtor definitions & thms"); val (((((((Izs, (Izs1, Izs1'))), (Izs2, Izs2')), xFs), yFs), init_phis), _) = lthy |> mk_Frees "z" Ts ||>> mk_Frees' "z1" Ts ||>> mk_Frees' "z2" Ts' ||>> mk_Frees "x" FTs ||>> mk_Frees "y" FTs' ||>> mk_Frees "P" (replicate n (mk_pred1T initT)); val phis = map2 retype_const_or_free (map mk_pred1T Ts) init_phis; val phi2s = map2 retype_const_or_free (map2 mk_pred2T Ts Ts') init_phis; val (ctor_induct_thm, induct_params) = let fun mk_prem phi ctor sets x = let fun mk_IH phi set z = let val prem = mk_Trueprop_mem (z, set $ x); val concl = HOLogic.mk_Trueprop (phi $ z); in Logic.all z (Logic.mk_implies (prem, concl)) end; val IHs = @{map 3} mk_IH phis (drop m sets) Izs; val concl = HOLogic.mk_Trueprop (phi $ (ctor $ x)); in Logic.all x (Logic.list_implies (IHs, concl)) end; val prems = @{map 4} mk_prem phis ctors FTs_setss xFs; fun mk_concl phi z = phi $ z; val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_concl phis Izs)); val goal = Logic.list_implies (prems, concl); val vars = Variable.add_free_names lthy goal []; in (Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_ctor_induct_tac ctxt m set_mapss init_induct_thm morE_thms mor_Abs_thm Rep_inverses Abs_inverses Reps) |> Thm.close_derivation \<^here>, rev (Term.add_tfrees goal [])) end; val cTs = map (SOME o Thm.ctyp_of lthy o TFree) induct_params; val weak_ctor_induct_thms = let fun insts i = (replicate (i - 1) TrueI) @ (asm_rl :: replicate (n - i) TrueI); in map (fn i => (ctor_induct_thm OF insts i) RS mk_conjunctN n i) ks end; val (ctor_induct2_thm, induct2_params) = let fun mk_prem phi ctor ctor' sets sets' x y = let fun mk_IH phi set set' z1 z2 = let val prem1 = mk_Trueprop_mem (z1, (set $ x)); val prem2 = mk_Trueprop_mem (z2, (set' $ y)); val concl = HOLogic.mk_Trueprop (phi $ z1 $ z2); in fold_rev Logic.all [z1, z2] (Logic.list_implies ([prem1, prem2], concl)) end; val IHs = @{map 5} mk_IH phi2s (drop m sets) (drop m sets') Izs1 Izs2; val concl = HOLogic.mk_Trueprop (phi $ (ctor $ x) $ (ctor' $ y)); in fold_rev Logic.all [x, y] (Logic.list_implies (IHs, concl)) end; val prems = @{map 7} mk_prem phi2s ctors ctor's FTs_setss FTs'_setss xFs yFs; fun mk_concl phi z1 z2 = phi $ z1 $ z2; val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (@{map 3} mk_concl phi2s Izs1 Izs2)); fun mk_t phi (z1, z1') (z2, z2') = Term.absfree z1' (HOLogic.mk_all (fst z2', snd z2', phi $ z1 $ z2)); val cts = @{map 3} (SOME o Thm.cterm_of lthy ooo mk_t) phi2s (Izs1 ~~ Izs1') (Izs2 ~~ Izs2'); val goal = Logic.list_implies (prems, concl); val vars = Variable.add_free_names lthy goal []; in (Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_ctor_induct2_tac ctxt cTs cts ctor_induct_thm weak_ctor_induct_thms) |> Thm.close_derivation \<^here>, rev (Term.add_tfrees goal [])) end; val timer = time (timer "induction"); fun mk_ctor_map_DEADID_thm ctor_inject map_id0 = trans OF [id_apply, iffD2 OF [ctor_inject, map_id0 RS sym]]; fun mk_ctor_map_unique_DEADID_thm () = let val (funs, algs) = HOLogic.conjuncts (HOLogic.dest_Trueprop (Thm.concl_of ctor_fold_unique_thm)) |> map_split HOLogic.dest_eq ||> snd o strip_comb o hd |> @{apply 2} (map (fst o dest_Var)); fun mk_fun_insts T ix = Thm.cterm_of lthy (Var (ix, T --> T)); val theta = (funs ~~ @{map 2} mk_fun_insts Ts funs) @ (algs ~~ map (Thm.cterm_of lthy) ctors); val ctor_fold_ctors = (ctor_fold_unique_thm OF map (fn thm => mk_trans @{thm id_o} (mk_sym (thm RS @{thm trans[OF arg_cong2[of _ _ _ _ "(\)", OF refl] o_id]}))) map_id0s) |> split_conj_thm |> map mk_sym; in infer_instantiate lthy theta ctor_fold_unique_thm |> unfold_thms lthy ctor_fold_ctors |> Morphism.thm (Local_Theory.target_morphism lthy) end; fun mk_ctor_Irel_DEADID_thm ctor_inject bnf = trans OF [ctor_inject, rel_eq_of_bnf bnf RS @{thm predicate2_eqD} RS sym]; val IphiTs = map2 mk_pred2T passiveAs passiveBs; val Ipsi1Ts = map2 mk_pred2T passiveAs passiveCs; val Ipsi2Ts = map2 mk_pred2T passiveCs passiveBs; val activephiTs = map2 mk_pred2T activeAs activeBs; val activeIphiTs = map2 mk_pred2T Ts Ts'; val rels = map2 (fn Ds => mk_rel_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs; (*register new datatypes as BNFs*) val (timer, Ibnfs, (ctor_Imap_o_thms, ctor_Imap_thms), ctor_Imap_unique_thm, ctor_Iset_thmss', ctor_Irel_thms, Ibnf_notes, lthy) = if m = 0 then (timer, replicate n DEADID_bnf, map_split (`(mk_pointfree2 lthy)) (map2 mk_ctor_map_DEADID_thm ctor_inject_thms map_ids), mk_ctor_map_unique_DEADID_thm (), replicate n [], map2 mk_ctor_Irel_DEADID_thm ctor_inject_thms bnfs, [], lthy) else let val fTs = map2 (curry op -->) passiveAs passiveBs; val uTs = map2 (curry op -->) Ts Ts'; val ((((fs, fs'), (AFss, AFss')), (ys, ys')), _) = lthy |> mk_Frees' "f" fTs ||>> mk_Freess' "z" setFTss ||>> mk_Frees' "y" passiveAs; val map_FTFT's = map2 (fn Ds => mk_map_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs; fun mk_passive_maps ATs BTs Ts = map2 (fn Ds => mk_map_of_bnf Ds (ATs @ Ts) (BTs @ Ts)) Dss bnfs; fun mk_map_fold_arg fs Ts ctor fmap = HOLogic.mk_comp (ctor, Term.list_comb (fmap, fs @ map HOLogic.id_const Ts)); fun mk_map Ts fs Ts' ctors mk_maps = mk_fold Ts (map2 (mk_map_fold_arg fs Ts') ctors (mk_maps Ts')); val pmapsABT' = mk_passive_maps passiveAs passiveBs; val fs_maps = map (mk_map Ts fs Ts' ctor's pmapsABT') ks; val ls = 1 upto m; val setsss = map (mk_setss o mk_set_Ts) passiveAs; fun mk_col l T z z' sets = let fun mk_UN set = mk_Union T $ (set $ z); in Term.absfree z' (mk_union (nth sets (l - 1) $ z, Library.foldl1 mk_union (map mk_UN (drop m sets)))) end; val colss = @{map 5} (fn l => fn T => @{map 3} (mk_col l T)) ls passiveAs AFss AFss' setsss; val setss_by_range = map (fn cols => map (mk_fold Ts cols) ks) colss; val setss_by_bnf = transpose setss_by_range; val set_bss = map (flat o map2 (fn B => fn b => if member (op =) deads (TFree B) then [] else [b]) resBs) set_bss0; val ctor_witss = let val witss = map2 (fn Ds => fn bnf => mk_wits_of_bnf (replicate (nwits_of_bnf bnf) Ds) (replicate (nwits_of_bnf bnf) (passiveAs @ Ts)) bnf) Dss bnfs; fun close_wit (I, wit) = fold_rev Term.absfree (map (nth ys') I) wit; fun wit_apply (arg_I, arg_wit) (fun_I, fun_wit) = (union (op =) arg_I fun_I, fun_wit $ arg_wit); fun gen_arg support i = if i < m then [([i], nth ys i)] else maps (mk_wit support (nth ctors (i - m)) (i - m)) (nth support (i - m)) and mk_wit support ctor i (I, wit) = let val args = map (gen_arg (nth_map i (remove (op =) (I, wit)) support)) I; in (args, [([], wit)]) |-> fold (map_product wit_apply) |> map (apsnd (fn t => ctor $ t)) |> minimize_wits end; in @{map 3} (fn ctor => fn i => map close_wit o minimize_wits o maps (mk_wit witss ctor i)) ctors (0 upto n - 1) witss end; val (lthy, sbd0, sbd0_card_order, sbd0_Cinfinite, set_sbd0ss) = if n = 1 then (lthy, hd bd0s, hd bd0_card_orders, hd bd0_Cinfinites, set_bd0ss) else let val sum_bd0 = Library.foldr1 (uncurry mk_csum) bd0s; val sum_bd0T = fst (dest_relT (fastype_of sum_bd0)); val (sum_bd0T_params, sum_bd0T_params') = `(map TFree) (Term.add_tfreesT sum_bd0T []); val sbd0T_bind = mk_internal_b (sum_bdTN ^ "0"); val ((sbd0T_name, (sbd0T_glob_info, sbd0T_loc_info)), lthy) = typedef (sbd0T_bind, sum_bd0T_params', NoSyn) (HOLogic.mk_UNIV sum_bd0T) NONE (fn ctxt => EVERY' [rtac ctxt exI, rtac ctxt UNIV_I] 1) lthy; val sbd0T = Type (sbd0T_name, sum_bd0T_params); val Abs_sbd0T = Const (#Abs_name sbd0T_glob_info, sum_bd0T --> sbd0T); val sbd0_bind = mk_internal_b (sum_bdN ^ "0"); val sbd0_def_bind = (Thm.def_binding sbd0_bind, []); val sbd0_spec = mk_dir_image sum_bd0 Abs_sbd0T; val ((sbd0_free, (_, sbd0_def_free)), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> Local_Theory.define ((sbd0_bind, NoSyn), (sbd0_def_bind, sbd0_spec)) ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val sbd0_def = HOLogic.mk_obj_eq (Morphism.thm phi sbd0_def_free); val sbd0 = Const (fst (Term.dest_Const (Morphism.term phi sbd0_free)), mk_relT (`I sbd0T)); val Abs_sbd0T_inj = mk_Abs_inj_thm (#Abs_inject sbd0T_loc_info); val Abs_sbd0T_bij = mk_Abs_bij_thm lthy Abs_sbd0T_inj (#Abs_cases sbd0T_loc_info); val sum_Cinfinite = mk_sum_Cinfinite bd0_Cinfinites; val sum_Card_order = sum_Cinfinite RS conjunct2; val sum_card_order = mk_sum_card_order bd0_card_orders; val sbd0_ordIso = @{thm ssubst_Pair_rhs} OF [@{thm dir_image} OF [Abs_sbd0T_inj, sum_Card_order], sbd0_def]; val sbd0_Cinfinite = @{thm Cinfinite_cong} OF [sbd0_ordIso, sum_Cinfinite]; val sbd0_card_order = @{thm iffD2[OF arg_cong[of _ _ card_order]]} OF [sbd0_def, @{thm card_order_dir_image} OF [Abs_sbd0T_bij, sum_card_order]]; fun mk_set_sbd0 i bd0_Card_order bd0s = map (fn thm => @{thm ordLeq_ordIso_trans} OF [bd0_Card_order RS mk_ordLeq_csum n i thm, sbd0_ordIso]) bd0s; val set_sbd0ss = @{map 3} mk_set_sbd0 ks bd0_Card_orders set_bd0ss; in (lthy, sbd0, sbd0_card_order, sbd0_Cinfinite, set_sbd0ss) end; val (Ibnf_consts, lthy) = @{fold_map 9} (fn b => fn map_b => fn rel_b => fn pred_b => fn set_bs => fn mapx => fn sets => fn wits => fn T => fn lthy => define_bnf_consts Hardly_Inline (user_policy Note_Some lthy) false (SOME deads) map_b rel_b pred_b set_bs (((((((b, T), fold_rev Term.absfree fs' mapx), sets), sbd0), wits), NONE), NONE) lthy) bs map_bs rel_bs pred_bs set_bss fs_maps setss_by_bnf ctor_witss Ts lthy; val ((((((((((((((Izs, (Izs1, Izs1')), (Izs2, Izs2')), xFs), yFs))), Iphis), Ipsi1s), Ipsi2s), fs), fs_copy), us), (ys, ys')), _) = lthy |> mk_Frees "z" Ts ||>> mk_Frees' "z1" Ts ||>> mk_Frees' "z2" Ts' ||>> mk_Frees "x" FTs ||>> mk_Frees "y" FTs' ||>> mk_Frees "R" IphiTs ||>> mk_Frees "R" Ipsi1Ts ||>> mk_Frees "Q" Ipsi2Ts ||>> mk_Frees "f" fTs ||>> mk_Frees "f" fTs ||>> mk_Frees "u" uTs ||>> mk_Frees' "y" passiveAs; val (_, Iconsts, Iconst_defs, mk_Iconsts) = @{split_list 4} Ibnf_consts; val (_, Isetss, Ibds_Ds, Iwitss_Ds, _, _) = @{split_list 6} Iconsts; val (Imap_defs, Iset_defss, Ibd_defs, Iwit_defss, Irel_defs, Ipred_defs) = @{split_list 6} Iconst_defs; val (mk_Imaps_Ds, mk_It_Ds, _, mk_Irels_Ds, mk_Ipreds_Ds, _, _) = @{split_list 7} mk_Iconsts; val Irel_unabs_defs = map (fn def => mk_unabs_def m (HOLogic.mk_obj_eq def)) Irel_defs; val Ipred_unabs_defs = map (fn def => mk_unabs_def m (HOLogic.mk_obj_eq def)) Ipred_defs; val Iset_defs = flat Iset_defss; fun mk_Imaps As Bs = map (fn mk => mk deads As Bs) mk_Imaps_Ds; fun mk_Isetss As = map2 (fn mk => fn Isets => map (mk deads As) Isets) mk_It_Ds Isetss; val Ibds = map2 (fn mk => mk deads passiveAs) mk_It_Ds Ibds_Ds; val Iwitss = map2 (fn mk => fn Iwits => map (mk deads passiveAs o snd) Iwits) mk_It_Ds Iwitss_Ds; fun mk_Irels As Bs = map (fn mk => mk deads As Bs) mk_Irels_Ds; fun mk_Ipreds As = map (fn mk => mk deads As) mk_Ipreds_Ds; val Imaps = mk_Imaps passiveAs passiveBs; val fs_Imaps = map (fn m => Term.list_comb (m, fs)) Imaps; val fs_copy_Imaps = map (fn m => Term.list_comb (m, fs_copy)) Imaps; val (Isetss_by_range, Isetss_by_bnf) = `transpose (mk_Isetss passiveAs); val map_setss = map (fn T => map2 (fn Ds => mk_map_of_bnf Ds (passiveAs @ Ts) (mk_set_Ts T)) Dss bnfs) passiveAs; val timer = time (timer "bnf constants for the new datatypes"); val (ctor_Imap_thms, ctor_Imap_o_thms) = let fun mk_goal fs_map map ctor ctor' = mk_Trueprop_eq (HOLogic.mk_comp (fs_map, ctor), HOLogic.mk_comp (ctor', Term.list_comb (map, fs @ fs_Imaps))); val goals = @{map 4} mk_goal fs_Imaps map_FTFT's ctors ctor's; val maps = @{map 4} (fn goal => fn foldx => fn map_comp_id => fn map_cong0 => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Imap_defs THEN mk_map_tac ctxt m n foldx map_comp_id map_cong0)) |> Thm.close_derivation \<^here>) goals ctor_fold_thms map_comp_id_thms map_cong0s; in `(map (fn thm => thm RS @{thm comp_eq_dest})) maps end; val (ctor_Imap_unique_thms, ctor_Imap_unique_thm) = let fun mk_prem u map ctor ctor' = mk_Trueprop_eq (HOLogic.mk_comp (u, ctor), HOLogic.mk_comp (ctor', Term.list_comb (map, fs @ us))); val prems = @{map 4} mk_prem us map_FTFT's ctors ctor's; val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_eq) us fs_Imaps)); val vars = fold (Variable.add_free_names lthy) (goal :: prems) []; val unique = Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, goal)) (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Imap_defs THEN mk_ctor_map_unique_tac ctxt ctor_fold_unique_thm sym_map_comps) |> Thm.close_derivation \<^here>; in `split_conj_thm unique end; val timer = time (timer "map functions for the new datatypes"); val ctor_Iset_thmss = let fun mk_goal sets ctor set col map = mk_Trueprop_eq (HOLogic.mk_comp (set, ctor), HOLogic.mk_comp (col, Term.list_comb (map, passive_ids @ sets))); val goalss = @{map 3} (fn sets => @{map 4} (mk_goal sets) ctors sets) Isetss_by_range colss map_setss; val setss = map (map2 (fn foldx => fn goal => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Iset_defs THEN mk_set_tac ctxt foldx) |> Thm.close_derivation \<^here>) ctor_fold_thms) goalss; fun mk_simp_goal pas_set act_sets sets ctor z set = mk_Trueprop_eq (set $ (ctor $ z), mk_union (pas_set $ z, Library.foldl1 mk_union (map2 (fn X => mk_UNION (X $ z)) act_sets sets))); val simp_goalss = map2 (fn i => fn sets => @{map 4} (fn Fsets => mk_simp_goal (nth Fsets (i - 1)) (drop m Fsets) sets) FTs_setss ctors xFs sets) ls Isetss_by_range; val ctor_setss = @{map 3} (fn i => @{map 3} (fn set_nats => fn goal => fn set => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_ctor_set_tac ctxt set (nth set_nats (i - 1)) (drop m set_nats))) |> Thm.close_derivation \<^here>) set_mapss) ls simp_goalss setss; in ctor_setss end; fun mk_set_thms ctor_set = (@{thm xt1(3)} OF [ctor_set, @{thm Un_upper1}]) :: map (fn i => (@{thm xt1(3)} OF [ctor_set, @{thm Un_upper2}]) RS (mk_Un_upper n i RS subset_trans) RSN (2, @{thm UN_upper} RS subset_trans)) (1 upto n); val set_Iset_thmsss = transpose (map (map mk_set_thms) ctor_Iset_thmss); val timer = time (timer "set functions for the new datatypes"); val cxs = map (SOME o Thm.cterm_of lthy) Izs; val Isetss_by_range' = map (map (Term.subst_atomic_types (passiveAs ~~ passiveBs))) Isetss_by_range; val Iset_Imap0_thmss = let fun mk_set_map0 f map z set set' = HOLogic.mk_eq (mk_image f $ (set $ z), set' $ (map $ z)); fun mk_cphi f map z set set' = Thm.cterm_of lthy (Term.absfree (dest_Free z) (mk_set_map0 f map z set set')); val csetss = map (map (Thm.cterm_of lthy)) Isetss_by_range'; val cphiss = @{map 3} (fn f => fn sets => fn sets' => (@{map 4} (mk_cphi f) fs_Imaps Izs sets sets')) fs Isetss_by_range Isetss_by_range'; val inducts = map (fn cphis => Thm.instantiate' cTs (map SOME cphis @ cxs) ctor_induct_thm) cphiss; val goals = @{map 3} (fn f => fn sets => fn sets' => HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (@{map 4} (mk_set_map0 f) fs_Imaps Izs sets sets'))) fs Isetss_by_range Isetss_by_range'; fun mk_tac ctxt induct = mk_set_nat_tac ctxt m (rtac ctxt induct) set_mapss ctor_Imap_thms; val thms = @{map 5} (fn goal => fn csets => fn ctor_sets => fn induct => fn i => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_tac ctxt induct csets ctor_sets i)) |> Thm.close_derivation \<^here>) goals csetss ctor_Iset_thmss inducts ls; in map split_conj_thm thms end; val Iset_bd_thmss = let fun mk_set_bd z bd set = mk_ordLeq (mk_card_of (set $ z)) bd; fun mk_cphi z set = Thm.cterm_of lthy (Term.absfree (dest_Free z) (mk_set_bd z sbd0 set)); val cphiss = map (map2 mk_cphi Izs) Isetss_by_range; val inducts = map (fn cphis => Thm.instantiate' cTs (map SOME cphis @ cxs) ctor_induct_thm) cphiss; val goals = map (fn sets => HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (@{map 3} mk_set_bd Izs Ibds sets))) Isetss_by_range; fun mk_tac ctxt induct = mk_set_bd_tac ctxt m (rtac ctxt induct) sbd0_Cinfinite set_sbd0ss; val thms = @{map 4} (fn goal => fn ctor_sets => fn induct => fn i => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Ibd_defs THEN mk_tac ctxt induct ctor_sets i)) |> Thm.close_derivation \<^here>) goals ctor_Iset_thmss inducts ls; in map split_conj_thm thms end; val Imap_cong0_thms = let fun mk_prem z set f g y y' = mk_Ball (set $ z) (Term.absfree y' (HOLogic.mk_eq (f $ y, g $ y))); fun mk_map_cong0 sets z fmap gmap = HOLogic.mk_imp (Library.foldr1 HOLogic.mk_conj (@{map 5} (mk_prem z) sets fs fs_copy ys ys'), HOLogic.mk_eq (fmap $ z, gmap $ z)); fun mk_cphi sets z fmap gmap = Thm.cterm_of lthy (Term.absfree (dest_Free z) (mk_map_cong0 sets z fmap gmap)); val cphis = @{map 4} mk_cphi Isetss_by_bnf Izs fs_Imaps fs_copy_Imaps; val induct = Thm.instantiate' cTs (map SOME cphis @ cxs) ctor_induct_thm; val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (@{map 4} mk_map_cong0 Isetss_by_bnf Izs fs_Imaps fs_copy_Imaps)); val vars = Variable.add_free_names lthy goal []; val thm = Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_mcong_tac ctxt (rtac ctxt induct) set_Iset_thmsss map_cong0s ctor_Imap_thms) |> Thm.close_derivation \<^here>; in split_conj_thm thm end; val in_rels = map in_rel_of_bnf bnfs; val in_Irels = map (fn def => trans OF [def, @{thm OO_Grp_alt}] RS @{thm predicate2_eqD}) Irel_unabs_defs; val ctor_Iset_incl_thmss = map (map hd) set_Iset_thmsss; val ctor_set_Iset_incl_thmsss = map (transpose o map tl) set_Iset_thmsss; val ctor_Iset_thmss' = transpose ctor_Iset_thmss; val Irels = mk_Irels passiveAs passiveBs; val Ipreds = mk_Ipreds passiveAs; val Irelphis = map (fn rel => Term.list_comb (rel, Iphis)) Irels; val relphis = map (fn rel => Term.list_comb (rel, Iphis @ Irelphis)) rels; val Irelpsi1s = map (fn rel => Term.list_comb (rel, Ipsi1s)) (mk_Irels passiveAs passiveCs); val Irelpsi2s = map (fn rel => Term.list_comb (rel, Ipsi2s)) (mk_Irels passiveCs passiveBs); val Irelpsi12s = map (fn rel => Term.list_comb (rel, map2 (curry mk_rel_compp) Ipsi1s Ipsi2s)) Irels; val ctor_Irel_thms = let fun mk_goal xF yF ctor ctor' Irelphi relphi = mk_Trueprop_eq (Irelphi $ (ctor $ xF) $ (ctor' $ yF), relphi $ xF $ yF); val goals = @{map 6} mk_goal xFs yFs ctors ctor's Irelphis relphis; in @{map 12} (fn i => fn goal => fn in_rel => fn map_comp0 => fn map_cong0 => fn ctor_map => fn ctor_sets => fn ctor_inject => fn ctor_dtor => fn set_map0s => fn ctor_set_incls => fn ctor_set_set_inclss => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_ctor_rel_tac ctxt in_Irels i in_rel map_comp0 map_cong0 ctor_map ctor_sets ctor_inject ctor_dtor set_map0s ctor_set_incls ctor_set_set_inclss)) |> Thm.close_derivation \<^here>) ks goals in_rels map_comps map_cong0s ctor_Imap_thms ctor_Iset_thmss' ctor_inject_thms ctor_dtor_thms set_mapss ctor_Iset_incl_thmss ctor_set_Iset_incl_thmsss end; val le_Irel_OO_thm = let fun mk_le_Irel_OO Irelpsi1 Irelpsi2 Irelpsi12 Iz1 Iz2 = HOLogic.mk_imp (mk_rel_compp (Irelpsi1, Irelpsi2) $ Iz1 $ Iz2, Irelpsi12 $ Iz1 $ Iz2); val goals = @{map 5} mk_le_Irel_OO Irelpsi1s Irelpsi2s Irelpsi12s Izs1 Izs2; val cTs = map (SOME o Thm.ctyp_of lthy o TFree) induct2_params; val cxs = map (SOME o Thm.cterm_of lthy) (splice Izs1 Izs2); fun mk_cphi z1 z2 goal = SOME (Thm.cterm_of lthy (Term.absfree z1 (Term.absfree z2 goal))); val cphis = @{map 3} mk_cphi Izs1' Izs2' goals; val induct = Thm.instantiate' cTs (cphis @ cxs) ctor_induct2_thm; val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj goals); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_le_rel_OO_tac ctxt m induct ctor_nchotomy_thms ctor_Irel_thms rel_mono_strong0s le_rel_OOs) |> Thm.close_derivation \<^here> end; val timer = time (timer "helpers for BNF properties"); val map_id0_tacs = map (fn thm => fn ctxt => mk_map_id0_tac ctxt map_id0s thm) ctor_Imap_unique_thms; val map_comp0_tacs = map2 (fn thm => fn i => fn ctxt => mk_map_comp0_tac ctxt map_comps ctor_Imap_thms thm i) ctor_Imap_unique_thms ks; val map_cong0_tacs = map (fn thm => fn ctxt => mk_map_cong0_tac ctxt m thm) Imap_cong0_thms; val set_map0_tacss = map (map (fn thm => fn ctxt => mk_set_map0_tac ctxt thm)) (transpose Iset_Imap0_thmss); val bd_co_tacs = replicate n (fn ctxt => unfold_thms_tac ctxt Ibd_defs THEN rtac ctxt sbd0_card_order 1); val bd_cinf_tacs = replicate n (fn ctxt => unfold_thms_tac ctxt Ibd_defs THEN rtac ctxt (sbd0_Cinfinite RS conjunct1) 1); val set_bd_tacss = map (map (fn thm => fn ctxt => rtac ctxt thm 1)) (transpose Iset_bd_thmss); val le_rel_OO_tacs = map (fn i => fn ctxt => (rtac ctxt @{thm predicate2I} THEN' etac ctxt (le_Irel_OO_thm RS mk_conjunctN n i RS mp)) 1) ks; val rel_OO_Grp_tacs = map (fn def => fn ctxt => rtac ctxt def 1) Irel_unabs_defs; val pred_set_tacs = map (fn def => fn ctxt => rtac ctxt def 1) Ipred_unabs_defs; val tacss = @{map 10} zip_axioms map_id0_tacs map_comp0_tacs map_cong0_tacs set_map0_tacss bd_co_tacs bd_cinf_tacs set_bd_tacss le_rel_OO_tacs rel_OO_Grp_tacs pred_set_tacs; fun wit_tac ctxt = unfold_thms_tac ctxt (flat Iwit_defss) THEN mk_wit_tac ctxt n (flat ctor_Iset_thmss) (maps wit_thms_of_bnf bnfs); val (Ibnfs, lthy) = @{fold_map 6} (fn tacs => fn map_b => fn rel_b => fn pred_b => fn set_bs => fn consts => bnf_def Do_Inline (user_policy Note_Some) false I tacs wit_tac (SOME deads) map_b rel_b pred_b set_bs consts) tacss map_bs rel_bs pred_bs set_bss (((((((replicate n Binding.empty ~~ Ts) ~~ Imaps) ~~ Isetss_by_bnf) ~~ Ibds) ~~ Iwitss) ~~ map SOME Irels) ~~ map SOME Ipreds) lthy; val timer = time (timer "registered new datatypes as BNFs"); val ls' = if m = 1 then [0] else ls val Ibnf_common_notes = [(ctor_map_uniqueN, [ctor_Imap_unique_thm])] |> map (fn (thmN, thms) => ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])])); val Ibnf_notes = [(ctor_mapN, map single ctor_Imap_thms), (ctor_relN, map single ctor_Irel_thms), (ctor_set_inclN, ctor_Iset_incl_thmss), (ctor_set_set_inclN, map flat ctor_set_Iset_incl_thmsss)] @ map2 (fn i => fn thms => (mk_ctor_setN i, map single thms)) ls' ctor_Iset_thmss |> maps (fn (thmN, thmss) => map2 (fn b => fn thms => ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])])) bs thmss) in (timer, Ibnfs, (ctor_Imap_o_thms, ctor_Imap_thms), ctor_Imap_unique_thm, ctor_Iset_thmss', ctor_Irel_thms, Ibnf_common_notes @ Ibnf_notes, lthy) end; val ((((((xFs, yFs)), Iphis), activephis), activeIphis), _) = lthy |> mk_Frees "x" FTs ||>> mk_Frees "y" FTs' ||>> mk_Frees "R" IphiTs ||>> mk_Frees "S" activephiTs ||>> mk_Frees "IR" activeIphiTs; val ctor_fold_o_Imap_thms = mk_xtor_co_iter_o_map_thms Least_FP false m ctor_fold_unique_thm ctor_Imap_o_thms (map (mk_pointfree2 lthy) ctor_fold_thms) sym_map_comps map_cong0s; val Irels = if m = 0 then map HOLogic.eq_const Ts else map (mk_rel_of_bnf deads passiveAs passiveBs) Ibnfs; val Irel_induct_thm = mk_xtor_rel_co_induct_thm Least_FP rels activeIphis Irels Iphis xFs yFs ctors ctor's (fn {context = ctxt, prems = IHs} => mk_rel_induct_tac ctxt IHs m ctor_induct2_thm ks ctor_Irel_thms rel_mono_strong0s) lthy; val rels = map2 (fn Ds => mk_rel_of_bnf Ds allAs allBs') Dss bnfs; val ctor_fold_transfer_thms = mk_xtor_co_iter_transfer_thms Least_FP rels activephis activephis Irels Iphis (mk_folds passiveAs activeAs) (mk_folds passiveBs activeBs) (fn {context = ctxt, prems = _} => mk_fold_transfer_tac ctxt m Irel_induct_thm (map map_transfer_of_bnf bnfs) ctor_fold_thms) lthy; val timer = time (timer "relator induction"); fun mk_Ts As = map (typ_subst_atomic (passiveAs ~~ As)) Ts; val export = map (Morphism.term (Local_Theory.target_morphism lthy)) val ((recs, (ctor_rec_thms, ctor_rec_unique_thm, ctor_rec_o_Imap_thms, ctor_rec_transfer_thms)), lthy) = lthy |> derive_xtor_co_recs Least_FP external_bs mk_Ts (Dss, resDs) bnfs (export ctors) (export folds) ctor_fold_unique_thm ctor_fold_thms ctor_fold_transfer_thms ctor_Imap_thms ctor_Irel_thms (replicate n NONE); val timer = time (timer "recursor"); val common_notes = [(ctor_inductN, [ctor_induct_thm]), (ctor_induct2N, [ctor_induct2_thm]), (ctor_rel_inductN, [Irel_induct_thm])] |> map (fn (thmN, thms) => ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])])); val notes = [(ctor_dtorN, ctor_dtor_thms), (ctor_exhaustN, ctor_exhaust_thms), (ctor_foldN, ctor_fold_thms), (ctor_fold_o_mapN, ctor_fold_o_Imap_thms), (ctor_fold_transferN, ctor_fold_transfer_thms), (ctor_fold_uniqueN, ctor_fold_unique_thms), (ctor_injectN, ctor_inject_thms), (dtor_ctorN, dtor_ctor_thms), (dtor_exhaustN, dtor_exhaust_thms), (dtor_injectN, dtor_inject_thms)] |> map (apsnd (map single)) |> maps (fn (thmN, thmss) => map2 (fn b => fn thms => ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])])) bs thmss); val lthy' = lthy |> internals ? snd o Local_Theory.notes (common_notes @ notes @ Ibnf_notes); val fp_res = {Ts = Ts, bnfs = Ibnfs, pre_bnfs = bnfs, absT_infos = absT_infos, ctors = ctors, dtors = dtors, xtor_un_folds = folds, xtor_co_recs = export recs, xtor_co_induct = ctor_induct_thm, dtor_ctors = dtor_ctor_thms, ctor_dtors = ctor_dtor_thms, ctor_injects = ctor_inject_thms, dtor_injects = dtor_inject_thms, xtor_maps = ctor_Imap_thms, xtor_map_unique = ctor_Imap_unique_thm, xtor_setss = ctor_Iset_thmss', xtor_rels = ctor_Irel_thms, xtor_un_fold_thms = ctor_fold_thms, xtor_co_rec_thms = ctor_rec_thms, xtor_un_fold_unique = ctor_fold_unique_thm, xtor_co_rec_unique = ctor_rec_unique_thm, xtor_un_fold_o_maps = ctor_fold_o_Imap_thms, xtor_co_rec_o_maps = ctor_rec_o_Imap_thms, xtor_un_fold_transfers = ctor_fold_transfer_thms, xtor_co_rec_transfers = ctor_rec_transfer_thms, xtor_rel_co_induct = Irel_induct_thm, dtor_set_inducts = []}; in timer; (fp_res, lthy') end; val _ = Outer_Syntax.local_theory \<^command_keyword>\datatype\ "define inductive datatypes" (parse_co_datatype_cmd Least_FP construct_lfp); end; diff --git a/src/HOL/Tools/BNF/bnf_lfp_size.ML b/src/HOL/Tools/BNF/bnf_lfp_size.ML --- a/src/HOL/Tools/BNF/bnf_lfp_size.ML +++ b/src/HOL/Tools/BNF/bnf_lfp_size.ML @@ -1,400 +1,400 @@ (* Title: HOL/Tools/BNF/bnf_lfp_size.ML Author: Jasmin Blanchette, TU Muenchen Copyright 2014 Generation of size functions for datatypes. *) signature BNF_LFP_SIZE = sig val register_size: string -> string -> thm -> thm list -> thm list -> local_theory -> local_theory val register_size_global: string -> string -> thm -> thm list -> thm list -> theory -> theory val size_of: Proof.context -> string -> (string * (thm * thm list * thm list)) option val size_of_global: theory -> string -> (string * (thm * thm list * thm list)) option end; structure BNF_LFP_Size : BNF_LFP_SIZE = struct open BNF_Util open BNF_Tactics open BNF_Def open BNF_FP_Def_Sugar val size_N = "size_"; val sizeN = "size"; val size_genN = "size_gen"; val size_gen_o_mapN = "size_gen_o_map"; val size_neqN = "size_neq"; val nitpicksimp_attrs = @{attributes [nitpick_simp]}; val simp_attrs = @{attributes [simp]}; fun mk_plus_nat (t1, t2) = Const (\<^const_name>\Groups.plus\, HOLogic.natT --> HOLogic.natT --> HOLogic.natT) $ t1 $ t2; fun mk_to_natT T = T --> HOLogic.natT; fun mk_abs_zero_nat T = Term.absdummy T HOLogic.zero; fun mk_unabs_def_unused_0 n = funpow n (fn thm => thm RS @{thm fun_cong_unused_0} handle THM _ => thm RS fun_cong); structure Data = Generic_Data ( type T = (string * (thm * thm list * thm list)) Symtab.table; val empty = Symtab.empty; val extend = I fun merge data = Symtab.merge (K true) data; ); fun check_size_type thy T_name size_name = let val n = Sign.arity_number thy T_name; val As = map (fn s => TFree (s, \<^sort>\type\)) (Name.invent Name.context Name.aT n); val T = Type (T_name, As); val size_T = map mk_to_natT As ---> mk_to_natT T; val size_const = Const (size_name, size_T); in can (Thm.global_cterm_of thy) size_const orelse error ("Constant " ^ quote size_name ^ " registered as size function for " ^ quote T_name ^ " must have type\n" ^ quote (Syntax.string_of_typ_global thy size_T)) end; fun register_size T_name size_name overloaded_size_def size_simps size_gen_o_maps lthy = (check_size_type (Proof_Context.theory_of lthy) T_name size_name; Context.proof_map (Data.map (Symtab.update (T_name, (size_name, (overloaded_size_def, size_simps, size_gen_o_maps))))) lthy); fun register_size_global T_name size_name overloaded_size_def size_simps size_gen_o_maps thy = (check_size_type thy T_name size_name; Context.theory_map (Data.map (Symtab.update (T_name, (size_name, (overloaded_size_def, size_simps, size_gen_o_maps))))) thy); val size_of = Symtab.lookup o Data.get o Context.Proof; val size_of_global = Symtab.lookup o Data.get o Context.Theory; fun all_overloaded_size_defs_of ctxt = Symtab.fold (fn (_, (_, (overloaded_size_def, _, _))) => can (Logic.dest_equals o Thm.prop_of) overloaded_size_def ? cons overloaded_size_def) (Data.get (Context.Proof ctxt)) []; val size_gen_o_map_simps = @{thms inj_on_id snd_comp_apfst[simplified apfst_def]}; fun mk_size_gen_o_map_tac ctxt size_def rec_o_map inj_maps size_maps = unfold_thms_tac ctxt [size_def] THEN HEADGOAL (rtac ctxt (rec_o_map RS trans) THEN' asm_simp_tac (ss_only (inj_maps @ size_maps @ size_gen_o_map_simps) ctxt)) THEN IF_UNSOLVED (unfold_thms_tac ctxt @{thms id_def o_def} THEN HEADGOAL (rtac ctxt refl)); fun mk_size_neq ctxt cts exhaust sizes = HEADGOAL (rtac ctxt (infer_instantiate' ctxt (map SOME cts) exhaust)) THEN ALLGOALS (hyp_subst_tac ctxt) THEN unfold_thms_tac ctxt (@{thm neq0_conv} :: sizes) THEN ALLGOALS (REPEAT_DETERM o (rtac ctxt @{thm zero_less_Suc} ORELSE' rtac ctxt @{thm trans_less_add2})); fun generate_datatype_size (fp_sugars as ({T = Type (_, As), BT = Type (_, Bs), fp = Least_FP, fp_res = {bnfs = fp_bnfs, ...}, fp_nesting_bnfs, live_nesting_bnfs, fp_co_induct_sugar = SOME _, ...} : fp_sugar) :: _) lthy0 = let val data = Data.get (Context.Proof lthy0); val Ts = map #T fp_sugars val T_names = map (fst o dest_Type) Ts; val nn = length Ts; val B_ify = Term.typ_subst_atomic (As ~~ Bs); val recs = map (#co_rec o the o #fp_co_induct_sugar) fp_sugars; val rec_thmss = map (#co_rec_thms o the o #fp_co_induct_sugar) fp_sugars; val rec_Ts as rec_T1 :: _ = map fastype_of recs; val rec_arg_Ts = binder_fun_types rec_T1; val Cs = map body_type rec_Ts; val Cs_rho = map (rpair HOLogic.natT) Cs; val substCnatT = Term.subst_atomic_types Cs_rho; val f_Ts = map mk_to_natT As; val f_TsB = map mk_to_natT Bs; val num_As = length As; fun variant_names n pre = fst (Variable.variant_fixes (replicate n pre) lthy0); val f_names = variant_names num_As "f"; val fs = map2 (curry Free) f_names f_Ts; val fsB = map2 (curry Free) f_names f_TsB; val As_fs = As ~~ fs; val size_bs = map ((fn base => Binding.qualify false base (Binding.name (prefix size_N base))) o Long_Name.base_name) T_names; fun is_prod_C \<^type_name>\prod\ [_, T'] = member (op =) Cs T' | is_prod_C _ _ = false; fun mk_size_of_typ (T as TFree _) = pair (case AList.lookup (op =) As_fs T of SOME f => f | NONE => if member (op =) Cs T then Term.absdummy T (Bound 0) else mk_abs_zero_nat T) | mk_size_of_typ (T as Type (s, Ts)) = if is_prod_C s Ts then pair (snd_const T) else if exists (exists_subtype_in (As @ Cs)) Ts then (case Symtab.lookup data s of SOME (size_name, (_, _, size_gen_o_maps)) => let val (args, size_gen_o_mapss') = fold_map mk_size_of_typ Ts []; val size_T = map fastype_of args ---> mk_to_natT T; val size_const = Const (size_name, size_T); in append (size_gen_o_maps :: size_gen_o_mapss') #> pair (Term.list_comb (size_const, args)) end | _ => pair (mk_abs_zero_nat T)) else pair (mk_abs_zero_nat T); fun mk_size_of_arg t = mk_size_of_typ (fastype_of t) #>> (fn s => substCnatT (betapply (s, t))); fun is_recursive_or_plain_case Ts = exists (exists_subtype_in Cs) Ts orelse forall (not o exists_subtype_in As) Ts; (* We want the size function to enjoy the following properties: 1. The size of a list should coincide with its length. 2. All the nonrecursive constructors of a type should have the same size. 3. Each constructor through which nested recursion takes place should count as at least one in the generic size function. 4. The "size" function should be definable as "size_t (%_. 0) ... (%_. 0)", where "size_t" is the generic function. This explains the somewhat convoluted logic ahead. *) val base_case = if forall (is_recursive_or_plain_case o binder_types) rec_arg_Ts then HOLogic.zero else HOLogic.Suc_zero; fun mk_size_arg rec_arg_T = let val x_Ts = binder_types rec_arg_T; val m = length x_Ts; val x_names = variant_names m "x"; val xs = map2 (curry Free) x_names x_Ts; val (summands, size_gen_o_mapss) = fold_map mk_size_of_arg xs [] |>> remove (op =) HOLogic.zero; val sum = if null summands then base_case else foldl1 mk_plus_nat (summands @ [HOLogic.Suc_zero]); in append size_gen_o_mapss #> pair (fold_rev Term.lambda (map substCnatT xs) sum) end; fun mk_size_rhs recx = fold_map mk_size_arg rec_arg_Ts #>> (fn args => fold_rev Term.lambda fs (Term.list_comb (substCnatT recx, args))); val maybe_conceal_def_binding = Thm.def_binding #> not (Config.get lthy0 bnf_internals) ? Binding.concealed; val (size_rhss, nested_size_gen_o_mapss) = fold_map mk_size_rhs recs []; val size_Ts = map fastype_of size_rhss; val nested_size_gen_o_maps_complete = forall (not o null) nested_size_gen_o_mapss; val nested_size_gen_o_maps = fold (union Thm.eq_thm_prop) nested_size_gen_o_mapss []; val ((raw_size_consts, raw_size_defs), (lthy1, lthy1_old)) = lthy0 - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> apfst split_list o @{fold_map 2} (fn b => fn rhs => Local_Theory.define ((b, NoSyn), ((maybe_conceal_def_binding b, []), rhs)) #>> apsnd snd) size_bs size_rhss ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy1_old lthy1; val size_defs = map (Morphism.thm phi) raw_size_defs; val size_consts0 = map (Morphism.term phi) raw_size_consts; val size_consts = map2 retype_const_or_free size_Ts size_consts0; val size_constsB = map (Term.map_types B_ify) size_consts; val zeros = map mk_abs_zero_nat As; val overloaded_size_rhss = map (fn c => Term.list_comb (c, zeros)) size_consts; val overloaded_size_Ts = map fastype_of overloaded_size_rhss; val overloaded_size_consts = map (curry Const \<^const_name>\size\) overloaded_size_Ts; val overloaded_size_def_bs = map (maybe_conceal_def_binding o Binding.suffix_name "_overloaded") size_bs; fun define_overloaded_size def_b lhs0 rhs lthy = let val Free (c, _) = Syntax.check_term lthy lhs0; val ((_, (_, thm)), lthy') = lthy |> Local_Theory.define ((Binding.name c, NoSyn), ((def_b, []), rhs)); val thy_ctxt = Proof_Context.init_global (Proof_Context.theory_of lthy'); val thm' = singleton (Proof_Context.export lthy' thy_ctxt) thm; in (thm', lthy') end; val (overloaded_size_defs, lthy2) = lthy1 |> Local_Theory.background_theory_result (Class.instantiation (T_names, map dest_TFree As, [HOLogic.class_size]) #> @{fold_map 3} define_overloaded_size overloaded_size_def_bs overloaded_size_consts overloaded_size_rhss ##> Class.prove_instantiation_instance (fn ctxt => Class.intro_classes_tac ctxt []) ##> Local_Theory.exit_global); val size_defs' = map (mk_unabs_def (num_As + 1) o HOLogic.mk_obj_eq) size_defs; val size_defs_unused_0 = map (mk_unabs_def_unused_0 (num_As + 1) o HOLogic.mk_obj_eq) size_defs; val overloaded_size_defs' = map (mk_unabs_def 1 o HOLogic.mk_obj_eq) overloaded_size_defs; val nested_size_maps = map (mk_pointful lthy2) nested_size_gen_o_maps @ nested_size_gen_o_maps; val all_inj_maps = @{thm prod.inj_map} :: map inj_map_of_bnf (fp_bnfs @ fp_nesting_bnfs @ live_nesting_bnfs) |> distinct Thm.eq_thm_prop; fun derive_size_simp size_def' simp0 = (trans OF [size_def', simp0]) |> Simplifier.asm_full_simplify (ss_only (@{thms inj_on_convol_ident id_def o_def snd_conv} @ all_inj_maps @ nested_size_maps) lthy2) |> Local_Defs.fold lthy2 size_defs_unused_0; fun derive_overloaded_size_simp overloaded_size_def' simp0 = (trans OF [overloaded_size_def', simp0]) |> unfold_thms lthy2 @{thms add_0_left add_0_right} |> Local_Defs.fold lthy2 (overloaded_size_defs @ all_overloaded_size_defs_of lthy2); val size_simpss = map2 (map o derive_size_simp) size_defs' rec_thmss; val size_simps = flat size_simpss; val overloaded_size_simpss = map2 (map o derive_overloaded_size_simp) overloaded_size_defs' size_simpss; val overloaded_size_simps = flat overloaded_size_simpss; val size_thmss = map2 append size_simpss overloaded_size_simpss; val size_gen_thmss = size_simpss; fun rhs_is_zero thm = let val Const (trueprop, _) $ (Const (eq, _) $ _ $ rhs) = Thm.prop_of thm in trueprop = \<^const_name>\Trueprop\ andalso eq = \<^const_name>\HOL.eq\ andalso rhs = HOLogic.zero end; val size_neq_thmss = @{map 3} (fn fp_sugar => fn size => fn size_thms => if exists rhs_is_zero size_thms then [] else let val (xs, _) = mk_Frees "x" (binder_types (fastype_of size)) lthy2; val goal = HOLogic.mk_Trueprop (BNF_LFP_Util.mk_not_eq (list_comb (size, xs)) HOLogic.zero); val vars = Variable.add_free_names lthy2 goal []; val thm = Goal.prove_sorry lthy2 vars [] goal (fn {context = ctxt, ...} => mk_size_neq ctxt (map (Thm.cterm_of lthy2) xs) (#exhaust (#ctr_sugar (#fp_ctr_sugar fp_sugar))) size_thms) |> single |> map (Thm.close_derivation \<^here>); in thm end) fp_sugars overloaded_size_consts overloaded_size_simpss; val ABs = As ~~ Bs; val g_names = variant_names num_As "g"; val gs = map2 (curry Free) g_names (map (op -->) ABs); val liveness = map (op <>) ABs; val live_gs = AList.find (op =) (gs ~~ liveness) true; val live = length live_gs; val maps0 = map map_of_bnf fp_bnfs; val size_gen_o_map_thmss = if live = 0 then replicate nn [] else let val gmaps = map (fn map0 => Term.list_comb (mk_map live As Bs map0, live_gs)) maps0; val size_gen_o_map_conds = if exists (can Logic.dest_implies o Thm.prop_of) nested_size_gen_o_maps then map (HOLogic.mk_Trueprop o mk_inj) live_gs else []; val fsizes = map (fn size_constB => Term.list_comb (size_constB, fsB)) size_constsB; val size_gen_o_map_lhss = map2 (curry HOLogic.mk_comp) fsizes gmaps; val fgs = map2 (fn fB => fn g as Free (_, Type (_, [A, B])) => if A = B then fB else HOLogic.mk_comp (fB, g)) fsB gs; val size_gen_o_map_rhss = map (fn c => Term.list_comb (c, fgs)) size_consts; val size_gen_o_map_goals = map2 (fold_rev (fold_rev Logic.all) [fsB, gs] o curry Logic.list_implies size_gen_o_map_conds o HOLogic.mk_Trueprop oo curry HOLogic.mk_eq) size_gen_o_map_lhss size_gen_o_map_rhss; val rec_o_maps = fold_rev (curry (op @) o #co_rec_o_maps o the o #fp_co_induct_sugar) fp_sugars []; val size_gen_o_map_thmss = if nested_size_gen_o_maps_complete andalso forall (fn TFree (_, S) => S = \<^sort>\type\) As then @{map 3} (fn goal => fn size_def => fn rec_o_map => Goal.prove_sorry lthy2 [] [] goal (fn {context = ctxt, ...} => mk_size_gen_o_map_tac ctxt size_def rec_o_map all_inj_maps nested_size_maps) |> Thm.close_derivation \<^here> |> single) size_gen_o_map_goals size_defs rec_o_maps else replicate nn []; in size_gen_o_map_thmss end; val massage_multi_notes = maps (fn (thmN, thmss, attrs) => map2 (fn T_name => fn thms => ((Binding.qualify true (Long_Name.base_name T_name) (Binding.name thmN), attrs), [(thms, [])])) T_names thmss) #> filter_out (null o fst o hd o snd); val notes = [(sizeN, size_thmss, nitpicksimp_attrs @ simp_attrs), (size_genN, size_gen_thmss, []), (size_gen_o_mapN, size_gen_o_map_thmss, []), (size_neqN, size_neq_thmss, [])] |> massage_multi_notes; val (noted, lthy3) = lthy2 |> Spec_Rules.add Binding.empty Spec_Rules.equational size_consts size_simps |> Spec_Rules.add Binding.empty Spec_Rules.equational overloaded_size_consts overloaded_size_simps |> Code.declare_default_eqns (map (rpair true) (flat size_thmss)) (*Ideally, this would be issued only if the "code" plugin is enabled.*) |> Local_Theory.notes notes; val phi0 = substitute_noted_thm noted; in lthy3 |> Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => Data.map (@{fold 3} (fn T_name => fn Const (size_name, _) => fn overloaded_size_def => let val morph = Morphism.thm (phi0 $> phi) in Symtab.update (T_name, (size_name, (morph overloaded_size_def, map morph overloaded_size_simps, maps (map morph) size_gen_o_map_thmss))) end) T_names size_consts overloaded_size_defs)) end | generate_datatype_size _ lthy = lthy; val size_plugin = Plugin_Name.declare_setup \<^binding>\size\; val _ = Theory.setup (fp_sugars_interpretation size_plugin generate_datatype_size); end; diff --git a/src/HOL/Tools/BNF/bnf_lift.ML b/src/HOL/Tools/BNF/bnf_lift.ML --- a/src/HOL/Tools/BNF/bnf_lift.ML +++ b/src/HOL/Tools/BNF/bnf_lift.ML @@ -1,2123 +1,2123 @@ (* Title: HOL/Tools/BNF/bnf_lift.ML Author: Julian Biendarra, TU Muenchen Author: Basil Fürer, ETH Zurich Author: Joshua Schneider, ETH Zurich Author: Dmitriy Traytel, ETH Zurich Copyright 2015, 2019 Lifting of BNFs through typedefs. *) signature BNF_LIFT = sig datatype lift_bnf_option = Plugins_Option of Proof.context -> Plugin_Name.filter | No_Warn_Wits | No_Warn_Transfer val copy_bnf: (((lift_bnf_option list * (binding option * (string * sort option)) list) * string) * thm option) * (binding * binding * binding) -> local_theory -> local_theory val copy_bnf_cmd: (((lift_bnf_option list * (binding option * (string * string option)) list) * string) * (Facts.ref * Token.src list) option) * (binding * binding * binding) -> local_theory -> local_theory val lift_bnf: ((((lift_bnf_option list * (binding option * (string * sort option)) list) * string) * term list option) * thm list option) * (binding * binding * binding) -> ({context: Proof.context, prems: thm list} -> tactic) list -> local_theory -> local_theory val lift_bnf_cmd: ((((lift_bnf_option list * (binding option * (string * string option)) list) * string) * string list) * (Facts.ref * Token.src list) list option) * (binding * binding * binding) -> local_theory -> Proof.state end structure BNF_Lift : BNF_LIFT = struct open Ctr_Sugar_Tactics open BNF_Util open BNF_Comp open BNF_Def datatype lift_bnf_option = Plugins_Option of Proof.context -> Plugin_Name.filter | No_Warn_Wits | No_Warn_Transfer; datatype equiv_thm = Typedef | Quotient of thm (** Util **) fun last2 [x, y] = ([], (x, y)) | last2 (x :: xs) = last2 xs |>> (fn y => x :: y) | last2 [] = raise Match; fun strip3 thm = (case Term.strip_comb (HOLogic.dest_Trueprop (Thm.prop_of thm)) of (_, [x1, x2, x3]) => (x1, x2, x3) | _ => error "strip3: wrong number of arguments"); val mk_Free = yield_singleton o mk_Frees; fun TWICE t = t THEN' t; fun prove lthy fvars tm tac = Goal.prove_sorry lthy (map (fst o dest_Free) fvars) [] tm (fn {context, ...} => tac context); (** Term construction **) fun mk_relT aT bT = aT --> bT --> HOLogic.boolT; fun mk_relcompp r s = let val (rT, sT) = apply2 fastype_of (r, s); val ((xT, _), (_, zTs)) = apply2 dest_funT (rT, sT); val T = rT --> sT --> mk_relT xT (fst (dest_funT zTs)); in Const (@{const_name relcompp}, T) $ r $ s end; val mk_OO = fold_rev mk_relcompp; (* option from sum *) fun mk_MaybeT T = mk_sumT (HOLogic.unitT, T); fun mk_Nothing T = BNF_FP_Util.mk_Inl T HOLogic.unit; val Just_const = BNF_FP_Util.Inr_const HOLogic.unitT; fun mk_Just tm = Just_const (fastype_of tm) $ tm; fun Maybe_map_const T = let val (xT, yT) = dest_funT T in Const (@{const_name map_sum}, (HOLogic.unitT --> HOLogic.unitT) --> T --> mk_MaybeT xT --> mk_MaybeT yT) $ HOLogic.id_const HOLogic.unitT end; fun mk_Maybe_map tm = Maybe_map_const (fastype_of tm) $ tm; fun fromJust_const T = Const (@{const_name sum.projr}, mk_MaybeT T --> T) fun rel_Maybe_const T U = Const (@{const_name rel_sum}, (HOLogic.unitT --> HOLogic.unitT --> HOLogic.boolT) --> (T --> U --> HOLogic.boolT) --> mk_MaybeT T --> mk_MaybeT U --> HOLogic.boolT) $ HOLogic.eq_const HOLogic.unitT fun set_Maybe_const T = Const (@{const_name Basic_BNFs.setr}, mk_MaybeT T --> HOLogic.mk_setT T) fun Inf_const T = Const (@{const_name Inf}, HOLogic.mk_setT T --> T); fun Image_const T = let val relT = HOLogic.mk_setT (HOLogic.mk_prodT (T, T)); val setT = HOLogic.mk_setT T; in Const (@{const_name Image}, relT --> setT --> setT) end; fun bot_const T = Const (@{const_name bot}, T); fun mk_insert x S = Const (@{const_name Set.insert}, fastype_of x --> fastype_of S --> fastype_of S) $ x $ S; fun mk_vimage f s = let val (xT, yT) = dest_funT (fastype_of f) in Const (@{const_name vimage}, (xT --> yT) --> HOLogic.mk_setT yT --> HOLogic.mk_setT xT) $ f $ s end; fun mk_case_prod (x, y) tm = let val ((x, xT), (y, yT)) = apply2 dest_Free (x, y); val prodT = HOLogic.mk_prodT (xT, yT); in HOLogic.Collect_const prodT $ (Const (@{const_name case_prod}, (xT --> yT --> HOLogic.boolT) --> prodT --> HOLogic.boolT) $ absfree (x, xT) (absfree (y, yT) tm)) end; fun mk_Trueprop_implies (ps, c) = Logic.list_implies (map HOLogic.mk_Trueprop ps, HOLogic.mk_Trueprop c); fun mk_Collect (v, tm) = let val (n, T) = dest_Free v in HOLogic.mk_Collect (n, T, tm) end; (** witnesses **) fun prepare_wits is_quotient RepT wits opts alphas wits_F var_as var_as' sets lthy = let fun binder_types_until_eq V T = let fun strip (TU as Type ("fun", [T, U])) = if V = TU then [] else T :: strip U | strip T = if V = T then [] else error ("Bad type for witness: " ^ quote (Syntax.string_of_typ lthy T)); in strip T end; val Iwits = the_default wits_F (Option.map (map (`(map (fn T => find_index (fn U => T = U) alphas) o binder_types_until_eq RepT o fastype_of))) wits); val Iwits = if is_quotient then let val subsumed_Iwits = filter (fn (J, _) => exists (fn (I, _) => subset (op =) (I, J)) wits_F) Iwits; val _ = if null subsumed_Iwits orelse exists (fn No_Warn_Wits => true | _ => false) opts then () else let val (suffix1, suffix2, be) = (if length subsumed_Iwits = 1 then ("", "", "is") else ("s", "es", "are")) in subsumed_Iwits |> map (Syntax.pretty_typ lthy o fastype_of o snd) |> Pretty.big_list ("The following type" ^ suffix1 ^ " of nonemptiness witness" ^ suffix2 ^ " of the raw type's BNF " ^ be ^ " subsumed by the existing raw witnesses:") |> (fn pt => Pretty.chunks [pt, Pretty.para ("You do not need to lift these subsumed witnesses.")]) |> Pretty.string_of |> warning end; in filter_out (fn (J, _) => exists (fn (I, _) => subset (op =) (I, J)) wits_F) Iwits end else Iwits; val wit_goals = maps (BNF_Def.mk_wit_goals var_as var_as' sets) Iwits; val lost_wits = if is_quotient then [] else filter_out (fn (J, _) => exists (fn (I, _) => subset (op =) (I, J)) Iwits) wits_F; val _ = if null lost_wits orelse exists (fn No_Warn_Wits => true | _ => false) opts then () else let val what = (if is_quotient then "quotient type" else "typedef"); val (suffix1, suffix2, be) = (if length lost_wits = 1 then ("", "", "was") else ("s", "es", "were")) in lost_wits |> map (Syntax.pretty_typ lthy o fastype_of o snd) |> Pretty.big_list ("The following type" ^ suffix1 ^ " of nonemptiness witness" ^ suffix2 ^ " of the raw type's BNF " ^ be ^ " lost:") |> (fn pt => Pretty.chunks [pt, Pretty.para ("You can specify a liftable witness (e.g., a term of one of the above types\ \ that satisfies the " ^ what ^ "'s invariant)\ \ using the annotation [wits: ].")]) |> Pretty.string_of |> warning end; in (Iwits, wit_goals) end; (** transfer theorems **) fun mk_transfer_thms' bnf_F bnf_G name consts Tss crel_def pcrel_def defs lthy = let val live = length (#alphas Tss); val (pcrel_tm, crel_tm) = apply2 (Thm.prop_of #> Logic.dest_equals #> fst #> head_of) (pcrel_def, crel_def); val (var_Qs, var_Rs) = fold Variable.declare_typ (#alphas Tss @ #deads Tss) lthy |> mk_Frees "Q" (map2 mk_relT (#alphas Tss) (#betas Tss)) ||>> mk_Frees "R" (map2 mk_relT (#gammas Tss) (#deltas Tss)) |> fst; (* get the "pcrel :: args_raw => rep => abs \ bool" term and instantiate types *) val (args_raw, (rep, abs)) = pcrel_tm |> fastype_of |> binder_types |> last2; val thy = Proof_Context.theory_of lthy; val tyenv_match = Vartab.empty |> Sign.typ_match thy ((rep, #rep Tss)) |> Sign.typ_match thy ((abs, #abs Tss)); val args = map (Envir.subst_type tyenv_match) args_raw; val (pcrel_a, pcrel_b) = Envir.subst_term (tyenv_match, Vartab.empty) pcrel_tm |> `(subst_atomic_types (#alphas Tss @ #betas Tss ~~ #gammas Tss @ #deltas Tss)); (* match "crel :: {?a F} \ a G" with "rep_G :: {a F}" *) val tyenv_match = Vartab.empty |> Sign.typ_match thy (crel_tm |> fastype_of |> binder_types |> hd, #rep Tss); val crel_b = Envir.subst_term (tyenv_match, Vartab.empty) crel_tm |> subst_atomic_types (#alphas Tss ~~ #betas Tss) val crel_d = subst_atomic_types (#betas Tss ~~ #deltas Tss) crel_b; (* instantiate pcrel with Qs and Rs*) val dead_args = map binder_types args |> map (fn [T,U] => if T = U then SOME T else NONE | _ => NONE); val parametrize = let fun go (SOME T :: dTs) tms = HOLogic.eq_const T :: go dTs tms | go (_ :: dTs) (tm :: tms) = tm :: go dTs tms | go (_ :: dTs) tms = go dTs tms | go _ _ = []; in go dead_args end; val pcrel_Qs = list_comb (pcrel_b, parametrize var_Qs); val pcrel_Rs = list_comb (pcrel_a, parametrize var_Rs); (* get the order of the dead variables right *) val Ds0 = maps the_list dead_args; val permute_Ds = (mk_T_of_bnf Ds0 (#betas Tss) bnf_G, nth (binder_types (type_of pcrel_Qs)) 1) |> apply2 (fn Type (_, Ts) => Ts | _ => []) |> op~~ |> typ_subst_atomic; val Ds0 = map permute_Ds Ds0; (* terms for sets of the set_transfer thms *) val rel_sets_Q = @{map 3} (fn aT => fn bT => fn Q => mk_rel 1 [aT] [bT] @{term rel_set} $ Q) (#alphas Tss) (#betas Tss) var_Qs; (* rewrite rules for pcrel and BNF's relator: "pcrel Q = rel_F OO crel" *) fun mk_pcr_rel_F_eq Ts Us pcrel vs crel = let val thm = HOLogic.mk_Trueprop (HOLogic.mk_eq (pcrel, mk_relcompp (list_comb (mk_rel_of_bnf (#deads Tss) (Ts Tss) (Us Tss) bnf_F, vs)) crel)); fun tac ctxt = unfold_thms_tac ctxt (pcrel_def :: defs @ @{thm id_bnf_apply} :: Transfer.get_relator_eq ctxt) THEN (HEADGOAL (rtac ctxt refl)) in prove lthy vs thm tac |> mk_abs_def end; val pcr_rel_F_eqs = [mk_pcr_rel_F_eq #alphas #betas pcrel_Qs var_Qs crel_b, mk_pcr_rel_F_eq #gammas #deltas pcrel_Rs var_Rs crel_d]; (* create transfer-relations from term('s type) *) fun mk_transfer_rel' i tm = let fun go T' (n, q) = case T' of Type ("fun", [T as Type ("fun", _), U]) => go U (n+1, q) |>> mk_rel_fun (fst (go T (n, q))) | Type ("fun", [T, U]) => go T (n, q) |-> (fn x => fn st => go U st |>> mk_rel_fun x) | Type (@{type_name bool}, _) => (HOLogic.eq_const HOLogic.boolT, (n, q)) | Type (@{type_name set}, _) => (nth rel_sets_Q n, (n, q)) | Type _ => (if q then pcrel_Qs else pcrel_Rs, (n, not q)) | TFree _ => (nth (if q then var_Qs else var_Rs) n, (n, not q)) | _ => raise Match in go (fastype_of tm) (i, true) |> fst end; (* prove transfer rules *) fun prove_transfer_thm' i vars new_tm const = let val tm = HOLogic.mk_Trueprop (mk_transfer_rel' i new_tm $ #raw const $ new_tm); val tac = (fn ctxt => unfold_thms_tac ctxt (pcr_rel_F_eqs @ [crel_def]) THEN #tac const {Rs=var_Rs,Qs=var_Qs} ctxt); in prove lthy vars tm tac end; val prove_transfer_thm = prove_transfer_thm' 0; (* map transfer: "((Q ===> R) ===> pcr_G Q ===> pcr_G R) map_F' map_G" *) val map_G = mk_map_of_bnf Ds0 (#betas Tss) (#deltas Tss) bnf_G; val map_transfer = prove_transfer_thm (var_Qs @ var_Rs) map_G (#map consts); (* pred_transfer: "((Q ===> (=)) ===> pcr_G Q ===> (=)) pred_F' pred_G" *) val pred_G = mk_pred_of_bnf Ds0 (#betas Tss) bnf_G; val pred_transfer = #pred consts |> Option.map (fn p => ("pred", [prove_transfer_thm (var_Qs @ var_Rs) pred_G p])); (* rel_transfer: "((Q ===> R ===> (=)) ===> pcr_G Q ===> pcr_G R ===> (=)) rel_F' rel_G" *) val rel_G = mk_rel_of_bnf Ds0 (#betas Tss) (#deltas Tss) bnf_G; val rel_transfer = prove_transfer_thm (var_Qs @ var_Rs) rel_G (#rel consts); (* set_transfer: "(pcr_G Q ===> rel_set Q) set_F' set_G" *) val sets_G = mk_sets_of_bnf (replicate live Ds0) (replicate live (#betas Tss)) bnf_G; fun mk_set_transfer i set_G raw tac = prove_transfer_thm' i var_Qs set_G {raw=raw, tac=tac}; val sets_transfer = @{map 4} mk_set_transfer (0 upto (live-1)) sets_G (#raws (#sets consts)) (#tacs (#sets consts)); (* export transfer theorems *) val transform = Morphism.thm (Morphism.thm_morphism "BNF" (unfold_thms lthy defs)) |> map; val b = Binding.qualified_name name; val qualify = let val qs = Binding.path_of b; in fold_rev (fn (s, mand) => Binding.qualify mand s) qs end; fun mk_binding n = Binding.name (n ^ "_transfer_raw") |> Binding.qualify true (Binding.name_of b) |> qualify; val notes = [("map", [map_transfer]), ("rel", [rel_transfer])] @ the_list pred_transfer @ [("set", sets_transfer)] |> map (fn (n, thms) => ((mk_binding n, []), [(transform thms, @{attributes [transfer_rule]})])); in lthy |> Local_Theory.notes notes |> snd end; (* transfer theorems for map, pred (in case of a typedef), rel and sets *) fun mk_transfer_thms quiet bnf_F bnf_G name consts thm Tss defs lthy = let fun mk_crel_def quot_thm = (case thm of Quotient equiv => @{thm Quotient_crel_quotient} OF [quot_thm, equiv] | Typedef => hd ([quot_thm] RL @{thms Quotient_crel_typedef Quotient_crel_typecopy})); fun no_quotient _ = [Pretty.para ("No quotient theorem has been registered for " ^ Binding.name_of (name_of_bnf bnf_G) ^ "."), Pretty.para "Use setup_lifting to register a quotient or type definition theorem."]; fun wrong_quotient T lthy = [Pretty.para ("A wrong quotient theorem has been registered for " ^ Binding.name_of (name_of_bnf bnf_G) ^ "."), Pretty.para ("Expected raw type " ^ Pretty.string_of (Syntax.pretty_typ lthy (T_of_bnf bnf_F)) ^ " but the quotient theorem has raw type " ^ Pretty.string_of (Syntax.pretty_typ lthy T) ^ "."), Pretty.para "Use setup_lifting to register a different quotient or type definition theorem."]; fun pcr_why _ = [Pretty.para ("The pcr_" ^ Binding.name_of (name_of_bnf bnf_G) ^ " relator has not been defined.")]; fun warn_transfer why lthy = (Pretty.para "The transfer theorems can't be generated:" :: why lthy) |> Pretty.chunks |> Pretty.string_of |> warning |> K lthy; fun maybe_warn_transfer why = not quiet ? warn_transfer why; in case Lifting_Info.lookup_quotients lthy name of SOME {pcr_info, quot_thm} => (let val crel_def = mk_crel_def quot_thm; val rty = Lifting_Util.quot_thm_rty_qty quot_thm |> fst; val thy = Proof_Context.theory_of lthy; in if Sign.typ_instance thy (rty, T_of_bnf bnf_F) then (case pcr_info of SOME {pcrel_def, ...} => mk_transfer_thms' bnf_F bnf_G name consts Tss crel_def pcrel_def defs lthy | _ => maybe_warn_transfer pcr_why lthy) else maybe_warn_transfer (wrong_quotient rty) lthy end) | _ => maybe_warn_transfer no_quotient lthy end; (** typedef_bnf **) fun mk_typedef_transfer_tacs bnf_F bnf_G thms old_defs map_raw rel_raw pred_raw sets_raw = let val live = live_of_bnf bnf_G; val Abs_G_inverse = @{thm type_definition.Abs_inverse} OF [#typedef thms]; val Rep_G = @{thm type_definition.Rep} OF [#typedef thms]; fun common_tac addefs tac = (fn _ => fn ctxt => HEADGOAL (EVERY' [SELECT_GOAL (unfold_thms_tac ctxt addefs), REPEAT_DETERM o rtac ctxt rel_funI, SELECT_GOAL (unfold_thms_tac ctxt @{thms o_apply}), REPEAT_DETERM o eresolve_tac ctxt @{thms relcomppE exE conjE}, hyp_subst_tac ctxt]) THEN tac ctxt) fun map_tac ctxt = (HEADGOAL o EVERY') [rtac ctxt @{thm relcomppI}, etac ctxt (mk_rel_funDN (live+1) (map_transfer_of_bnf bnf_F)), REPEAT_DETERM_N live o assume_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt [Abs_G_inverse OF [#map_closed thms] OF [Rep_G]]), REPEAT_DETERM o rtac ctxt refl]; val map_tac = common_tac [#map old_defs] map_tac; fun rel_tac ctxt = HEADGOAL (etac ctxt (mk_rel_funDN (live+2) (rel_transfer_of_bnf bnf_F)) THEN' REPEAT_DETERM_N (live+1) o assume_tac ctxt); val rel_tac = common_tac (#rel old_defs :: @{thms vimage2p_def}) rel_tac; fun pred_tac ctxt = HEADGOAL (etac ctxt (mk_rel_funDN (live+1) (pred_transfer_of_bnf bnf_F)) THEN' REPEAT_DETERM_N live o (assume_tac ctxt)); val pred_tac = common_tac [#pred old_defs] pred_tac; fun set_tac set_transfer_thm ctxt = HEADGOAL (etac ctxt (rel_funD OF [set_transfer_thm])); fun mk_set_tac set_def set_transfer = common_tac [set_def] (set_tac set_transfer); val set_tacs = map2 mk_set_tac (#sets old_defs) (set_transfer_of_bnf bnf_F); in {map={raw=map_raw,tac=map_tac},rel={raw=rel_raw,tac=rel_tac}, sets={raws=sets_raw,tacs=set_tacs},pred=SOME{raw=pred_raw,tac=pred_tac}} end; fun typedef_bnf thm wits specs map_b rel_b pred_b opts lthy = let val plugins = get_first (fn Plugins_Option f => SOME (f lthy) | _ => NONE) (rev opts) |> the_default Plugin_Name.default_filter; (* extract Rep Abs F RepT AbsT *) val (Rep_G, Abs_G, F) = strip3 thm; val typ_Abs_G = dest_funT (fastype_of Abs_G); val RepT = fst typ_Abs_G; (* F *) val AbsT = snd typ_Abs_G; (* G *) val AbsT_name = fst (dest_Type AbsT); val tvs = AbsT |> dest_Type |> snd |> map (fst o dest_TVar); val alpha0s = map (TFree o snd) specs; val _ = length tvs = length alpha0s orelse error ("Expected " ^ string_of_int (length tvs) ^ " type argument(s) to " ^ quote AbsT_name); (* instantiate the new type variables newtvs to oldtvs *) val subst = subst_TVars (tvs ~~ alpha0s); val typ_subst = typ_subst_TVars (tvs ~~ alpha0s); val Rep_G = subst Rep_G; val Abs_G = subst Abs_G; val F = subst F; val RepT = typ_subst RepT; val AbsT = typ_subst AbsT; fun flatten_tyargs Ass = map dest_TFree alpha0s |> filter (fn T => exists (fn Ts => member op= Ts T) Ass); val Ds0 = filter (is_none o fst) specs |> map snd; (* get the bnf for RepT *) val ((bnf_F, (deads, alphas)), ((_, unfolds), lthy)) = bnf_of_typ true Dont_Inline (Binding.qualify true AbsT_name) flatten_tyargs [] Ds0 RepT ((empty_comp_cache, empty_unfolds), lthy); val set_bs = map (fn T => find_index (fn U => T = U) alpha0s) alphas |> map (the_default Binding.empty o fst o nth specs); val _ = (case alphas of [] => error "No live variables" | _ => ()); val defs = #map_unfolds unfolds @ flat (#set_unfoldss unfolds) @ #rel_unfolds unfolds @ #pred_unfolds unfolds; (* number of live variables *) val live = length alphas; (* state the three required properties *) val sorts = map Type.sort_of_atyp alphas; val names_lthy = fold Variable.declare_typ (alphas @ deads) lthy; val (((alphas', betas), betas'), names_lthy) = names_lthy |> mk_TFrees' sorts ||>> mk_TFrees' sorts ||>> mk_TFrees' sorts; val map_F = mk_map_of_bnf deads alphas betas bnf_F; val (typ_fs, typ_aF) = fastype_of map_F |> strip_typeN live ||> domain_type; val typ_pairs = map HOLogic.mk_prodT (alphas ~~ alphas'); val typ_subst_pair = typ_subst_atomic (alphas ~~ typ_pairs); val typ_pair = typ_subst_pair RepT; val subst_b = subst_atomic_types (alphas ~~ betas); val subst_a' = subst_atomic_types (alphas ~~ alphas'); val subst_pair = subst_atomic_types (alphas ~~ typ_pairs); val aF_set = F; val aF_set' = subst_a' F; val pairF_set = subst_pair F; val bF_set = subst_b F; val map_F_fst = mk_map_of_bnf deads typ_pairs alphas bnf_F; val map_F_snd = mk_map_of_bnf deads typ_pairs alphas' bnf_F val sets_F_pairs = mk_sets_of_bnf (replicate live deads) (replicate live typ_pairs) bnf_F val wits_F = mk_wits_of_bnf (replicate (nwits_of_bnf bnf_F) deads) (replicate (nwits_of_bnf bnf_F) alphas) bnf_F; (* val map_closed_F = @{term "\f x. x \ F \ map_F f x \ F"}; *) val (var_fs, names_lthy) = mk_Frees "f" typ_fs names_lthy; val (var_x, names_lthy) = mk_Frees "x" [typ_aF] names_lthy |>> the_single; val mem_x = HOLogic.mk_Trueprop (HOLogic.mk_mem (var_x, aF_set)); val map_f = list_comb (map_F, var_fs); val mem_map = HOLogic.mk_Trueprop (HOLogic.mk_mem (map_f $ var_x, bF_set)); val imp_map = Logic.mk_implies (mem_x, mem_map); val map_closed_F = fold_rev Logic.all var_fs (Logic.all var_x imp_map); (* val zip_closed_F = @{term "\z. \map_F fst z \ F; map_F snd z \ F\ \ \z' \ F. set_F z' \ set_F z \ map_F fst z' = map_F fst z \ map_F snd z' = map_F snd z"}; *) val (var_z, names_lthy) = mk_Free "z" typ_pair names_lthy; val (var_z', names_lthy) = mk_Free "z'" typ_pair names_lthy; val (pairs, names_lthy) = mk_Frees "tmp" typ_pairs names_lthy; fun mk_map mfs f z = Term.list_comb (mfs, map (fst o Term.strip_comb o f) pairs) $ z; fun mk_set var = map (fn t => t $ var) sets_F_pairs; val (map_fst', map_fst) = apply2 (mk_map map_F_fst HOLogic.mk_fst) (var_z', var_z); val (map_snd', map_snd) = apply2 (mk_map map_F_snd HOLogic.mk_snd) (var_z', var_z); val mem_map_fst = HOLogic.mk_Trueprop (HOLogic.mk_mem (map_fst, aF_set)); val mem_map_snd = HOLogic.mk_Trueprop (HOLogic.mk_mem (map_snd, aF_set')); val ex_conj = foldr1 HOLogic.mk_conj (map2 mk_leq (mk_set var_z') (mk_set var_z) @ [HOLogic.mk_eq (map_fst', map_fst), HOLogic.mk_eq (map_snd', map_snd)]); val zip_concl = HOLogic.mk_Trueprop (mk_Bex pairF_set (absfree (dest_Free var_z') ex_conj)); val zip_closed_F = Logic.all var_z (Logic.list_implies ([mem_map_fst, mem_map_snd], zip_concl)); (* val wit_closed_F = @{term "wit_F a \ F"}; *) val (var_as, names_lthy) = mk_Frees "a" alphas names_lthy; val (var_bs, _) = mk_Frees "a" alphas names_lthy; val sets = mk_sets_of_bnf (replicate live deads) (replicate live alphas) bnf_F; val (Iwits, wit_goals) = prepare_wits false RepT wits opts alphas wits_F var_as var_bs sets lthy; val wit_closed_Fs = Iwits |> map (fn (I, wit_F) => let val vars = map (nth var_as) I; val wit_a = list_comb (wit_F, vars); in fold_rev Logic.all vars (HOLogic.mk_Trueprop (HOLogic.mk_mem (wit_a, aF_set))) end); val goals = [map_closed_F, zip_closed_F] @ wit_closed_Fs @ (case wits of NONE => [] | _ => wit_goals); fun after_qed ([map_closed_thm] :: [zip_closed_thm] :: wit_thmss) lthy = let val (wit_closed_thms, wit_thms) = (case wits of NONE => (map the_single wit_thmss, wit_thms_of_bnf bnf_F) | _ => chop (length wit_closed_Fs) (map the_single wit_thmss)) (* construct map set bd rel wit *) (* val map_G = @{term "\f. Abs_G o map_F f o Rep_G"}; *) val Abs_Gb = subst_b Abs_G; val map_G = fold_rev lambda var_fs (HOLogic.mk_comp (HOLogic.mk_comp (Abs_Gb, map_f), Rep_G)); val map_raw = fold_rev lambda var_fs map_f; (* val sets_G = [@{term "set_F o Rep_G"}]; *) val sets_F = mk_sets_of_bnf (replicate live deads) (replicate live alphas) bnf_F; val sets_G = map (fn set_F => HOLogic.mk_comp (set_F, Rep_G)) sets_F; (* val bd_G = @{term "bd_F"}; *) val bd_F = mk_bd_of_bnf deads alphas bnf_F; val bd_G = bd_F; (* val rel_G = @{term "\R. BNF_Def.vimage2p Rep_G Rep_G (rel_F R)"}; *) val rel_F = mk_rel_of_bnf deads alphas betas bnf_F; val (typ_Rs, _) = strip_typeN live (fastype_of rel_F); val (var_Rs, names_lthy) = mk_Frees "R" typ_Rs lthy; val Rep_Gb = subst_b Rep_G; val rel_G = fold_rev absfree (map dest_Free var_Rs) (mk_vimage2p Rep_G Rep_Gb $ list_comb (rel_F, var_Rs)); val rel_raw = fold_rev absfree (map dest_Free var_Rs) (list_comb (rel_F, var_Rs)); (* val pred_G = @{term "\P. pred_F P o Rep_G"}; *) val pred_F = mk_pred_of_bnf deads alphas bnf_F; val (typ_Ps, _) = strip_typeN live (fastype_of pred_F); val (var_Ps, names_lthy) = mk_Frees "P" typ_Ps names_lthy; val pred_G = fold_rev absfree (map dest_Free var_Ps) (HOLogic.mk_comp (list_comb (pred_F, var_Ps), Rep_G)); val pred_raw = fold_rev absfree (map dest_Free var_Ps) (list_comb (pred_F, var_Ps)); (* val wits_G = [@{term "Abs_G o wit_F"}]; *) val (var_as, _) = mk_Frees "a" alphas names_lthy; val wits_G = map (fn (I, wit_F) => let val vs = map (nth var_as) I; in fold_rev absfree (map dest_Free vs) (Abs_G $ (list_comb (wit_F, vs))) end) Iwits; (* tactics *) val Rep_thm = thm RS @{thm type_definition.Rep}; val Abs_inverse_thm = thm RS @{thm type_definition.Abs_inverse}; val Abs_inject_thm = thm RS @{thm type_definition.Abs_inject}; val Rep_cases_thm = thm RS @{thm type_definition.Rep_cases}; val Rep_inverse_thm = thm RS @{thm type_definition.Rep_inverse}; fun map_id0_tac ctxt = HEADGOAL (EVERY' [rtac ctxt ext, SELECT_GOAL (unfold_thms_tac ctxt [map_id0_of_bnf bnf_F, id_apply, o_apply, Rep_inverse_thm]), rtac ctxt refl]); fun map_comp0_tac ctxt = HEADGOAL (EVERY' [rtac ctxt ext, SELECT_GOAL (unfold_thms_tac ctxt [map_comp0_of_bnf bnf_F, o_apply, Rep_thm RS (map_closed_thm RS Abs_inverse_thm)]), rtac ctxt refl]); fun map_cong0_tac ctxt = HEADGOAL (EVERY' ([SELECT_GOAL (unfold_thms_tac ctxt [o_apply]), rtac ctxt (([Rep_thm RS map_closed_thm, Rep_thm RS map_closed_thm] MRS Abs_inject_thm) RS iffD2), rtac ctxt (map_cong0_of_bnf bnf_F)] @ replicate live (Goal.assume_rule_tac ctxt))); val set_map0s_tac = map (fn set_map => fn ctxt => HEADGOAL (EVERY' [rtac ctxt ext, SELECT_GOAL (unfold_thms_tac ctxt [set_map, o_apply, Rep_thm RS (map_closed_thm RS Abs_inverse_thm)]), rtac ctxt refl])) (set_map_of_bnf bnf_F); fun card_order_bd_tac ctxt = HEADGOAL (rtac ctxt (bd_card_order_of_bnf bnf_F)); fun cinfinite_bd_tac ctxt = HEADGOAL (rtac ctxt (bd_cinfinite_of_bnf bnf_F)); val set_bds_tac = map (fn set_bd => fn ctxt => HEADGOAL (EVERY' [SELECT_GOAL (unfold_thms_tac ctxt [o_apply]), rtac ctxt set_bd])) (set_bd_of_bnf bnf_F); fun le_rel_OO_tac ctxt = HEADGOAL (EVERY' [rtac ctxt @{thm vimage2p_relcompp_mono}, rtac ctxt ((rel_OO_of_bnf bnf_F RS sym) RS @{thm ord_eq_le_trans}), rtac ctxt @{thm order_refl}]); fun rel_OO_Grp_tac ctxt = HEADGOAL (EVERY' ([SELECT_GOAL (REPEAT_DETERM (HEADGOAL (rtac ctxt ext))), SELECT_GOAL (unfold_thms_tac ctxt [@{thm OO_Grp_alt}, mem_Collect_eq, o_apply, @{thm vimage2p_def}, in_rel_of_bnf bnf_F, Bex_def, mem_Collect_eq]), rtac ctxt iffI, SELECT_GOAL (REPEAT_DETERM (HEADGOAL (eresolve_tac ctxt [exE, conjE]))), forward_tac ctxt [zip_closed_thm OF (replicate 2 (Rep_thm RSN (2, @{thm ssubst_mem})))], assume_tac ctxt, SELECT_GOAL (REPEAT_DETERM (HEADGOAL (eresolve_tac ctxt [bexE, conjE]))), etac ctxt Rep_cases_thm, hyp_subst_tac ctxt, SELECT_GOAL (REPEAT_DETERM (HEADGOAL (rtac ctxt exI))), rtac ctxt conjI] @ replicate live (EVERY' [TRY o rtac ctxt conjI, etac ctxt @{thm order_trans}, assume_tac ctxt]) @ [SELECT_GOAL (REPEAT_DETERM (HEADGOAL (rtac ctxt conjI))), REPEAT_DETERM_N 2 o EVERY' [rtac ctxt (trans OF [iffD2 OF [Abs_inject_thm OF [map_closed_thm OF [Rep_thm], Rep_thm]], Rep_inverse_thm]), etac ctxt trans, assume_tac ctxt], SELECT_GOAL (REPEAT_DETERM (HEADGOAL (eresolve_tac ctxt [exE, conjE]))), rtac ctxt exI, rtac ctxt conjI] @ replicate (live-1) (rtac ctxt conjI THEN' assume_tac ctxt) @ [assume_tac ctxt, rtac ctxt conjI, REPEAT_DETERM_N 2 o EVERY' [rtac ctxt (iffD1 OF [Abs_inject_thm OF [map_closed_thm OF [Rep_thm], Rep_thm]]), etac ctxt (Rep_inverse_thm RS sym RSN (2, trans))]])); fun pred_set_tac ctxt = HEADGOAL (EVERY' [rtac ctxt (pred_set_of_bnf bnf_F RS @{thm arg_cong[of _ _ "\f. f \ _"]} RS trans), SELECT_GOAL (unfold_thms_tac ctxt (@{thms Ball_comp_iff conj_comp_iff})), rtac ctxt refl]); fun wit_tac ctxt = HEADGOAL (EVERY' (map (fn thm => (EVERY' [SELECT_GOAL (unfold_thms_tac ctxt (o_apply :: (wit_closed_thms RL [Abs_inverse_thm]))), dtac ctxt thm, assume_tac ctxt])) wit_thms)); val tactics = [map_id0_tac, map_comp0_tac, map_cong0_tac] @ set_map0s_tac @ [card_order_bd_tac, cinfinite_bd_tac] @ set_bds_tac @ [le_rel_OO_tac, rel_OO_Grp_tac, pred_set_tac]; val (bnf_G, lthy) = bnf_def Dont_Inline (user_policy Note_Some) true I tactics wit_tac NONE map_b rel_b pred_b set_bs (((((((Binding.empty, AbsT), map_G), sets_G), bd_G), wits_G), SOME rel_G), SOME pred_G) lthy; val old_defs = {sets = set_defs_of_bnf bnf_G, map = map_def_of_bnf bnf_G, rel = rel_def_of_bnf bnf_G, pred = pred_def_of_bnf bnf_G}; val unfold_morphism = Morphism.thm_morphism "BNF" (unfold_thms lthy defs); val (bnf_G, lthy) = morph_bnf_defs unfold_morphism bnf_G |> (fn bnf => note_bnf_defs bnf lthy); val quiet = exists (fn No_Warn_Transfer => true | _ => false) opts; val transfer_consts = mk_typedef_transfer_tacs bnf_F bnf_G {map_closed=map_closed_thm,typedef=thm} old_defs map_raw rel_raw pred_raw sets_F; in lthy |> BNF_Def.register_bnf plugins AbsT_name bnf_G |> mk_transfer_thms quiet bnf_F bnf_G AbsT_name transfer_consts Typedef {abs=typ_subst_atomic (alphas ~~ alphas') AbsT, rep=RepT, Ds0=map TFree Ds0, deads = deads, alphas=alphas, betas=alphas', gammas=betas, deltas=betas'} defs end | after_qed _ _ = raise Match; in (goals, after_qed, defs, lthy) end; (** quotient_bnf **) fun mk_quotient_transfer_tacs bnf_F Tss live qthms thms set_F'_thmss old_defs inst_REL_pos_distrI map_raw rel_raw sets_raw = let fun common_tac ctxt addefs = unfold_thms_tac ctxt (#REL qthms :: addefs) THEN (REPEAT_DETERM o HEADGOAL) (rtac ctxt rel_funI); (* quotient.map_transfer tactic *) val map_F_transfer = map_transfer_of_bnf bnf_F |> mk_rel_funDN (live+1); fun map_transfer_q _ ctxt = common_tac ctxt (#map old_defs :: @{thms o_def}) THEN (HEADGOAL o EVERY') [REPEAT_DETERM o eresolve_tac ctxt @{thms relcomppE}, rtac ctxt @{thm relcomppI[rotated]}, hyp_subst_tac ctxt THEN' EVERY' (map (rtac ctxt) [#rel_abs qthms, #map_F_rsp thms, (#rep_abs_rsp qthms), (#reflp qthms)]), hyp_subst_tac ctxt, rtac ctxt map_F_transfer, REPEAT_DETERM_N (live+1) o assume_tac ctxt]; (* quotient.rel_transfer tactic *) val rel_F_maps = rel_map_of_bnf bnf_F; val rel_F_map_iffD2s = map (fn thm => thm RS @{thm iffD2}) rel_F_maps; fun inst_REL_pos_distrI_order_refls vs aTs bTs ctxt = inst_REL_pos_distrI live vs aTs bTs ctxt OF (replicate (live+1) asm_rl @ replicate live @{thm order_refl}); fun rel_transfer_q {Qs, Rs} ctxt = EVERY [common_tac ctxt [#rel old_defs, @{thm vimage2p_def}], HEADGOAL (rtac ctxt iffI), (REPEAT_DETERM o ALLGOALS) (eresolve_tac ctxt @{thms exE conjE relcomppE} ORELSE' hyp_subst_tac ctxt), (HEADGOAL o EVERY') [REPEAT_DETERM o dtac ctxt @{thm rel_fun_rel_OO1}, rtac ctxt (inst_REL_pos_distrI 0 (map mk_conversep Qs) (#betas Tss) (#alphas Tss) ctxt), rtac ctxt @{thm relcomppI}, rtac ctxt (#symp qthms), rtac ctxt (#map_F_rsp thms), rtac ctxt (#rep_abs_rsp qthms), rtac ctxt (#reflp qthms), rtac ctxt @{thm relcomppI}, rtac ctxt (rel_flip_of_bnf bnf_F RS @{thm iffD1}), rtac ctxt (nth rel_F_map_iffD2s 0), rtac ctxt (nth rel_F_map_iffD2s 1), etac ctxt (#rel_monoD_rotated thms)], (REPEAT_DETERM_N live o HEADGOAL o EVERY') [rtac ctxt @{thm predicate2I}, rtac ctxt @{thm conversepI}, rtac ctxt @{thm Basic_BNFs.rel_sum_simps(4)[THEN iffD2]}, etac ctxt @{thm conversepI}], (HEADGOAL o EVERY') [rtac ctxt (inst_REL_pos_distrI_order_refls Rs (#gammas Tss) (#deltas Tss) ctxt), (SELECT_GOAL o REPEAT_DETERM o HEADGOAL) (etac ctxt @{thm relcomppI}), rtac ctxt @{thm relcomppI[rotated]}, rtac ctxt (#map_F_rsp thms), rtac ctxt (#rep_abs_rsp qthms OF [#reflp qthms]), SELECT_GOAL (unfold_thms_tac ctxt (@{thms rel_sum_simps} @ rel_F_maps)), assume_tac ctxt], (REPEAT_DETERM_N (2*live) o HEADGOAL) (rtac ctxt @{thm rel_sum_eq2_nonempty} ORELSE' rtac ctxt @{thm rel_sum_eq3_nonempty}), (REPEAT_DETERM_N live) (unfold_thms_tac ctxt @{thms sum.rel_compp[symmetric] eq_OO} THEN HEADGOAL (etac ctxt @{thm sum.rel_mono[OF order_refl]})), (HEADGOAL o EVERY') [(SELECT_GOAL o REPEAT_DETERM o HEADGOAL) (dtac ctxt @{thm rel_fun_rel_OO2}), rtac ctxt (inst_REL_pos_distrI 0 Qs (#alphas Tss) (#betas Tss) ctxt), rtac ctxt @{thm relcomppI}, rtac ctxt (#reflp qthms), rtac ctxt @{thm relcomppI}, rtac ctxt (nth rel_F_map_iffD2s 0), rtac ctxt (nth rel_F_map_iffD2s 1), etac ctxt (#rel_monoD_rotated thms)], (REPEAT_DETERM_N live o HEADGOAL o EVERY') [rtac ctxt @{thm predicate2I}, rtac ctxt @{thm rel_sum.intros(2)}, assume_tac ctxt], (HEADGOAL o EVERY') [rtac ctxt (inst_REL_pos_distrI_order_refls (map mk_conversep Rs) (#deltas Tss) (#gammas Tss) ctxt), rtac ctxt @{thm relcomppI}, etac ctxt (rotate_prems 1 (#transp qthms)), rtac ctxt (#map_F_rsp thms), rtac ctxt (#rep_abs_rsp qthms OF [#reflp qthms]), etac ctxt @{thm relcomppI}, rtac ctxt @{thm relcomppI}, etac ctxt (#transp qthms), rtac ctxt (#symp qthms), rtac ctxt (#map_F_rsp thms), rtac ctxt (#rep_abs_rsp qthms), rtac ctxt (#reflp qthms), rtac ctxt @{thm relcomppI[rotated]}, rtac ctxt (#reflp qthms), rtac ctxt (rel_flip_of_bnf bnf_F RS @{thm iffD1}), rtac ctxt (nth rel_F_map_iffD2s 0), rtac ctxt (nth rel_F_map_iffD2s 1), etac ctxt (#rel_monoD_rotated thms)], (REPEAT_DETERM_N live o HEADGOAL o EVERY') [rtac ctxt @{thm predicate2I}, rtac ctxt @{thm conversepI}, rtac ctxt @{thm rel_sum.intros(2)}, etac ctxt @{thm conversepI}], (REPEAT_DETERM_N (2*live) o HEADGOAL) (rtac ctxt @{thm rel_sum_eq2_nonempty} ORELSE' rtac ctxt @{thm rel_sum_eq3_nonempty}), (REPEAT_DETERM_N live o EVERY) [unfold_thms_tac ctxt @{thms sum.rel_compp[symmetric] eq_OO}, HEADGOAL (etac ctxt @{thm sum.rel_mono[OF order_refl]})]]; (* quotient.set_transfer tactics *) fun set_transfer_q set_G_def set_F'_thms _ ctxt = let val set_F'_rsp = mk_rel_funDN 1 (#set_F'_respect set_F'_thms) in common_tac ctxt (set_G_def :: @{thms o_def}) THEN (HEADGOAL o EVERY') [etac ctxt @{thm relcomppE}, hyp_subst_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt [set_F'_rsp OF [#rep_abs qthms] OF [#reflp qthms], @{thm rel_set_def}]), dtac ctxt (#rel_F_rel_F' thms), rtac ctxt conjI] THEN (REPEAT_DETERM_N 2 o HEADGOAL o EVERY') [SELECT_GOAL (unfold_thms_tac ctxt [#rel_F'_set thms]), REPEAT_DETERM o eresolve_tac ctxt [exE, conjE], REPEAT_DETERM o dtac ctxt (mk_sym set_F'_rsp), SELECT_GOAL (unfold_thms_tac ctxt [#set_map_F' set_F'_thms]), rtac ctxt ballI, dtac ctxt @{thm equalityD1[THEN subsetD]}, assume_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt @{thms image_iff}), etac ctxt bexE, dtac ctxt set_mp, assume_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt @{thms mem_Collect_eq case_prod_beta}), rtac ctxt bexI, hyp_subst_tac ctxt, assume_tac ctxt, etac ctxt @{thm hypsubst}, etac ctxt @{thm imageI}, assume_tac ctxt] end; in {map={raw=map_raw, tac=map_transfer_q}, rel={raw=rel_raw, tac=rel_transfer_q}, sets={raws=sets_raw,tacs=map2 set_transfer_q (#sets old_defs) set_F'_thmss}, pred=NONE} end; fun quotient_bnf (equiv_thm, quot_thm) wits specs map_b rel_b pred_b opts lthy = let (* extract rep_G and abs_G *) val (equiv_rel, abs_G, rep_G) = strip3 quot_thm; val (repT, absT) = dest_funT (fastype_of abs_G); (* ("?'a F", "?'a G") *) val absT_name = fst (dest_Type absT); val tvs = map (fst o dest_TVar) (snd (dest_Type absT)); val _ = length tvs = length specs orelse error ("Expected " ^ string_of_int (length tvs) ^ " type argument" ^ (if (length tvs) = 1 then "" else "s") ^ " to " ^ quote absT_name); (* instantiate TVars *) val alpha0s = map (TFree o snd) specs; val typ_subst = typ_subst_TVars (tvs ~~ alpha0s); val (repT, absT) = apply2 typ_subst (repT, absT); (* get the bnf for RepT *) val Ds0 = filter (is_none o fst) specs |> map snd; fun flatten_tyargs Ass = map dest_TFree alpha0s |> filter (fn T => exists (fn Ts => member op= Ts T) Ass); val ((bnf_F, (deads, alphas)), ((_, unfolds), lthy)) = bnf_of_typ true Dont_Inline (Binding.qualify true absT_name) flatten_tyargs [] Ds0 repT ((empty_comp_cache, empty_unfolds), lthy); val live = length alphas; val _ = (if live = 0 then error "No live variables" else ()); val defs = #map_unfolds unfolds @ flat (#set_unfoldss unfolds) @ #rel_unfolds unfolds; val set_bs = map (fn T => find_index (fn U => T = U) alpha0s) alphas |> map (the_default Binding.empty o fst o nth specs); (* create and instantiate all the needed type variables *) val subst = subst_TVars (tvs ~~ alpha0s); val (abs_G, rep_G) = apply2 subst (abs_G, rep_G); val sorts = map Type.sort_of_atyp alphas; val (((betas, gammas), deltas), names_lthy) = fold Variable.declare_typ (alphas @ deads) lthy |> mk_TFrees' sorts ||>> mk_TFrees' sorts ||>> mk_TFrees' sorts; fun subst_Ts tm Ts = subst_atomic_types (alphas ~~ Ts) tm; val subst_b = subst_atomic_types (alphas ~~ betas); val subst_Maybe = subst_atomic_types o map (swap o `mk_MaybeT); val equiv_rel_a = subst equiv_rel; val map_F = mk_map_of_bnf deads alphas betas bnf_F; val rel_F_ab = mk_rel_of_bnf deads alphas betas bnf_F; val rel_F_bc = mk_rel_of_bnf deads betas gammas bnf_F; val rel_F_ac = mk_rel_of_bnf deads alphas gammas bnf_F; val rel_F_option = mk_rel_of_bnf deads (map mk_MaybeT alphas) (map mk_MaybeT betas) bnf_F; val sets_F = mk_sets_of_bnf (replicate live deads) (replicate live alphas) bnf_F; val wits_F = mk_wits_of_bnf (replicate (nwits_of_bnf bnf_F) deads) (replicate (nwits_of_bnf bnf_F) alphas) bnf_F; val (typ_fs, (typ_aF, typ_bF)) = strip_typeN live (fastype_of map_F) ||> dest_funT; val typ_MaybeF = typ_subst_atomic (alphas ~~ map mk_MaybeT alphas) typ_aF; val typ_a_sets = map HOLogic.mk_setT alphas; val typ_pairs = map HOLogic.mk_prodT (alphas ~~ betas); val typ_fs' = map (typ_subst_atomic (map (swap o `mk_MaybeT) betas)) typ_fs; (* create all the needed variables *) val ((((((((((((((((((((((var_Ps, var_Qs), var_Rs), var_x), var_x'), var_y), var_y'), var_mx), var_As), var_As'), var_Ss), var_Bs), var_as), var_as'), var_bs), var_bs'), var_R), var_fs), var_fs'), var_gs), var_gs'), var_z), var_ts) = names_lthy |> mk_Frees "Ps" (map2 mk_relT alphas betas) ||>> mk_Frees "Qs" (map2 mk_relT betas gammas) ||>> mk_Frees "Rs" (map2 mk_relT alphas gammas) ||>> mk_Free "x" typ_aF ||>> mk_Free "x'" typ_aF ||>> mk_Free "y" typ_bF ||>> mk_Free "y'" (typ_subst_atomic (alphas ~~ gammas) typ_aF) ||>> mk_Free "mx" typ_MaybeF ||>> mk_Frees "As" typ_a_sets ||>> mk_Frees "As'" typ_a_sets ||>> mk_Frees "Ss" (map HOLogic.mk_setT typ_a_sets) ||>> mk_Frees "Bs" (map HOLogic.mk_setT betas) ||>> mk_Frees "as" alphas ||>> mk_Frees "as'" alphas ||>> mk_Frees "bs" betas ||>> mk_Frees "bs'" betas ||>> mk_Free "R" (typ_aF --> typ_bF --> HOLogic.boolT) ||>> mk_Frees "fs" typ_fs ||>> mk_Frees "fs'" typ_fs' ||>> mk_Frees "gs" typ_fs ||>> mk_Frees "gs'" typ_fs' ||>> mk_Free "z" (typ_subst_atomic (alphas ~~ typ_pairs) typ_aF) ||>> mk_Frees "ts" typ_pairs |> fst; (* create local definitions `b = tm` with n arguments *) fun suffix tm s = (dest_Const tm |> fst |> Long_Name.base_name) ^ s; fun define lthy b n tm = let val b = Binding.qualify true absT_name (Binding.qualified_name b); val ((tm, (_, def)), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> Local_Theory.define_internal (((Binding.concealed b, NoSyn), (Binding.empty_atts, tm))) ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val tm = Term.subst_atomic_types (map (`(Morphism.typ phi)) (alphas @ betas @ gammas @ map TFree Ds0)) (Morphism.term phi tm); val def = mk_unabs_def n (HOLogic.mk_obj_eq (Morphism.thm phi def)); in ({def=def, tm=tm}, lthy) end; (* internally use REL, not the user-provided definition *) val (REL, lthy) = define lthy "REL" 0 equiv_rel_a; val REL_def = sym RS eq_reflection OF [#def REL]; fun REL_rewr_all ctxt thm = Conv.fconv_rule (Conv.top_conv (fn _ => Conv.try_conv (Conv.rewr_conv REL_def)) ctxt) thm; val equiv_rel_a' = equiv_rel_a; val equiv_rel_a = #tm REL; val (equiv_rel_b, equiv_rel_c) = apply2 (subst_Ts equiv_rel_a) (betas, gammas); (* rel_pos_distr: @{term "\A B. A \\ B \ bot \ REL \\ rel_F A \\ REL \\ rel_F B \\ REL \ REL \\ rel_F (A \\ B) \\ REL"} *) fun compp_not_bot comp aT cT = let val T = mk_relT aT cT; val mk_eq = HOLogic.eq_const T; in HOLogic.mk_not (mk_eq $ comp $ bot_const T) end; val ab_comps = map2 mk_relcompp var_Ps var_Qs; val ne_comps = (@{map 3} compp_not_bot ab_comps alphas gammas); val ab_prem = foldr1 HOLogic.mk_conj ne_comps; val REL_pos_distrI_tm = let val le_relcomps = map2 mk_leq ab_comps var_Rs; val assm = mk_OO [equiv_rel_a, list_comb (rel_F_ab, var_Ps), equiv_rel_b, list_comb (rel_F_bc, var_Qs)] equiv_rel_c; val concl = mk_OO [equiv_rel_a, list_comb (rel_F_ac, var_Rs)] equiv_rel_c; in mk_Trueprop_implies ([assm $ var_x $ var_y'] @ ne_comps @ le_relcomps, concl $ var_x $ var_y') end; val ab_concl = mk_leq (mk_OO [list_comb (rel_F_ab, var_Ps), equiv_rel_b] (list_comb (rel_F_bc, var_Qs))) (mk_OO [equiv_rel_a, list_comb (rel_F_ac, ab_comps)] (equiv_rel_c)); val ab_imp = Logic.mk_implies (apply2 HOLogic.mk_Trueprop (ab_prem, ab_concl)); val rel_pos_distr = fold_rev Logic.all (var_Ps @ var_Qs) ab_imp; (* {(x, y) . REL x y} *) fun mk_rel_pairs rel = mk_case_prod (var_x, var_x') (rel $ var_x $ var_x') val rel_pairs = mk_rel_pairs equiv_rel_a; (* rel_Inter: \S. \ S \ {}; \S \ {} \ \ (\A\S. {(x, y). REL x y} `` {x. set_F x \ A}) \ {(x, y). REL x y} `` {x. set_F x \ \S} *) fun rel_Inter_from_set_F (var_A, var_S) set_F = let val typ_aset = fastype_of var_A; (* \S *) val inter_S = Inf_const typ_aset $ var_S; (* [S \ {}, \S \ {}] *) fun not_empty x = let val ty = fastype_of x in HOLogic.mk_not (HOLogic.mk_eq (x, bot_const ty)) end; val prems = map (HOLogic.mk_Trueprop o not_empty) [var_S, inter_S]; (* {x. set_F x \ A} *) val setF_sub_A = mk_in [var_A] [set_F] typ_aF; (* {x. set_F x \ \S} *) val setF_sub_S = mk_in [inter_S] [set_F] typ_aF; val lhs = Inf_const (HOLogic.mk_setT typ_aF) $ (mk_image (absfree (dest_Free var_A) (Image_const typ_aF $ rel_pairs $ setF_sub_A)) $ var_S); val rhs = Image_const typ_aF $ rel_pairs $ setF_sub_S; val concl = HOLogic.mk_Trueprop (mk_leq lhs rhs); in Logic.all var_S (Logic.list_implies (prems, concl)) end; val rel_Inters = map2 rel_Inter_from_set_F (var_As ~~ var_Ss) sets_F; (* map_F_Just = map_F Just *) val map_F_Just = let val option_tys = map mk_MaybeT alphas; val somes = map Just_const alphas; in list_comb (subst_atomic_types (betas ~~ option_tys) map_F, somes) end; fun mk_set_F'_tm typ_a set_F = let val typ_aset = HOLogic.mk_setT typ_a; (* set_F' x = (\y\{y. REL (map_F Just x) y}. UNION (set_F y) set_Maybe) *) val sbind = mk_UNION (subst_Maybe alphas set_F $ var_mx) (set_Maybe_const typ_a); val collection = HOLogic.Collect_const typ_MaybeF $ absfree (dest_Free var_mx) (subst_Maybe alphas equiv_rel_a $ (map_F_Just $ var_x) $ var_mx); val set_F'_tm = lambda var_x (Inf_const typ_aset $ (mk_image (absfree (dest_Free var_mx) sbind) $ collection)); in set_F'_tm end val sets = mk_sets_of_bnf (replicate live deads) (replicate live alphas) bnf_F; val sets' = map2 mk_set_F'_tm alphas sets; val (Iwits, wit_goals) = prepare_wits true repT wits opts alphas wits_F var_as var_as' sets' lthy; val goals = rel_pos_distr :: rel_Inters @ (case wits of NONE => [] | _ => wit_goals); val plugins = get_first (fn Plugins_Option f => SOME (f lthy) | _ => NONE) (rev opts) |> the_default Plugin_Name.default_filter; fun after_qed thmss lthy = (case thmss of [rel_pos_distr_thm0] :: thmss => let val equiv_thm' = REL_rewr_all lthy equiv_thm; val rel_pos_distr_thm = @{thm equivp_add_relconj} OF [equiv_thm', equiv_thm', rel_pos_distr_thm0]; val (rel_Inter_thms, wit_thmss) = apply2 (fn f => flat (f live thmss)) (take, drop); (* construct map_G, sets_G, bd_G, rel_G and wits_G *) (* map_G f = abs_G o map_F f o rep_G *) val map_G = fold_rev lambda var_fs (HOLogic.mk_comp (HOLogic.mk_comp (subst_Ts abs_G betas, list_comb (map_F, var_fs)), rep_G)); val map_raw = fold_rev lambda var_fs (list_comb (map_F, var_fs)) |> subst_atomic_types (betas ~~ gammas); (* Define set_G and the three auxiliary definitions (set_F', F_in, F_in') *) fun mk_set_G var_A set_F lthy = let val typ_a = HOLogic.dest_setT (fastype_of var_A); val set_F'_tm = mk_set_F'_tm typ_a set_F val (set_F', lthy) = define lthy (suffix set_F "'") 1 set_F'_tm; (* set_G = set_F' o rep_G *) val set_G = HOLogic.mk_comp (#tm set_F', rep_G); (* F_in A = {x. set_F x \ A} *) val F_in_tm = lambda var_A (mk_in [var_A] [set_F] typ_aF); val (F_in, lthy) = define lthy (suffix set_F "_in") 1 F_in_tm; (* F_in' A = map_F Inr -` ({(x, y). REL x y} `` F_in (insert (Inl ()) (Inr ` A))) *) val F_in' = lambda var_A (mk_vimage map_F_Just (Image_const typ_MaybeF $ subst_Maybe alphas rel_pairs $ (subst_Maybe alphas (#tm F_in) $ mk_insert (mk_Nothing typ_a) (mk_image (Just_const typ_a) $ var_A)))); val (F_in', lthy) = define lthy (suffix set_F "_in'") 1 F_in'; in ((set_G, {set_F'=set_F', F_in=F_in, F_in'=F_in'}), lthy) end; val ((sets_G, set_F'_aux_defs), lthy) = @{fold_map 2} mk_set_G var_As sets_F lthy |>> split_list; (* bd_G = bd_F *) val bd_G = mk_bd_of_bnf deads alphas bnf_F; (* rel_F' A = BNF_Def.vimage2p (map_F Just) (map_F Just) ((\) OO rel_F (rel_Maybe A) OO (\)) *) val rel_Maybes = @{map 3} (fn v => fn aT => fn bT => rel_Maybe_const aT bT $ v); val rel_F'_tm = let val equiv_equiv_rel_option = subst_Ts equiv_rel_a' o map mk_MaybeT in mk_vimage2p map_F_Just (subst_atomic_types (alphas ~~ betas) map_F_Just) $ mk_OO [equiv_equiv_rel_option alphas, list_comb (rel_F_option, rel_Maybes var_Ps alphas betas)] (equiv_equiv_rel_option betas) end; val (rel_F', lthy) = define lthy (suffix rel_F_ab "'") (live+2) (fold_rev lambda var_Ps rel_F'_tm); (* rel_G A = vimage2p rep_G rep_G (rel_F' A) *) val rel_G = fold_rev lambda var_Ps (mk_vimage2p rep_G (subst_Ts rep_G betas) $ rel_F'_tm); val rel_raw = fold_rev lambda var_Ps rel_F'_tm |> subst_atomic_types (betas ~~ gammas); (* auxiliary lemmas *) val bd_card_order = bd_card_order_of_bnf bnf_F; val bd_cinfinite = bd_cinfinite_of_bnf bnf_F; val in_rel = in_rel_of_bnf bnf_F; val map_F_comp = map_comp_of_bnf bnf_F; val map_F_comp0 = map_comp0_of_bnf bnf_F; val map_F_cong = map_cong_of_bnf bnf_F; val map_F_id0 = map_id0_of_bnf bnf_F; val map_F_id = map_id_of_bnf bnf_F; val rel_conversep = rel_conversep_of_bnf bnf_F; val rel_flip = rel_flip_of_bnf bnf_F; val rel_eq_onp = rel_eq_onp_of_bnf bnf_F; val rel_Grp = rel_Grp_of_bnf bnf_F; val rel_OO = rel_OO_of_bnf bnf_F; val rel_map = rel_map_of_bnf bnf_F; val rel_refl_strong = rel_refl_strong_of_bnf bnf_F; val set_bd_thms = set_bd_of_bnf bnf_F; val set_map_thms = set_map_of_bnf bnf_F; (* map_F_respect: @{term "((=) ===> REL ===> REL) map_F map_F"} *) val map_F_respect = HOLogic.mk_Trueprop (fold_rev mk_rel_fun (map2 (fn xT => fn yT => HOLogic.eq_const (xT --> yT)) alphas betas @ [equiv_rel_a]) (equiv_rel_b) $ map_F $ map_F); fun map_F_respect_tac ctxt = HEADGOAL (EVERY' [REPEAT_DETERM_N (live + 1) o rtac ctxt @{thm rel_funI}, hyp_subst_tac ctxt, rtac ctxt (BNF_FP_Util.split_conj_prems live rel_pos_distr_thm0 OF replicate live @{thm Grp_conversep_nonempty} RS rev_mp), rtac ctxt impI, dtac ctxt @{thm predicate2D}, etac ctxt @{thm relcomppI2[rotated]}, rtac ctxt (rel_Grp RS @{thm predicate2_eqD} RS iffD2), rtac ctxt @{thm GrpI[OF refl]}, REPEAT_DETERM o resolve_tac ctxt [CollectI, conjI, subset_UNIV], rtac ctxt (rel_flip RS iffD2), rtac ctxt (rel_Grp RS @{thm predicate2_eqD} RS iffD2), rtac ctxt @{thm GrpI[OF refl]}, REPEAT_DETERM o resolve_tac ctxt [CollectI, conjI, subset_UNIV], SELECT_GOAL (unfold_thms_tac ctxt (rel_eq_onp :: @{thms Grp_conversep_eq_onp})), etac ctxt @{thm predicate2D[OF rel_conj_eq_onp, rotated]}, rtac ctxt equiv_thm']); val map_F_respect_thm = prove lthy [] map_F_respect map_F_respect_tac; val rel_funD = mk_rel_funDN (live+1); val map_F_rsp = (rel_funD map_F_respect_thm) OF (replicate live refl); fun map_F_rsp_of tms ctxt = (infer_instantiate' ctxt (NONE :: NONE :: map (SOME o Thm.cterm_of ctxt) tms) map_F_rsp) val qthms = let fun equivp_THEN thm = REL_rewr_all lthy equiv_thm RS thm; fun Quotient3_THEN thm = REL_rewr_all lthy quot_thm RS thm; in {abs_rep = Quotient3_THEN @{thm Quotient3_abs_rep}, rel_abs = Quotient3_THEN @{thm Quotient3_rel_abs}, rep_abs = Quotient3_THEN @{thm Quotient3_rep_abs}, rep_reflp = Quotient3_THEN @{thm Quotient3_rep_reflp}, rep_abs_rsp = Quotient3_THEN @{thm rep_abs_rsp}, reflp = equivp_THEN @{thm equivp_reflp}, symp = equivp_THEN @{thm equivp_symp}, transp = equivp_THEN @{thm equivp_transp}, REL = REL_def} end; (* lemma REL_OO_REL_left: REL OO REL OO R = REL OO R *) val REL_OO_REL_left_thm = let val tm = mk_Trueprop_eq (mk_OO [equiv_rel_a, equiv_rel_a] var_R, mk_relcompp equiv_rel_a var_R) fun tac ctxt = HEADGOAL (EVERY' [rtac ctxt ext, rtac ctxt ext, rtac ctxt iffI, TWICE (etac ctxt @{thm relcomppE}), rtac ctxt @{thm relcomppI}, etac ctxt (#transp qthms), TWICE (assume_tac ctxt), etac ctxt @{thm relcomppE}, etac ctxt @{thm relcomppI}, rtac ctxt @{thm relcomppI}, rtac ctxt (#reflp qthms), assume_tac ctxt]); in prove lthy [var_R] tm tac end; (* Generate theorems related to the setters *) val map_F_fs = list_comb (map_F, var_fs); (* aset aset asetset bset typ_b typ_b *) fun mk_set_F'_thmss (((((var_A, var_A'), var_S), var_B), var_b), var_b') set_F {set_F', F_in, F_in'} rel_Inter_thm set_map_F_thm (idx, vf) = let val (var_f, var_fs') = case vf of (f :: fs) => (f, fs) | _ => error "won't happen"; val typ_a = fastype_of var_f |> dest_funT |> fst; val typ_b = fastype_of var_b; val (typ_asetset, typ_aset) = `HOLogic.mk_setT (fastype_of var_A); val map_F_fs_x = map_F_fs $ var_x; (* F_in'_mono: A \ B \ F_in' A \ F_in' B *) val F_in'_mono_tm = mk_Trueprop_implies ([mk_leq var_A var_A'], mk_leq (#tm F_in' $ var_A) (#tm F_in' $ var_A')); fun F_in'_mono_tac ctxt = unfold_thms_tac ctxt [#def F_in', #def F_in] THEN HEADGOAL (EVERY' [rtac ctxt subsetI, etac ctxt vimageE, etac ctxt @{thm ImageE}, etac ctxt CollectE, etac ctxt CollectE, dtac ctxt @{thm case_prodD}, hyp_subst_tac ctxt, rtac ctxt (vimageI OF [refl]), rtac ctxt @{thm ImageI}, rtac ctxt CollectI, rtac ctxt @{thm case_prodI}, assume_tac ctxt ORELSE' rtac ctxt refl, rtac ctxt CollectI, etac ctxt subset_trans, etac ctxt (@{thm insert_mono} OF [@{thm image_mono}])]); val F_in'_mono_thm = prove lthy [var_A, var_A'] F_in'_mono_tm F_in'_mono_tac; (* F_in'_Inter: F_in' (\S) = (\A\S. F_in' A) *) val F_in'_Inter_tm = mk_Trueprop_eq ((#tm F_in' $ (Inf_const typ_aset $ var_S)), (Inf_const (HOLogic.mk_setT typ_aF) $ (mk_image (#tm F_in') $ var_S))); fun F_in'_Inter_tac ctxt = Local_Defs.unfold_tac ctxt [#def F_in', #def F_in] THEN HEADGOAL (rtac ctxt (infer_instantiate' ctxt [SOME (Thm.cterm_of ctxt (HOLogic.mk_eq (var_S, bot_const typ_asetset)))] @{thm case_split}) THEN' EVERY' [ hyp_subst_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt @{thms Inter_empty INT_empty UNIV_sum_unit_conv}), rtac ctxt @{thm set_eqI}, rtac ctxt iffI, rtac ctxt UNIV_I, rtac ctxt (vimageI OF [refl]), rtac ctxt @{thm ImageI}, rtac ctxt CollectI, rtac ctxt @{thm case_prodI}, rtac ctxt (#reflp qthms), rtac ctxt CollectI, rtac ctxt subset_UNIV, etac ctxt @{thm exE[OF ex_in_conv[THEN iffD2]]}, EqSubst.eqsubst_tac ctxt [0] @{thms image_INT[of _ UNIV _ id, simplified]}, rtac ctxt @{thm inj_Inr}, assume_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt @{thms INT_extend_simps vimage_INT[symmetric]}), rtac ctxt @{thm arg_cong2[where f=vimage, OF refl]}, rtac ctxt equalityI, rtac ctxt subsetI, rtac ctxt @{thm InterI}, etac ctxt imageE, etac ctxt @{thm ImageE}, etac ctxt CollectE, etac ctxt CollectE, dtac ctxt @{thm case_prodD}, hyp_subst_tac ctxt, rtac ctxt @{thm ImageI[OF CollectI]}, etac ctxt @{thm case_prodI} ORELSE' (SELECT_GOAL (unfold_thms_tac ctxt @{thms prod.case}) THEN' rtac ctxt @{thm refl}), rtac ctxt CollectI, etac ctxt subset_trans, etac ctxt @{thm INT_lower[OF imageI]}, rtac ctxt (@{thm subset_trans} OF [asm_rl, rel_Inter_thm]), K (unfold_thms_tac ctxt @{thms image_image}), rtac ctxt subset_refl, K (unfold_thms_tac ctxt @{thms INT_extend_simps ex_in_conv[symmetric]}), rtac ctxt exI, rtac ctxt imageI, assume_tac ctxt, rtac ctxt exI, rtac ctxt @{thm InterI}, etac ctxt imageE, hyp_subst_tac ctxt, rtac ctxt @{thm insertI1}]); val F_in'_Inter_thm = prove lthy [var_S] F_in'_Inter_tm F_in'_Inter_tac; (* set_F'_respect: (REL ===> (=)) set_F' set_F' *) val set_F'_respect_tm = HOLogic.mk_Trueprop (mk_rel_fun equiv_rel_a (HOLogic.eq_const typ_aset) $ #tm set_F' $ #tm set_F'); fun set_F'_respect_tac ctxt = unfold_thms_tac ctxt (#def set_F' :: @{thms rel_fun_def}) THEN HEADGOAL (EVERY' [TWICE (rtac ctxt allI), rtac ctxt impI, dtac ctxt (map_F_rsp_of (map Just_const alphas) ctxt), rtac ctxt @{thm INF_cong}, rtac ctxt @{thm Collect_eqI}, rtac ctxt iffI, etac ctxt (#transp qthms OF [#symp qthms]), assume_tac ctxt, etac ctxt (#transp qthms), assume_tac ctxt, rtac ctxt refl]); (* F_in'_alt2: F_in' A = {x. set_F' x \ A} *) val F_in'_alt2_tm = mk_Trueprop_eq (#tm F_in' $ var_A, mk_in [var_A] [#tm set_F'] typ_aF); fun F_in'_alt2_tac ctxt = HEADGOAL (rtac ctxt equalityI THEN' (Subgoal.FOCUS o K) (unfold_thms_tac ctxt (map #def [set_F', F_in', F_in])) ctxt THEN' EVERY' [rtac ctxt subsetI, rtac ctxt CollectI, rtac ctxt subsetI, dtac ctxt vimageD, etac ctxt @{thm ImageE}, etac ctxt CollectE, etac ctxt CollectE, dtac ctxt @{thm case_prodD}, dtac ctxt @{thm InterD}, rtac ctxt @{thm imageI[OF CollectI]}, etac ctxt (#symp qthms), etac ctxt @{thm UnionE}, etac ctxt imageE, hyp_subst_tac ctxt, etac ctxt @{thm subset_lift_sum_unitD}, etac ctxt @{thm setr.cases}, hyp_subst_tac ctxt, assume_tac ctxt]) THEN unfold_thms_tac ctxt [#def set_F'] THEN (HEADGOAL o EVERY') [rtac ctxt subsetI, etac ctxt CollectE, etac ctxt (subsetD OF [F_in'_mono_thm]), EqSubst.eqsubst_tac ctxt [0] [F_in'_Inter_thm], rtac ctxt @{thm InterI}] THEN REPEAT_DETERM (HEADGOAL (etac ctxt @{thm imageE} THEN' hyp_subst_tac ctxt)) THEN (HEADGOAL o EVERY') [etac ctxt CollectE, SELECT_GOAL (unfold_thms_tac ctxt (map #def [F_in', F_in])), rtac ctxt @{thm vimageI[OF refl]}, rtac ctxt @{thm ImageI}, rtac ctxt CollectI, rtac ctxt @{thm case_prodI}, etac ctxt (#symp qthms), rtac ctxt CollectI, rtac ctxt subsetI, rtac ctxt @{thm sum_insert_Inl_unit}, assume_tac ctxt, hyp_subst_tac ctxt, rtac ctxt imageI, rtac ctxt @{thm UnionI}, rtac ctxt imageI, assume_tac ctxt, rtac ctxt @{thm setr.intros[OF refl]}]; val F_in'_alt2_thm = prove lthy [var_A] F_in'_alt2_tm F_in'_alt2_tac; (* set_F'_alt: set_F' x = \{A. x \ F_in' A} *) val set_F'_alt_tm = mk_Trueprop_eq (#tm set_F' $ var_x, Inf_const typ_aset $ mk_Collect (var_A, HOLogic.mk_mem (var_x, #tm F_in' $ var_A))); fun set_F'_alt_tac ctxt = unfold_thms_tac ctxt [F_in'_alt2_thm] THEN HEADGOAL (EVERY' [rtac ctxt @{thm set_eqI}, rtac ctxt iffI, rtac ctxt @{thm InterI}, etac ctxt CollectE, etac ctxt CollectE, dtac ctxt subsetD, assume_tac ctxt, assume_tac ctxt, etac ctxt @{thm InterD}, rtac ctxt CollectI, rtac ctxt CollectI, rtac ctxt subset_refl]); val set_F'_alt_thm = prove lthy [var_x] set_F'_alt_tm set_F'_alt_tac; (* map_F_in_F_in'I: x \ F_in' B \ map_F f x \ F_in' (f ` B) *) val map_F_in_F_in'I_tm = mk_Trueprop_implies ([HOLogic.mk_mem (var_x, #tm F_in' $ var_A')], HOLogic.mk_mem (map_F_fs_x, subst_b (#tm F_in') $ (mk_image var_f $ var_A'))); fun map_F_in_F_in'I_tac ctxt = Local_Defs.unfold_tac ctxt ([#def F_in', #def F_in, Bex_def] @ @{thms vimage_def Image_iff}) THEN HEADGOAL (EVERY' [etac ctxt @{thm CollectE}, etac ctxt exE, etac ctxt conjE, etac ctxt @{thm CollectE}, etac ctxt @{thm CollectE}, dtac ctxt @{thm case_prodD}, rtac ctxt @{thm CollectI}, rtac ctxt exI, rtac ctxt @{thm conjI[rotated]}, rtac ctxt @{thm CollectI}, rtac ctxt @{thm case_prodI}, dtac ctxt (map_F_rsp_of (map mk_Maybe_map var_fs) ctxt), SELECT_GOAL (unfold_thms_tac ctxt (map_F_comp :: @{thms o_def map_sum.simps})), assume_tac ctxt, rtac ctxt CollectI, SELECT_GOAL (unfold_thms_tac ctxt set_map_thms), etac ctxt @{thm image_map_sum_unit_subset}]); val map_F_in_F_in'I_thm = prove lthy (var_A' :: var_x :: var_fs) map_F_in_F_in'I_tm map_F_in_F_in'I_tac; (* REL_preimage_eq: C \ range f \ {} \ {(a, b). REL a b} `` {x. set_F x \ f -` C} = map_F f -` {(a, b). REL a b} `` {x. set_F x \ C} *) val REL_preimage_eq_tm = mk_Trueprop_implies ([HOLogic.mk_not (HOLogic.mk_eq (HOLogic.mk_binop @{const_name inf} (var_B, mk_image var_f $ HOLogic.mk_UNIV typ_a), bot_const (HOLogic.mk_setT typ_b)))], HOLogic.mk_eq (Image_const typ_aF $ rel_pairs $ mk_in [mk_vimage var_f var_B] [set_F] typ_aF, mk_vimage map_F_fs (Image_const typ_bF $ subst_b rel_pairs $ mk_in [var_B] [subst_b set_F] typ_bF))); (* Bs \ range fs \ {} \ set_F xb \ Bs \ REL xb (map_F fs x) \ x \ {(x, x'). REL x x'} `` {x. set_F x \ fs -` Bs} *) fun subgoal_tac {context = ctxt, params, ...} = let val (x, y) = case params of [(_, x), _, (_, y)] => (x, y) | _ => error "won't happen"; val cond = HOLogic.mk_conj (apply2 HOLogic.mk_mem ((var_b, var_B), (var_b', var_B))); (* ["\x y. x \ B \ y \ B", "(Grp UNIV f_1)\\"] *) val cvars = var_fs |> maps (fn f => let val fT = fastype_of f in map (SOME o Thm.cterm_of ctxt) [if f = var_f then fold_rev lambda [var_b, var_b'] cond else HOLogic.eq_const (range_type fT), mk_conversep (mk_Grp (HOLogic.mk_UNIV (domain_type fT)) f)] end); val rel_pos_distr_thm_inst = infer_instantiate' ctxt (cvars @ [SOME y,SOME x]) (@{thm predicate2D} OF [rel_pos_distr_thm]); (* GrpI[of "map_F f1 .. fN" x, OF refl CollectI, OF "B1 \ UNIV \ ... \ Bn \ UNIV"] *) fun subset_UNIVs n = fold (fn a => fn b => conjI OF [a, b]) (replicate (n-1) @{thm subset_UNIV}) @{thm subset_UNIV}; val GrpI_inst = infer_instantiate' ctxt (map SOME [Thm.cterm_of ctxt map_F_fs, x]) @{thm GrpI} OF [refl, CollectI] OF [subset_UNIVs live]; in EVERY [ HEADGOAL (Method.insert_tac ctxt [rel_pos_distr_thm_inst]), unfold_thms_tac ctxt [rel_conversep, rel_OO, rel_Grp], HEADGOAL (etac ctxt @{thm meta_impE}), REPEAT_DETERM_N (live-1) (HEADGOAL (rtac ctxt @{thm conjI[rotated]})), REPEAT_DETERM_N live (HEADGOAL (etac ctxt @{thm relcompp_mem_Grp_neq_bot} ORELSE' rtac ctxt @{thm relcompp_eq_Grp_neq_bot})), HEADGOAL (EVERY' [etac ctxt @{thm meta_impE}, rtac ctxt @{thm relcomppI}, rtac ctxt (#reflp qthms), rtac ctxt @{thm relcomppI}, rtac ctxt rel_refl_strong]), REPEAT_DETERM_N idx (HEADGOAL (rtac ctxt refl)), HEADGOAL (rtac ctxt conjI THEN' TWICE (etac ctxt subsetD THEN' assume_tac ctxt)), REPEAT_DETERM_N (live-idx-1) (HEADGOAL (rtac ctxt refl)), HEADGOAL (EVERY' [rtac ctxt @{thm relcomppI}, assume_tac ctxt, rtac ctxt @{thm relcomppI}, rtac ctxt @{thm conversepI}, rtac ctxt GrpI_inst, rtac ctxt (#reflp qthms), etac ctxt @{thm relcomppE}, etac ctxt @{thm relcomppE}, etac ctxt @{thm relcomppE}, etac ctxt @{thm conversepE}, etac ctxt @{thm GrpE}, hyp_subst_tac ctxt, rtac ctxt @{thm ImageI}, rtac ctxt CollectI, rtac ctxt @{thm case_prodI}, assume_tac ctxt, EqSubst.eqsubst_asm_tac ctxt [1] rel_map, EqSubst.eqsubst_asm_tac ctxt [1] [in_rel_of_bnf bnf_F], etac ctxt exE, etac ctxt CollectE, etac ctxt conjE, etac ctxt conjE, etac ctxt CollectE, hyp_subst_tac ctxt, rtac ctxt CollectI]), unfold_thms_tac ctxt set_map_thms, HEADGOAL (rtac ctxt (subsetI OF [vimageI] OF [refl]) THEN' etac ctxt @{thm imageE} THEN' hyp_subst_tac ctxt), REPEAT_DETERM_N 6 (HEADGOAL (etac ctxt Drule.thin_rl)), REPEAT_DETERM_N (live-1) (HEADGOAL (etac ctxt conjE)), HEADGOAL (EVERY' [dtac ctxt subsetD, assume_tac ctxt, etac ctxt CollectE]), unfold_thms_tac ctxt @{thms split_beta}, HEADGOAL (etac ctxt conjunct2)] end; fun REL_preimage_eq_tac ctxt = HEADGOAL (EVERY' [rtac ctxt @{thm set_eqI}, rtac ctxt iffI, etac ctxt @{thm ImageE}, etac ctxt CollectE, etac ctxt CollectE, dtac ctxt @{thm case_prodD}, rtac ctxt (vimageI OF [refl]), rtac ctxt @{thm ImageI}, rtac ctxt CollectI, rtac ctxt @{thm case_prodI}, etac ctxt map_F_rsp, rtac ctxt CollectI, EqSubst.eqsubst_tac ctxt [0] [set_map_F_thm], etac ctxt @{thm subset_vimage_image_subset}, etac ctxt vimageE, etac ctxt @{thm ImageE}, TWICE (etac ctxt CollectE), dtac ctxt @{thm case_prodD}, hyp_subst_tac ctxt, Subgoal.FOCUS_PARAMS subgoal_tac ctxt]); val REL_preimage_eq_thm = prove lthy (var_B :: var_fs) REL_preimage_eq_tm REL_preimage_eq_tac; (* set_map_F': set_F' (map_F f x) = f ` set_F' x *) val set_map_F'_tm = mk_Trueprop_eq (subst_b (#tm set_F') $ map_F_fs_x, mk_image var_f $ (#tm set_F' $ var_x)); fun set_map_F'_tac ctxt = HEADGOAL (EVERY' [rtac ctxt @{thm set_eqI}, rtac ctxt iffI, EqSubst.eqsubst_asm_tac ctxt [0] [set_F'_alt_thm], etac ctxt @{thm InterD}, rtac ctxt CollectI, rtac ctxt map_F_in_F_in'I_thm, SELECT_GOAL (unfold_thms_tac ctxt [F_in'_alt2_thm]), rtac ctxt CollectI, rtac ctxt subset_refl, SELECT_GOAL (unfold_thms_tac ctxt [set_F'_alt_thm]), rtac ctxt @{thm InterI}, etac ctxt imageE, etac ctxt CollectE, hyp_subst_tac ctxt, etac ctxt @{thm vimageD[OF InterD]}, rtac ctxt CollectI]) THEN (* map_F f x \ F_in' X \ x \ F_in' (f -` X) *) HEADGOAL (Subgoal.FOCUS_PARAMS (fn {context = ctxt, params, ...} => let val X = nth params 1 |> snd |> Thm.term_of; val Inr_img = mk_image (Just_const (HOLogic.dest_setT (fastype_of X))) $ X; fun cvars_of ctxt = map (SOME o Thm.cterm_of ctxt); val cut_thm = infer_instantiate' ctxt (cvars_of ctxt [Inr_img, var_f]) @{thm insert_Inl_int_map_sum_unit}; val preimage_thm = infer_instantiate' ctxt (cvars_of ctxt (filter (fn f => var_f <> f) var_fs |> map mk_Maybe_map)) (cut_thm RS REL_preimage_eq_thm); in EVERY [ unfold_thms_tac ctxt (map #def [F_in', F_in] @ preimage_thm :: map_F_comp :: @{thms lift_sum_unit_vimage_commute vimage_comp o_def map_sum.simps}), unfold_thms_tac ctxt [@{thm o_def[symmetric]}, map_F_comp0], Local_Defs.fold_tac ctxt @{thms vimage_comp}, HEADGOAL (etac ctxt (vimageI OF [refl]))] end) ctxt); (* set_F'_subset: set_F' x \ set_F x *) val set_F'_subset_tm = HOLogic.mk_Trueprop (mk_leq (#tm set_F' $ var_x) (set_F $ var_x)); fun set_F'_subset_tac ctxt = let val int_e_thm = infer_instantiate' ctxt (replicate 3 NONE @ [SOME (Thm.cterm_of ctxt (map_F_Just $ var_x))]) @{thm INT_E}; in HEADGOAL (EVERY' [SELECT_GOAL (Local_Defs.unfold_tac ctxt [#def set_F']), rtac ctxt subsetI, etac ctxt int_e_thm, SELECT_GOAL (unfold_thms_tac ctxt [set_map_F_thm]), etac ctxt @{thm UN_E}, etac ctxt imageE, hyp_subst_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt @{thms sum_set_simps singleton_iff}), hyp_subst_tac ctxt, assume_tac ctxt, etac ctxt notE, rtac ctxt CollectI, rtac ctxt (#reflp qthms)]) end; in ({F_in'_mono = F_in'_mono_thm, F_in'_Inter = F_in'_Inter_thm, set_F'_respect = prove lthy [] set_F'_respect_tm set_F'_respect_tac, F_in'_alt2 = F_in'_alt2_thm, set_F'_alt = set_F'_alt_thm, map_F_in_F_in'I = map_F_in_F_in'I_thm, set_map_F' = prove lthy (var_x :: var_fs) set_map_F'_tm set_map_F'_tac, set_F'_subset = prove lthy [var_x] set_F'_subset_tm set_F'_subset_tac, set_F'_def = #def set_F', F_in_def = #def F_in, F_in'_def = #def F_in'}, (idx + 1, var_fs')) end; val set_F'_thmss = @{fold_map 5} mk_set_F'_thmss (var_As ~~ var_As' ~~ var_Ss ~~ var_Bs ~~ var_bs ~~ var_bs') sets_F set_F'_aux_defs rel_Inter_thms set_map_thms (0, var_fs) |> fst; (* F_in'D: x \ F_in' A \ \a\A. f a = g a \ REL (map_F f x) (map_F g x) *) fun rel_map_F_fs_map_F_gs subst fs gs = subst equiv_rel_b $ (list_comb (subst map_F, fs) $ var_x) $ (list_comb (subst map_F, gs) $ var_x); val F_in'D_thm = let fun mk_prem var_a var_Aset {F_in', ...} var_f var_g = [HOLogic.mk_mem (var_x, #tm F_in' $ var_Aset), mk_Ball var_Aset ((absfree (dest_Free var_a)) (HOLogic.mk_eq (var_f $ var_a, var_g $ var_a)))]; val prems = @{map 5} mk_prem var_as var_As set_F'_aux_defs var_fs' var_gs'; val F_in'D_tm = mk_Trueprop_implies (flat prems, rel_map_F_fs_map_F_gs (subst_Maybe betas) var_fs' var_gs'); fun map_F_rsp_of_case_sums fs ctxt = map_F_rsp_of (@{map 2} (fn f => fn T => BNF_FP_Util.mk_case_sum (Term.absdummy HOLogic.unitT (mk_Nothing T), f)) fs betas) ctxt; fun mk_var_fgs n = take n var_gs' @ drop n var_fs'; fun F_in'D_tac ctxt = EVERY (unfold_thms_tac ctxt (maps (fn {F_in'_def, F_in_def, ...} => [F_in'_def, F_in_def]) set_F'_thmss) :: map (REPEAT_DETERM_N live o HEADGOAL) [etac ctxt vimageE, etac ctxt @{thm ImageE}, etac ctxt CollectE THEN' etac ctxt CollectE, dtac ctxt @{thm case_prodD}] @ HEADGOAL (hyp_subst_tac ctxt THEN' rotate_tac (~live)) :: map (fn i => (HEADGOAL o EVERY') [if i < live then rtac ctxt (#transp qthms) else K all_tac, Ctr_Sugar_Tactics.select_prem_tac ctxt live (dresolve_tac ctxt [asm_rl]) i, rtac ctxt (#transp qthms), dtac ctxt (map_F_rsp_of_case_sums (mk_var_fgs (i-1)) ctxt), SELECT_GOAL (unfold_thms_tac ctxt (map_F_comp :: @{thms case_sum_o_inj(2)})), etac ctxt (#symp qthms), dtac ctxt (map_F_rsp_of_case_sums (mk_var_fgs i) ctxt), SELECT_GOAL (unfold_thms_tac ctxt (map_F_comp :: @{thms case_sum_o_inj(2)})), EqSubst.eqsubst_tac ctxt [0] [map_F_cong OF replicate i refl @ asm_rl :: replicate (live-i) refl], rtac ctxt @{thm sum.case_cong[OF refl refl]}, etac ctxt bspec, hyp_subst_tac ctxt, etac ctxt @{thm subset_lift_sum_unitD}, assume_tac ctxt, assume_tac ctxt]) (1 upto live)); in prove lthy (var_x :: var_As @ var_fs' @ var_gs') F_in'D_tm F_in'D_tac end; (* map_F_cong': (\a. a \ set_F' x \ f a = g a) \ REL (map_F f x) (map_F g x) *) val map_F_cong'_thm = let fun mk_prem {set_F', ...} var_a var_f var_g = Logic.all var_a (mk_Trueprop_implies ([HOLogic.mk_mem (var_a, #tm set_F' $ var_x)], HOLogic.mk_eq (var_f $ var_a, var_g $ var_a))); val map_F_cong'_tm = Logic.list_implies (@{map 4} mk_prem set_F'_aux_defs var_as var_fs var_gs, HOLogic.mk_Trueprop (rel_map_F_fs_map_F_gs I var_fs var_gs)); fun Just_o_fun bT f = HOLogic.mk_comp (Just_const bT, f); fun map_F_Just_o_funs fs = list_comb (subst_Maybe betas map_F, map2 Just_o_fun betas fs) $ var_x; fun map_F_cong'_tac ctxt = let val map_F_respect_inst = map_F_rsp |> infer_instantiate' ctxt (map (SOME o Thm.cterm_of ctxt) (map map_F_Just_o_funs [var_fs, var_gs] @ map fromJust_const betas)) |> Local_Defs.unfold ctxt (map_F_comp :: @{thms o_assoc Fun.o_apply[where f=projr and g=Inr, unfolded sum.sel] id_def[symmetric]}) |> Local_Defs.unfold ctxt @{thms id_comp}; in HEADGOAL (rtac ctxt map_F_respect_inst THEN' rtac ctxt F_in'D_thm) THEN EVERY (map (fn {F_in'_alt2, ...} => unfold_thms_tac ctxt [F_in'_alt2] THEN HEADGOAL (EVERY' [rtac ctxt CollectI, rtac ctxt subset_refl, rtac ctxt ballI, SELECT_GOAL (unfold_thms_tac ctxt [o_apply]), rtac ctxt @{thm arg_cong[where f=Inr]}, asm_full_simp_tac ctxt])) set_F'_thmss) end; in prove lthy (var_x :: var_fs @ var_gs) map_F_cong'_tm map_F_cong'_tac end; (* rel_F'_set: "rel_F' P x y \ (\z. set_F' z \ {(x, y). P x y} \ REL (map_F fst z) x \ REL (map_F snd z) y)" *) val rel_F'_set_thm = let val lhs = list_comb (#tm rel_F', var_Ps) $ var_x $ var_y; fun mk_subset_A var_a var_b var_P {set_F', ...} = let val collect_A = mk_case_prod (var_a, var_b) (var_P $ var_a $ var_b); in mk_leq (subst_atomic_types (alphas ~~ typ_pairs) (#tm set_F') $ var_z) collect_A end; val subset_As = @{map 4} mk_subset_A var_as var_bs var_Ps set_F'_aux_defs; fun mk_map mfs f z = Term.list_comb (mfs, map (fst o Term.strip_comb o f) var_ts) $ z; val map_F_fst = mk_map_of_bnf deads typ_pairs alphas bnf_F; val map_F_snd = mk_map_of_bnf deads typ_pairs betas bnf_F; val map_fst = equiv_rel_a $ (mk_map map_F_fst HOLogic.mk_fst var_z) $ var_x; val map_snd = equiv_rel_b $ (mk_map map_F_snd HOLogic.mk_snd var_z) $ var_y; val rhs = let val (z, T) = dest_Free var_z in HOLogic.mk_exists (z, T, fold_rev (fn a => fn b => HOLogic.mk_conj (a, b)) (subset_As @ [map_fst]) map_snd) end; val rel_F'_set_tm = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)); val maybePairsTs = map HOLogic.mk_prodT (map mk_MaybeT alphas ~~ map mk_MaybeT betas) fun mk_map_prod_projr aT bT = let val (mabT, (maT, mbT)) = `HOLogic.mk_prodT (apply2 mk_MaybeT (aT, bT)); val map_prod_const = Const (@{const_name map_prod}, (maT --> aT) --> (mbT --> bT) --> mabT --> HOLogic.mk_prodT (aT, bT)); in map_prod_const $ fromJust_const aT $ fromJust_const bT end; fun exI_OF_tac ctxt tm = rtac ctxt (infer_instantiate' ctxt (NONE :: [SOME (Thm.cterm_of ctxt tm)]) exI); (* REL (map_F Inr x) (map_F fst z) \ REL (map_F snd z) (map_F Inr y) \ set_F z \ {(x, y). rel_sum (=) P x y} \ \z. set_F' z \ {(x, y). P x y} \ REL (map_F fst z) x \ REL (map_F snd z) y *) fun subgoal1_tac {context = ctxt, params, ...} = let val z = (case params of (_ :: _ :: (_, ct) :: _) => Thm.term_of ct | _ => error "won't happen"); val map_F_projr_z = list_comb (mk_map_of_bnf deads maybePairsTs typ_pairs bnf_F, map2 mk_map_prod_projr alphas betas) $ z; in HEADGOAL (exI_OF_tac ctxt map_F_projr_z) THEN HEADGOAL (EVERY' (maps (fn {set_F'_subset, set_F'_respect, set_map_F', ...} => [rtac ctxt conjI, dtac ctxt (set_F'_subset RS @{thm order_trans}), TWICE (dtac ctxt (set_F'_respect RS @{thm rel_funD})), SELECT_GOAL (unfold_thms_tac ctxt [set_map_F']), etac ctxt @{thm in_rel_sum_in_image_projr}, TWICE (assume_tac ctxt)]) set_F'_thmss)) THEN HEADGOAL (EVERY' (map (fn Ts => FIRST' [dtac ctxt (map_F_rsp_of (map fromJust_const Ts) ctxt), etac ctxt sym , assume_tac ctxt]) [alphas, betas])) THEN unfold_thms_tac ctxt (map_F_comp :: @{thms fst_comp_map_prod snd_comp_map_prod comp_projr_Inr} @ [map_F_id]) THEN HEADGOAL (rtac ctxt conjI) THEN HEADGOAL (etac ctxt (#symp qthms) THEN' assume_tac ctxt ORELSE' (EVERY' (maps (fn Ts => [dtac ctxt (map_F_rsp_of (map fromJust_const Ts) ctxt), SELECT_GOAL (unfold_thms_tac ctxt (map_F_comp :: @{thms fst_comp_map_prod snd_comp_map_prod comp_projr_Inr} @ [map_F_id])), assume_tac ctxt]) [alphas, betas]))) end; (* set_F' z \ {(x, y). P x y} \ REL (map_F fst z) x \ REL (map_F snd z) y \ \b. REL (map_F Inr x) b \ (\ba. rel_F (rel_sum (=) P) b ba \ REL ba (map_F Inr y)) *) fun subgoal2_tac {context = ctxt, params, ...} = let val z = (case params of ((_, ct) :: _) => Thm.term_of ct | _ => error "won't happen"); fun exI_map_Ifs_tac mk_proj Ts = exI_OF_tac ctxt (list_comb (mk_map_of_bnf deads typ_pairs (map mk_MaybeT Ts) bnf_F, @{map 3} (fn var_t => fn {set_F', ...} => fn T => lambda var_t (BNF_FP_Util.mk_If (HOLogic.mk_mem (var_t, subst_Ts (#tm set_F') typ_pairs $ z)) (mk_Just (mk_proj var_t)) (mk_Nothing T))) var_ts set_F'_aux_defs Ts) $ z) fun mk_REL_trans_map_F n = (rotate_prems n (#transp qthms) OF [rel_funD map_F_respect_thm] OF (replicate live refl @ [#symp qthms])); in HEADGOAL (EVERY' [exI_map_Ifs_tac HOLogic.mk_fst alphas, rtac ctxt conjI, etac ctxt (mk_REL_trans_map_F 0)]) THEN unfold_thms_tac ctxt [map_F_comp, @{thm o_def}] THEN HEADGOAL (rtac ctxt map_F_cong'_thm) THEN REPEAT_DETERM_N live (HEADGOAL (etac ctxt @{thm if_P[symmetric]})) THEN HEADGOAL (EVERY' [exI_map_Ifs_tac HOLogic.mk_snd betas, rtac ctxt conjI]) THEN unfold_thms_tac ctxt rel_map THEN HEADGOAL (rtac ctxt rel_refl_strong) THEN REPEAT_DETERM_N live (HEADGOAL (etac ctxt @{thm subset_rel_sumI})) THEN HEADGOAL (etac ctxt (mk_REL_trans_map_F 1 OF [#symp qthms])) THEN unfold_thms_tac ctxt [map_F_comp, @{thm o_def}] THEN HEADGOAL (rtac ctxt map_F_cong'_thm) THEN REPEAT_DETERM_N live (HEADGOAL (etac ctxt @{thm if_P})) end; fun rel_F'_set_tac ctxt = EVERY ([unfold_thms_tac ctxt (#def rel_F' :: #REL qthms :: @{thms vimage2p_def relcompp_apply}), HEADGOAL (rtac ctxt iffI), (HEADGOAL o TWICE) (etac ctxt exE THEN' etac ctxt conjE), HEADGOAL (EVERY' [dtac ctxt (in_rel RS iffD1), etac ctxt exE, TWICE (etac ctxt conjE), etac ctxt CollectE, hyp_subst_tac ctxt]), (REPEAT_DETERM_N (live-1) o HEADGOAL) (etac ctxt conjE), HEADGOAL (Subgoal.FOCUS_PARAMS subgoal1_tac ctxt THEN' etac ctxt exE), (REPEAT_DETERM_N (live+1) o HEADGOAL) (etac ctxt conjE), HEADGOAL (Subgoal.FOCUS_PARAMS subgoal2_tac ctxt)]); in prove lthy (var_x :: var_y :: var_Ps) rel_F'_set_tm rel_F'_set_tac end; (* tactics *) (* map_G_id0: abs_G \ map_F id \ rep_G = id *) fun map_G_id0_tac ctxt = HEADGOAL (EVERY' [SELECT_GOAL (unfold_thms_tac ctxt [@{thm fun_eq_iff}, o_apply, map_F_id0, id_apply, #abs_rep qthms]), rtac ctxt allI, rtac ctxt refl]); (* map_G (g \ f) = map_G g \ map_G f *) fun map_G_comp0_tac ctxt = HEADGOAL (EVERY' [rtac ctxt ext, rtac ctxt sym, SELECT_GOAL (unfold_thms_tac ctxt [o_apply, map_F_comp0]), rtac ctxt (#rel_abs qthms), rtac ctxt map_F_rsp, rtac ctxt (#rep_abs qthms), rtac ctxt (#reflp qthms)]); (* map_G_cong: (\z. z \ set_G x \ f z = g z) \ map_G f x = map_G g x *) fun map_G_cong_tac ctxt = EVERY [Local_Defs.fold_tac ctxt (map #set_F'_def set_F'_thmss), unfold_thms_tac ctxt [o_apply], HEADGOAL (rtac ctxt (#rel_abs qthms) THEN' rtac ctxt map_F_cong'_thm), REPEAT_DETERM_N live (HEADGOAL (asm_full_simp_tac ctxt))]; (* set_G_map0_G: set_G \ map_G f = f ` set_G *) fun mk_set_G_map0_G_tac thms ctxt = HEADGOAL (rtac ctxt ext) THEN EVERY [unfold_thms_tac ctxt [o_apply], Local_Defs.fold_tac ctxt [#set_F'_def thms]] THEN HEADGOAL (EVERY' (map (rtac ctxt) [trans OF [#set_map_F' thms RS sym, sym] RS sym, @{thm rel_funD} OF [#set_F'_respect thms], #rep_abs qthms, map_F_rsp, #rep_reflp qthms])); (* bd_card_order: card_order bd_F *) fun bd_card_order_tac ctxt = HEADGOAL (rtac ctxt bd_card_order); (* bd_cinfinite: BNF_Cardinal_Arithmetic.cinfinite bd_F *) fun bd_cinfinite_tac ctxt = HEADGOAL (rtac ctxt bd_cinfinite); (*target: ordLeq3 (card_of (set_F' (rep_G x_))) bd_F*) fun mk_set_G_bd_tac thms set_bd_thm ctxt = EVERY [Local_Defs.fold_tac ctxt [#set_F'_def thms], unfold_thms_tac ctxt [o_apply], HEADGOAL (rtac ctxt (@{thm ordLeq_transitive} OF [@{thm card_of_mono1} OF [#set_F'_subset thms], set_bd_thm]))]; (* rel_compp: rel_G R OO rel_G S \ rel_G (R OO S) *) fun rel_compp_tac ctxt = EVERY [unfold_thms_tac ctxt [#REL qthms], HEADGOAL (TWICE (rtac ctxt @{thm vimage2p_relcompp_mono})), (unfold_thms_tac ctxt (REL_OO_REL_left_thm :: @{thms relcompp_assoc})), (unfold_thms_tac ctxt [Local_Defs.unfold ctxt @{thms eq_OO} (infer_instantiate' ctxt [HOLogic.eq_const HOLogic.unitT |> Thm.cterm_of ctxt |> SOME] @{thm sum.rel_compp})]), HEADGOAL (rtac ctxt rel_pos_distr_thm), unfold_thms_tac ctxt @{thms fun_eq_iff bot_apply bot_bool_def not_all eq_False not_not OO_def}, REPEAT_DETERM (HEADGOAL (resolve_tac ctxt [exI, conjI, @{thm rel_sum.intros(1)}, refl]))]; (* rel_G R_ = (\x y. \z. set_G z \ {(x, y). R x y} \ map_G fst z = x \ map_G snd z = y) *) fun rel_compp_Grp_tac ctxt = let val _ = () in EVERY [Local_Defs.fold_tac ctxt (@{thm Grp_def} :: map #set_F'_def set_F'_thmss), unfold_thms_tac ctxt [o_apply, @{thm mem_Collect_eq}, @{thm OO_Grp_alt}, @{thm vimage2p_def}], Local_Defs.fold_tac ctxt [Local_Defs.unfold ctxt @{thms vimage2p_def} (#def rel_F')], unfold_thms_tac ctxt [rel_F'_set_thm], HEADGOAL (TWICE (rtac ctxt ext)), HEADGOAL (rtac ctxt iffI), REPEAT_DETERM (ALLGOALS (eresolve_tac ctxt [exE, conjE])), HEADGOAL (rtac ctxt exI), REPEAT_FIRST (resolve_tac ctxt [conjI]), HEADGOAL (EVERY' (maps (fn {set_F'_respect, ...} => [etac ctxt @{thm subset_trans[rotated]}, rtac ctxt equalityD1, rtac ctxt (@{thm rel_funD} OF [set_F'_respect]), rtac ctxt (#rep_abs qthms), rtac ctxt (#reflp qthms)]) set_F'_thmss)), (HEADGOAL o TWICE o EVERY') [rtac ctxt (trans OF [asm_rl, #abs_rep qthms]), rtac ctxt (#rel_abs qthms), etac ctxt (rotate_prems 1 (#transp qthms)), rtac ctxt map_F_rsp, rtac ctxt (#rep_abs qthms), rtac ctxt (#reflp qthms) ], HEADGOAL (rtac ctxt exI THEN' rtac ctxt conjI), (REPEAT_DETERM_N live o HEADGOAL o EVERY') [assume_tac ctxt, rtac ctxt conjI], (HEADGOAL o TWICE o EVERY') [ hyp_subst_tac ctxt, rtac ctxt (#rep_abs_rsp qthms), rtac ctxt map_F_rsp, rtac ctxt (#rep_reflp qthms)]] end; fun pred_G_set_G_tac ctxt = HEADGOAL (rtac ctxt refl); val tactics = map_G_id0_tac :: map_G_comp0_tac :: map_G_cong_tac :: map mk_set_G_map0_G_tac set_F'_thmss @ bd_card_order_tac :: bd_cinfinite_tac :: map2 mk_set_G_bd_tac set_F'_thmss set_bd_thms @ rel_compp_tac :: rel_compp_Grp_tac :: [pred_G_set_G_tac]; (* val wits_G = [abs_G o wit_F] *) val (wits_G, wit_thms) = let val wit_F_thmss = map (map2 (fn set_F' => fn wit => (#set_F'_subset set_F' RS set_mp RS wit) |> unfold_thms lthy [#set_F'_def set_F']) set_F'_thmss) (wit_thmss_of_bnf bnf_F); val (wits_F, wit_F_thmss) = split_list (filter_out (fn ((J, _), _) => exists (fn (I, _) => subset (op =) (I, J)) Iwits) (wits_F ~~ wit_F_thmss)); fun mk_wit (I, wit) = let val vars = (map (fn n => nth var_as n) I) in fold_rev lambda vars (abs_G $ list_comb (wit, vars)) end; in (map mk_wit (Iwits @ wits_F), wit_thmss @ flat wit_F_thmss) end; fun mk_wit_tacs ({set_F'_def, set_F'_respect, ...} :: set_F'_thmss) (w :: ws) ctxt = EVERY [unfold_thms_tac ctxt [@{thm o_def}, set_F'_respect RS @{thm rel_funD} OF [#rep_abs qthms OF [(#reflp qthms)]]], unfold_thms_tac ctxt [set_F'_def], HEADGOAL (etac ctxt w)] THEN mk_wit_tacs set_F'_thmss ws ctxt | mk_wit_tacs [] ws ctxt = mk_wit_tacs set_F'_thmss ws ctxt | mk_wit_tacs _ _ _ = all_tac; val (bnf_G, lthy) = bnf_def Dont_Inline (user_policy Note_Some) true I tactics (mk_wit_tacs [] wit_thms) NONE map_b rel_b pred_b set_bs (((((((Binding.empty, absT), map_G), sets_G), bd_G), wits_G), SOME rel_G), NONE) lthy; val old_defs = {sets = set_defs_of_bnf bnf_G, map = map_def_of_bnf bnf_G, rel = rel_def_of_bnf bnf_G}; val set_F'_defs = map (mk_abs_def o #set_F'_def) set_F'_thmss; val unfold_morphism = Morphism.thm_morphism "BNF" (unfold_thms lthy (defs @ #def REL :: set_F'_defs)); val (bnf_G, lthy) = morph_bnf_defs unfold_morphism bnf_G |> (fn bnf => note_bnf_defs bnf lthy); (* auxiliary lemmas transfer for transfer *) val rel_monoD_rotated = rotate_prems ~1 (rel_mono_of_bnf bnf_F RS @{thm predicate2D}); val REL_pos_distrI = let fun tac ctxt = EVERY [HEADGOAL (dtac ctxt (rotate_prems ~1 (rel_pos_distr_thm RS @{thm predicate2D}))), (REPEAT_DETERM o HEADGOAL) (rtac ctxt conjI ORELSE' assume_tac ctxt), (REPEAT_DETERM o HEADGOAL) (etac ctxt @{thm relcomppE}), HEADGOAL (dtac ctxt rel_monoD_rotated), (REPEAT_DETERM o HEADGOAL) (assume_tac ctxt ORELSE' rtac ctxt @{thm relcomppI})]; in prove lthy (var_x :: var_y' :: var_Ps @ var_Qs @ var_Rs) REL_pos_distrI_tm tac end; val rel_F_rel_F' = let val rel_F = mk_rel_of_bnf deads alphas betas bnf_F; val rel_F_rel_F'_tm = (rel_F, #tm rel_F') |> apply2 (fn R => HOLogic.mk_Trueprop (list_comb (R, var_Ps) $ var_x $ var_y)) |> Logic.mk_implies; fun rel_F_rel_F'_tac ctxt = EVERY [HEADGOAL (dtac ctxt (in_rel_of_bnf bnf_F RS iffD1)), unfold_thms_tac ctxt (rel_F'_set_thm :: @{thms mem_Collect_eq}), (REPEAT_DETERM o HEADGOAL) (eresolve_tac ctxt [exE, conjE]), HEADGOAL (rtac ctxt exI), HEADGOAL (EVERY' (maps (fn thms => [rtac ctxt conjI, rtac ctxt subsetI, dtac ctxt (set_mp OF [#set_F'_subset thms]), dtac ctxt subsetD, assume_tac ctxt, assume_tac ctxt]) set_F'_thmss)), (REPEAT_DETERM o HEADGOAL) (rtac ctxt conjI ORELSE' hyp_subst_tac ctxt THEN' rtac ctxt (#reflp qthms))] in prove lthy (var_x :: var_y :: var_Ps) rel_F_rel_F'_tm rel_F_rel_F'_tac end; fun inst_REL_pos_distrI n vs aTs bTs ctxt = infer_instantiate' ctxt (replicate n NONE @ (rel_Maybes vs aTs bTs |> map (SOME o Thm.cterm_of ctxt))) REL_pos_distrI; val Tss = {abs = typ_subst_atomic (alphas ~~ betas) absT, rep = repT, Ds0 = map TFree Ds0, deads = deads, alphas = alphas, betas = betas, gammas = gammas, deltas = deltas}; val thms = {map_F_rsp = map_F_rsp, rel_F'_def = #def rel_F', rel_F_rel_F' = rel_F_rel_F', rel_F'_set = rel_F'_set_thm, rel_monoD_rotated = rel_monoD_rotated} val transfer_consts = mk_quotient_transfer_tacs bnf_F Tss live qthms thms set_F'_thmss old_defs inst_REL_pos_distrI map_raw rel_raw (map (#tm o #set_F') set_F'_aux_defs); val quiet = exists (fn No_Warn_Transfer => true | _ => false) opts; in lthy |> BNF_Def.register_bnf plugins absT_name bnf_G |> mk_transfer_thms quiet bnf_F bnf_G absT_name transfer_consts (Quotient equiv_thm) Tss (defs @ #def REL :: set_F'_defs) end | _ => raise Match); in (goals, after_qed, #def REL :: defs, lthy) end; (** main commands **) local fun prepare_common prepare_name prepare_sort prepare_term prepare_thm (((((plugins, raw_specs), raw_absT_name), raw_wits), xthms_opt), (map_b, rel_b, pred_b)) lthy = let val absT_name = prepare_name lthy raw_absT_name; fun bad_input input = Pretty.chunks [Pretty.para ("Expected theorem(s) of either form:"), Pretty.item [Pretty.enum "," "" "" [Syntax.pretty_term lthy @{term "Quotient R Abs Rep T"}, Syntax.pretty_term lthy @{term "reflp R"}]], Pretty.item [Syntax.pretty_term lthy @{term "Quotient R Abs Rep T"}], Pretty.item [Syntax.pretty_term lthy @{term "type_definition Rep Abs A"}], Pretty.para ("Got"), Pretty.enum "," "" "" (map (Thm.pretty_thm lthy) input)] |> Pretty.string_of |> error; fun no_refl qthm = Pretty.chunks [Pretty.para ("Could not find a reflexivity rule for the Quotient theorem:"), Pretty.item [Thm.pretty_thm lthy qthm], Pretty.para ("Try supplying a reflexivity theorem manually or registering it in setup_lifting.")] |> Pretty.string_of |> error; fun find_equiv_thm_via_Quotient qthm = let val refl_thms = Lifting_Info.get_reflexivity_rules lthy |> map (unfold_thms lthy @{thms reflp_eq[symmetric]}); in (case refl_thms RL [qthm RS @{thm Quotient_reflp_imp_equivp}] of [] => no_refl qthm | thm :: _ => thm) end; val (lift_thm, equiv_thm) = (case Option.map (prepare_thm lthy) xthms_opt of SOME (thms as [qthm, _]) => (case try (fn thms => @{thm Quotient_reflp_imp_equivp} OF thms) thms of SOME equiv_thm => (qthm RS @{thm Quotient_Quotient3}, Quotient equiv_thm) | NONE => bad_input thms) | SOME [thm] => (case try (fn thm => thm RS @{thm Quotient_Quotient3}) thm of SOME qthm => (qthm, Quotient (find_equiv_thm_via_Quotient thm)) | NONE => if can (fn thm => thm RS @{thm type_definition.Rep}) thm then (thm, Typedef) else bad_input [thm]) | NONE => (case Lifting_Info.lookup_quotients lthy absT_name of SOME {quot_thm = qthm, ...} => (case [qthm] RL @{thms Quotient_eq_onp_typedef Quotient_eq_onp_type_copy} of thm :: _ => (thm, Typedef) | _ => (qthm RS @{thm Quotient_Quotient3}, Quotient (find_equiv_thm_via_Quotient qthm))) | NONE => (Typedef.get_info lthy absT_name |> hd |> snd |> #type_definition, Typedef)) | SOME thms => bad_input thms); val wits = (Option.map o map) (prepare_term lthy) raw_wits; val specs = map (apsnd (apsnd (the_default @{sort type} o Option.map (prepare_sort lthy)))) raw_specs; val which_bnf = (case equiv_thm of Quotient thm => quotient_bnf (thm, lift_thm) | Typedef => typedef_bnf lift_thm); in which_bnf wits specs map_b rel_b pred_b plugins lthy end; fun prepare_lift_bnf prepare_name prepare_sort prepare_term prepare_thm = (fn (goals, after_qed, definitions, lthy) => lthy |> Proof.theorem NONE after_qed (map (single o rpair []) goals) |> Proof.refine_singleton (Method.Basic (fn ctxt => SIMPLE_METHOD (unfold_thms_tac ctxt definitions))) |> Proof.refine_singleton (Method.primitive_text (K I))) oo prepare_common prepare_name prepare_sort prepare_term prepare_thm o apfst (apfst (apsnd SOME)); fun prepare_solve prepare_name prepare_typ prepare_sort prepare_thm tacs = (fn (goals, after_qed, definitions, lthy) => lthy |> after_qed (map2 (fn goal => fn tac => [Goal.prove_sorry lthy [] [] goal (fn (ctxtprems as {context = ctxt, prems = _}) => unfold_thms_tac ctxt definitions THEN tac ctxtprems)]) goals (tacs (length goals)))) oo prepare_common prepare_name prepare_typ prepare_sort prepare_thm; in val lift_bnf_cmd = prepare_lift_bnf (fst o dest_Type oo Proof_Context.read_type_name {proper = true, strict = false}) Syntax.read_sort Syntax.read_term Attrib.eval_thms; fun lift_bnf args tacs = prepare_solve (K I) (K I) (K I) (K I) (K tacs) args; fun copy_bnf_tac {context = ctxt, prems = _} = REPEAT_DETERM (resolve_tac ctxt [bexI, conjI, UNIV_I, refl, subset_refl] 1); val copy_bnf = apfst (apfst (rpair NONE)) #> apfst (apsnd (Option.map single)) #> prepare_solve (K I) (K I) (K I) (K I) (fn n => replicate n copy_bnf_tac); val copy_bnf_cmd = apfst (apfst (rpair NONE)) #> apfst (apsnd (Option.map single)) #> prepare_solve (fst o dest_Type oo Proof_Context.read_type_name {proper = true, strict = false}) Syntax.read_sort Syntax.read_term Attrib.eval_thms (fn n => replicate n copy_bnf_tac); end; (** outer syntax **) local (* parsers *) val parse_wits = @{keyword "["} |-- (Parse.name --| @{keyword ":"} -- Scan.repeat Parse.term >> (fn ("wits", Ts) => Ts | (s, _) => error ("Unknown label " ^ quote s ^ " (expected \"wits\")"))) --| @{keyword "]"} || Scan.succeed []; fun parse_common_opts p = Scan.optional (@{keyword "("} |-- Parse.list1 (Parse.group (K "option") (Scan.first (p :: [Plugin_Name.parse_filter >> Plugins_Option, Parse.reserved "no_warn_transfer" >> K No_Warn_Transfer]))) --| @{keyword ")"}) []; val parse_lift_opts = Parse.reserved "no_warn_wits" >> K No_Warn_Wits |> parse_common_opts; val parse_copy_opts = parse_common_opts Scan.fail; val parse_xthm = Scan.option (Parse.reserved "via" |-- Parse.thm); val parse_xthms = Scan.option (Parse.reserved "via" |-- Parse.thms1); in (* exposed commands *) val _ = Outer_Syntax.local_theory_to_proof @{command_keyword lift_bnf} "register a quotient type/subtype of a bounded natural functor (BNF) as a BNF" ((parse_lift_opts -- parse_type_args_named_constrained -- Parse.type_const -- parse_wits -- parse_xthms -- parse_map_rel_pred_bindings) >> lift_bnf_cmd); val _ = Outer_Syntax.local_theory @{command_keyword copy_bnf} "register a type copy of a bounded natural functor (BNF) as a BNF" ((parse_copy_opts -- parse_type_args_named_constrained -- Parse.type_const -- parse_xthm -- parse_map_rel_pred_bindings) >> copy_bnf_cmd); end; end; diff --git a/src/HOL/Tools/BNF/bnf_util.ML b/src/HOL/Tools/BNF/bnf_util.ML --- a/src/HOL/Tools/BNF/bnf_util.ML +++ b/src/HOL/Tools/BNF/bnf_util.ML @@ -1,480 +1,480 @@ (* Title: HOL/Tools/BNF/bnf_util.ML Author: Dmitriy Traytel, TU Muenchen Copyright 2012 Library for bounded natural functors. *) signature BNF_UTIL = sig include CTR_SUGAR_UTIL include BNF_FP_REC_SUGAR_UTIL val transfer_plugin: string val unflatt: 'a list list list -> 'b list -> 'b list list list val unflattt: 'a list list list list -> 'b list -> 'b list list list list val mk_TFreess: int list -> Proof.context -> typ list list * Proof.context val mk_Freesss: string -> typ list list list -> Proof.context -> term list list list * Proof.context val mk_Freessss: string -> typ list list list list -> Proof.context -> term list list list list * Proof.context val nonzero_string_of_int: int -> string val binder_fun_types: typ -> typ list val body_fun_type: typ -> typ val strip_fun_type: typ -> typ list * typ val strip_typeN: int -> typ -> typ list * typ val mk_pred2T: typ -> typ -> typ val mk_relT: typ * typ -> typ val dest_relT: typ -> typ * typ val dest_pred2T: typ -> typ * typ val mk_sumT: typ * typ -> typ val ctwo: term val fst_const: typ -> term val snd_const: typ -> term val Id_const: typ -> term val enforce_type: Proof.context -> (typ -> typ) -> typ -> term -> term val mk_Ball: term -> term -> term val mk_Bex: term -> term -> term val mk_Card_order: term -> term val mk_Field: term -> term val mk_Gr: term -> term -> term val mk_Grp: term -> term -> term val mk_UNION: term -> term -> term val mk_Union: typ -> term val mk_card_binop: string -> (typ * typ -> typ) -> term -> term -> term val mk_card_of: term -> term val mk_card_order: term -> term val mk_cexp: term -> term -> term val mk_cinfinite: term -> term val mk_collect: term list -> typ -> term val mk_converse: term -> term val mk_conversep: term -> term val mk_cprod: term -> term -> term val mk_csum: term -> term -> term val mk_dir_image: term -> term -> term val mk_eq_onp: term -> term val mk_rel_fun: term -> term -> term val mk_image: term -> term val mk_in: term list -> term list -> typ -> term val mk_inj: term -> term val mk_leq: term -> term -> term val mk_ordLeq: term -> term -> term val mk_rel_comp: term * term -> term val mk_rel_compp: term * term -> term val mk_vimage2p: term -> term -> term val mk_reflp: term -> term val mk_symp: term -> term val mk_transp: term -> term val mk_union: term * term -> term (*parameterized terms*) val mk_nthN: int -> term -> int -> term (*parameterized thms*) val prod_injectD: thm val prod_injectI: thm val ctrans: thm val id_apply: thm val meta_mp: thm val meta_spec: thm val o_apply: thm val rel_funD: thm val rel_funI: thm val set_mp: thm val set_rev_mp: thm val subset_UNIV: thm val mk_conjIN: int -> thm val mk_nthI: int -> int -> thm val mk_nth_conv: int -> int -> thm val mk_ordLeq_csum: int -> int -> thm -> thm val mk_pointful: Proof.context -> thm -> thm val mk_rel_funDN: int -> thm -> thm val mk_rel_funDN_rotated: int -> thm -> thm val mk_sym: thm -> thm val mk_trans: thm -> thm -> thm val mk_UnIN: int -> int -> thm val mk_Un_upper: int -> int -> thm val is_refl_bool: term -> bool val is_refl: thm -> bool val is_concl_refl: thm -> bool val no_refl: thm list -> thm list val no_reflexive: thm list -> thm list val parse_type_args_named_constrained: (binding option * (string * string option)) list parser val parse_map_rel_pred_bindings: (binding * binding * binding) parser val typedef: binding * (string * sort) list * mixfix -> term -> (binding * binding) option -> (Proof.context -> tactic) -> local_theory -> (string * Typedef.info) * local_theory end; structure BNF_Util : BNF_UTIL = struct open Ctr_Sugar_Util open BNF_FP_Rec_Sugar_Util val transfer_plugin = Plugin_Name.declare_setup \<^binding>\transfer\; (* Library proper *) fun unfla0 xs = fold_map (fn _ => fn (c :: cs) => (c, cs)) xs; fun unflat0 xss = fold_map unfla0 xss; fun unflatt0 xsss = fold_map unflat0 xsss; fun unflattt0 xssss = fold_map unflatt0 xssss; fun unflatt xsss = fst o unflatt0 xsss; fun unflattt xssss = fst o unflattt0 xssss; val parse_type_arg_constrained = Parse.type_ident -- Scan.option (\<^keyword>\::\ |-- Parse.!!! Parse.sort); val parse_type_arg_named_constrained = (Parse.reserved "dead" >> K NONE || parse_opt_binding_colon >> SOME) -- parse_type_arg_constrained; val parse_type_args_named_constrained = parse_type_arg_constrained >> (single o pair (SOME Binding.empty)) || \<^keyword>\(\ |-- Parse.!!! (Parse.list1 parse_type_arg_named_constrained --| \<^keyword>\)\) || Scan.succeed []; val parse_map_rel_pred_binding = Parse.name --| \<^keyword>\:\ -- Parse.binding; val no_map_rel = (Binding.empty, Binding.empty, Binding.empty); fun extract_map_rel_pred ("map", m) = (fn (_, r, p) => (m, r, p)) | extract_map_rel_pred ("rel", r) = (fn (m, _, p) => (m, r, p)) | extract_map_rel_pred ("pred", p) = (fn (m, r, _) => (m, r, p)) | extract_map_rel_pred (s, _) = error ("Unknown label " ^ quote s ^ " (expected \"map\" or \"rel\")"); val parse_map_rel_pred_bindings = \<^keyword>\for\ |-- Scan.repeat parse_map_rel_pred_binding >> (fn ps => fold extract_map_rel_pred ps no_map_rel) || Scan.succeed no_map_rel; fun typedef (b, Ts, mx) set opt_morphs tac lthy = let (*Work around loss of qualification in "typedef" axioms by replicating it in the name*) val b' = fold_rev Binding.prefix_name (map (suffix "_" o fst) (Binding.path_of b)) b; val default_bindings = Typedef.default_bindings (Binding.concealed b'); val bindings = (case opt_morphs of NONE => default_bindings | SOME (Rep_name, Abs_name) => {Rep_name = Binding.concealed Rep_name, Abs_name = Binding.concealed Abs_name, type_definition_name = #type_definition_name default_bindings}); val ((name, info), (lthy, lthy_old)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> Typedef.add_typedef {overloaded = false} (b', Ts, mx) set (SOME bindings) tac ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; in ((name, Typedef.transform_info phi info), lthy) end; (* Term construction *) (** Fresh variables **) fun nonzero_string_of_int 0 = "" | nonzero_string_of_int n = string_of_int n; val mk_TFreess = fold_map mk_TFrees; fun mk_Freesss x Tsss = @{fold_map 2} mk_Freess (mk_names (length Tsss) x) Tsss; fun mk_Freessss x Tssss = @{fold_map 2} mk_Freesss (mk_names (length Tssss) x) Tssss; (** Types **) (*maps [T1,...,Tn]--->T to ([T1,T2,...,Tn], T)*) fun strip_typeN 0 T = ([], T) | strip_typeN n (Type (\<^type_name>\fun\, [T, T'])) = strip_typeN (n - 1) T' |>> cons T | strip_typeN _ T = raise TYPE ("strip_typeN", [T], []); (*maps [T1,...,Tn]--->T-->U to ([T1,T2,...,Tn], T-->U), where U is not a function type*) fun strip_fun_type T = strip_typeN (num_binder_types T - 1) T; val binder_fun_types = fst o strip_fun_type; val body_fun_type = snd o strip_fun_type; fun mk_pred2T T U = mk_predT [T, U]; val mk_relT = HOLogic.mk_setT o HOLogic.mk_prodT; val dest_relT = HOLogic.dest_prodT o HOLogic.dest_setT; val dest_pred2T = apsnd Term.domain_type o Term.dest_funT; fun mk_sumT (LT, RT) = Type (\<^type_name>\Sum_Type.sum\, [LT, RT]); (** Constants **) fun fst_const T = Const (\<^const_name>\fst\, T --> fst (HOLogic.dest_prodT T)); fun snd_const T = Const (\<^const_name>\snd\, T --> snd (HOLogic.dest_prodT T)); fun Id_const T = Const (\<^const_name>\Id\, mk_relT (T, T)); (** Operators **) fun enforce_type ctxt get_T T t = Term.subst_TVars (tvar_subst (Proof_Context.theory_of ctxt) [get_T (fastype_of t)] [T]) t; fun mk_converse R = let val RT = dest_relT (fastype_of R); val RST = mk_relT (snd RT, fst RT); in Const (\<^const_name>\converse\, fastype_of R --> RST) $ R end; fun mk_rel_comp (R, S) = let val RT = fastype_of R; val ST = fastype_of S; val RST = mk_relT (fst (dest_relT RT), snd (dest_relT ST)); in Const (\<^const_name>\relcomp\, RT --> ST --> RST) $ R $ S end; fun mk_Gr A f = let val ((AT, BT), FT) = `dest_funT (fastype_of f); in Const (\<^const_name>\Gr\, HOLogic.mk_setT AT --> FT --> mk_relT (AT, BT)) $ A $ f end; fun mk_conversep R = let val RT = dest_pred2T (fastype_of R); val RST = mk_pred2T (snd RT) (fst RT); in Const (\<^const_name>\conversep\, fastype_of R --> RST) $ R end; fun mk_rel_compp (R, S) = let val RT = fastype_of R; val ST = fastype_of S; val RST = mk_pred2T (fst (dest_pred2T RT)) (snd (dest_pred2T ST)); in Const (\<^const_name>\relcompp\, RT --> ST --> RST) $ R $ S end; fun mk_Grp A f = let val ((AT, BT), FT) = `dest_funT (fastype_of f); in Const (\<^const_name>\Grp\, HOLogic.mk_setT AT --> FT --> mk_pred2T AT BT) $ A $ f end; fun mk_image f = let val (T, U) = dest_funT (fastype_of f); in Const (\<^const_name>\image\, (T --> U) --> HOLogic.mk_setT T --> HOLogic.mk_setT U) $ f end; fun mk_Ball X f = Const (\<^const_name>\Ball\, fastype_of X --> fastype_of f --> HOLogic.boolT) $ X $ f; fun mk_Bex X f = Const (\<^const_name>\Bex\, fastype_of X --> fastype_of f --> HOLogic.boolT) $ X $ f; fun mk_UNION X f = let val (T, U) = dest_funT (fastype_of f); in Const (\<^const_name>\Sup\, HOLogic.mk_setT U --> U) $ (Const (\<^const_name>\image\, (T --> U) --> fastype_of X --> HOLogic.mk_setT U) $ f $ X) end; fun mk_Union T = Const (\<^const_name>\Sup\, HOLogic.mk_setT (HOLogic.mk_setT T) --> HOLogic.mk_setT T); val mk_union = HOLogic.mk_binop \<^const_name>\sup\; fun mk_Field r = let val T = fst (dest_relT (fastype_of r)); in Const (\<^const_name>\Field\, mk_relT (T, T) --> HOLogic.mk_setT T) $ r end; fun mk_card_order bd = let val T = fastype_of bd; val AT = fst (dest_relT T); in Const (\<^const_name>\card_order_on\, HOLogic.mk_setT AT --> T --> HOLogic.boolT) $ HOLogic.mk_UNIV AT $ bd end; fun mk_Card_order bd = let val T = fastype_of bd; val AT = fst (dest_relT T); in Const (\<^const_name>\card_order_on\, HOLogic.mk_setT AT --> T --> HOLogic.boolT) $ mk_Field bd $ bd end; fun mk_cinfinite bd = Const (\<^const_name>\cinfinite\, fastype_of bd --> HOLogic.boolT) $ bd; fun mk_ordLeq t1 t2 = HOLogic.mk_mem (HOLogic.mk_prod (t1, t2), Const (\<^const_name>\ordLeq\, mk_relT (fastype_of t1, fastype_of t2))); fun mk_card_of A = let val AT = fastype_of A; val T = HOLogic.dest_setT AT; in Const (\<^const_name>\card_of\, AT --> mk_relT (T, T)) $ A end; fun mk_dir_image r f = let val (T, U) = dest_funT (fastype_of f); in Const (\<^const_name>\dir_image\, mk_relT (T, T) --> (T --> U) --> mk_relT (U, U)) $ r $ f end; fun mk_rel_fun R S = let val ((RA, RB), RT) = `dest_pred2T (fastype_of R); val ((SA, SB), ST) = `dest_pred2T (fastype_of S); in Const (\<^const_name>\rel_fun\, RT --> ST --> mk_pred2T (RA --> SA) (RB --> SB)) $ R $ S end; (*FIXME: "x"?*) (*(nth sets i) must be of type "T --> 'ai set"*) fun mk_in As sets T = let fun in_single set A = let val AT = fastype_of A; in Const (\<^const_name>\less_eq\, AT --> AT --> HOLogic.boolT) $ (set $ Free ("x", T)) $ A end; in if null sets then HOLogic.mk_UNIV T else HOLogic.mk_Collect ("x", T, foldr1 (HOLogic.mk_conj) (map2 in_single sets As)) end; fun mk_inj t = let val T as Type (\<^type_name>\fun\, [domT, _]) = fastype_of t in Const (\<^const_name>\inj_on\, T --> HOLogic.mk_setT domT --> HOLogic.boolT) $ t $ HOLogic.mk_UNIV domT end; fun mk_leq t1 t2 = Const (\<^const_name>\less_eq\, fastype_of t1 --> fastype_of t2 --> HOLogic.boolT) $ t1 $ t2; fun mk_card_binop binop typop t1 t2 = let val (T1, relT1) = `(fst o dest_relT) (fastype_of t1); val (T2, relT2) = `(fst o dest_relT) (fastype_of t2); in Const (binop, relT1 --> relT2 --> mk_relT (typop (T1, T2), typop (T1, T2))) $ t1 $ t2 end; val mk_csum = mk_card_binop \<^const_name>\csum\ mk_sumT; val mk_cprod = mk_card_binop \<^const_name>\cprod\ HOLogic.mk_prodT; val mk_cexp = mk_card_binop \<^const_name>\cexp\ (op --> o swap); val ctwo = \<^term>\ctwo\; fun mk_collect xs defT = let val T = (case xs of [] => defT | (x::_) => fastype_of x); in Const (\<^const_name>\collect\, HOLogic.mk_setT T --> T) $ (HOLogic.mk_set T xs) end; fun mk_vimage2p f g = let val (T1, T2) = dest_funT (fastype_of f); val (U1, U2) = dest_funT (fastype_of g); in Const (\<^const_name>\vimage2p\, (T1 --> T2) --> (U1 --> U2) --> mk_pred2T T2 U2 --> mk_pred2T T1 U1) $ f $ g end; fun mk_eq_onp P = let val T = domain_type (fastype_of P); in Const (\<^const_name>\eq_onp\, (T --> HOLogic.boolT) --> T --> T --> HOLogic.boolT) $ P end; fun mk_pred name R = Const (name, uncurry mk_pred2T (fastype_of R |> dest_pred2T) --> HOLogic.boolT) $ R; val mk_reflp = mk_pred \<^const_name>\reflp\; val mk_symp = mk_pred \<^const_name>\symp\; val mk_transp = mk_pred \<^const_name>\transp\; fun mk_trans thm1 thm2 = trans OF [thm1, thm2]; fun mk_sym thm = thm RS sym; val prod_injectD = @{thm iffD1[OF prod.inject]}; val prod_injectI = @{thm iffD2[OF prod.inject]}; val ctrans = @{thm ordLeq_transitive}; val id_apply = @{thm id_apply}; val meta_mp = @{thm meta_mp}; val meta_spec = @{thm meta_spec}; val o_apply = @{thm o_apply}; val rel_funD = @{thm rel_funD}; val rel_funI = @{thm rel_funI}; val set_mp = @{thm set_mp}; val set_rev_mp = @{thm set_rev_mp}; val subset_UNIV = @{thm subset_UNIV}; fun mk_pointful ctxt thm = unfold_thms ctxt [o_apply] (thm RS fun_cong); fun mk_nthN 1 t 1 = t | mk_nthN _ t 1 = HOLogic.mk_fst t | mk_nthN 2 t 2 = HOLogic.mk_snd t | mk_nthN n t m = mk_nthN (n - 1) (HOLogic.mk_snd t) (m - 1); fun mk_nth_conv n m = let fun thm b = if b then @{thm fstI} else @{thm sndI}; fun mk_nth_conv _ 1 1 = refl | mk_nth_conv _ _ 1 = @{thm fst_conv} | mk_nth_conv _ 2 2 = @{thm snd_conv} | mk_nth_conv b _ 2 = @{thm snd_conv} RS thm b | mk_nth_conv b n m = mk_nth_conv false (n - 1) (m - 1) RS thm b; in mk_nth_conv (not (m = n)) n m end; fun mk_nthI 1 1 = @{thm TrueE[OF TrueI]} | mk_nthI n m = fold (curry op RS) (replicate (m - 1) @{thm sndI}) (if m = n then @{thm TrueE[OF TrueI]} else @{thm fstI}); fun mk_conjIN 1 = @{thm TrueE[OF TrueI]} | mk_conjIN n = mk_conjIN (n - 1) RSN (2, conjI); fun mk_ordLeq_csum 1 1 thm = thm | mk_ordLeq_csum _ 1 thm = @{thm ordLeq_transitive} OF [thm, @{thm ordLeq_csum1}] | mk_ordLeq_csum 2 2 thm = @{thm ordLeq_transitive} OF [thm, @{thm ordLeq_csum2}] | mk_ordLeq_csum n m thm = @{thm ordLeq_transitive} OF [mk_ordLeq_csum (n - 1) (m - 1) thm, @{thm ordLeq_csum2[OF Card_order_csum]}]; fun mk_rel_funDN n = funpow n (fn thm => thm RS rel_funD); val mk_rel_funDN_rotated = rotate_prems ~1 oo mk_rel_funDN; local fun mk_Un_upper' 0 = @{thm subset_refl} | mk_Un_upper' 1 = @{thm Un_upper1} | mk_Un_upper' k = Library.foldr (op RS o swap) (replicate (k - 1) @{thm subset_trans[OF Un_upper1]}, @{thm Un_upper1}); in fun mk_Un_upper 1 1 = @{thm subset_refl} | mk_Un_upper n 1 = mk_Un_upper' (n - 2) RS @{thm subset_trans[OF Un_upper1]} | mk_Un_upper n m = mk_Un_upper' (n - m) RS @{thm subset_trans[OF Un_upper2]}; end; local fun mk_UnIN' 0 = @{thm UnI2} | mk_UnIN' m = mk_UnIN' (m - 1) RS @{thm UnI1}; in fun mk_UnIN 1 1 = @{thm TrueE[OF TrueI]} | mk_UnIN n 1 = Library.foldr1 (op RS o swap) (replicate (n - 1) @{thm UnI1}) | mk_UnIN n m = mk_UnIN' (n - m) end; fun is_refl_bool t = op aconv (HOLogic.dest_eq t) handle TERM _ => false; fun is_refl_prop t = op aconv (HOLogic.dest_eq (HOLogic.dest_Trueprop t)) handle TERM _ => false; val is_refl = is_refl_prop o Thm.prop_of; val is_concl_refl = is_refl_prop o Logic.strip_imp_concl o Thm.prop_of; val no_refl = filter_out is_refl; val no_reflexive = filter_out Thm.is_reflexive; end; diff --git a/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML b/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML --- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML +++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML @@ -1,1270 +1,1270 @@ (* Title: HOL/Tools/Ctr_Sugar/ctr_sugar.ML Author: Jasmin Blanchette, TU Muenchen Author: Martin Desharnais, TU Muenchen Copyright 2012, 2013 Wrapping existing freely generated type's constructors. *) signature CTR_SUGAR = sig datatype ctr_sugar_kind = Datatype | Codatatype | Record | Unknown type ctr_sugar = {kind: ctr_sugar_kind, T: typ, ctrs: term list, casex: term, discs: term list, selss: term list list, exhaust: thm, nchotomy: thm, injects: thm list, distincts: thm list, case_thms: thm list, case_cong: thm, case_cong_weak: thm, case_distribs: thm list, split: thm, split_asm: thm, disc_defs: thm list, disc_thmss: thm list list, discIs: thm list, disc_eq_cases: thm list, sel_defs: thm list, sel_thmss: thm list list, distinct_discsss: thm list list list, exhaust_discs: thm list, exhaust_sels: thm list, collapses: thm list, expands: thm list, split_sels: thm list, split_sel_asms: thm list, case_eq_ifs: thm list}; val morph_ctr_sugar: morphism -> ctr_sugar -> ctr_sugar val transfer_ctr_sugar: theory -> ctr_sugar -> ctr_sugar val ctr_sugar_of: Proof.context -> string -> ctr_sugar option val ctr_sugar_of_global: theory -> string -> ctr_sugar option val ctr_sugars_of: Proof.context -> ctr_sugar list val ctr_sugars_of_global: theory -> ctr_sugar list val ctr_sugar_of_case: Proof.context -> string -> ctr_sugar option val ctr_sugar_of_case_global: theory -> string -> ctr_sugar option val ctr_sugar_interpretation: string -> (ctr_sugar -> local_theory -> local_theory) -> theory -> theory val interpret_ctr_sugar: (string -> bool) -> ctr_sugar -> local_theory -> local_theory val register_ctr_sugar_raw: ctr_sugar -> local_theory -> local_theory val register_ctr_sugar: (string -> bool) -> ctr_sugar -> local_theory -> local_theory val default_register_ctr_sugar_global: (string -> bool) -> ctr_sugar -> theory -> theory val mk_half_pairss: 'a list * 'a list -> ('a * 'a) list list val join_halves: int -> 'a list list -> 'a list list -> 'a list * 'a list list list val mk_ctr: typ list -> term -> term val mk_case: typ list -> typ -> term -> term val mk_disc_or_sel: typ list -> term -> term val name_of_ctr: term -> string val name_of_disc: term -> string val dest_ctr: Proof.context -> string -> term -> term * term list val dest_case: Proof.context -> string -> typ list -> term -> (ctr_sugar * term list * term list) option type ('c, 'a) ctr_spec = (binding * 'c) * 'a list val disc_of_ctr_spec: ('c, 'a) ctr_spec -> binding val ctr_of_ctr_spec: ('c, 'a) ctr_spec -> 'c val args_of_ctr_spec: ('c, 'a) ctr_spec -> 'a list val code_plugin: string type ctr_options = (string -> bool) * bool type ctr_options_cmd = (Proof.context -> string -> bool) * bool val fake_local_theory_for_sel_defaults: (binding * typ) list -> Proof.context -> Proof.context val free_constructors: ctr_sugar_kind -> ({prems: thm list, context: Proof.context} -> tactic) list list -> ((ctr_options * binding) * (term, binding) ctr_spec list) * term list -> local_theory -> ctr_sugar * local_theory val free_constructors_cmd: ctr_sugar_kind -> ((((Proof.context -> Plugin_Name.filter) * bool) * binding) * ((binding * string) * binding list) list) * string list -> Proof.context -> Proof.state val default_ctr_options: ctr_options val default_ctr_options_cmd: ctr_options_cmd val parse_bound_term: (binding * string) parser val parse_ctr_options: ctr_options_cmd parser val parse_ctr_spec: 'c parser -> 'a parser -> ('c, 'a) ctr_spec parser val parse_sel_default_eqs: string list parser end; structure Ctr_Sugar : CTR_SUGAR = struct open Ctr_Sugar_Util open Ctr_Sugar_Tactics open Ctr_Sugar_Code datatype ctr_sugar_kind = Datatype | Codatatype | Record | Unknown; type ctr_sugar = {kind: ctr_sugar_kind, T: typ, ctrs: term list, casex: term, discs: term list, selss: term list list, exhaust: thm, nchotomy: thm, injects: thm list, distincts: thm list, case_thms: thm list, case_cong: thm, case_cong_weak: thm, case_distribs: thm list, split: thm, split_asm: thm, disc_defs: thm list, disc_thmss: thm list list, discIs: thm list, disc_eq_cases: thm list, sel_defs: thm list, sel_thmss: thm list list, distinct_discsss: thm list list list, exhaust_discs: thm list, exhaust_sels: thm list, collapses: thm list, expands: thm list, split_sels: thm list, split_sel_asms: thm list, case_eq_ifs: thm list}; fun morph_ctr_sugar phi ({kind, T, ctrs, casex, discs, selss, exhaust, nchotomy, injects, distincts, case_thms, case_cong, case_cong_weak, case_distribs, split, split_asm, disc_defs, disc_thmss, discIs, disc_eq_cases, sel_defs, sel_thmss, distinct_discsss, exhaust_discs, exhaust_sels, collapses, expands, split_sels, split_sel_asms, case_eq_ifs} : ctr_sugar) = {kind = kind, T = Morphism.typ phi T, ctrs = map (Morphism.term phi) ctrs, casex = Morphism.term phi casex, discs = map (Morphism.term phi) discs, selss = map (map (Morphism.term phi)) selss, exhaust = Morphism.thm phi exhaust, nchotomy = Morphism.thm phi nchotomy, injects = map (Morphism.thm phi) injects, distincts = map (Morphism.thm phi) distincts, case_thms = map (Morphism.thm phi) case_thms, case_cong = Morphism.thm phi case_cong, case_cong_weak = Morphism.thm phi case_cong_weak, case_distribs = map (Morphism.thm phi) case_distribs, split = Morphism.thm phi split, split_asm = Morphism.thm phi split_asm, disc_defs = map (Morphism.thm phi) disc_defs, disc_thmss = map (map (Morphism.thm phi)) disc_thmss, discIs = map (Morphism.thm phi) discIs, disc_eq_cases = map (Morphism.thm phi) disc_eq_cases, sel_defs = map (Morphism.thm phi) sel_defs, sel_thmss = map (map (Morphism.thm phi)) sel_thmss, distinct_discsss = map (map (map (Morphism.thm phi))) distinct_discsss, exhaust_discs = map (Morphism.thm phi) exhaust_discs, exhaust_sels = map (Morphism.thm phi) exhaust_sels, collapses = map (Morphism.thm phi) collapses, expands = map (Morphism.thm phi) expands, split_sels = map (Morphism.thm phi) split_sels, split_sel_asms = map (Morphism.thm phi) split_sel_asms, case_eq_ifs = map (Morphism.thm phi) case_eq_ifs}; val transfer_ctr_sugar = morph_ctr_sugar o Morphism.transfer_morphism; structure Data = Generic_Data ( type T = (Position.T * ctr_sugar) Symtab.table; val empty = Symtab.empty; val extend = I; fun merge data : T = Symtab.merge (K true) data; ); fun ctr_sugar_of_generic context = Option.map (transfer_ctr_sugar (Context.theory_of context) o #2) o Symtab.lookup (Data.get context); fun ctr_sugars_of_generic context = Symtab.fold (cons o transfer_ctr_sugar (Context.theory_of context) o #2 o #2) (Data.get context) []; fun ctr_sugar_of_case_generic context s = find_first (fn {casex = Const (s', _), ...} => s' = s | _ => false) (ctr_sugars_of_generic context); val ctr_sugar_of = ctr_sugar_of_generic o Context.Proof; val ctr_sugar_of_global = ctr_sugar_of_generic o Context.Theory; val ctr_sugars_of = ctr_sugars_of_generic o Context.Proof; val ctr_sugars_of_global = ctr_sugars_of_generic o Context.Theory; val ctr_sugar_of_case = ctr_sugar_of_case_generic o Context.Proof; val ctr_sugar_of_case_global = ctr_sugar_of_case_generic o Context.Theory; structure Ctr_Sugar_Plugin = Plugin(type T = ctr_sugar); fun ctr_sugar_interpretation name f = Ctr_Sugar_Plugin.interpretation name (fn ctr_sugar => fn lthy => f (transfer_ctr_sugar (Proof_Context.theory_of lthy) ctr_sugar) lthy); val interpret_ctr_sugar = Ctr_Sugar_Plugin.data; fun register_ctr_sugar_raw (ctr_sugar as {T = Type (name, _), ...}) = Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => fn context => let val pos = Position.thread_data () in Data.map (Symtab.update (name, (pos, morph_ctr_sugar phi ctr_sugar))) context end); fun register_ctr_sugar plugins ctr_sugar = register_ctr_sugar_raw ctr_sugar #> interpret_ctr_sugar plugins ctr_sugar; fun default_register_ctr_sugar_global plugins (ctr_sugar as {T = Type (name, _), ...}) thy = let val tab = Data.get (Context.Theory thy); val pos = Position.thread_data (); in if Symtab.defined tab name then thy else thy |> Context.theory_map (Data.put (Symtab.update_new (name, (pos, ctr_sugar)) tab)) |> Named_Target.theory_map (Ctr_Sugar_Plugin.data plugins ctr_sugar) end; val is_prefix = "is_"; val un_prefix = "un_"; val not_prefix = "not_"; fun mk_unN 1 1 suf = un_prefix ^ suf | mk_unN _ l suf = un_prefix ^ suf ^ string_of_int l; val caseN = "case"; val case_congN = "case_cong"; val case_eq_ifN = "case_eq_if"; val collapseN = "collapse"; val discN = "disc"; val disc_eq_caseN = "disc_eq_case"; val discIN = "discI"; val distinctN = "distinct"; val distinct_discN = "distinct_disc"; val exhaustN = "exhaust"; val exhaust_discN = "exhaust_disc"; val expandN = "expand"; val injectN = "inject"; val nchotomyN = "nchotomy"; val selN = "sel"; val exhaust_selN = "exhaust_sel"; val splitN = "split"; val split_asmN = "split_asm"; val split_selN = "split_sel"; val split_sel_asmN = "split_sel_asm"; val splitsN = "splits"; val split_selsN = "split_sels"; val case_cong_weak_thmsN = "case_cong_weak"; val case_distribN = "case_distrib"; val cong_attrs = @{attributes [cong]}; val dest_attrs = @{attributes [dest]}; val safe_elim_attrs = @{attributes [elim!]}; val iff_attrs = @{attributes [iff]}; val inductsimp_attrs = @{attributes [induct_simp]}; val nitpicksimp_attrs = @{attributes [nitpick_simp]}; val simp_attrs = @{attributes [simp]}; fun unflat_lookup eq xs ys = map (fn xs' => permute_like_unique eq xs xs' ys); fun mk_half_pairss' _ ([], []) = [] | mk_half_pairss' indent (x :: xs, _ :: ys) = indent @ fold_rev (cons o single o pair x) ys (mk_half_pairss' ([] :: indent) (xs, ys)); fun mk_half_pairss p = mk_half_pairss' [[]] p; fun join_halves n half_xss other_half_xss = (splice (flat half_xss) (flat other_half_xss), map2 (map2 append) (Library.chop_groups n half_xss) (transpose (Library.chop_groups n other_half_xss))); fun mk_undefined T = Const (\<^const_name>\undefined\, T); fun mk_ctr Ts t = let val Type (_, Ts0) = body_type (fastype_of t) in subst_nonatomic_types (Ts0 ~~ Ts) t end; fun mk_case Ts T t = let val (Type (_, Ts0), body) = strip_type (fastype_of t) |>> List.last in subst_nonatomic_types ((body, T) :: (Ts0 ~~ Ts)) t end; fun mk_disc_or_sel Ts t = subst_nonatomic_types (snd (Term.dest_Type (domain_type (fastype_of t))) ~~ Ts) t; val name_of_ctr = name_of_const "constructor" body_type; fun name_of_disc t = (case head_of t of Abs (_, _, \<^const>\Not\ $ (t' $ Bound 0)) => Long_Name.map_base_name (prefix not_prefix) (name_of_disc t') | Abs (_, _, Const (\<^const_name>\HOL.eq\, _) $ Bound 0 $ t') => Long_Name.map_base_name (prefix is_prefix) (name_of_disc t') | Abs (_, _, \<^const>\Not\ $ (Const (\<^const_name>\HOL.eq\, _) $ Bound 0 $ t')) => Long_Name.map_base_name (prefix (not_prefix ^ is_prefix)) (name_of_disc t') | t' => name_of_const "discriminator" (perhaps (try domain_type)) t'); val base_name_of_ctr = Long_Name.base_name o name_of_ctr; fun dest_ctr ctxt s t = let val (f, args) = Term.strip_comb t in (case ctr_sugar_of ctxt s of SOME {ctrs, ...} => (case find_first (can (fo_match ctxt f)) ctrs of SOME f' => (f', args) | NONE => raise Fail "dest_ctr") | NONE => raise Fail "dest_ctr") end; fun dest_case ctxt s Ts t = (case Term.strip_comb t of (Const (c, _), args as _ :: _) => (case ctr_sugar_of ctxt s of SOME (ctr_sugar as {casex = Const (case_name, _), discs = discs0, selss = selss0, ...}) => if case_name = c then let val n = length discs0 in if n < length args then let val (branches, obj :: leftovers) = chop n args; val discs = map (mk_disc_or_sel Ts) discs0; val selss = map (map (mk_disc_or_sel Ts)) selss0; val conds = map (rapp obj) discs; val branch_argss = map (fn sels => map (rapp obj) sels @ leftovers) selss; val branches' = map2 (curry Term.betapplys) branches branch_argss; in SOME (ctr_sugar, conds, branches') end else NONE end else NONE | _ => NONE) | _ => NONE); fun const_or_free_name (Const (s, _)) = Long_Name.base_name s | const_or_free_name (Free (s, _)) = s | const_or_free_name t = raise TERM ("const_or_free_name", [t]) fun extract_sel_default ctxt t = let fun malformed () = error ("Malformed selector default value equation: " ^ Syntax.string_of_term ctxt t); val ((sel, (ctr, vars)), rhs) = fst (Term.replace_dummy_patterns (Syntax.check_term ctxt t) 0) |> HOLogic.dest_eq |>> (Term.dest_comb #>> const_or_free_name ##> (Term.strip_comb #>> (Term.dest_Const #> fst))) handle TERM _ => malformed (); in if forall (is_Free orf is_Var) vars andalso not (has_duplicates (op aconv) vars) then ((ctr, sel), fold_rev Term.lambda vars rhs) else malformed () end; (* Ideally, we would enrich the context with constants rather than free variables. *) fun fake_local_theory_for_sel_defaults sel_bTs = Proof_Context.allow_dummies #> Proof_Context.add_fixes (map (fn (b, T) => (b, SOME T, NoSyn)) sel_bTs) #> snd; type ('c, 'a) ctr_spec = (binding * 'c) * 'a list; fun disc_of_ctr_spec ((disc, _), _) = disc; fun ctr_of_ctr_spec ((_, ctr), _) = ctr; fun args_of_ctr_spec (_, args) = args; val code_plugin = Plugin_Name.declare_setup \<^binding>\code\; fun prepare_free_constructors kind prep_plugins prep_term ((((raw_plugins, discs_sels), raw_case_binding), ctr_specs), sel_default_eqs) no_defs_lthy = let val plugins = prep_plugins no_defs_lthy raw_plugins; (* TODO: sanity checks on arguments *) val raw_ctrs = map ctr_of_ctr_spec ctr_specs; val raw_disc_bindings = map disc_of_ctr_spec ctr_specs; val raw_sel_bindingss = map args_of_ctr_spec ctr_specs; val n = length raw_ctrs; val ks = 1 upto n; val _ = n > 0 orelse error "No constructors specified"; val ctrs0 = map (prep_term no_defs_lthy) raw_ctrs; val (fcT_name, As0) = (case body_type (fastype_of (hd ctrs0)) of Type T' => T' | _ => error "Expected type constructor in body type of constructor"); val _ = forall ((fn Type (T_name, _) => T_name = fcT_name | _ => false) o body_type o fastype_of) (tl ctrs0) orelse error "Constructors not constructing same type"; val fc_b_name = Long_Name.base_name fcT_name; val fc_b = Binding.name fc_b_name; fun qualify mandatory = Binding.qualify mandatory fc_b_name; val (unsorted_As, [B, C]) = no_defs_lthy |> variant_tfrees (map (fst o dest_TFree_or_TVar) As0) ||> fst o mk_TFrees 2; val As = map2 (resort_tfree_or_tvar o snd o dest_TFree_or_TVar) As0 unsorted_As; val fcT = Type (fcT_name, As); val ctrs = map (mk_ctr As) ctrs0; val ctr_Tss = map (binder_types o fastype_of) ctrs; val ms = map length ctr_Tss; fun can_definitely_rely_on_disc k = not (Binding.is_empty (nth raw_disc_bindings (k - 1))) orelse nth ms (k - 1) = 0; fun can_rely_on_disc k = can_definitely_rely_on_disc k orelse (k = 1 andalso not (can_definitely_rely_on_disc 2)); fun should_omit_disc_binding k = n = 1 orelse (n = 2 andalso can_rely_on_disc (3 - k)); val equal_binding = \<^binding>\=\; fun is_disc_binding_valid b = not (Binding.is_empty b orelse Binding.eq_name (b, equal_binding)); val standard_disc_binding = Binding.name o prefix is_prefix o base_name_of_ctr; val disc_bindings = raw_disc_bindings |> @{map 4} (fn k => fn m => fn ctr => fn disc => qualify false (if Binding.is_empty disc then if m = 0 then equal_binding else if should_omit_disc_binding k then disc else standard_disc_binding ctr else if Binding.eq_name (disc, standard_binding) then standard_disc_binding ctr else disc)) ks ms ctrs0; fun standard_sel_binding m l = Binding.name o mk_unN m l o base_name_of_ctr; val sel_bindingss = @{map 3} (fn ctr => fn m => map2 (fn l => fn sel => qualify false (if Binding.is_empty sel orelse Binding.eq_name (sel, standard_binding) then standard_sel_binding m l ctr else sel)) (1 upto m) o pad_list Binding.empty m) ctrs0 ms raw_sel_bindingss; val add_bindings = Variable.add_fixes (distinct (op =) (filter Symbol_Pos.is_identifier (map Binding.name_of (disc_bindings @ flat sel_bindingss)))) #> snd; val case_Ts = map (fn Ts => Ts ---> B) ctr_Tss; val (((((((((u, exh_y), xss), yss), fs), gs), w), (p, p'))), _) = no_defs_lthy |> add_bindings |> yield_singleton (mk_Frees fc_b_name) fcT ||>> yield_singleton (mk_Frees "y") fcT (* for compatibility with "datatype_realizer.ML" *) ||>> mk_Freess "x" ctr_Tss ||>> mk_Freess "y" ctr_Tss ||>> mk_Frees "f" case_Ts ||>> mk_Frees "g" case_Ts ||>> yield_singleton (mk_Frees "z") B ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "P") HOLogic.boolT; val q = Free (fst p', mk_pred1T B); val xctrs = map2 (curry Term.list_comb) ctrs xss; val yctrs = map2 (curry Term.list_comb) ctrs yss; val xfs = map2 (curry Term.list_comb) fs xss; val xgs = map2 (curry Term.list_comb) gs xss; (* TODO: Eta-expension is for compatibility with the old datatype package (but it also provides nicer names). Consider removing. *) val eta_fs = map2 (fold_rev Term.lambda) xss xfs; val eta_gs = map2 (fold_rev Term.lambda) xss xgs; val case_binding = qualify false (if Binding.is_empty raw_case_binding orelse Binding.eq_name (raw_case_binding, standard_binding) then Binding.prefix_name (caseN ^ "_") fc_b else raw_case_binding); fun mk_case_disj xctr xf xs = list_exists_free xs (HOLogic.mk_conj (HOLogic.mk_eq (u, xctr), HOLogic.mk_eq (w, xf))); val case_rhs = fold_rev (fold_rev Term.lambda) [fs, [u]] (Const (\<^const_name>\The\, (B --> HOLogic.boolT) --> B) $ Term.lambda w (Library.foldr1 HOLogic.mk_disj (@{map 3} mk_case_disj xctrs xfs xss))); val ((raw_case, (_, raw_case_def)), (lthy, lthy_old)) = no_defs_lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> Local_Theory.define ((case_binding, NoSyn), ((Binding.concealed (Thm.def_binding case_binding), []), case_rhs)) ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val case_def = Morphism.thm phi raw_case_def; val case0 = Morphism.term phi raw_case; val casex = mk_case As B case0; val casexC = mk_case As C case0; val casexBool = mk_case As HOLogic.boolT case0; fun mk_uu_eq () = HOLogic.mk_eq (u, u); val exist_xs_u_eq_ctrs = map2 (fn xctr => fn xs => list_exists_free xs (HOLogic.mk_eq (u, xctr))) xctrs xss; val unique_disc_no_def = TrueI; (*arbitrary marker*) val alternate_disc_no_def = FalseE; (*arbitrary marker*) fun alternate_disc_lhs get_udisc k = HOLogic.mk_not (let val b = nth disc_bindings (k - 1) in if is_disc_binding_valid b then get_udisc b (k - 1) else nth exist_xs_u_eq_ctrs (k - 1) end); val no_discs_sels = not discs_sels andalso forall (forall Binding.is_empty) (raw_disc_bindings :: raw_sel_bindingss) andalso null sel_default_eqs; val (all_sels_distinct, discs, selss, disc_defs, sel_defs, sel_defss, lthy) = if no_discs_sels then (true, [], [], [], [], [], lthy) else let val all_sel_bindings = flat sel_bindingss; val num_all_sel_bindings = length all_sel_bindings; val uniq_sel_bindings = distinct Binding.eq_name all_sel_bindings; val all_sels_distinct = (length uniq_sel_bindings = num_all_sel_bindings); val sel_binding_index = if all_sels_distinct then 1 upto num_all_sel_bindings else map (fn b => find_index (curry Binding.eq_name b) uniq_sel_bindings) all_sel_bindings; val all_proto_sels = flat (@{map 3} (fn k => fn xs => map (pair k o pair xs)) ks xss xss); val sel_infos = AList.group (op =) (sel_binding_index ~~ all_proto_sels) |> sort (int_ord o apply2 fst) |> map snd |> curry (op ~~) uniq_sel_bindings; val sel_bindings = map fst sel_infos; val sel_defaults = if null sel_default_eqs then [] else let val sel_Ts = map (curry (op -->) fcT o fastype_of o snd o snd o hd o snd) sel_infos; val fake_lthy = fake_local_theory_for_sel_defaults (sel_bindings ~~ sel_Ts) no_defs_lthy; in map (extract_sel_default fake_lthy o prep_term fake_lthy) sel_default_eqs end; fun disc_free b = Free (Binding.name_of b, mk_pred1T fcT); fun disc_spec b exist_xs_u_eq_ctr = mk_Trueprop_eq (disc_free b $ u, exist_xs_u_eq_ctr); fun alternate_disc k = Term.lambda u (alternate_disc_lhs (K o rapp u o disc_free) (3 - k)); fun mk_sel_case_args b proto_sels T = @{map 3} (fn Const (c, _) => fn Ts => fn k => (case AList.lookup (op =) proto_sels k of NONE => (case filter (curry (op =) (c, Binding.name_of b) o fst) sel_defaults of [] => fold_rev (Term.lambda o curry Free Name.uu) Ts (mk_undefined T) | [(_, t)] => t | _ => error "Multiple default values for selector/constructor pair") | SOME (xs, x) => fold_rev Term.lambda xs x)) ctrs ctr_Tss ks; fun sel_spec b proto_sels = let val _ = (case duplicates (op =) (map fst proto_sels) of k :: _ => error ("Duplicate selector name " ^ quote (Binding.name_of b) ^ " for constructor " ^ quote (Syntax.string_of_term lthy (nth ctrs (k - 1)))) | [] => ()) val T = (case distinct (op =) (map (fastype_of o snd o snd) proto_sels) of [T] => T | T :: T' :: _ => error ("Inconsistent range type for selector " ^ quote (Binding.name_of b) ^ ": " ^ quote (Syntax.string_of_typ lthy T) ^ " vs. " ^ quote (Syntax.string_of_typ lthy T'))); in mk_Trueprop_eq (Free (Binding.name_of b, fcT --> T) $ u, Term.list_comb (mk_case As T case0, mk_sel_case_args b proto_sels T) $ u) end; fun unflat_selss xs = unflat_lookup Binding.eq_name sel_bindings xs sel_bindingss; val (((raw_discs, raw_disc_defs), (raw_sels, raw_sel_defs)), (lthy', lthy)) = lthy - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> apfst split_list o @{fold_map 3} (fn k => fn exist_xs_u_eq_ctr => fn b => if Binding.is_empty b then if n = 1 then pair (Term.lambda u (mk_uu_eq ()), unique_disc_no_def) else pair (alternate_disc k, alternate_disc_no_def) else if Binding.eq_name (b, equal_binding) then pair (Term.lambda u exist_xs_u_eq_ctr, refl) else Specification.definition (SOME (b, NONE, NoSyn)) [] [] ((Thm.def_binding b, []), disc_spec b exist_xs_u_eq_ctr) #>> apsnd snd) ks exist_xs_u_eq_ctrs disc_bindings ||>> apfst split_list o fold_map (fn (b, proto_sels) => Specification.definition (SOME (b, NONE, NoSyn)) [] [] ((Thm.def_binding b, []), sel_spec b proto_sels) #>> apsnd snd) sel_infos ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy lthy'; val disc_defs = map (Morphism.thm phi) raw_disc_defs; val sel_defs = map (Morphism.thm phi) raw_sel_defs; val sel_defss = unflat_selss sel_defs; val discs0 = map (Morphism.term phi) raw_discs; val selss0 = unflat_selss (map (Morphism.term phi) raw_sels); val discs = map (mk_disc_or_sel As) discs0; val selss = map (map (mk_disc_or_sel As)) selss0; in (all_sels_distinct, discs, selss, disc_defs, sel_defs, sel_defss, lthy') end; fun mk_imp_p Qs = Logic.list_implies (Qs, HOLogic.mk_Trueprop p); val exhaust_goal = let fun mk_prem xctr xs = fold_rev Logic.all xs (mk_imp_p [mk_Trueprop_eq (exh_y, xctr)]) in fold_rev Logic.all [p, exh_y] (mk_imp_p (map2 mk_prem xctrs xss)) end; val inject_goalss = let fun mk_goal _ _ [] [] = [] | mk_goal xctr yctr xs ys = [fold_rev Logic.all (xs @ ys) (mk_Trueprop_eq (HOLogic.mk_eq (xctr, yctr), Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_eq) xs ys)))]; in @{map 4} mk_goal xctrs yctrs xss yss end; val half_distinct_goalss = let fun mk_goal ((xs, xc), (xs', xc')) = fold_rev Logic.all (xs @ xs') (HOLogic.mk_Trueprop (HOLogic.mk_not (HOLogic.mk_eq (xc, xc')))); in map (map mk_goal) (mk_half_pairss (`I (xss ~~ xctrs))) end; val goalss = [exhaust_goal] :: inject_goalss @ half_distinct_goalss; fun after_qed ([exhaust_thm] :: thmss) lthy = let val ((((((((u, u'), (xss, xss')), fs), gs), h), v), p), _) = lthy |> add_bindings |> yield_singleton (apfst (op ~~) oo mk_Frees' fc_b_name) fcT ||>> mk_Freess' "x" ctr_Tss ||>> mk_Frees "f" case_Ts ||>> mk_Frees "g" case_Ts ||>> yield_singleton (mk_Frees "h") (B --> C) ||>> yield_singleton (mk_Frees (fc_b_name ^ "'")) fcT ||>> yield_singleton (mk_Frees "P") HOLogic.boolT; val xfs = map2 (curry Term.list_comb) fs xss; val xgs = map2 (curry Term.list_comb) gs xss; val fcase = Term.list_comb (casex, fs); val ufcase = fcase $ u; val vfcase = fcase $ v; val eta_fcase = Term.list_comb (casex, eta_fs); val eta_gcase = Term.list_comb (casex, eta_gs); val eta_ufcase = eta_fcase $ u; val eta_vgcase = eta_gcase $ v; fun mk_uu_eq () = HOLogic.mk_eq (u, u); val uv_eq = mk_Trueprop_eq (u, v); val ((inject_thms, inject_thmss), half_distinct_thmss) = chop n thmss |>> `flat; val rho_As = map (fn (T, U) => (dest_TVar T, Thm.ctyp_of lthy U)) (map Logic.varifyT_global As ~~ As); fun inst_thm t thm = Thm.instantiate' [] [SOME (Thm.cterm_of lthy t)] (Thm.instantiate (rho_As, []) (Drule.zero_var_indexes thm)); val uexhaust_thm = inst_thm u exhaust_thm; val exhaust_cases = map base_name_of_ctr ctrs; val other_half_distinct_thmss = map (map (fn thm => thm RS not_sym)) half_distinct_thmss; val (distinct_thms, (distinct_thmsss', distinct_thmsss)) = join_halves n half_distinct_thmss other_half_distinct_thmss ||> `transpose; val nchotomy_thm = let val goal = HOLogic.mk_Trueprop (HOLogic.mk_all (fst u', snd u', Library.foldr1 HOLogic.mk_disj exist_xs_u_eq_ctrs)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_nchotomy_tac ctxt n exhaust_thm) |> Thm.close_derivation \<^here> end; val case_thms = let val goals = @{map 3} (fn xctr => fn xf => fn xs => fold_rev Logic.all (fs @ xs) (mk_Trueprop_eq (fcase $ xctr, xf))) xctrs xfs xss; in @{map 4} (fn k => fn goal => fn injects => fn distinctss => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => mk_case_tac ctxt n k case_def injects distinctss) |> Thm.close_derivation \<^here>) ks goals inject_thmss distinct_thmsss end; val (case_cong_thm, case_cong_weak_thm) = let fun mk_prem xctr xs xf xg = fold_rev Logic.all xs (Logic.mk_implies (mk_Trueprop_eq (v, xctr), mk_Trueprop_eq (xf, xg))); val goal = Logic.list_implies (uv_eq :: @{map 4} mk_prem xctrs xss xfs xgs, mk_Trueprop_eq (eta_ufcase, eta_vgcase)); val weak_goal = Logic.mk_implies (uv_eq, mk_Trueprop_eq (ufcase, vfcase)); val vars = Variable.add_free_names lthy goal []; val weak_vars = Variable.add_free_names lthy weak_goal []; in (Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_case_cong_tac ctxt uexhaust_thm case_thms), Goal.prove_sorry lthy weak_vars [] weak_goal (fn {context = ctxt, prems = _} => etac ctxt arg_cong 1)) |> apply2 (Thm.close_derivation \<^here>) end; val split_lhs = q $ ufcase; fun mk_split_conjunct xctr xs f_xs = list_all_free xs (HOLogic.mk_imp (HOLogic.mk_eq (u, xctr), q $ f_xs)); fun mk_split_disjunct xctr xs f_xs = list_exists_free xs (HOLogic.mk_conj (HOLogic.mk_eq (u, xctr), HOLogic.mk_not (q $ f_xs))); fun mk_split_goal xctrs xss xfs = mk_Trueprop_eq (split_lhs, Library.foldr1 HOLogic.mk_conj (@{map 3} mk_split_conjunct xctrs xss xfs)); fun mk_split_asm_goal xctrs xss xfs = mk_Trueprop_eq (split_lhs, HOLogic.mk_not (Library.foldr1 HOLogic.mk_disj (@{map 3} mk_split_disjunct xctrs xss xfs))); fun prove_split selss goal = Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_split_tac ctxt uexhaust_thm case_thms selss inject_thmss distinct_thmsss)) |> Thm.close_derivation \<^here>; fun prove_split_asm asm_goal split_thm = Variable.add_free_names lthy asm_goal [] |> (fn vars => Goal.prove_sorry lthy vars [] asm_goal (fn {context = ctxt, ...} => mk_split_asm_tac ctxt split_thm)) |> Thm.close_derivation \<^here>; val (split_thm, split_asm_thm) = let val goal = mk_split_goal xctrs xss xfs; val asm_goal = mk_split_asm_goal xctrs xss xfs; val thm = prove_split (replicate n []) goal; val asm_thm = prove_split_asm asm_goal thm; in (thm, asm_thm) end; val (sel_defs, all_sel_thms, sel_thmss, nontriv_disc_defs, disc_thmss, nontriv_disc_thmss, discI_thms, nontriv_discI_thms, distinct_disc_thms, distinct_disc_thmsss, exhaust_disc_thms, exhaust_sel_thms, all_collapse_thms, safe_collapse_thms, expand_thms, split_sel_thms, split_sel_asm_thms, case_eq_if_thms, disc_eq_case_thms) = if no_discs_sels then ([], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], []) else let val udiscs = map (rapp u) discs; val uselss = map (map (rapp u)) selss; val usel_ctrs = map2 (curry Term.list_comb) ctrs uselss; val usel_fs = map2 (curry Term.list_comb) fs uselss; val vdiscs = map (rapp v) discs; val vselss = map (map (rapp v)) selss; fun make_sel_thm xs' case_thm sel_def = zero_var_indexes (Variable.gen_all lthy (Drule.rename_bvars' (map (SOME o fst) xs') (Drule.forall_intr_vars (case_thm RS (sel_def RS trans))))); val sel_thmss = @{map 3} (map oo make_sel_thm) xss' case_thms sel_defss; fun has_undefined_rhs thm = (case snd (HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of thm))) of Const (\<^const_name>\undefined\, _) => true | _ => false); val all_sel_thms = (if all_sels_distinct andalso null sel_default_eqs then flat sel_thmss else map_product (fn s => fn (xs', c) => make_sel_thm xs' c s) sel_defs (xss' ~~ case_thms)) |> filter_out has_undefined_rhs; fun mk_unique_disc_def () = let val m = the_single ms; val goal = mk_Trueprop_eq (mk_uu_eq (), the_single exist_xs_u_eq_ctrs); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_unique_disc_def_tac ctxt m uexhaust_thm) |> Thm.close_derivation \<^here> end; fun mk_alternate_disc_def k = let val goal = mk_Trueprop_eq (alternate_disc_lhs (K (nth udiscs)) (3 - k), nth exist_xs_u_eq_ctrs (k - 1)); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_alternate_disc_def_tac ctxt k (nth disc_defs (2 - k)) (nth distinct_thms (2 - k)) uexhaust_thm) |> Thm.close_derivation \<^here> end; val has_alternate_disc_def = exists (fn def => Thm.eq_thm_prop (def, alternate_disc_no_def)) disc_defs; val nontriv_disc_defs = disc_defs |> filter_out (member Thm.eq_thm_prop [unique_disc_no_def, alternate_disc_no_def, refl]); val disc_defs' = map2 (fn k => fn def => if Thm.eq_thm_prop (def, unique_disc_no_def) then mk_unique_disc_def () else if Thm.eq_thm_prop (def, alternate_disc_no_def) then mk_alternate_disc_def k else def) ks disc_defs; val discD_thms = map (fn def => def RS iffD1) disc_defs'; val discI_thms = map2 (fn m => fn def => funpow m (fn thm => exI RS thm) (def RS iffD2)) ms disc_defs'; val not_discI_thms = map2 (fn m => fn def => funpow m (fn thm => allI RS thm) (unfold_thms lthy @{thms not_ex} (def RS @{thm ssubst[of _ _ Not]}))) ms disc_defs'; val (disc_thmss', disc_thmss) = let fun mk_thm discI _ [] = refl RS discI | mk_thm _ not_discI [distinct] = distinct RS not_discI; fun mk_thms discI not_discI distinctss = map (mk_thm discI not_discI) distinctss; in @{map 3} mk_thms discI_thms not_discI_thms distinct_thmsss' |> `transpose end; val nontriv_disc_thmss = map2 (fn b => if is_disc_binding_valid b then I else K []) disc_bindings disc_thmss; fun is_discI_triv b = (n = 1 andalso Binding.is_empty b) orelse Binding.eq_name (b, equal_binding); val nontriv_discI_thms = flat (map2 (fn b => if is_discI_triv b then K [] else single) disc_bindings discI_thms); val (distinct_disc_thms, (distinct_disc_thmsss', distinct_disc_thmsss)) = let fun mk_goal [] = [] | mk_goal [((_, udisc), (_, udisc'))] = [Logic.all u (Logic.mk_implies (HOLogic.mk_Trueprop udisc, HOLogic.mk_Trueprop (HOLogic.mk_not udisc')))]; fun prove tac goal = Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => tac ctxt) |> Thm.close_derivation \<^here>; val half_pairss = mk_half_pairss (`I (ms ~~ discD_thms ~~ udiscs)); val half_goalss = map mk_goal half_pairss; val half_thmss = @{map 3} (fn [] => K (K []) | [goal] => fn [(((m, discD), _), _)] => fn disc_thm => [prove (fn ctxt => mk_half_distinct_disc_tac ctxt m discD disc_thm) goal]) half_goalss half_pairss (flat disc_thmss'); val other_half_goalss = map (mk_goal o map swap) half_pairss; val other_half_thmss = map2 (map2 (fn thm => prove (fn ctxt => mk_other_half_distinct_disc_tac ctxt thm))) half_thmss other_half_goalss; in join_halves n half_thmss other_half_thmss ||> `transpose |>> has_alternate_disc_def ? K [] end; val exhaust_disc_thm = let fun mk_prem udisc = mk_imp_p [HOLogic.mk_Trueprop udisc]; val goal = fold_rev Logic.all [p, u] (mk_imp_p (map mk_prem udiscs)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_exhaust_disc_tac ctxt n exhaust_thm discI_thms) |> Thm.close_derivation \<^here> end; val (safe_collapse_thms, all_collapse_thms) = let fun mk_goal m udisc usel_ctr = let val prem = HOLogic.mk_Trueprop udisc; val concl = mk_Trueprop_eq ((usel_ctr, u) |> m = 0 ? swap); in (prem aconv concl, Logic.all u (Logic.mk_implies (prem, concl))) end; val (trivs, goals) = @{map 3} mk_goal ms udiscs usel_ctrs |> split_list; val thms = @{map 5} (fn m => fn discD => fn sel_thms => fn triv => fn goal => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => mk_collapse_tac ctxt m discD sel_thms ORELSE HEADGOAL (assume_tac ctxt)) |> Thm.close_derivation \<^here> |> not triv ? perhaps (try (fn thm => refl RS thm))) ms discD_thms sel_thmss trivs goals; in (map_filter (fn (true, _) => NONE | (false, thm) => SOME thm) (trivs ~~ thms), thms) end; val swapped_all_collapse_thms = map2 (fn m => fn thm => if m = 0 then thm else thm RS sym) ms all_collapse_thms; val exhaust_sel_thm = let fun mk_prem usel_ctr = mk_imp_p [mk_Trueprop_eq (u, usel_ctr)]; val goal = fold_rev Logic.all [p, u] (mk_imp_p (map mk_prem usel_ctrs)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_exhaust_sel_tac ctxt n exhaust_disc_thm swapped_all_collapse_thms) |> Thm.close_derivation \<^here> end; val expand_thm = let fun mk_prems k udisc usels vdisc vsels = (if k = n then [] else [mk_Trueprop_eq (udisc, vdisc)]) @ (if null usels then [] else [Logic.list_implies (if n = 1 then [] else map HOLogic.mk_Trueprop [udisc, vdisc], HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_eq) usels vsels)))]); val goal = Library.foldr Logic.list_implies (@{map 5} mk_prems ks udiscs uselss vdiscs vselss, uv_eq); val uncollapse_thms = map2 (fn thm => fn [] => thm | _ => thm RS sym) all_collapse_thms uselss; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_expand_tac ctxt n ms (inst_thm u exhaust_disc_thm) (inst_thm v exhaust_disc_thm) uncollapse_thms distinct_disc_thmsss distinct_disc_thmsss') |> Thm.close_derivation \<^here> end; val (split_sel_thm, split_sel_asm_thm) = let val zss = map (K []) xss; val goal = mk_split_goal usel_ctrs zss usel_fs; val asm_goal = mk_split_asm_goal usel_ctrs zss usel_fs; val thm = prove_split sel_thmss goal; val asm_thm = prove_split_asm asm_goal thm; in (thm, asm_thm) end; val case_eq_if_thm = let val goal = mk_Trueprop_eq (ufcase, mk_IfN B udiscs usel_fs); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_case_eq_if_tac ctxt n uexhaust_thm case_thms disc_thmss' sel_thmss) |> Thm.close_derivation \<^here> end; val disc_eq_case_thms = let fun const_of_bool b = if b then \<^const>\True\ else \<^const>\False\; fun mk_case_args n = map_index (fn (k, argTs) => fold_rev Term.absdummy argTs (const_of_bool (n = k))) ctr_Tss; val goals = map_index (fn (n, udisc) => mk_Trueprop_eq (udisc, list_comb (casexBool, mk_case_args n) $ u)) udiscs; val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_disc_eq_case_tac ctxt (Thm.cterm_of ctxt u) exhaust_thm (flat nontriv_disc_thmss) distinct_thms case_thms) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end; in (sel_defs, all_sel_thms, sel_thmss, nontriv_disc_defs, disc_thmss, nontriv_disc_thmss, discI_thms, nontriv_discI_thms, distinct_disc_thms, distinct_disc_thmsss, [exhaust_disc_thm], [exhaust_sel_thm], all_collapse_thms, safe_collapse_thms, [expand_thm], [split_sel_thm], [split_sel_asm_thm], [case_eq_if_thm], disc_eq_case_thms) end; val case_distrib_thm = let val args = @{map 2} (fn f => fn argTs => let val (args, _) = mk_Frees "x" argTs lthy in fold_rev Term.lambda args (h $ list_comb (f, args)) end) fs ctr_Tss; val goal = mk_Trueprop_eq (h $ ufcase, list_comb (casexC, args) $ u); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_case_distrib_tac ctxt (Thm.cterm_of ctxt u) exhaust_thm case_thms) |> Thm.close_derivation \<^here> end; val exhaust_case_names_attr = Attrib.internal (K (Rule_Cases.case_names exhaust_cases)); val cases_type_attr = Attrib.internal (K (Induct.cases_type fcT_name)); val nontriv_disc_eq_thmss = map (map (fn th => th RS @{thm eq_False[THEN iffD2]} handle THM _ => th RS @{thm eq_True[THEN iffD2]})) nontriv_disc_thmss; val anonymous_notes = [(map (fn th => th RS notE) distinct_thms, safe_elim_attrs), (flat nontriv_disc_eq_thmss, nitpicksimp_attrs)] |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])])); val notes = [(caseN, case_thms, nitpicksimp_attrs @ simp_attrs), (case_congN, [case_cong_thm], []), (case_cong_weak_thmsN, [case_cong_weak_thm], cong_attrs), (case_distribN, [case_distrib_thm], []), (case_eq_ifN, case_eq_if_thms, []), (collapseN, safe_collapse_thms, if ms = [0] then [] else simp_attrs), (discN, flat nontriv_disc_thmss, simp_attrs), (disc_eq_caseN, disc_eq_case_thms, []), (discIN, nontriv_discI_thms, []), (distinctN, distinct_thms, simp_attrs @ inductsimp_attrs), (distinct_discN, distinct_disc_thms, dest_attrs), (exhaustN, [exhaust_thm], [exhaust_case_names_attr, cases_type_attr]), (exhaust_discN, exhaust_disc_thms, [exhaust_case_names_attr]), (exhaust_selN, exhaust_sel_thms, [exhaust_case_names_attr]), (expandN, expand_thms, []), (injectN, inject_thms, iff_attrs @ inductsimp_attrs), (nchotomyN, [nchotomy_thm], []), (selN, all_sel_thms, nitpicksimp_attrs @ simp_attrs), (splitN, [split_thm], []), (split_asmN, [split_asm_thm], []), (split_selN, split_sel_thms, []), (split_sel_asmN, split_sel_asm_thms, []), (split_selsN, split_sel_thms @ split_sel_asm_thms, []), (splitsN, [split_thm, split_asm_thm], [])] |> filter_out (null o #2) |> map (fn (thmN, thms, attrs) => ((qualify true (Binding.name thmN), attrs), [(thms, [])])); val (noted, lthy') = lthy |> Spec_Rules.add Binding.empty Spec_Rules.equational [casex] case_thms |> fold (uncurry (Spec_Rules.add Binding.empty Spec_Rules.equational)) (AList.group (eq_list (op aconv)) (map (`(single o lhs_head_of)) all_sel_thms)) |> fold (uncurry (Spec_Rules.add Binding.empty Spec_Rules.equational)) (filter_out (null o snd) (map single discs ~~ nontriv_disc_eq_thmss)) |> Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => Case_Translation.register (Morphism.term phi casex) (map (Morphism.term phi) ctrs)) |> plugins code_plugin ? (Code.declare_default_eqns (map (rpair true) (flat nontriv_disc_eq_thmss @ case_thms @ all_sel_thms)) #> Local_Theory.declaration {syntax = false, pervasive = false} (fn phi => Context.mapping (add_ctr_code fcT_name (map (Morphism.typ phi) As) (map (dest_Const o Morphism.term phi) ctrs) (Morphism.fact phi inject_thms) (Morphism.fact phi distinct_thms) (Morphism.fact phi case_thms)) I)) |> Local_Theory.notes (anonymous_notes @ notes) (* for "datatype_realizer.ML": *) |>> name_noted_thms fcT_name exhaustN; val ctr_sugar = {kind = kind, T = fcT, ctrs = ctrs, casex = casex, discs = discs, selss = selss, exhaust = exhaust_thm, nchotomy = nchotomy_thm, injects = inject_thms, distincts = distinct_thms, case_thms = case_thms, case_cong = case_cong_thm, case_cong_weak = case_cong_weak_thm, case_distribs = [case_distrib_thm], split = split_thm, split_asm = split_asm_thm, disc_defs = nontriv_disc_defs, disc_thmss = disc_thmss, discIs = discI_thms, disc_eq_cases = disc_eq_case_thms, sel_defs = sel_defs, sel_thmss = sel_thmss, distinct_discsss = distinct_disc_thmsss, exhaust_discs = exhaust_disc_thms, exhaust_sels = exhaust_sel_thms, collapses = all_collapse_thms, expands = expand_thms, split_sels = split_sel_thms, split_sel_asms = split_sel_asm_thms, case_eq_ifs = case_eq_if_thms} |> morph_ctr_sugar (substitute_noted_thm noted); in (ctr_sugar, lthy' |> register_ctr_sugar plugins ctr_sugar) end; in (goalss, after_qed, lthy) end; fun free_constructors kind tacss = (fn (goalss, after_qed, lthy) => map2 (map2 (Thm.close_derivation \<^here> oo Goal.prove_sorry lthy [] [])) goalss tacss |> (fn thms => after_qed thms lthy)) oo prepare_free_constructors kind (K I) (K I); fun free_constructors_cmd kind = (fn (goalss, after_qed, lthy) => Proof.theorem NONE (snd oo after_qed) (map (map (rpair [])) goalss) lthy) oo prepare_free_constructors kind Plugin_Name.make_filter Syntax.read_term; val parse_bound_term = Parse.binding --| \<^keyword>\:\ -- Parse.term; type ctr_options = Plugin_Name.filter * bool; type ctr_options_cmd = (Proof.context -> Plugin_Name.filter) * bool; val default_ctr_options : ctr_options = (Plugin_Name.default_filter, false); val default_ctr_options_cmd : ctr_options_cmd = (K Plugin_Name.default_filter, false); val parse_ctr_options = Scan.optional (\<^keyword>\(\ |-- Parse.list1 (Plugin_Name.parse_filter >> (apfst o K) || Parse.reserved "discs_sels" >> (apsnd o K o K true)) --| \<^keyword>\)\ >> (fn fs => fold I fs default_ctr_options_cmd)) default_ctr_options_cmd; fun parse_ctr_spec parse_ctr parse_arg = parse_opt_binding_colon -- parse_ctr -- Scan.repeat parse_arg; val parse_ctr_specs = Parse.enum1 "|" (parse_ctr_spec Parse.term Parse.binding); val parse_sel_default_eqs = Scan.optional (\<^keyword>\where\ |-- Parse.enum1 "|" Parse.prop) []; val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\free_constructors\ "register an existing freely generated type's constructors" (parse_ctr_options -- Parse.binding --| \<^keyword>\for\ -- parse_ctr_specs -- parse_sel_default_eqs >> free_constructors_cmd Unknown); (** external views **) (* document antiquotations *) local fun antiquote_setup binding co = Thy_Output.antiquotation_pretty_source_embedded binding ((Scan.ahead (Scan.lift Parse.not_eof) >> Token.pos_of) -- Args.type_name {proper = true, strict = true}) (fn ctxt => fn (pos, type_name) => let fun err () = error ("Bad " ^ Binding.name_of binding ^ ": " ^ quote type_name ^ Position.here pos); in (case ctr_sugar_of ctxt type_name of NONE => err () | SOME {kind, T = T0, ctrs = ctrs0, ...} => let val _ = if co = (kind = Codatatype) then () else err (); val T = Logic.unvarifyT_global T0; val ctrs = map Logic.unvarify_global ctrs0; val pretty_typ_bracket = Syntax.pretty_typ (Config.put pretty_priority 1001 ctxt); fun pretty_ctr ctr = Pretty.block (Pretty.breaks (Syntax.pretty_term ctxt ctr :: map pretty_typ_bracket (binder_types (fastype_of ctr)))); in Pretty.block (Pretty.keyword1 (Binding.name_of binding) :: Pretty.brk 1 :: Syntax.pretty_typ ctxt T :: Pretty.str " =" :: Pretty.brk 1 :: flat (separate [Pretty.brk 1, Pretty.str "| "] (map (single o pretty_ctr) ctrs))) end) end); in val _ = Theory.setup (antiquote_setup \<^binding>\datatype\ false #> antiquote_setup \<^binding>\codatatype\ true); end; (* theory export *) val _ = Export_Theory.setup_presentation (fn context => fn thy => let val parents = map (Data.get o Context.Theory) (Theory.parents_of thy); val datatypes = (Data.get (Context.Theory thy), []) |-> Symtab.fold (fn (name, (pos, {kind, T, ctrs, ...})) => if kind = Record orelse exists (fn tab => Symtab.defined tab name) parents then I else let val pos_properties = Thy_Info.adjust_pos_properties context pos; val typ = Logic.unvarifyT_global T; val constrs = map Logic.unvarify_global ctrs; val typargs = rev (fold Term.add_tfrees (Logic.mk_type typ :: constrs) []); val constructors = map (fn t => (t, Term.type_of t)) constrs; in cons (pos_properties, (name, (kind = Codatatype, (typargs, (typ, constructors))))) end); in if null datatypes then () else Export_Theory.export_body thy "datatypes" let open XML.Encode Term_XML.Encode in list (pair properties (pair string (pair bool (pair (list (pair string sort)) (pair typ (list (pair (term (Sign.consts_of thy)) typ))))))) datatypes end end); end; diff --git a/src/HOL/Tools/Lifting/lifting_def.ML b/src/HOL/Tools/Lifting/lifting_def.ML --- a/src/HOL/Tools/Lifting/lifting_def.ML +++ b/src/HOL/Tools/Lifting/lifting_def.ML @@ -1,676 +1,676 @@ (* Title: HOL/Tools/Lifting/lifting_def.ML Author: Ondrej Kuncar Definitions for constants on quotient types. *) signature LIFTING_DEF = sig datatype code_eq = UNKNOWN_EQ | NONE_EQ | ABS_EQ | REP_EQ type lift_def val rty_of_lift_def: lift_def -> typ val qty_of_lift_def: lift_def -> typ val rhs_of_lift_def: lift_def -> term val lift_const_of_lift_def: lift_def -> term val def_thm_of_lift_def: lift_def -> thm val rsp_thm_of_lift_def: lift_def -> thm val abs_eq_of_lift_def: lift_def -> thm val rep_eq_of_lift_def: lift_def -> thm option val code_eq_of_lift_def: lift_def -> code_eq val transfer_rules_of_lift_def: lift_def -> thm list val morph_lift_def: morphism -> lift_def -> lift_def val inst_of_lift_def: Proof.context -> typ -> lift_def -> lift_def val mk_lift_const_of_lift_def: typ -> lift_def -> term type config = { notes: bool } val map_config: (bool -> bool) -> config -> config val default_config: config val generate_parametric_transfer_rule: Proof.context -> thm -> thm -> thm val add_lift_def: config -> binding * mixfix -> typ -> term -> thm -> thm list -> local_theory -> lift_def * local_theory val prepare_lift_def: (binding * mixfix -> typ -> term -> thm -> thm list -> Proof.context -> lift_def * local_theory) -> binding * mixfix -> typ -> term -> thm list -> local_theory -> term option * (thm -> Proof.context -> lift_def * local_theory) val gen_lift_def: (binding * mixfix -> typ -> term -> thm -> thm list -> local_theory -> lift_def * local_theory) -> binding * mixfix -> typ -> term -> (Proof.context -> tactic) -> thm list -> local_theory -> lift_def * local_theory val lift_def: config -> binding * mixfix -> typ -> term -> (Proof.context -> tactic) -> thm list -> local_theory -> lift_def * local_theory val can_generate_code_cert: thm -> bool end structure Lifting_Def: LIFTING_DEF = struct open Lifting_Util infix 0 MRSL datatype code_eq = UNKNOWN_EQ | NONE_EQ | ABS_EQ | REP_EQ datatype lift_def = LIFT_DEF of { rty: typ, qty: typ, rhs: term, lift_const: term, def_thm: thm, rsp_thm: thm, abs_eq: thm, rep_eq: thm option, code_eq: code_eq, transfer_rules: thm list }; fun rep_lift_def (LIFT_DEF lift_def) = lift_def; val rty_of_lift_def = #rty o rep_lift_def; val qty_of_lift_def = #qty o rep_lift_def; val rhs_of_lift_def = #rhs o rep_lift_def; val lift_const_of_lift_def = #lift_const o rep_lift_def; val def_thm_of_lift_def = #def_thm o rep_lift_def; val rsp_thm_of_lift_def = #rsp_thm o rep_lift_def; val abs_eq_of_lift_def = #abs_eq o rep_lift_def; val rep_eq_of_lift_def = #rep_eq o rep_lift_def; val code_eq_of_lift_def = #code_eq o rep_lift_def; val transfer_rules_of_lift_def = #transfer_rules o rep_lift_def; fun mk_lift_def rty qty rhs lift_const def_thm rsp_thm abs_eq rep_eq code_eq transfer_rules = LIFT_DEF {rty = rty, qty = qty, rhs = rhs, lift_const = lift_const, def_thm = def_thm, rsp_thm = rsp_thm, abs_eq = abs_eq, rep_eq = rep_eq, code_eq = code_eq, transfer_rules = transfer_rules }; fun map_lift_def f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 (LIFT_DEF {rty = rty, qty = qty, rhs = rhs, lift_const = lift_const, def_thm = def_thm, rsp_thm = rsp_thm, abs_eq = abs_eq, rep_eq = rep_eq, code_eq = code_eq, transfer_rules = transfer_rules }) = LIFT_DEF {rty = f1 rty, qty = f2 qty, rhs = f3 rhs, lift_const = f4 lift_const, def_thm = f5 def_thm, rsp_thm = f6 rsp_thm, abs_eq = f7 abs_eq, rep_eq = f8 rep_eq, code_eq = f9 code_eq, transfer_rules = f10 transfer_rules } fun morph_lift_def phi = let val mtyp = Morphism.typ phi val mterm = Morphism.term phi val mthm = Morphism.thm phi in map_lift_def mtyp mtyp mterm mterm mthm mthm mthm (Option.map mthm) I (map mthm) end fun mk_inst_of_lift_def qty lift_def = Vartab.empty |> Type.raw_match (qty_of_lift_def lift_def, qty) fun mk_lift_const_of_lift_def qty lift_def = Envir.subst_term_types (mk_inst_of_lift_def qty lift_def) (lift_const_of_lift_def lift_def) fun inst_of_lift_def ctxt qty lift_def = let val instT = Vartab.fold (fn (a, (S, T)) => cons ((a, S), Thm.ctyp_of ctxt T)) (mk_inst_of_lift_def qty lift_def) [] val phi = Morphism.instantiate_morphism (instT, []) in morph_lift_def phi lift_def end (* Config *) type config = { notes: bool }; fun map_config f1 { notes = notes } = { notes = f1 notes } val default_config = { notes = true }; (* Reflexivity prover *) fun mono_eq_prover ctxt prop = let val refl_rules = Lifting_Info.get_reflexivity_rules ctxt val transfer_rules = Transfer.get_transfer_raw ctxt fun main_tac i = (REPEAT_ALL_NEW (DETERM o resolve_tac ctxt refl_rules) THEN_ALL_NEW (REPEAT_ALL_NEW (DETERM o resolve_tac ctxt transfer_rules))) i in SOME (Goal.prove ctxt [] [] prop (K (main_tac 1))) handle ERROR _ => NONE end fun try_prove_refl_rel ctxt rel = let fun mk_ge_eq x = let val T = fastype_of x in Const (\<^const_name>\less_eq\, T --> T --> HOLogic.boolT) $ (Const (\<^const_name>\HOL.eq\, T)) $ x end; val goal = HOLogic.mk_Trueprop (mk_ge_eq rel); in mono_eq_prover ctxt goal end; fun try_prove_reflexivity ctxt prop = let val cprop = Thm.cterm_of ctxt prop val rule = @{thm ge_eq_refl} val concl_pat = Drule.strip_imp_concl (Thm.cprop_of rule) val insts = Thm.first_order_match (concl_pat, cprop) val rule = Drule.instantiate_normalize insts rule val prop = hd (Thm.prems_of rule) in case mono_eq_prover ctxt prop of SOME thm => SOME (thm RS rule) | NONE => NONE end (* Generates a parametrized transfer rule. transfer_rule - of the form T t f parametric_transfer_rule - of the form par_R t' t Result: par_T t' f, after substituing op= for relations in par_R that relate a type constructor to the same type constructor, it is a merge of (par_R' OO T) t' f using Lifting_Term.merge_transfer_relations *) fun generate_parametric_transfer_rule ctxt0 transfer_rule parametric_transfer_rule = let fun preprocess ctxt thm = let val tm = (strip_args 2 o HOLogic.dest_Trueprop o Thm.concl_of) thm; val param_rel = (snd o dest_comb o fst o dest_comb) tm; val free_vars = Term.add_vars param_rel []; fun make_subst (xi, typ) subst = let val [rty, rty'] = binder_types typ in if Term.is_TVar rty andalso Term.is_Type rty' then (xi, Thm.cterm_of ctxt (HOLogic.eq_const rty')) :: subst else subst end; val inst_thm = infer_instantiate ctxt (fold make_subst free_vars []) thm; in Conv.fconv_rule ((Conv.concl_conv (Thm.nprems_of inst_thm) o HOLogic.Trueprop_conv o Conv.fun2_conv o Conv.arg1_conv) (Raw_Simplifier.rewrite ctxt false (Transfer.get_sym_relator_eq ctxt))) inst_thm end fun inst_relcomppI ctxt ant1 ant2 = let val t1 = (HOLogic.dest_Trueprop o Thm.concl_of) ant1 val t2 = (HOLogic.dest_Trueprop o Thm.prop_of) ant2 val fun1 = Thm.cterm_of ctxt (strip_args 2 t1) val args1 = map (Thm.cterm_of ctxt) (get_args 2 t1) val fun2 = Thm.cterm_of ctxt (strip_args 2 t2) val args2 = map (Thm.cterm_of ctxt) (get_args 1 t2) val relcomppI = Drule.incr_indexes2 ant1 ant2 @{thm relcomppI} val vars = map #1 (rev (Term.add_vars (Thm.prop_of relcomppI) [])) in infer_instantiate ctxt (vars ~~ ([fun1] @ args1 @ [fun2] @ args2)) relcomppI end fun zip_transfer_rules ctxt thm = let fun mk_POS ty = Const (\<^const_name>\POS\, ty --> ty --> HOLogic.boolT) val rel = (Thm.dest_fun2 o Thm.dest_arg o Thm.cprop_of) thm val typ = Thm.typ_of_cterm rel val POS_const = Thm.cterm_of ctxt (mk_POS typ) val var = Thm.cterm_of ctxt (Var (("X", Thm.maxidx_of_cterm rel + 1), typ)) val goal = Thm.apply (Thm.cterm_of ctxt HOLogic.Trueprop) (Thm.apply (Thm.apply POS_const rel) var) in [Lifting_Term.merge_transfer_relations ctxt goal, thm] MRSL @{thm POS_apply} end val thm = inst_relcomppI ctxt0 parametric_transfer_rule transfer_rule OF [parametric_transfer_rule, transfer_rule] val preprocessed_thm = preprocess ctxt0 thm val (fixed_thm, ctxt1) = ctxt0 |> yield_singleton (apfst snd oo Variable.import true) preprocessed_thm val assms = cprems_of fixed_thm val add_transfer_rule = Thm.attribute_declaration Transfer.transfer_add val (prems, ctxt2) = ctxt1 |> fold_map Thm.assume_hyps assms val ctxt3 = ctxt2 |> Context.proof_map (fold add_transfer_rule prems) val zipped_thm = fixed_thm |> undisch_all |> zip_transfer_rules ctxt3 |> implies_intr_list assms |> singleton (Variable.export ctxt3 ctxt0) in zipped_thm end fun print_generate_transfer_info msg = let val error_msg = cat_lines ["Generation of a parametric transfer rule failed.", (Pretty.string_of (Pretty.block [Pretty.str "Reason:", Pretty.brk 2, msg]))] in error error_msg end fun map_ter _ x [] = x | map_ter f _ xs = map f xs fun generate_transfer_rules lthy quot_thm rsp_thm def_thm par_thms = let val transfer_rule = ([quot_thm, rsp_thm, def_thm] MRSL @{thm Quotient_to_transfer}) |> Lifting_Term.parametrize_transfer_rule lthy in (map_ter (generate_parametric_transfer_rule lthy transfer_rule) [transfer_rule] par_thms handle Lifting_Term.MERGE_TRANSFER_REL msg => (print_generate_transfer_info msg; [transfer_rule])) end (* Generation of the code certificate from the rsp theorem *) fun get_body_types (Type ("fun", [_, U]), Type ("fun", [_, V])) = get_body_types (U, V) | get_body_types (U, V) = (U, V) fun get_binder_types (Type ("fun", [T, U]), Type ("fun", [V, W])) = (T, V) :: get_binder_types (U, W) | get_binder_types _ = [] fun get_binder_types_by_rel (Const (\<^const_name>\rel_fun\, _) $ _ $ S) (Type ("fun", [T, U]), Type ("fun", [V, W])) = (T, V) :: get_binder_types_by_rel S (U, W) | get_binder_types_by_rel _ _ = [] fun get_body_type_by_rel (Const (\<^const_name>\rel_fun\, _) $ _ $ S) (Type ("fun", [_, U]), Type ("fun", [_, V])) = get_body_type_by_rel S (U, V) | get_body_type_by_rel _ (U, V) = (U, V) fun get_binder_rels (Const (\<^const_name>\rel_fun\, _) $ R $ S) = R :: get_binder_rels S | get_binder_rels _ = [] fun force_rty_type ctxt rty rhs = let val thy = Proof_Context.theory_of ctxt val rhs_schematic = singleton (Variable.polymorphic ctxt) rhs val rty_schematic = fastype_of rhs_schematic val match = Sign.typ_match thy (rty_schematic, rty) Vartab.empty in Envir.subst_term_types match rhs_schematic end fun unabs_def ctxt def = let val (_, rhs) = Thm.dest_equals (Thm.cprop_of def) fun dest_abs (Abs (var_name, T, _)) = (var_name, T) | dest_abs tm = raise TERM("get_abs_var",[tm]) val (var_name, T) = dest_abs (Thm.term_of rhs) val (new_var_names, ctxt') = Variable.variant_fixes [var_name] ctxt val refl_thm = Thm.reflexive (Thm.cterm_of ctxt' (Free (hd new_var_names, T))) in Thm.combination def refl_thm |> singleton (Proof_Context.export ctxt' ctxt) end fun unabs_all_def ctxt def = let val (_, rhs) = Thm.dest_equals (Thm.cprop_of def) val xs = strip_abs_vars (Thm.term_of rhs) in fold (K (unabs_def ctxt)) xs def end val map_fun_unfolded = @{thm map_fun_def[abs_def]} |> unabs_def \<^context> |> unabs_def \<^context> |> Local_Defs.unfold0 \<^context> [@{thm comp_def}] fun unfold_fun_maps ctm = let fun unfold_conv ctm = case (Thm.term_of ctm) of Const (\<^const_name>\map_fun\, _) $ _ $ _ => (Conv.arg_conv unfold_conv then_conv Conv.rewr_conv map_fun_unfolded) ctm | _ => Conv.all_conv ctm in (Conv.fun_conv unfold_conv) ctm end fun unfold_fun_maps_beta ctm = let val try_beta_conv = Conv.try_conv (Thm.beta_conversion false) in (unfold_fun_maps then_conv try_beta_conv) ctm end fun prove_rel ctxt rsp_thm (rty, qty) = let val ty_args = get_binder_types (rty, qty) fun disch_arg args_ty thm = let val quot_thm = Lifting_Term.prove_quot_thm ctxt args_ty in [quot_thm, thm] MRSL @{thm apply_rsp''} end in fold disch_arg ty_args rsp_thm end exception CODE_CERT_GEN of string fun simplify_code_eq ctxt def_thm = Local_Defs.unfold0 ctxt [@{thm o_apply}, @{thm map_fun_def}, @{thm id_apply}] def_thm (* quot_thm - quotient theorem (Quotient R Abs Rep T). returns: whether the Lifting package is capable to generate code for the abstract type represented by quot_thm *) fun can_generate_code_cert quot_thm = case quot_thm_rel quot_thm of Const (\<^const_name>\HOL.eq\, _) => true | Const (\<^const_name>\eq_onp\, _) $ _ => true | _ => false fun generate_rep_eq ctxt def_thm rsp_thm (rty, qty) = let val unfolded_def = Conv.fconv_rule (Conv.arg_conv unfold_fun_maps_beta) def_thm val unabs_def = unabs_all_def ctxt unfolded_def in if body_type rty = body_type qty then SOME (simplify_code_eq ctxt (HOLogic.mk_obj_eq unabs_def)) else let val quot_thm = Lifting_Term.prove_quot_thm ctxt (get_body_types (rty, qty)) val rel_fun = prove_rel ctxt rsp_thm (rty, qty) val rep_abs_thm = [quot_thm, rel_fun] MRSL @{thm Quotient_rep_abs_eq} in case mono_eq_prover ctxt (hd (Thm.prems_of rep_abs_thm)) of SOME mono_eq_thm => let val rep_abs_eq = mono_eq_thm RS rep_abs_thm val rep = Thm.cterm_of ctxt (quot_thm_rep quot_thm) val rep_refl = HOLogic.mk_obj_eq (Thm.reflexive rep) val repped_eq = [rep_refl, HOLogic.mk_obj_eq unabs_def] MRSL @{thm cong} val code_cert = [repped_eq, rep_abs_eq] MRSL trans in SOME (simplify_code_eq ctxt code_cert) end | NONE => NONE end end fun generate_abs_eq ctxt def_thm rsp_thm quot_thm = let val abs_eq_with_assms = let val (rty, qty) = quot_thm_rty_qty quot_thm val rel = quot_thm_rel quot_thm val ty_args = get_binder_types_by_rel rel (rty, qty) val body_type = get_body_type_by_rel rel (rty, qty) val quot_ret_thm = Lifting_Term.prove_quot_thm ctxt body_type val rep_abs_folded_unmapped_thm = let val rep_id = [quot_thm, def_thm] MRSL @{thm Quotient_Rep_eq} val ctm = Thm.dest_equals_lhs (Thm.cprop_of rep_id) val unfolded_maps_eq = unfold_fun_maps ctm val t1 = [quot_thm, def_thm, rsp_thm] MRSL @{thm Quotient_rep_abs_fold_unmap} val prems_pat = (hd o Drule.cprems_of) t1 val insts = Thm.first_order_match (prems_pat, Thm.cprop_of unfolded_maps_eq) in unfolded_maps_eq RS (Drule.instantiate_normalize insts t1) end in rep_abs_folded_unmapped_thm |> fold (fn _ => fn thm => thm RS @{thm rel_funD2}) ty_args |> (fn x => x RS (@{thm Quotient_rel_abs2} OF [quot_ret_thm])) end val prem_rels = get_binder_rels (quot_thm_rel quot_thm); val proved_assms = prem_rels |> map (try_prove_refl_rel ctxt) |> map_index (apfst (fn x => x + 1)) |> filter (is_some o snd) |> map (apsnd the) |> map (apsnd (fn assm => assm RS @{thm ge_eq_refl})) val abs_eq = fold_rev (fn (i, assm) => fn thm => assm RSN (i, thm)) proved_assms abs_eq_with_assms in simplify_code_eq ctxt abs_eq end fun register_code_eq_thy abs_eq_thm opt_rep_eq_thm (rty, qty) thy = let fun no_abstr (t $ u) = no_abstr t andalso no_abstr u | no_abstr (Abs (_, _, t)) = no_abstr t | no_abstr (Const (name, _)) = not (Code.is_abstr thy name) | no_abstr _ = true fun is_valid_eq eqn = can (Code.assert_eqn thy) (mk_meta_eq eqn, true) andalso no_abstr (Thm.prop_of eqn) fun is_valid_abs_eq abs_eq = can (Code.assert_abs_eqn thy NONE) (mk_meta_eq abs_eq) in if is_valid_eq abs_eq_thm then (ABS_EQ, Code.declare_default_eqns_global [(abs_eq_thm, true)] thy) else let val (rty_body, qty_body) = get_body_types (rty, qty) in if rty_body = qty_body then (REP_EQ, Code.declare_default_eqns_global [(the opt_rep_eq_thm, true)] thy) else if is_some opt_rep_eq_thm andalso is_valid_abs_eq (the opt_rep_eq_thm) then (REP_EQ, Code.declare_abstract_eqn_global (the opt_rep_eq_thm) thy) else (NONE_EQ, thy) end end local fun no_no_code ctxt (rty, qty) = if same_type_constrs (rty, qty) then forall (no_no_code ctxt) (Targs rty ~~ Targs qty) else if Term.is_Type qty then if Lifting_Info.is_no_code_type ctxt (Tname qty) then false else let val (rty', rtyq) = Lifting_Term.instantiate_rtys ctxt (rty, qty) val (rty's, rtyqs) = (Targs rty', Targs rtyq) in forall (no_no_code ctxt) (rty's ~~ rtyqs) end else true fun encode_code_eq ctxt abs_eq opt_rep_eq (rty, qty) = let fun mk_type typ = typ |> Logic.mk_type |> Thm.cterm_of ctxt |> Drule.mk_term in Conjunction.intr_balanced [abs_eq, (the_default TrueI opt_rep_eq), mk_type rty, mk_type qty] end exception DECODE fun decode_code_eq thm = if Thm.nprems_of thm > 0 then raise DECODE else let val [abs_eq, rep_eq, rty, qty] = Conjunction.elim_balanced 4 thm val opt_rep_eq = if Thm.eq_thm_prop (rep_eq, TrueI) then NONE else SOME rep_eq fun dest_type typ = typ |> Drule.dest_term |> Thm.term_of |> Logic.dest_type in (abs_eq, opt_rep_eq, (dest_type rty, dest_type qty)) end structure Data = Generic_Data ( type T = code_eq option val empty = NONE val extend = I fun merge _ = NONE ); fun register_encoded_code_eq thm thy = let val (abs_eq_thm, opt_rep_eq_thm, (rty, qty)) = decode_code_eq thm val (code_eq, thy) = register_code_eq_thy abs_eq_thm opt_rep_eq_thm (rty, qty) thy in Context.theory_map (Data.put (SOME code_eq)) thy end handle DECODE => thy val register_code_eq_attribute = Thm.declaration_attribute (fn thm => Context.mapping (register_encoded_code_eq thm) I) val register_code_eq_attrib = Attrib.internal (K register_code_eq_attribute) in fun register_code_eq abs_eq_thm opt_rep_eq_thm (rty, qty) lthy = let val encoded_code_eq = encode_code_eq lthy abs_eq_thm opt_rep_eq_thm (rty, qty) in if no_no_code lthy (rty, qty) then let val lthy' = lthy |> (#2 oo Local_Theory.note) ((Binding.empty, [register_code_eq_attrib]), [encoded_code_eq]) val opt_code_eq = Data.get (Context.Theory (Proof_Context.theory_of lthy')) val code_eq = if is_some opt_code_eq then the opt_code_eq else UNKNOWN_EQ (* UNKNOWN_EQ means that we are in a locale and we do not know which code equation is going to be used. This is going to be resolved at the point when an interpretation of the locale is executed. *) val lthy'' = lthy' |> Local_Theory.declaration {syntax = false, pervasive = true} (K (Data.put NONE)) in (code_eq, lthy'') end else (NONE_EQ, lthy) end end (* Defines an operation on an abstract type in terms of a corresponding operation on a representation type. var - a binding and a mixfix of the new constant being defined qty - an abstract type of the new constant rhs - a term representing the new constant on the raw level rsp_thm - a respectfulness theorem in the internal tagged form (like '(R ===> R ===> R) f f'), i.e. "(Lifting_Term.equiv_relation (fastype_of rhs, qty)) $ rhs $ rhs" par_thms - a parametricity theorem for rhs *) fun add_lift_def (config: config) (binding, mx) qty rhs rsp_thm par_thms lthy0 = let val rty = fastype_of rhs val quot_thm = Lifting_Term.prove_quot_thm lthy0 (rty, qty) val absrep_trm = quot_thm_abs quot_thm val rty_forced = (domain_type o fastype_of) absrep_trm val forced_rhs = force_rty_type lthy0 rty_forced rhs val lhs = Free (Binding.name_of binding, qty) val prop = Logic.mk_equals (lhs, absrep_trm $ forced_rhs) val (_, prop') = Local_Defs.cert_def lthy0 (K []) prop val (_, newrhs) = Local_Defs.abs_def prop' val var = ((#notes config = false ? Binding.concealed) binding, mx) val def_name = Thm.make_def_binding (#notes config) (#1 var) val ((lift_const, (_ , def_thm)), lthy1) = lthy0 |> Local_Theory.define (var, ((def_name, []), newrhs)) val transfer_rules = generate_transfer_rules lthy1 quot_thm rsp_thm def_thm par_thms val abs_eq_thm = generate_abs_eq lthy1 def_thm rsp_thm quot_thm val opt_rep_eq_thm = generate_rep_eq lthy1 def_thm rsp_thm (rty_forced, qty) fun notes names = let val lhs_name = Binding.reset_pos (#1 var) val rsp_thmN = Binding.qualify_name true lhs_name "rsp" val abs_eq_thmN = Binding.qualify_name true lhs_name "abs_eq" val rep_eq_thmN = Binding.qualify_name true lhs_name "rep_eq" val transfer_ruleN = Binding.qualify_name true lhs_name "transfer" val notes = [(rsp_thmN, [], [rsp_thm]), (transfer_ruleN, @{attributes [transfer_rule]}, transfer_rules), (abs_eq_thmN, [], [abs_eq_thm])] @ (case opt_rep_eq_thm of SOME rep_eq_thm => [(rep_eq_thmN, [], [rep_eq_thm])] | NONE => []) in if names then map (fn (name, attrs, thms) => ((name, []), [(thms, attrs)])) notes else map_filter (fn (_, attrs, thms) => if null attrs then NONE else SOME (Binding.empty_atts, [(thms, attrs)])) notes end val (code_eq, lthy2) = lthy1 |> register_code_eq abs_eq_thm opt_rep_eq_thm (rty_forced, qty) val lift_def = mk_lift_def rty_forced qty newrhs lift_const def_thm rsp_thm abs_eq_thm opt_rep_eq_thm code_eq transfer_rules in lthy2 - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> Local_Theory.notes (notes (#notes config)) |> snd |> `(fn lthy => morph_lift_def (Local_Theory.target_morphism lthy) lift_def) ||> Local_Theory.close_target end (* This is not very cheap way of getting the rules but we have only few active liftings in the current setting *) fun get_cr_pcr_eqs ctxt = let fun collect (data : Lifting_Info.quotient) l = if is_some (#pcr_info data) then ((Thm.symmetric o safe_mk_meta_eq o Thm.transfer' ctxt o #pcr_cr_eq o the o #pcr_info) data :: l) else l val table = Lifting_Info.get_quotients ctxt in Symtab.fold (fn (_, data) => fn l => collect data l) table [] end fun prepare_lift_def add_lift_def var qty rhs par_thms ctxt = let val rsp_rel = Lifting_Term.equiv_relation ctxt (fastype_of rhs, qty) val rty_forced = (domain_type o fastype_of) rsp_rel; val forced_rhs = force_rty_type ctxt rty_forced rhs; val cr_to_pcr_conv = HOLogic.Trueprop_conv (Conv.fun2_conv (Raw_Simplifier.rewrite ctxt false (get_cr_pcr_eqs ctxt))) val (prsp_tm, rsp_prsp_eq) = HOLogic.mk_Trueprop (rsp_rel $ forced_rhs $ forced_rhs) |> Thm.cterm_of ctxt |> cr_to_pcr_conv |> `Thm.concl_of |>> Logic.dest_equals |>> snd val to_rsp = rsp_prsp_eq RS Drule.equal_elim_rule2 val opt_proven_rsp_thm = try_prove_reflexivity ctxt prsp_tm fun after_qed internal_rsp_thm = add_lift_def var qty rhs (internal_rsp_thm RS to_rsp) par_thms in case opt_proven_rsp_thm of SOME thm => (NONE, K (after_qed thm)) | NONE => (SOME prsp_tm, after_qed) end fun gen_lift_def add_lift_def var qty rhs tac par_thms lthy = let val (goal, after_qed) = prepare_lift_def add_lift_def var qty rhs par_thms lthy in case goal of SOME goal => let val rsp_thm = Goal.prove_sorry lthy [] [] goal (tac o #context) |> Thm.close_derivation \<^here> in after_qed rsp_thm lthy end | NONE => after_qed Drule.dummy_thm lthy end val lift_def = gen_lift_def o add_lift_def end (* structure *) diff --git a/src/HOL/Tools/Lifting/lifting_def_code_dt.ML b/src/HOL/Tools/Lifting/lifting_def_code_dt.ML --- a/src/HOL/Tools/Lifting/lifting_def_code_dt.ML +++ b/src/HOL/Tools/Lifting/lifting_def_code_dt.ML @@ -1,830 +1,830 @@ (* Title: HOL/Tools/Lifting/lifting_def_code_dt.ML Author: Ondrej Kuncar Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype; lift_definition command *) signature LIFTING_DEF_CODE_DT = sig type rep_isom_data val isom_of_rep_isom_data: rep_isom_data -> term val transfer_of_rep_isom_data: rep_isom_data -> thm val bundle_name_of_rep_isom_data: rep_isom_data -> string val pointer_of_rep_isom_data: rep_isom_data -> string type code_dt val rty_of_code_dt: code_dt -> typ val qty_of_code_dt: code_dt -> typ val wit_of_code_dt: code_dt -> term val wit_thm_of_code_dt: code_dt -> thm val rep_isom_data_of_code_dt: code_dt -> rep_isom_data option val morph_code_dt: morphism -> code_dt -> code_dt val mk_witness_of_code_dt: typ -> code_dt -> term val mk_rep_isom_of_code_dt: typ -> code_dt -> term option val code_dt_of: Proof.context -> typ * typ -> code_dt option val code_dt_of_global: theory -> typ * typ -> code_dt option val all_code_dt_of: Proof.context -> code_dt list val all_code_dt_of_global: theory -> code_dt list type config_code_dt = { code_dt: bool, lift_config: Lifting_Def.config } val default_config_code_dt: config_code_dt val add_lift_def_code_dt: config_code_dt -> binding * mixfix -> typ -> term -> thm -> thm list -> local_theory -> Lifting_Def.lift_def * local_theory val lift_def_code_dt: config_code_dt -> binding * mixfix -> typ -> term -> (Proof.context -> tactic) -> thm list -> local_theory -> Lifting_Def.lift_def * local_theory val lift_def_cmd: string list * (binding * string option * mixfix) * string * (Facts.ref * Token.src list) list -> local_theory -> Proof.state end structure Lifting_Def_Code_Dt: LIFTING_DEF_CODE_DT = struct open Ctr_Sugar_Util BNF_Util BNF_FP_Util BNF_FP_Def_Sugar Lifting_Def Lifting_Util infix 0 MRSL (** data structures **) (* all type variables in qty are in rty *) datatype rep_isom_data = REP_ISOM of { isom: term, transfer: thm, bundle_name: string, pointer: string } fun isom_of_rep_isom_data (REP_ISOM rep_isom) = #isom rep_isom; fun transfer_of_rep_isom_data (REP_ISOM rep_isom) = #transfer rep_isom; fun bundle_name_of_rep_isom_data (REP_ISOM rep_isom) = #bundle_name rep_isom; fun pointer_of_rep_isom_data (REP_ISOM rep_isom) = #pointer rep_isom; datatype code_dt = CODE_DT of { rty: typ, qty: typ, wit: term, wit_thm: thm, rep_isom_data: rep_isom_data option }; fun rty_of_code_dt (CODE_DT code_dt) = #rty code_dt; fun qty_of_code_dt (CODE_DT code_dt) = #qty code_dt; fun wit_of_code_dt (CODE_DT code_dt) = #wit code_dt; fun wit_thm_of_code_dt (CODE_DT code_dt) = #wit_thm code_dt; fun rep_isom_data_of_code_dt (CODE_DT code_dt) = #rep_isom_data code_dt; fun ty_alpha_equiv (T, U) = Type.raw_instance (T, U) andalso Type.raw_instance (U, T); fun code_dt_eq c = (ty_alpha_equiv o apply2 rty_of_code_dt) c andalso (ty_alpha_equiv o apply2 qty_of_code_dt) c; fun term_of_code_dt code_dt = code_dt |> `rty_of_code_dt ||> qty_of_code_dt |> HOLogic.mk_prodT |> Net.encode_type |> single; (* modulo renaming, typ must contain TVars *) fun is_code_dt_of_type (rty, qty) code_dt = code_dt |> `rty_of_code_dt ||> qty_of_code_dt |> HOLogic.mk_prodT |> curry ty_alpha_equiv (HOLogic.mk_prodT (rty, qty)); fun mk_rep_isom_data isom transfer bundle_name pointer = REP_ISOM { isom = isom, transfer = transfer, bundle_name = bundle_name, pointer = pointer} fun mk_code_dt rty qty wit wit_thm rep_isom_data = CODE_DT { rty = rty, qty = qty, wit = wit, wit_thm = wit_thm, rep_isom_data = rep_isom_data }; fun map_rep_isom_data f1 f2 f3 f4 (REP_ISOM { isom = isom, transfer = transfer, bundle_name = bundle_name, pointer = pointer }) = REP_ISOM { isom = f1 isom, transfer = f2 transfer, bundle_name = f3 bundle_name, pointer = f4 pointer }; fun map_code_dt f1 f2 f3 f4 f5 f6 f7 f8 (CODE_DT {rty = rty, qty = qty, wit = wit, wit_thm = wit_thm, rep_isom_data = rep_isom_data}) = CODE_DT {rty = f1 rty, qty = f2 qty, wit = f3 wit, wit_thm = f4 wit_thm, rep_isom_data = Option.map (map_rep_isom_data f5 f6 f7 f8) rep_isom_data}; fun update_rep_isom isom transfer binding pointer i = mk_code_dt (rty_of_code_dt i) (qty_of_code_dt i) (wit_of_code_dt i) (wit_thm_of_code_dt i) (SOME (mk_rep_isom_data isom transfer binding pointer)) fun morph_code_dt phi = let val mty = Morphism.typ phi val mterm = Morphism.term phi val mthm = Morphism.thm phi in map_code_dt mty mty mterm mthm mterm mthm I I end val transfer_code_dt = morph_code_dt o Morphism.transfer_morphism; structure Data = Generic_Data ( type T = code_dt Item_Net.T val empty = Item_Net.init code_dt_eq term_of_code_dt val extend = I val merge = Item_Net.merge ); fun code_dt_of_generic context (rty, qty) = let val typ = HOLogic.mk_prodT (rty, qty) val prefiltred = Item_Net.retrieve_matching (Data.get context) (Net.encode_type typ) in prefiltred |> filter (is_code_dt_of_type (rty, qty)) |> map (transfer_code_dt (Context.theory_of context)) |> find_first (fn _ => true) end; fun code_dt_of ctxt (rty, qty) = let val sch_rty = Logic.type_map (singleton (Variable.polymorphic ctxt)) rty val sch_qty = Logic.type_map (singleton (Variable.polymorphic ctxt)) qty in code_dt_of_generic (Context.Proof ctxt) (sch_rty, sch_qty) end; fun code_dt_of_global thy (rty, qty) = let val sch_rty = Logic.varifyT_global rty val sch_qty = Logic.varifyT_global qty in code_dt_of_generic (Context.Theory thy) (sch_rty, sch_qty) end; fun all_code_dt_of_generic context = Item_Net.content (Data.get context) |> map (transfer_code_dt (Context.theory_of context)); val all_code_dt_of = all_code_dt_of_generic o Context.Proof; val all_code_dt_of_global = all_code_dt_of_generic o Context.Theory; fun update_code_dt code_dt = - Local_Theory.open_target #> snd + Local_Theory.open_target #> Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => Data.map (Item_Net.update (morph_code_dt phi code_dt))) #> Local_Theory.close_target fun mk_match_of_code_dt qty code_dt = Vartab.empty |> Type.raw_match (qty_of_code_dt code_dt, qty) |> Vartab.dest |> map (fn (x, (S, T)) => (TVar (x, S), T)); fun mk_witness_of_code_dt qty code_dt = Term.subst_atomic_types (mk_match_of_code_dt qty code_dt) (wit_of_code_dt code_dt) fun mk_rep_isom_of_code_dt qty code_dt = Option.map (isom_of_rep_isom_data #> Term.subst_atomic_types (mk_match_of_code_dt qty code_dt)) (rep_isom_data_of_code_dt code_dt) (** unique name for a type **) fun var_name name sort = if sort = \<^sort>\{type}\ orelse sort = [] then ["x" ^ name] else "x" ^ name :: "x_" :: sort @ ["x_"]; fun concat_Tnames (Type (name, ts)) = name :: maps concat_Tnames ts | concat_Tnames (TFree (name, sort)) = var_name name sort | concat_Tnames (TVar ((name, _), sort)) = var_name name sort; fun unique_Tname (rty, qty) = let val Tnames = map Long_Name.base_name (concat_Tnames rty @ ["x_x"] @ concat_Tnames qty); in fold (Binding.qualify false) (tl Tnames) (Binding.name (hd Tnames)) end; (** witnesses **) fun mk_undefined T = Const (\<^const_name>\undefined\, T); fun mk_witness quot_thm = let val wit_thm = quot_thm RS @{thm type_definition_Quotient_not_empty_witness} val wit = quot_thm_rep quot_thm $ mk_undefined (quot_thm_rty_qty quot_thm |> snd) in (wit, wit_thm) end (** config **) type config_code_dt = { code_dt: bool, lift_config: config } val default_config_code_dt = { code_dt = false, lift_config = default_config } (** Main code **) val ld_no_notes = { notes = false } fun comp_lift_error _ _ = error "Composition of abstract types has not been implemented yet." fun lift qty (quot_thm, (lthy, rel_eq_onps)) = let val quot_thm = Lifting_Term.force_qty_type lthy qty quot_thm val (rty, qty) = quot_thm_rty_qty quot_thm; in if is_none (code_dt_of lthy (rty, qty)) then let val (wit, wit_thm) = (mk_witness quot_thm handle THM _ => error ("code_dt: " ^ quote (Tname qty) ^ " was not defined as a subtype.")) val code_dt = mk_code_dt rty qty wit wit_thm NONE in (quot_thm, (update_code_dt code_dt lthy, rel_eq_onps)) end else (quot_thm, (lthy, rel_eq_onps)) end; fun case_tac rule = Subgoal.FOCUS_PARAMS (fn {context = ctxt, params, ...} => HEADGOAL (rtac ctxt (infer_instantiate' ctxt [SOME (snd (hd params))] rule))); fun bundle_name_of_bundle_binding binding phi context = Name_Space.full_name (Name_Space.naming_of context) (Morphism.binding phi binding); fun prove_schematic_quot_thm actions ctxt = Lifting_Term.prove_schematic_quot_thm actions (Lifting_Info.get_quotients ctxt) ctxt fun prove_code_dt (rty, qty) lthy = let val (fold_quot_thm: (local_theory * thm list) Lifting_Term.fold_quot_thm) = { constr = constr, lift = lift, comp_lift = comp_lift_error }; in prove_schematic_quot_thm fold_quot_thm lthy (rty, qty) (lthy, []) |> snd end and add_lift_def_code_dt config var qty rhs rsp_thm par_thms lthy = let fun binop_conv2 cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2 fun ret_rel_conv conv ctm = case (Thm.term_of ctm) of Const (\<^const_name>\rel_fun\, _) $ _ $ _ => binop_conv2 Conv.all_conv conv ctm | _ => conv ctm fun R_conv rel_eq_onps = Transfer.top_sweep_rewr_conv @{thms eq_onp_top_eq_eq[symmetric, THEN eq_reflection]} then_conv Transfer.bottom_rewr_conv rel_eq_onps val (ret_lift_def, lthy1) = add_lift_def (#lift_config config) var qty rhs rsp_thm par_thms lthy in if (not (#code_dt config) orelse (code_eq_of_lift_def ret_lift_def <> NONE_EQ) andalso (code_eq_of_lift_def ret_lift_def <> UNKNOWN_EQ)) (* Let us try even in case of UNKNOWN_EQ. If this leads to problems, the user can always say that they do not want this workaround. *) then (ret_lift_def, lthy1) else let val lift_def = inst_of_lift_def lthy1 qty ret_lift_def val rty = rty_of_lift_def lift_def val rty_ret = body_type rty val qty_ret = body_type qty val (lthy2, rel_eq_onps) = prove_code_dt (rty_ret, qty_ret) lthy1 val code_dt = code_dt_of lthy2 (rty_ret, qty_ret) in if is_none code_dt orelse is_none (rep_isom_data_of_code_dt (the code_dt)) then (ret_lift_def, lthy2) else let val code_dt = the code_dt val rhs = dest_comb (rhs_of_lift_def lift_def) |> snd val rep_isom_data = code_dt |> rep_isom_data_of_code_dt |> the val pointer = pointer_of_rep_isom_data rep_isom_data val quot_active = Lifting_Info.lookup_restore_data lthy2 pointer |> the |> #quotient |> #quot_thm |> Lifting_Info.lookup_quot_thm_quotients lthy2 |> is_some val qty_code_dt_bundle_name = bundle_name_of_rep_isom_data rep_isom_data val rep_isom = mk_rep_isom_of_code_dt qty_ret code_dt |> the val lthy3 = if quot_active then lthy2 else Bundle.includes [qty_code_dt_bundle_name] lthy2 fun qty_isom_of_rep_isom rep = rep |> dest_Const |> snd |> domain_type val qty_isom = qty_isom_of_rep_isom rep_isom val f'_var = (Binding.suffix_name "_aux" (fst var), NoSyn); val f'_qty = strip_type qty |> fst |> rpair qty_isom |> op ---> val f'_rsp_rel = Lifting_Term.equiv_relation lthy3 (rty, f'_qty); val rsp = rsp_thm_of_lift_def lift_def val rel_eq_onps_conv = HOLogic.Trueprop_conv (Conv.fun2_conv (ret_rel_conv (R_conv rel_eq_onps))) val rsp_norm = Conv.fconv_rule rel_eq_onps_conv rsp val f'_rsp_goal = HOLogic.mk_Trueprop (f'_rsp_rel $ rhs $ rhs); val f'_rsp = Goal.prove_sorry lthy3 [] [] f'_rsp_goal (fn {context = ctxt, prems = _} => HEADGOAL (CONVERSION (rel_eq_onps_conv) THEN' rtac ctxt rsp_norm)) |> Thm.close_derivation \<^here> val (f'_lift_def, lthy4) = add_lift_def ld_no_notes f'_var f'_qty rhs f'_rsp [] lthy3 val f'_lift_def = inst_of_lift_def lthy4 f'_qty f'_lift_def val f'_lift_const = mk_lift_const_of_lift_def f'_qty f'_lift_def val (args, args_ctxt) = mk_Frees "x" (binder_types qty) lthy4 val f_alt_def_goal_lhs = list_comb (lift_const_of_lift_def lift_def, args); val f_alt_def_goal_rhs = rep_isom $ list_comb (f'_lift_const, args); val f_alt_def_goal = HOLogic.mk_Trueprop (HOLogic.mk_eq (f_alt_def_goal_lhs, f_alt_def_goal_rhs)); fun f_alt_def_tac ctxt i = EVERY' [Transfer.gen_frees_tac [] ctxt, DETERM o Transfer.transfer_tac true ctxt, SELECT_GOAL (Local_Defs.unfold0_tac ctxt [id_apply]), rtac ctxt refl] i; val rep_isom_transfer = transfer_of_rep_isom_data rep_isom_data val (_, transfer_ctxt) = args_ctxt |> Proof_Context.note_thms "" (Binding.empty_atts, [([rep_isom_transfer], [Transfer.transfer_add])]) val f_alt_def = Goal.prove_sorry transfer_ctxt [] [] f_alt_def_goal (fn {context = goal_ctxt, ...} => HEADGOAL (f_alt_def_tac goal_ctxt)) |> Thm.close_derivation \<^here> |> singleton (Variable.export transfer_ctxt lthy4) val lthy5 = lthy4 |> Local_Theory.note ((Binding.empty, @{attributes [code]}), [f_alt_def]) |> snd (* if processing a mutual datatype (there is a cycle!) the corresponding quotient will be needed later and will be forgotten later *) |> (if quot_active then I else Lifting_Setup.lifting_forget pointer) in (ret_lift_def, lthy5) end end end and mk_rep_isom qty_isom_bundle (rty, qty, qty_isom) lthy0 = let (* logical definition of qty qty_isom isomorphism *) val uTname = unique_Tname (rty, qty) fun eq_onp_to_top_tac ctxt = SELECT_GOAL (Local_Defs.unfold0_tac ctxt (@{thm eq_onp_top_eq_eq[symmetric]} :: Lifting_Info.get_relator_eq_onp_rules ctxt)) fun lift_isom_tac ctxt = HEADGOAL (eq_onp_to_top_tac ctxt THEN' (rtac ctxt @{thm id_transfer})); val (rep_isom_lift_def, lthy1) = lthy0 - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> lift_def ld_no_notes (Binding.qualify_name true uTname "Rep_isom", NoSyn) (qty_isom --> qty) (HOLogic.id_const rty) lift_isom_tac [] |>> inst_of_lift_def lthy0 (qty_isom --> qty); val (abs_isom, lthy2) = lthy1 |> lift_def ld_no_notes (Binding.qualify_name true uTname "Abs_isom", NoSyn) (qty --> qty_isom) (HOLogic.id_const rty) lift_isom_tac [] |>> mk_lift_const_of_lift_def (qty --> qty_isom); val rep_isom = lift_const_of_lift_def rep_isom_lift_def val pointer = Lifting_Setup.pointer_of_bundle_binding lthy2 qty_isom_bundle fun code_dt phi context = code_dt_of lthy2 (rty, qty) |> the |> update_rep_isom rep_isom (transfer_rules_of_lift_def rep_isom_lift_def |> hd) (bundle_name_of_bundle_binding qty_isom_bundle phi context) pointer; val lthy3 = lthy2 |> Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => fn context => Data.map (Item_Net.update (morph_code_dt phi (code_dt phi context))) context) |> Local_Theory.close_target (* in order to make the qty qty_isom isomorphism executable we have to define discriminators and selectors for qty_isom *) val (rty_name, typs) = dest_Type rty val (_, qty_typs) = dest_Type qty val fp = BNF_FP_Def_Sugar.fp_sugar_of lthy3 rty_name val fp = if is_some fp then the fp else error ("code_dt: " ^ quote rty_name ^ " is not a datatype.") val ctr_sugar = fp |> #fp_ctr_sugar |> #ctr_sugar val ctrs = map (Ctr_Sugar.mk_ctr typs) (#ctrs ctr_sugar); val qty_ctrs = map (Ctr_Sugar.mk_ctr qty_typs) (#ctrs ctr_sugar); val ctr_Tss = map (dest_Const #> snd #> binder_types) ctrs; val qty_ctr_Tss = map (dest_Const #> snd #> binder_types) qty_ctrs; val n = length ctrs; val ks = 1 upto n; val (xss, _) = mk_Freess "x" ctr_Tss lthy3; fun sel_retT (rty' as Type (s, rtys'), qty' as Type (s', qtys')) = if (rty', qty') = (rty, qty) then qty_isom else (if s = s' then Type (s, map sel_retT (rtys' ~~ qtys')) else qty') | sel_retT (_, qty') = qty'; val sel_retTs = map2 (map2 (sel_retT oo pair)) ctr_Tss qty_ctr_Tss fun lazy_prove_code_dt (rty, qty) lthy = if is_none (code_dt_of lthy (rty, qty)) then prove_code_dt (rty, qty) lthy |> fst else lthy; val lthy4 = fold2 (fold2 (lazy_prove_code_dt oo pair)) ctr_Tss sel_retTs lthy3 val sel_argss = @{map 4} (fn k => fn xs => @{map 2} (fn x => fn qty_ret => (k, qty_ret, (xs, x)))) ks xss xss sel_retTs; fun mk_sel_casex (_, _, (_, x)) = Ctr_Sugar.mk_case typs (x |> dest_Free |> snd) (#casex ctr_sugar); val dis_casex = Ctr_Sugar.mk_case typs HOLogic.boolT (#casex ctr_sugar); fun mk_sel_case_args lthy ctr_Tss ks (k, qty_ret, (xs, x)) = let val T = x |> dest_Free |> snd; fun gen_undef_wit Ts wits = case code_dt_of lthy (T, qty_ret) of SOME code_dt => (fold_rev (Term.lambda o curry Free Name.uu) Ts (mk_witness_of_code_dt qty_ret code_dt), wit_thm_of_code_dt code_dt :: wits) | NONE => (fold_rev (Term.lambda o curry Free Name.uu) Ts (mk_undefined T), wits) in @{fold_map 2} (fn Ts => fn k' => fn wits => (if k = k' then (fold_rev Term.lambda xs x, wits) else gen_undef_wit Ts wits)) ctr_Tss ks [] end; fun mk_sel_rhs arg = let val (sel_rhs, wits) = mk_sel_case_args lthy4 ctr_Tss ks arg in (arg |> #2, wits, list_comb (mk_sel_casex arg, sel_rhs)) end; fun mk_dis_case_args args k = map (fn (k', arg) => (if k = k' then fold_rev Term.lambda arg \<^const>\True\ else fold_rev Term.lambda arg \<^const>\False\)) args; val sel_rhs = map (map mk_sel_rhs) sel_argss val dis_rhs = map (fn k => list_comb (dis_casex, mk_dis_case_args (ks ~~ xss) k)) ks val dis_qty = qty_isom --> HOLogic.boolT; val dis_names = map (fn k => Binding.qualify_name true uTname ("dis" ^ string_of_int k)) ks; val (diss, lthy5) = @{fold_map 2} (fn b => fn rhs => fn lthy => lift_def ld_no_notes (b, NoSyn) dis_qty rhs (K all_tac) [] lthy |>> mk_lift_const_of_lift_def dis_qty) dis_names dis_rhs lthy4 val unfold_lift_sel_rsp = @{lemma "(\x. P1 x \ P2 (f x)) \ (rel_fun (eq_onp P1) (eq_onp P2)) f f" by (simp add: eq_onp_same_args rel_fun_eq_onp_rel)} fun lift_sel_tac exhaust_rule dt_rules wits ctxt i = (Method.insert_tac ctxt wits THEN' eq_onp_to_top_tac ctxt THEN' (* normalize *) rtac ctxt unfold_lift_sel_rsp THEN' case_tac exhaust_rule ctxt THEN_ALL_NEW ( EVERY' [hyp_subst_tac ctxt, (* does not kill wits because = was rewritten to eq_onp top *) Raw_Simplifier.rewrite_goal_tac ctxt (map safe_mk_meta_eq dt_rules), REPEAT_DETERM o etac ctxt conjE, assume_tac ctxt])) i val pred_simps = Transfer.lookup_pred_data lthy5 (Tname rty) |> the |> Transfer.pred_simps val sel_tac = lift_sel_tac (#exhaust ctr_sugar) (#case_thms ctr_sugar @ pred_simps) val sel_names = map (fn (k, xs) => map (fn k' => Binding.qualify_name true uTname ("sel" ^ string_of_int k ^ string_of_int k')) (1 upto length xs)) (ks ~~ ctr_Tss); val (selss, lthy6) = @{fold_map 2} (@{fold_map 2} (fn b => fn (qty_ret, wits, rhs) => fn lthy => lift_def_code_dt { code_dt = true, lift_config = ld_no_notes } (b, NoSyn) (qty_isom --> qty_ret) rhs (HEADGOAL o sel_tac wits) [] lthy |>> mk_lift_const_of_lift_def (qty_isom --> qty_ret))) sel_names sel_rhs lthy5 (* now we can execute the qty qty_isom isomorphism *) fun mk_type_definition newT oldT RepC AbsC A = let val typedefC = Const (\<^const_name>\type_definition\, (newT --> oldT) --> (oldT --> newT) --> HOLogic.mk_setT oldT --> HOLogic.boolT); in typedefC $ RepC $ AbsC $ A end; val typedef_goal = mk_type_definition qty_isom qty rep_isom abs_isom (HOLogic.mk_UNIV qty) |> HOLogic.mk_Trueprop; fun typ_isom_tac ctxt i = EVERY' [ SELECT_GOAL (Local_Defs.unfold0_tac ctxt @{thms type_definition_def}), DETERM o Transfer.transfer_tac true ctxt, SELECT_GOAL (Local_Defs.unfold0_tac ctxt @{thms eq_onp_top_eq_eq}) (* normalize *), Raw_Simplifier.rewrite_goal_tac ctxt (map safe_mk_meta_eq @{thms id_apply simp_thms Ball_def}), rtac ctxt TrueI] i; val (_, transfer_ctxt) = Proof_Context.note_thms "" (Binding.empty_atts, [(@{thms right_total_UNIV_transfer}, [Transfer.transfer_add]), (@{thms Domain_eq_top}, [Transfer.transfer_domain_add])]) lthy6; val quot_thm_isom = Goal.prove_sorry transfer_ctxt [] [] typedef_goal (fn {context = goal_ctxt, ...} => typ_isom_tac goal_ctxt 1) |> Thm.close_derivation \<^here> |> singleton (Variable.export transfer_ctxt lthy6) |> (fn thm => @{thm UNIV_typedef_to_Quotient} OF [thm, @{thm reflexive}]) val qty_isom_name = Tname qty_isom; val quot_isom_rep = let val (quotients : Lifting_Term.quotients) = Symtab.insert (Lifting_Info.quotient_eq) (qty_isom_name, {quot_thm = quot_thm_isom, pcr_info = NONE}) Symtab.empty val id_actions = { constr = K I, lift = K I, comp_lift = K I } in fn ctxt => fn (rty, qty) => Lifting_Term.prove_schematic_quot_thm id_actions quotients ctxt (rty, qty) () |> fst |> Lifting_Term.force_qty_type ctxt qty |> quot_thm_rep end; val (x, x_ctxt) = yield_singleton (mk_Frees "x") qty_isom lthy6; fun mk_ctr ctr ctr_Ts sels = let val sel_ret_Ts = map (dest_Const #> snd #> body_type) sels; fun rep_isom lthy t (rty, qty) = let val rep = quot_isom_rep lthy (rty, qty) in if is_Const rep andalso (rep |> dest_Const |> fst) = \<^const_name>\id\ then t else rep $ t end; in @{fold 3} (fn sel => fn ctr_T => fn sel_ret_T => fn ctr => ctr $ rep_isom x_ctxt (sel $ x) (ctr_T, sel_ret_T)) sels ctr_Ts sel_ret_Ts ctr end; (* stolen from Metis *) exception BREAK_LIST fun break_list (x :: xs) = (x, xs) | break_list _ = raise BREAK_LIST val (ctr, ctrs) = qty_ctrs |> rev |> break_list; val (ctr_Ts, ctr_Tss) = qty_ctr_Tss |> rev |> break_list; val (sel, rselss) = selss |> rev |> break_list; val rdiss = rev diss |> tl; val first_ctr = mk_ctr ctr ctr_Ts sel; fun mk_If_ctr dis ctr ctr_Ts sel elsex = mk_If (dis$x) (mk_ctr ctr ctr_Ts sel) elsex; val rhs = @{fold 4} mk_If_ctr rdiss ctrs ctr_Tss rselss first_ctr; val rep_isom_code_goal = HOLogic.mk_Trueprop (HOLogic.mk_eq (rep_isom$x, rhs)); local val rep_isom_code_tac_rules = map safe_mk_meta_eq @{thms refl id_apply if_splits simp_thms} in fun rep_isom_code_tac (ctr_sugar:Ctr_Sugar.ctr_sugar) ctxt i = let val exhaust = ctr_sugar |> #exhaust val cases = ctr_sugar |> #case_thms val map_ids = fp |> #fp_nesting_bnfs |> map BNF_Def.map_id0_of_bnf val simp_rules = map safe_mk_meta_eq (cases @ map_ids) @ rep_isom_code_tac_rules in EVERY' [Transfer.gen_frees_tac [] ctxt, DETERM o (Transfer.transfer_tac true ctxt), case_tac exhaust ctxt THEN_ALL_NEW EVERY' [hyp_subst_tac ctxt, Raw_Simplifier.rewrite_goal_tac ctxt simp_rules, rtac ctxt TrueI ]] i end end (* stolen from bnf_fp_n2m.ML *) fun force_typ ctxt T = Term.map_types Type_Infer.paramify_vars #> Type.constraint T #> singleton (Type_Infer_Context.infer_types ctxt); (* The following tests that types in rty have corresponding arities imposed by constraints of the datatype fp. Otherwise rep_isom_code_tac could fail (especially transfer in it) is such a way that it is not easy to infer the problem with sorts. *) val _ = yield_singleton (mk_Frees "x") (#T fp) x_ctxt |> fst |> force_typ x_ctxt qty val rep_isom_code = Goal.prove_sorry x_ctxt [] [] rep_isom_code_goal (fn {context = goal_ctxt, ...} => rep_isom_code_tac ctr_sugar goal_ctxt 1) |> Thm.close_derivation \<^here> |> singleton(Variable.export x_ctxt lthy6) in lthy6 |> snd o Local_Theory.note ((Binding.empty, @{attributes [code]}), [rep_isom_code]) |> Lifting_Setup.lifting_forget pointer |> pair (selss, diss, rep_isom_code) end and constr qty (quot_thm, (lthy0, rel_eq_onps)) = let val quot_thm = Lifting_Term.force_qty_type lthy0 qty quot_thm val (rty, qty) = quot_thm_rty_qty quot_thm val rty_name = Tname rty; val pred_data = Transfer.lookup_pred_data lthy0 rty_name val pred_data = if is_some pred_data then the pred_data else error ("code_dt: " ^ quote rty_name ^ " is not a datatype.") val rel_eq_onp = safe_mk_meta_eq (Transfer.rel_eq_onp pred_data); val rel_eq_onps = insert Thm.eq_thm rel_eq_onp rel_eq_onps val R_conv = Transfer.top_sweep_rewr_conv @{thms eq_onp_top_eq_eq[symmetric, THEN eq_reflection]} then_conv Conv.rewr_conv rel_eq_onp val quot_thm = Conv.fconv_rule(HOLogic.Trueprop_conv (Quotient_R_conv R_conv)) quot_thm; in if is_none (code_dt_of lthy0 (rty, qty)) then let val non_empty_pred = quot_thm RS @{thm type_definition_Quotient_not_empty} val pred = quot_thm_rel quot_thm |> dest_comb |> snd; val (pred, lthy1) = lthy0 - |> Local_Theory.open_target |> snd + |> Local_Theory.open_target |> yield_singleton (Variable.import_terms true) pred; val TFrees = Term.add_tfreesT qty [] fun non_empty_typedef_tac non_empty_pred ctxt i = (Method.insert_tac ctxt [non_empty_pred] THEN' SELECT_GOAL (Local_Defs.unfold0_tac ctxt [mem_Collect_eq]) THEN' assume_tac ctxt) i val uTname = unique_Tname (rty, qty) val Tdef_set = HOLogic.mk_Collect ("x", rty, pred $ Free("x", rty)); val ((_, tcode_dt), lthy2) = lthy1 |> conceal_naming_result (typedef (Binding.concealed uTname, TFrees, NoSyn) Tdef_set NONE (fn lthy => HEADGOAL (non_empty_typedef_tac non_empty_pred lthy))); val type_definition_thm = tcode_dt |> snd |> #type_definition; val qty_isom = tcode_dt |> fst |> #abs_type; val (binding, lthy3) = lthy2 |> conceal_naming_result (Lifting_Setup.setup_by_typedef_thm {notes = false} type_definition_thm) ||> Local_Theory.close_target val (wit, wit_thm) = mk_witness quot_thm; val code_dt = mk_code_dt rty qty wit wit_thm NONE; val lthy4 = lthy3 |> update_code_dt code_dt |> mk_rep_isom binding (rty, qty, qty_isom) |> snd in (quot_thm, (lthy4, rel_eq_onps)) end else (quot_thm, (lthy0, rel_eq_onps)) end and lift_def_code_dt config = gen_lift_def (add_lift_def_code_dt config) (** from parsed parameters to the config record **) fun map_config_code_dt f1 f2 ({code_dt = code_dt, lift_config = lift_config}: config_code_dt) = {code_dt = f1 code_dt, lift_config = f2 lift_config} fun update_config_code_dt nval = map_config_code_dt (K nval) I val config_flags = [("code_dt", update_config_code_dt true)] fun evaluate_params params = let fun eval_param param config = case AList.lookup (op =) config_flags param of SOME update => update config | NONE => error ("Unknown parameter: " ^ (quote param)) in fold eval_param params default_config_code_dt end (** lift_definition command. It opens a proof of a corresponding respectfulness theorem in a user-friendly, readable form. Then add_lift_def_code_dt is called internally. **) local val eq_onp_assms_tac_fixed_rules = map (Transfer.prep_transfer_domain_thm \<^context>) [@{thm pcr_Domainp_total}, @{thm pcr_Domainp_par_left_total}, @{thm pcr_Domainp_par}, @{thm pcr_Domainp}] in fun mk_readable_rsp_thm_eq tm ctxt = let val ctm = Thm.cterm_of ctxt tm fun assms_rewr_conv tactic rule ct = let fun prove_extra_assms thm = let val assms = cprems_of thm fun finish thm = if Thm.no_prems thm then SOME (Goal.conclude thm) else NONE fun prove ctm = Option.mapPartial finish (SINGLE tactic (Goal.init ctm)) in map_interrupt prove assms end fun cconl_of thm = Drule.strip_imp_concl (Thm.cprop_of thm) fun lhs_of thm = fst (Thm.dest_equals (cconl_of thm)) fun rhs_of thm = snd (Thm.dest_equals (cconl_of thm)) val rule1 = Thm.incr_indexes (Thm.maxidx_of_cterm ct + 1) rule; val lhs = lhs_of rule1; val rule2 = Thm.rename_boundvars (Thm.term_of lhs) (Thm.term_of ct) rule1; val rule3 = Thm.instantiate (Thm.match (lhs, ct)) rule2 handle Pattern.MATCH => raise CTERM ("assms_rewr_conv", [lhs, ct]); val proved_assms = prove_extra_assms rule3 in case proved_assms of SOME proved_assms => let val rule3 = proved_assms MRSL rule3 val rule4 = if lhs_of rule3 aconvc ct then rule3 else let val ceq = Thm.dest_fun2 (Thm.cprop_of rule3) in rule3 COMP Thm.trivial (Thm.mk_binop ceq ct (rhs_of rule3)) end in Thm.transitive rule4 (Thm.beta_conversion true (rhs_of rule4)) end | NONE => Conv.no_conv ct end fun assms_rewrs_conv tactic rules = Conv.first_conv (map (assms_rewr_conv tactic) rules) fun simp_arrows_conv ctm = let val unfold_conv = Conv.rewrs_conv [@{thm rel_fun_eq_eq_onp[THEN eq_reflection]}, @{thm rel_fun_eq_onp_rel[THEN eq_reflection]}, @{thm rel_fun_eq[THEN eq_reflection]}, @{thm rel_fun_eq_rel[THEN eq_reflection]}, @{thm rel_fun_def[THEN eq_reflection]}] fun binop_conv2 cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2 val eq_onp_assms_tac_rules = @{thm left_unique_OO} :: eq_onp_assms_tac_fixed_rules @ (Transfer.get_transfer_raw ctxt) val intro_top_rule = @{thm eq_onp_top_eq_eq[symmetric, THEN eq_reflection]} val kill_tops = Transfer.top_sweep_rewr_conv [@{thm eq_onp_top_eq_eq[THEN eq_reflection]}] val eq_onp_assms_tac = (CONVERSION kill_tops THEN' TRY o REPEAT_ALL_NEW (resolve_tac ctxt eq_onp_assms_tac_rules) THEN_ALL_NEW (DETERM o Transfer.eq_tac ctxt)) 1 val relator_eq_onp_conv = Conv.bottom_conv (K (Conv.try_conv (assms_rewrs_conv eq_onp_assms_tac (intro_top_rule :: Lifting_Info.get_relator_eq_onp_rules ctxt)))) ctxt then_conv kill_tops val relator_eq_conv = Conv.bottom_conv (K (Conv.try_conv (Conv.rewrs_conv (Transfer.get_relator_eq ctxt)))) ctxt in case (Thm.term_of ctm) of Const (\<^const_name>\rel_fun\, _) $ _ $ _ => (binop_conv2 simp_arrows_conv simp_arrows_conv then_conv unfold_conv) ctm | _ => (relator_eq_onp_conv then_conv relator_eq_conv) ctm end val unfold_ret_val_invs = Conv.bottom_conv (K (Conv.try_conv (Conv.rewr_conv @{thm eq_onp_same_args[THEN eq_reflection]}))) ctxt val unfold_inv_conv = Conv.top_sweep_conv (K (Conv.rewr_conv @{thm eq_onp_def[THEN eq_reflection]})) ctxt val simp_conv = HOLogic.Trueprop_conv (Conv.fun2_conv simp_arrows_conv) val univq_conv = Conv.rewr_conv @{thm HOL.all_simps(6)[symmetric, THEN eq_reflection]} val univq_prenex_conv = Conv.top_conv (K (Conv.try_conv univq_conv)) ctxt val beta_conv = Thm.beta_conversion true val eq_thm = (simp_conv then_conv univq_prenex_conv then_conv beta_conv then_conv unfold_ret_val_invs then_conv unfold_inv_conv) ctm in Object_Logic.rulify ctxt (eq_thm RS Drule.equal_elim_rule2) end end fun rename_to_tnames ctxt term = let fun all_typs (Const (\<^const_name>\Pure.all\, _) $ Abs (_, T, t)) = T :: all_typs t | all_typs _ = [] fun rename (Const (\<^const_name>\Pure.all\, T1) $ Abs (_, T2, t)) (new_name :: names) = (Const (\<^const_name>\Pure.all\, T1) $ Abs (new_name, T2, rename t names)) | rename t _ = t val (fixed_def_t, _) = yield_singleton (Variable.importT_terms) term ctxt val new_names = Old_Datatype_Prop.make_tnames (all_typs fixed_def_t) in rename term new_names end fun quot_thm_err ctxt (rty, qty) pretty_msg = let val error_msg = cat_lines ["Lifting failed for the following types:", Pretty.string_of (Pretty.block [Pretty.str "Raw type:", Pretty.brk 2, Syntax.pretty_typ ctxt rty]), Pretty.string_of (Pretty.block [Pretty.str "Abstract type:", Pretty.brk 2, Syntax.pretty_typ ctxt qty]), "", (Pretty.string_of (Pretty.block [Pretty.str "Reason:", Pretty.brk 2, pretty_msg]))] in error error_msg end fun check_rty_err ctxt (rty_schematic, rty_forced) (raw_var, rhs_raw) = let val (_, ctxt') = Proof_Context.read_var raw_var ctxt val rhs = Syntax.read_term ctxt' rhs_raw val error_msg = cat_lines ["Lifting failed for the following term:", Pretty.string_of (Pretty.block [Pretty.str "Term:", Pretty.brk 2, Syntax.pretty_term ctxt rhs]), Pretty.string_of (Pretty.block [Pretty.str "Type:", Pretty.brk 2, Syntax.pretty_typ ctxt rty_schematic]), "", (Pretty.string_of (Pretty.block [Pretty.str "Reason:", Pretty.brk 2, Pretty.str "The type of the term cannot be instantiated to", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt rty_forced), Pretty.str "."]))] in error error_msg end fun lift_def_cmd (params, raw_var, rhs_raw, par_xthms) lthy0 = let val config = evaluate_params params val ((binding, SOME qty, mx), lthy1) = Proof_Context.read_var raw_var lthy0 val var = (binding, mx) val rhs = Syntax.read_term lthy1 rhs_raw val par_thms = Attrib.eval_thms lthy1 par_xthms val (goal, after_qed) = lthy1 |> prepare_lift_def (add_lift_def_code_dt config) var qty rhs par_thms val (goal, after_qed) = case goal of NONE => (goal, K (after_qed Drule.dummy_thm)) | SOME prsp_tm => let val readable_rsp_thm_eq = mk_readable_rsp_thm_eq prsp_tm lthy1 val (readable_rsp_tm, _) = Logic.dest_implies (Thm.prop_of readable_rsp_thm_eq) val readable_rsp_tm_tnames = rename_to_tnames lthy1 readable_rsp_tm fun after_qed' [[thm]] lthy = let val internal_rsp_thm = Goal.prove lthy [] [] prsp_tm (fn {context = goal_ctxt, ...} => rtac goal_ctxt readable_rsp_thm_eq 1 THEN Proof_Context.fact_tac goal_ctxt [thm] 1) in after_qed internal_rsp_thm lthy end in (SOME readable_rsp_tm_tnames, after_qed') end fun after_qed_with_err_handling thmss ctxt = (after_qed thmss ctxt handle Lifting_Term.QUOT_THM (rty, qty, msg) => quot_thm_err lthy1 (rty, qty) msg) handle Lifting_Term.CHECK_RTY (rty_schematic, rty_forced) => check_rty_err lthy1 (rty_schematic, rty_forced) (raw_var, rhs_raw); in lthy1 |> Proof.theorem NONE (snd oo after_qed_with_err_handling) [map (rpair []) (the_list goal)] end fun lift_def_cmd_with_err_handling (params, (raw_var, rhs_raw, par_xthms)) lthy = (lift_def_cmd (params, raw_var, rhs_raw, par_xthms) lthy handle Lifting_Term.QUOT_THM (rty, qty, msg) => quot_thm_err lthy (rty, qty) msg) handle Lifting_Term.CHECK_RTY (rty_schematic, rty_forced) => check_rty_err lthy (rty_schematic, rty_forced) (raw_var, rhs_raw); val parse_param = Parse.name val parse_params = Scan.optional (Args.parens (Parse.list parse_param)) []; (* command syntax *) val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\lift_definition\ "definition for constants over the quotient type" (parse_params -- (((Parse.binding -- (\<^keyword>\::\ |-- (Parse.typ >> SOME) -- Parse.opt_mixfix') >> Scan.triple2) -- (\<^keyword>\is\ |-- Parse.term) -- Scan.optional (\<^keyword>\parametric\ |-- Parse.!!! Parse.thms1) []) >> Scan.triple1) >> lift_def_cmd_with_err_handling); end diff --git a/src/Pure/Isar/local_theory.ML b/src/Pure/Isar/local_theory.ML --- a/src/Pure/Isar/local_theory.ML +++ b/src/Pure/Isar/local_theory.ML @@ -1,415 +1,418 @@ (* Title: Pure/Isar/local_theory.ML Author: Makarius Local theory operations, with abstract target context. *) type local_theory = Proof.context; type generic_theory = Context.generic; structure Attrib = struct type binding = binding * Token.src list; type thms = (thm list * Token.src list) list; type fact = binding * thms; end; structure Locale = struct type registration = {inst: string * morphism, mixin: (morphism * bool) option, export: morphism}; end; signature LOCAL_THEORY = sig type operations val assert: local_theory -> local_theory val level: Proof.context -> int val assert_bottom: local_theory -> local_theory val mark_brittle: local_theory -> local_theory val assert_nonbrittle: local_theory -> local_theory val map_contexts: (int -> Proof.context -> Proof.context) -> local_theory -> local_theory val background_naming_of: local_theory -> Name_Space.naming val map_background_naming: (Name_Space.naming -> Name_Space.naming) -> local_theory -> local_theory val restore_background_naming: local_theory -> local_theory -> local_theory val full_name: local_theory -> binding -> string val new_group: local_theory -> local_theory val reset_group: local_theory -> local_theory val standard_morphism: local_theory -> Proof.context -> morphism val standard_morphism_theory: local_theory -> morphism val standard_form: local_theory -> Proof.context -> (morphism -> 'a) -> 'a val raw_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory val raw_theory: (theory -> theory) -> local_theory -> local_theory val background_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory val background_theory: (theory -> theory) -> local_theory -> local_theory val target_of: local_theory -> Proof.context val target: (Proof.context -> Proof.context) -> local_theory -> local_theory val target_morphism: local_theory -> morphism val propagate_ml_env: generic_theory -> generic_theory val operations_of: local_theory -> operations val define: (binding * mixfix) * (Attrib.binding * term) -> local_theory -> (term * (string * thm)) * local_theory val define_internal: (binding * mixfix) * (Attrib.binding * term) -> local_theory -> (term * (string * thm)) * local_theory val note: Attrib.binding * thm list -> local_theory -> (string * thm list) * local_theory val notes: Attrib.fact list -> local_theory -> (string * thm list) list * local_theory val notes_kind: string -> Attrib.fact list -> local_theory -> (string * thm list) list * local_theory val abbrev: Syntax.mode -> (binding * mixfix) * term -> local_theory -> (term * term) * local_theory val declaration: {syntax: bool, pervasive: bool} -> declaration -> local_theory -> local_theory val theory_registration: Locale.registration -> local_theory -> local_theory val locale_dependency: Locale.registration -> local_theory -> local_theory val pretty: local_theory -> Pretty.T list val add_thms_dynamic: binding * (Context.generic -> thm list) -> local_theory -> local_theory val set_defsort: sort -> local_theory -> local_theory val type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> local_theory -> local_theory val notation: bool -> Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory val type_alias: binding -> string -> local_theory -> local_theory val const_alias: binding -> string -> local_theory -> local_theory val init: {background_naming: Name_Space.naming, exit: local_theory -> Proof.context} -> operations -> Proof.context -> local_theory val exit: local_theory -> Proof.context val exit_global: local_theory -> theory val exit_result: (morphism -> 'a -> 'b) -> 'a * local_theory -> 'b * Proof.context val exit_result_global: (morphism -> 'a -> 'b) -> 'a * local_theory -> 'b * theory val init_target: {background_naming: Name_Space.naming, after_close: local_theory -> local_theory} -> operations -> local_theory -> Binding.scope * local_theory - val open_target: local_theory -> Binding.scope * local_theory + val begin_target: local_theory -> Binding.scope * local_theory + val open_target: local_theory -> local_theory val close_target: local_theory -> local_theory val subtarget: (local_theory -> local_theory) -> local_theory -> local_theory val subtarget_result: (morphism -> 'a -> 'b) -> (local_theory -> 'a * local_theory) -> local_theory -> 'b * local_theory end; signature PRIVATE_LOCAL_THEORY = sig include LOCAL_THEORY val reset: local_theory -> local_theory end structure Local_Theory: PRIVATE_LOCAL_THEORY = struct (** local theory data **) (* type lthy *) type operations = {define: bool -> (binding * mixfix) * (Attrib.binding * term) -> local_theory -> (term * (string * thm)) * local_theory, notes: string -> Attrib.fact list -> local_theory -> (string * thm list) list * local_theory, abbrev: Syntax.mode -> (binding * mixfix) * term -> local_theory -> (term * term) * local_theory, declaration: {syntax: bool, pervasive: bool} -> declaration -> local_theory -> local_theory, theory_registration: Locale.registration -> local_theory -> local_theory, locale_dependency: Locale.registration -> local_theory -> local_theory, pretty: local_theory -> Pretty.T list}; type lthy = {background_naming: Name_Space.naming, operations: operations, after_close: local_theory -> local_theory, exit: local_theory -> Proof.context, brittle: bool, target: Proof.context}; fun make_lthy (background_naming, operations, after_close, exit, brittle, target) : lthy = {background_naming = background_naming, operations = operations, after_close = after_close, exit = exit, brittle = brittle, target = target}; (* context data *) structure Data = Proof_Data ( type T = lthy list; fun init _ = []; ); (* nested structure *) val level = length o Data.get; (*1: main target at bottom, >= 2: nested target context*) fun assert lthy = if level lthy = 0 then error "Missing local theory context" else lthy; fun assert_bottom lthy = let val _ = assert lthy; in if level lthy > 1 then error "Not at bottom of local theory nesting" else lthy end; fun assert_not_bottom lthy = let val _ = assert lthy; in if level lthy = 1 then error "Already at bottom of local theory nesting" else lthy end; val bottom_of = List.last o Data.get o assert; val top_of = hd o Data.get o assert; fun map_top f = assert #> Data.map (fn {background_naming, operations, after_close, exit, brittle, target} :: parents => make_lthy (f (background_naming, operations, after_close, exit, brittle, target)) :: parents); fun reset lthy = #target (top_of lthy) |> Data.put (Data.get lthy); fun map_contexts f lthy = let val n = level lthy in lthy |> (Data.map o map_index) (fn (i, {background_naming, operations, after_close, exit, brittle, target}) => make_lthy (background_naming, operations, after_close, exit, brittle, target |> Context_Position.set_visible false |> f (n - i - 1) |> Context_Position.restore_visible target)) |> f n end; (* brittle context -- implicit for nested structures *) fun mark_brittle lthy = if level lthy = 1 then map_top (fn (background_naming, operations, after_close, exit, _, target) => (background_naming, operations, after_close, exit, true, target)) lthy else lthy; fun assert_nonbrittle lthy = if #brittle (top_of lthy) then error "Brittle local theory context" else lthy; (* naming for background theory *) val background_naming_of = #background_naming o top_of; fun map_background_naming f = map_top (fn (background_naming, operations, after_close, exit, brittle, target) => (f background_naming, operations, after_close, exit, brittle, target)); val restore_background_naming = map_background_naming o K o background_naming_of; val full_name = Name_Space.full_name o background_naming_of; val new_group = map_background_naming Name_Space.new_group; val reset_group = map_background_naming Name_Space.reset_group; (* standard morphisms *) fun standard_morphism lthy ctxt = Proof_Context.norm_export_morphism lthy ctxt $> Morphism.binding_morphism "Local_Theory.standard_binding" (Name_Space.transform_binding (Proof_Context.naming_of lthy)); fun standard_morphism_theory lthy = standard_morphism lthy (Proof_Context.init_global (Proof_Context.theory_of lthy)); fun standard_form lthy ctxt x = Morphism.form (Morphism.transform (standard_morphism lthy ctxt) x); (* background theory *) fun raw_theory_result f lthy = let val (res, thy') = f (Proof_Context.theory_of lthy); val lthy' = map_contexts (K (Proof_Context.transfer thy')) lthy; in (res, lthy') end; fun raw_theory f = #2 o raw_theory_result (f #> pair ()); fun background_theory_result f lthy = let val naming = background_naming_of lthy |> Name_Space.transform_naming (Proof_Context.naming_of lthy); in lthy |> raw_theory_result (fn thy => thy |> Sign.map_naming (K naming) |> f ||> Sign.restore_naming thy) end; fun background_theory f = #2 o background_theory_result (f #> pair ()); (* target contexts *) val target_of = #target o bottom_of; fun target f lthy = let val ctxt = target_of lthy; val ctxt' = ctxt |> Context_Position.set_visible false |> f |> Context_Position.restore_visible ctxt; val thy' = Proof_Context.theory_of ctxt'; in map_contexts (fn 0 => K ctxt' | _ => Proof_Context.transfer thy') lthy end; fun target_morphism lthy = standard_morphism lthy (target_of lthy); fun propagate_ml_env (context as Context.Proof lthy) = let val inherit = ML_Env.inherit [context] in lthy |> background_theory (Context.theory_map inherit) |> map_contexts (K (Context.proof_map inherit)) |> Context.Proof end | propagate_ml_env context = context; (** operations **) val operations_of = #operations o top_of; fun operation f lthy = f (operations_of lthy) lthy; fun operation2 f x y = operation (fn ops => f ops x y); (* primitives *) val pretty = operation #pretty; val abbrev = operation2 #abbrev; val define = operation2 #define false; val define_internal = operation2 #define true; val notes_kind = operation2 #notes; val declaration = operation2 #declaration; fun theory_registration registration = assert_bottom #> operation (fn ops => #theory_registration ops registration); fun locale_dependency registration = assert_bottom #> operation (fn ops => #locale_dependency ops registration); (* theorems *) val notes = notes_kind ""; fun note (a, ths) = notes [(a, [(ths, [])])] #>> the_single; fun add_thms_dynamic (binding, f) lthy = lthy |> background_theory_result (fn thy => thy |> Global_Theory.add_thms_dynamic' (Sign.inherit_naming thy lthy) (binding, f)) |-> (fn name => map_contexts (fn _ => fn ctxt => Proof_Context.transfer_facts (Proof_Context.theory_of ctxt) ctxt) #> declaration {syntax = false, pervasive = false} (fn phi => let val binding' = Morphism.binding phi binding in Context.mapping (Global_Theory.alias_fact binding' name) (Proof_Context.alias_fact binding' name) end)); (* default sort *) fun set_defsort S = declaration {syntax = true, pervasive = false} (K (Context.mapping (Sign.set_defsort S) (Proof_Context.set_defsort S))); (* notation *) fun type_notation add mode raw_args lthy = let val args = map (apfst (Logic.type_map (Assumption.export_term lthy (target_of lthy)))) raw_args; val args' = map (apsnd Mixfix.reset_pos) args; val _ = lthy |> Context_Position.is_visible lthy ? Proof_Context.type_notation add mode args; in declaration {syntax = true, pervasive = false} (Proof_Context.generic_type_notation add mode args') lthy end; fun notation add mode raw_args lthy = let val args = map (apfst (Assumption.export_term lthy (target_of lthy))) raw_args val args' = map (apsnd Mixfix.reset_pos) args; val _ = lthy |> Context_Position.is_visible lthy ? Proof_Context.notation add mode args; in declaration {syntax = true, pervasive = false} (Proof_Context.generic_notation add mode args') lthy end; (* name space aliases *) fun syntax_alias global_alias local_alias b name = declaration {syntax = true, pervasive = false} (fn phi => let val b' = Morphism.binding phi b in Context.mapping (global_alias b' name) (local_alias b' name) end); val type_alias = syntax_alias Sign.type_alias Proof_Context.type_alias; val const_alias = syntax_alias Sign.const_alias Proof_Context.const_alias; (** manage targets **) (* main target *) fun init {background_naming, exit} operations target = target |> Data.map (fn [] => [make_lthy (background_naming, operations, I, exit, false, target)] | _ => error "Local theory already initialized"); val exit_of = #exit o bottom_of; fun exit lthy = exit_of lthy lthy; val exit_global = Proof_Context.theory_of o exit; fun exit_result decl (x, lthy) = let val ctxt = exit lthy; val phi = standard_morphism lthy ctxt; in (decl phi x, ctxt) end; fun exit_result_global decl (x, lthy) = let val thy = exit_global lthy; val thy_ctxt = Proof_Context.init_global thy; val phi = standard_morphism lthy thy_ctxt; in (decl phi x, thy) end; (* nested targets *) fun init_target {background_naming, after_close} operations lthy = let val _ = assert lthy; val after_close' = Proof_Context.restore_naming lthy #> after_close; val (scope, target) = Proof_Context.new_scope lthy; val entry = make_lthy (background_naming, operations, after_close', exit_of lthy, true, target); val lthy' = Data.map (cons entry) target; in (scope, lthy') end; -fun open_target lthy = +fun begin_target lthy = init_target {background_naming = background_naming_of lthy, after_close = I} (operations_of lthy) lthy; +val open_target = begin_target #> #2; + fun close_target lthy = let val _ = assert_not_bottom lthy; val ({after_close, ...} :: rest) = Data.get lthy; in lthy |> Data.put rest |> reset |> after_close end; -fun subtarget body = open_target #> #2 #> body #> close_target; +fun subtarget body = open_target #> body #> close_target; fun subtarget_result decl body lthy = let - val (x, inner_lthy) = lthy |> open_target |> #2 |> body; + val (x, inner_lthy) = lthy |> open_target |> body; val lthy' = inner_lthy |> close_target; val phi = Proof_Context.export_morphism inner_lthy lthy'; in (decl phi x, lthy') end; end;