# HG changeset patch # User blanchet # Date 1458646777 -3600 # Node ID 0701f25fac39e30ce7032293af762180a076df89 # Parent 9bfcbab7cd9903112be28337deb539f8e9aecf81 moved 'corec' from ssh://hg@bitbucket.org/jasmin_blanchette/nonprim-corec to Isabelle diff -r 9bfcbab7cd99 -r 0701f25fac39 src/HOL/Library/BNF_Corec.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/BNF_Corec.thy Tue Mar 22 12:39:37 2016 +0100 @@ -0,0 +1,213 @@ +(* Title: HOL/Library/BNF_Corec.thy + Author: Jasmin Blanchette, Inria, LORIA, MPII + Author: Aymeric Bouzy, Ecole polytechnique + Author: Dmitriy Traytel, ETH Zurich + Copyright 2015, 2016 + +Generalized corecursor sugar ("corec" and friends). +*) + +chapter {* Generalized Corecursor Sugar (corec and friends) *} + +theory BNF_Corec +imports Main +keywords + "corec" :: thy_decl and + "corecursive" :: thy_goal and + "friend_of_corec" :: thy_goal and + "coinduction_upto" :: thy_decl +begin + +lemma obj_distinct_prems: "P \ P \ Q \ P \ Q" + by auto + +lemma inject_refine: "g (f x) = x \ g (f y) = y \ f x = f y \ x = y" + by (metis (no_types)) + +lemma convol_apply: "BNF_Def.convol f g x = (f x, g x)" + unfolding convol_def .. + +lemma Grp_UNIV_id: "BNF_Def.Grp UNIV id = (op =)" + unfolding BNF_Def.Grp_def by auto + +lemma sum_comp_cases: + assumes "f o Inl = g o Inl" and "f o Inr = g o Inr" + shows "f = g" +proof (rule ext) + fix a show "f a = g a" + using assms unfolding comp_def fun_eq_iff by (cases a) auto +qed + +lemma case_sum_Inl_Inr_L: "case_sum (f o Inl) (f o Inr) = f" + by (metis case_sum_expand_Inr') + +lemma eq_o_InrI: "\g o Inl = h; case_sum h f = g\ \ f = g o Inr" + by (auto simp: fun_eq_iff split: sum.splits) + +lemma id_bnf_o: "BNF_Composition.id_bnf \ f = f" + unfolding BNF_Composition.id_bnf_def by (rule o_def) + +lemma o_id_bnf: "f \ BNF_Composition.id_bnf = f" + unfolding BNF_Composition.id_bnf_def by (rule o_def) + +lemma if_True_False: + "(if P then True else Q) \ P \ Q" + "(if P then False else Q) \ \ P \ Q" + "(if P then Q else True) \ \ P \ Q" + "(if P then Q else False) \ P \ Q" + by auto + +lemma if_distrib_fun: "(if c then f else g) x = (if c then f x else g x)" + by simp + + +section \Coinduction\ + +lemma eq_comp_compI: "a o b = f o x \ x o c = id \ f = a o (b o c)" + unfolding fun_eq_iff by simp + +lemma self_bounded_weaken_left: "(a :: 'a :: semilattice_inf) \ inf a b \ a \ b" + by (erule le_infE) + +lemma self_bounded_weaken_right: "(a :: 'a :: semilattice_inf) \ inf b a \ a \ b" + by (erule le_infE) + +lemma symp_iff: "symp R \ R = R^--1" + by (metis antisym conversep.cases conversep_le_swap predicate2I symp_def) + +lemma equivp_inf: "\equivp R; equivp S\ \ equivp (inf R S)" + unfolding equivp_def inf_fun_def inf_bool_def by metis + +lemma vimage2p_rel_prod: + "(\x y. rel_prod R S (BNF_Def.convol f1 g1 x) (BNF_Def.convol f2 g2 y)) = + (inf (BNF_Def.vimage2p f1 f2 R) (BNF_Def.vimage2p g1 g2 S))" + unfolding vimage2p_def rel_prod.simps convol_def by auto + +lemma predicate2I_obj: "(\x y. P x y \ Q x y) \ P \ Q" + by auto + +lemma predicate2D_obj: "P \ Q \ P x y \ Q x y" + by auto + +locale cong = + fixes rel :: "('a \ 'a \ bool) \ ('b \ 'b \ bool)" + and eval :: "'b \ 'a" + and retr :: "('a \ 'a \ bool) \ ('a \ 'a \ bool)" + assumes rel_mono: "\R S. R \ S \ rel R \ rel S" + and equivp_retr: "\R. equivp R \ equivp (retr R)" + and retr_eval: "\R x y. \(rel_fun (rel R) R) eval eval; rel (inf R (retr R)) x y\ \ + retr R (eval x) (eval y)" +begin + +definition cong :: "('a \ 'a \ bool) \ bool" where + "cong R \ equivp R \ (rel_fun (rel R) R) eval eval" + +lemma cong_retr: "cong R \ cong (inf R (retr R))" + unfolding cong_def + by (auto simp: rel_fun_def dest: predicate2D[OF rel_mono, rotated] + intro: equivp_inf equivp_retr retr_eval) + +lemma cong_equivp: "cong R \ equivp R" + unfolding cong_def by simp + +definition gen_cong :: "('a \ 'a \ bool) \ 'a \ 'a \ bool" where + "gen_cong R j1 j2 \ \R'. R \ R' \ cong R' \ R' j1 j2" + +lemma gen_cong_reflp[intro, simp]: "x = y \ gen_cong R x y" + unfolding gen_cong_def by (auto dest: cong_equivp equivp_reflp) + +lemma gen_cong_symp[intro]: "gen_cong R x y \ gen_cong R y x" + unfolding gen_cong_def by (auto dest: cong_equivp equivp_symp) + +lemma gen_cong_transp[intro]: "gen_cong R x y \ gen_cong R y z \ gen_cong R x z" + unfolding gen_cong_def by (auto dest: cong_equivp equivp_transp) + +lemma equivp_gen_cong: "equivp (gen_cong R)" + by (intro equivpI reflpI sympI transpI) auto + +lemma leq_gen_cong: "R \ gen_cong R" + unfolding gen_cong_def[abs_def] by auto + +lemmas imp_gen_cong[intro] = predicate2D[OF leq_gen_cong] + +lemma gen_cong_minimal: "\R \ R'; cong R'\ \ gen_cong R \ R'" + unfolding gen_cong_def[abs_def] by (rule predicate2I) metis + +lemma congdd_base_gen_congdd_base_aux: + "rel (gen_cong R) x y \ R \ R' \ cong R' \ R' (eval x) (eval y)" + by (force simp: rel_fun_def gen_cong_def cong_def dest: spec[of _ R'] predicate2D[OF rel_mono, rotated -1, of _ _ _ R']) + +lemma cong_gen_cong: "cong (gen_cong R)" +proof - + { fix R' x y + have "rel (gen_cong R) x y \ R \ R' \ cong R' \ R' (eval x) (eval y)" + by (force simp: rel_fun_def gen_cong_def cong_def dest: spec[of _ R'] + predicate2D[OF rel_mono, rotated -1, of _ _ _ R']) + } + then show "cong (gen_cong R)" by (auto simp: equivp_gen_cong rel_fun_def gen_cong_def cong_def) +qed + +lemma gen_cong_eval_rel_fun: + "(rel_fun (rel (gen_cong R)) (gen_cong R)) eval eval" + using cong_gen_cong[of R] unfolding cong_def by simp + +lemma gen_cong_eval: + "rel (gen_cong R) x y \ gen_cong R (eval x) (eval y)" + by (erule rel_funD[OF gen_cong_eval_rel_fun]) + +lemma gen_cong_idem: "gen_cong (gen_cong R) = gen_cong R" + by (simp add: antisym cong_gen_cong gen_cong_minimal leq_gen_cong) + +lemma gen_cong_rho: + "\ = eval o f \ rel (gen_cong R) (f x) (f y) \ gen_cong R (\ x) (\ y)" + by (simp add: gen_cong_eval) +lemma coinduction: + assumes coind: "\R. R \ retr R \ R \ op =" + assumes cih: "R \ retr (gen_cong R)" + shows "R \ op =" + apply (rule order_trans[OF leq_gen_cong mp[OF spec[OF coind]]]) + apply (rule self_bounded_weaken_left[OF gen_cong_minimal]) + apply (rule inf_greatest[OF leq_gen_cong cih]) + apply (rule cong_retr[OF cong_gen_cong]) + done + +end + +lemma rel_sum_case_sum: + "rel_fun (rel_sum R S) T (case_sum f1 g1) (case_sum f2 g2) = (rel_fun R T f1 f2 \ rel_fun S T g1 g2)" + by (auto simp: rel_fun_def rel_sum.simps split: sum.splits) + +context + fixes rel eval rel' eval' retr emb + assumes base: "cong rel eval retr" + and step: "cong rel' eval' retr" + and emb: "eval' o emb = eval" + and emb_transfer: "rel_fun (rel R) (rel' R) emb emb" +begin + +interpretation base: cong rel eval retr by (rule base) +interpretation step: cong rel' eval' retr by (rule step) + +lemma gen_cong_emb: "base.gen_cong R \ step.gen_cong R" +proof (rule base.gen_cong_minimal[OF step.leq_gen_cong]) + note step.gen_cong_eval_rel_fun[transfer_rule] emb_transfer[transfer_rule] + have "(rel_fun (rel (step.gen_cong R)) (step.gen_cong R)) eval eval" + unfolding emb[symmetric] by transfer_prover + then show "base.cong (step.gen_cong R)" + by (auto simp: base.cong_def step.equivp_gen_cong) +qed + +end + +ML_file "../Tools/BNF/bnf_gfp_grec_tactics.ML" +ML_file "../Tools/BNF/bnf_gfp_grec.ML" +ML_file "../Tools/BNF/bnf_gfp_grec_sugar_util.ML" +ML_file "../Tools/BNF/bnf_gfp_grec_sugar_tactics.ML" +ML_file "../Tools/BNF/bnf_gfp_grec_sugar.ML" +ML_file "../Tools/BNF/bnf_gfp_grec_unique_sugar.ML" + +method_setup corec_unique = \ + Scan.succeed (SIMPLE_METHOD' o BNF_GFP_Grec_Unique_Sugar.corec_unique_tac) +\ "prove uniqueness of corecursive equation" + +end diff -r 9bfcbab7cd99 -r 0701f25fac39 src/HOL/Library/Library.thy --- a/src/HOL/Library/Library.thy Tue Mar 22 12:39:37 2016 +0100 +++ b/src/HOL/Library/Library.thy Tue Mar 22 12:39:37 2016 +0100 @@ -5,6 +5,7 @@ BigO Bit BNF_Axiomatization + BNF_Corec Boolean_Algebra Bourbaki_Witt_Fixpoint Char_ord diff -r 9bfcbab7cd99 -r 0701f25fac39 src/HOL/Tools/BNF/bnf_gfp_grec.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/BNF/bnf_gfp_grec.ML Tue Mar 22 12:39:37 2016 +0100 @@ -0,0 +1,3234 @@ +(* 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 enforce_type: Proof.context -> (typ -> typ) -> typ -> 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 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_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.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.map_background_naming (mk_internal true lthy Name_Space.concealed) (*TODO check*) + |> primrec [(b, NONE, NoSyn)] (map (pair Attrib.empty_binding) 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_Def.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 inline_pre_bnfs f lthy = + lthy + |> Config.put typedef_threshold ~1 + |> f + |> Config.put typedef_threshold (Config.get lthy typedef_threshold); + +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 T_name = Local_Theory.full_name lthy T_b; + + val tyargs = map2 (fn alive => pair (if alive then SOME Binding.empty else NONE) + o rpair @{sort type}) (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*) + |> inline_pre_bnfs (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 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 => pair (if alive then SOME Binding.empty else NONE) + o rpair @{sort type}) (fp_alives @ [true]) As; + 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*) + |> inline_pre_bnfs (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 + 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_sorry (*FIXME*) (*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 + 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_sorry (*FIXME*) (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} => + mk_dtor_transfer_tac ctxt dtor_rel_thm)) + |> Thm.close_derivation + 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_sorry (*FIXME*) (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} => + mk_transfer_by_transfer_prover_tac ctxt [const_def] rel_eqs transfers)) + |> Thm.close_derivation + 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_sorry (*FIXME*) (*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 + 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_sorry (*FIXME*) (*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 + 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_sorry (*FIXME*) (*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 + 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 + 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 + 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 + 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 + 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 + 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 + end; + +fun derive_eval_thm ctxt dtor_inject dtor_unfold_thm eval_def = + (trans OF [iffD2 OF [dtor_inject, eval_def RS meta_eq_to_obj_eq RS fun_cong], dtor_unfold_thm]) + |> unfold_thms ctxt [o_apply, eval_def RS Drule.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 + 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 + 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 + 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 + 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 + 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 + 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 + 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 [Drule.asm_rl, corecU_pointfree]) + OF [Drule.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 + 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 + 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 + 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 + 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 + 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 + 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 + 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 + 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 + 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 + 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 + 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 + 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 + 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_fp_sugar = fp_sugar_of ctxt @{type_name prod}; + val prod_bnf = #fp_bnf prod_fp_sugar; + + 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 (cutSsig_def RS meta_eq_to_obj_eq) 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 + 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 + 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 + 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 = fold_thms 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 = fold_thms ctxt [cong_def] (thm OF [cong_locale]); + fun instance_of_old_gen thm = fold_thms 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 (algrho_def RS meta_eq_to_obj_eq); + + val gen_cong_emb = + (@{thm gen_cong_emb} OF [old_cong_locale, cong_locale, eval_embL, embL_transfer]) + |> fold_thms 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]) + |> fold_thms lthy [old1_cong_def, cong_def]; + val emb2 = (@{thm gen_cong_emb} OF [old2_cong_locale, cong_locale, eval_embLR, embLR_transfer]) + |> fold_thms 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 num_fp_tyargs = length fpT_args0; + val fp_alives = liveness_of_fp_bnf num_fp_tyargs fp_bnf; + + val (((Es, Fs0), [Y, Z]), names_lthy) = lthy + |> mk_TFrees num_fp_tyargs + ||>> mk_TFrees num_fp_tyargs + ||>> 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 = #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 dead_ssig_rel = mk_rel 1 As Bs (rel_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_uniques 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 = #fp_co_induct_sugar old_ssig_fp_sugar; + val ssig_fp_induct_sugar = #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 dead_ssig_rel = mk_rel 1 res_As res_Bs (rel_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 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_uniques 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 {pre_bnf = live_pre_bnf, fp_bnf = live_fp_bnf, fp_res, ...} = + checked_fp_sugar_of lthy fpT_name; + val num_fp_tyargs = length res_Ds; + val live_fp_alives = liveness_of_fp_bnf num_fp_tyargs live_fp_bnf; + + val ((Ds, [Y, Z]), names_lthy) = lthy + |> mk_TFrees num_fp_tyargs + ||>> 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 = #fp_co_induct_sugar old1_ssig_fp_sugar; + val old2_ssig_fp_induct_sugar = #fp_co_induct_sugar old2_ssig_fp_sugar; + val ssig_fp_induct_sugar = #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_uniques 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 SOME expr = + find_first ((fn {fpT, ...} => Sign.typ_instance thy (res_T, fpT)) o corec_ad_of_expr) exprs; + 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 {pre_bnf, fp_bnf = live_fp_bnf, fp_res, ...} = checked_fp_sugar_of lthy fpT_name; + val num_fp_tyargs = length res_Ds; + val live_fp_alives = liveness_of_fp_bnf num_fp_tyargs live_fp_bnf; + + val fpT_name = fst (dest_Type res_T); + + 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 num_fp_tyargs + ||>> 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)) ^ " used like that in friend")); + + 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 -r 9bfcbab7cd99 -r 0701f25fac39 src/HOL/Tools/BNF/bnf_gfp_grec_sugar.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/BNF/bnf_gfp_grec_sugar.ML Tue Mar 22 12:39:37 2016 +0100 @@ -0,0 +1,2412 @@ +(* 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 -> typ -> thm -> + thm + + val corec_cmd: corec_option list -> (binding * string option * mixfix) list * string -> + local_theory -> local_theory + val corecursive_cmd: 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_Rec_Sugar +open BNF_GFP_Util +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 cong_intros_friendN = "cong_intros_friend"; +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 transfer_rule_attrs = @{attributes [transfer_rule]}; + +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; + +fun unexpected_corec_call ctxt eqns t = + error_at ctxt eqns ("Unexpected corecursive call in " ^ quote (Syntax.string_of_term ctxt t)); +fun unsupported_case_around_corec_call ctxt eqns t = + error_at ctxt eqns ("Unsupported corecursive call under case expression " ^ + quote (Syntax.string_of_term ctxt t) ^ "\n(Define " ^ + quote (Syntax.string_of_typ ctxt (domain_type (fastype_of t))) ^ + " with discriminators and selectors to circumvent this limitation.)"); + +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_tab: thm list Symtab.table}; + +fun morph_coinduct_extra phi ({coinduct, coinduct_attrs, cong_intro_tab} : coinduct_extra) = + {coinduct = Morphism.thm phi coinduct, coinduct_attrs = coinduct_attrs, + cong_intro_tab = Symtab.map (K (Morphism.fact phi)) cong_intro_tab}; + +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_tab #> Symtab.dest #> maps snd); +fun get_cong_intros fpT_name name context = + coinduct_extras_of_generic context fpT_name + |> maps (#cong_intro_tab #> (fn tab => Symtab.lookup_list tab 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, 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, _])) = 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 dtor = nth dtors fp_res_index; + 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 + end; + +fun derive_case_trivial ctxt fpT_name = + let + val SOME {casex, exhaust, case_thms, ...} = ctr_sugar_of ctxt fpT_name; + + val fpT0 as Type (_, As0) = domain_type (body_fun_type (fastype_of casex)); + + val (As, _) = ctxt + |> mk_TFrees (length 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' = Drule.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 + end; + +fun mk_abs_rep_transfers ctxt fpT_name = + [mk_abs_transfer ctxt fpT_name, mk_rep_transfer ctxt fpT_name] + handle Fail _ => []; + +fun set_transfer_rule_attrs thms = + snd o Local_Theory.notes [((Binding.empty, []), [(thms, transfer_rule_attrs)])]; + +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; + +(*TODO: Merge with primcorec "case_of"*) +fun case_of ctxt fcT_name = + (case ctr_sugar_of ctxt fcT_name of + SOME {casex = Const (s, _), ...} => SOME s + | _ => NONE); + +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, ...} res_T 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 + end; + +fun derive_unique ctxt phi code_goal + {sig_fp_sugars, ssig_fp_sugar, eval, eval_simps, all_algLam_algs, corecUU_unique, ...} + (res_T as Type (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 + 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 + 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_sorry (*FIXME*) (*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 + 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; + + 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_sorry (*FIXME*) (*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 + 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 = + BNF_Def.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; + + 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 + 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 + Symtab.make_list (names ~~ thms) + end; + +fun derive_coinduct ctxt (fpT as Type (fpT_name, _)) 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 = T0, 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] [fpT]; + 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 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] ctxt + |> 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 + error (quote (Syntax.string_of_term ctxt fun_t) ^ + " is not the function currently being defined"); + + fun check_args_are_vars () = + let + fun is_ok_Free_or_Var (Free (s, _)) = not (String.isSuffix inner_fp_suffix s) + | is_ok_Free_or_Var (Var _) = true + | is_ok_Free_or_Var _ = false; + + fun is_valid arg = + (is_ok_Free_or_Var arg andalso not (member (op aconv) fun_frees arg)) orelse + error ("Argument " ^ quote (Syntax.string_of_term ctxt arg) ^ " is not free"); + in + forall is_valid arg_ts + end; + + fun check_no_duplicate_arg () = + (case duplicates (op =) arg_ts of + [] => () + | arg :: _ => error ("Repeated argument: " ^ quote (Syntax.string_of_term ctxt arg))); + + fun check_no_other_frees () = + let + val known_frees = fun_frees @ arg_ts; + + fun check_free (t as Free (s, _)) = + Variable.is_fixed ctxt s orelse member (op aconv) known_frees t orelse + error ("Unexpected variable: " ^ quote s) + | check_free _ = false; + in + Term.exists_subterm check_free rhs + end; + in + check_no_duplicate_arg (); + check_fun_name (); + check_args_are_vars (); + 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 _ => error "Expected equation"; + + 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) + |> type_of + |> dest_funT + |-> BNF_GFP_Grec_Sugar_Util.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; + (* FIXME: This looks suspicious *) + val Us = map (fpT_to ssig_T) (snd (dest_Type (fastype_of1 (bound_Us, mapped_arg')))); + 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 + | 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 (fold (curry op @) 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_polymorphic_const 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 = maybe_const_type lthy fun_x in + if 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); + val t' = build_function_after_encapsulation fun_t fun_t' params arg_ts arg_ts'; + in + (case try type_of1 (bound_Us, t') of + SOME _ => + (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') + | 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_polymorphic_const] + 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 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 + error ("Unallowed corecursive call in condition " ^ quote (Syntax.string_of_term 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-lhs argument " ^ + 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-lhs argument " ^ + 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 + error ("Constructor expected on right-hand side, " ^ + quote (Syntax.string_of_term lthy fun_t) ^ " found instead") + 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 fun_T k_T + ssig_T ssig_fp_sugar friend_parse_info fun_t parsed_eq lthy = + let + val Type (fpT_name, _) = body_type fun_T; + + 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 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 name 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.name name), NONE, NoSyn)] + [(apfst Binding.concealed Attrib.empty_binding, 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 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 = (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 + 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_tab = 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_tab = cong_intro_tab}; + 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 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_spec raw_fixes [(Attrib.empty_binding, raw_eq)] lthy); + + val _ = Sign.of_sort (Proof_Context.theory_of lthy) (fun_T, @{sort type}) orelse + error ("Type of " ^ Binding.print b ^ " contains top sort"); + + 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 lthy = 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 fun_T k_T + ssig_T ssig_fp_sugar friend_parse_info fun_t parsed_eq' + |>> pair prepared + end; + + val ((prepareds, (rho_datas, transfer_goal_datas)), lthy) = + if friend then extract_rho lthy |>> (apfst single ##> (apfst single #> apsnd single)) + else (([], ([], [])), lthy); + + val ((buffer, corec_infos), lthy) = + if friend then + ((#13 (the_single prepareds), []), lthy) + else + corec_info_of res_T lthy + ||> no_base ? update_coinduct_cong_intross_dynamic fpT_name + |>> (fn info as {buffer, ...} => (buffer, [info])); + + val corec_parse_info = corec_parse_info_of lthy arg_Ts res_T buffer; + + val explored_eq = + explore_corec_equation lthy true friend fun_name fun_free corec_parse_info res_T parsed_eq; + + val (((inner_fp_triple, termin_goals), corecUU_arg), lthy) = + build_corecUU_arg_and_goals (not long_cmd) fun_free explored_eq lthy; + + fun def_fun (inner_fp_elims0, inner_fp_inducts0, inner_fp_simps0) const_transfers + rho_transfers_foldeds lthy = + 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, lthy) = lthy |> (if friend then register_friend #>> single else pair []); + val (corec_info as {corecUU = corecUU0, ...}, lthy) = + (case corec_infos of + [] => corec_info_of res_T lthy + | [info] => (info, lthy)); + + val def_rhs = mk_corec_fun_def_rhs lthy arg_Ts corecUU0 corecUU_arg; + val def = ((b, mx), ((Binding.concealed (Thm.def_binding b), []), def_rhs)); + + val ((fun_t0, (_, fun_def0)), (lthy, lthy_old)) = lthy + |> Local_Theory.open_target |> snd + |> Local_Theory.define def + ||> `Local_Theory.close_target; + + val parsed_eq = parse_corec_equation lthy [fun_free] eq; + val views0 = generate_views lthy eq fun_free parsed_eq; + + val lthy' = lthy |> fold Variable.declare_typ (res_T :: arg_Ts); + val phi = Proof_Context.export_morphism lthy_old lthy'; + + val fun_t = Morphism.term phi fun_t0; (* FIXME: shadows "fun_t" -- identical? *) + 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_t 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 lthy = lthy |> friend ? derive_and_note_friend_extra_theorems; + + val code_thm = derive_code lthy inner_fp_simps code_goal corec_info res_T fun_t 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 lthy phi (#1 views0) corec_info res_T 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_tab, coinduct, coinduct_attrs}]), lthy) = lthy + |> derive_and_update_coinduct_cong_intross [corec_info]; + val cong_intros_pairs = Symtab.dest cong_intro_tab; + + val code_attrs = if plugins code_plugin then [Code.add_default_eqn_attrib] else []; + + 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], code_attrs @ 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 lthy 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 + lthy +(* TODO: + |> Spec_Rules.add Spec_Rules.Equational ([fun_t0], flat sel_thmss) + |> Spec_Rules.add Spec_Rules.Equational ([fun_t0], flat ctr_thmss) +*) + |> Spec_Rules.add Spec_Rules.Equational ([fun_t0], [code_thm]) + |> Local_Theory.notes (anonymous_notes @ notes) + |> snd + end; + + fun prove_transfer_goal ctxt goal = + Variable.add_free_names ctxt goal [] + |> (fn vars => Goal.prove_sorry (*FIXME*) (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} => + HEADGOAL (Transfer.transfer_prover_tac ctxt))) + |> Thm.close_derivation; + + 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 lthy) 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'))), lthy) + end; + +fun corec_cmd 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 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 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 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, default_fun_T0) = + Proof_Context.read_const {proper = true, strict = false} lthy raw_fun_name; + val fun_T = + (case raw_fun_T_opt of + SOME raw_T => Syntax.read_typ lthy raw_T + | NONE => singleton (freeze_types lthy []) default_fun_T0); + + val fun_t = Const (fun_name, fun_T); + val fun_b = Binding.name (Long_Name.base_name fun_name); + + val fake_lthy = lthy |> Proof_Context.add_const_constraint (fun_name, SOME fun_T) + handle TYPE (msg, _, _) => error msg; + + val code_goal = Syntax.read_prop fake_lthy raw_eq; + + val (arg_Ts, res_T as Type (fpT_name, _)) = strip_type fun_T; + + val no_base = has_no_corec_info lthy fpT_name; + val lthy = lthy |> no_base ? setup_base fpT_name; + + val lthy = lthy |> 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), 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 = parse_corec_equation lthy [] code_goal; + + val (((rho, rho_def), (const_transfer_goals, rho_transfer_goal)), lthy) = + extract_rho_return_transfer_goals fun_b version dead_pre_bnf dead_k_bnf Y Z preT fun_T k_T + ssig_T ssig_fp_sugar friend_parse_info fun_t parsed_eq lthy; + + fun register_friend_extra_and_note_thms code_goal code_thm const_transfers k_T friend_info + lthy = + let + val (corec_info, lthy) = corec_info_of res_T lthy; + + val fun_free = Free (Binding.name_of fun_b, fun_T); + + fun freeze_fun (t as Const (s, _)) = if s = fun_name then fun_free else t + | freeze_fun t = t; + + val eq = Term.map_aterms freeze_fun code_goal; + val parsed_eq = parse_corec_equation lthy [fun_free] eq; + + val corec_parse_info = corec_parse_info_of lthy arg_Ts res_T buffer; + val explored_eq = explore_corec_equation lthy 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 lthy; + + val eq_corecUU = derive_eq_corecUU lthy corec_info fun_t corecUU_arg code_thm; + val (eq_algrho, algrho_eq) = derive_eq_algrho lthy corec_info friend_info fun_t k_T + code_goal const_transfers rho_def eq_corecUU; + + val ((_, [{cong_intro_tab, coinduct, coinduct_attrs}]), lthy) = lthy + |> 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 = Symtab.dest cong_intro_tab; + + val unique = derive_unique lthy Morphism.identity code_goal corec_info res_T eq_corecUU; + + val notes = + [(cong_intros_friendN, maps snd cong_intros_pairs, []), + (codeN, [code_thm], []), + (coinductN, [coinduct], coinduct_attrs), + (uniqueN, [unique], [])] @ + map (fn (thmN, thms) => (thmN, thms, [])) cong_intros_pairs @ + (if Config.get lthy 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 + lthy + |> Local_Theory.notes notes |> snd + end; + + val (rho_transfer_goal', unprime_rho_transfer_and_fold) = + prime_rho_transfer_goal lthy fpT_name rho_def rho_transfer_goal; + in + lthy + |> 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, ...}, lthy) = lthy + |> corec_info_of fpT; + val lthy = lthy |> no_base ? setup_base fpT_name; + + val ((changed, [{cong_intro_tab, coinduct, coinduct_attrs}]), lthy) = lthy + |> derive_and_update_coinduct_cong_intross [corec_info]; + val lthy = lthy |> (changed orelse no_base) ? update_coinduct_cong_intross_dynamic fpT_name; + val cong_intros_pairs = Symtab.dest cong_intro_tab; + + 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 + lthy |> 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.fixes --| Parse.where_ -- Parse.prop) + >> uncurry corec_cmd); + +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.fixes --| Parse.where_ -- Parse.prop) + >> uncurry corecursive_cmd); + +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 (Parse.$$$ "::" |-- 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 --| Parse.$$$ ":" -- Parse.typ >> coinduction_upto_cmd); + +val _ = Theory.setup (Theory.at_begin consolidate_global); + +end; diff -r 9bfcbab7cd99 -r 0701f25fac39 src/HOL/Tools/BNF/bnf_gfp_grec_sugar_tactics.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/BNF/bnf_gfp_grec_sugar_tactics.ML Tue Mar 22 12:39:37 2016 +0100 @@ -0,0 +1,217 @@ +(* Title: HOL/Tools/BNF/bnf_gfp_grec_tactics.ML + Author: Jasmin Blanchette, Inria, LORIA, MPII + Copyright 2016 + +Tactics for generalized corecursor sugar. +*) + +signature BNF_GFP_GREC_SUGAR_TACTICS = +sig + val rho_transfer_simps: thm list + + val mk_case_dtor_tac: Proof.context -> term -> thm -> thm -> thm list -> thm -> thm list -> tactic + val mk_cong_intro_ctr_or_friend_tac: Proof.context -> thm -> thm list -> thm -> tactic + val mk_code_tac: Proof.context -> int -> term list -> term -> term -> thm -> thm -> thm list -> + thm list -> thm list -> thm list -> thm -> thm -> thm list -> thm list -> thm -> thm list -> + thm list -> thm list -> thm list -> thm list -> thm list -> thm -> tactic + val mk_eq_algrho_tac: Proof.context -> term list -> term -> term -> term -> term -> term -> thm -> + thm -> thm list -> thm list -> thm list -> thm list -> thm -> thm -> thm -> thm list -> + thm list -> thm list -> thm -> thm list -> thm list -> thm list -> thm -> thm -> thm -> thm -> + thm list -> thm list -> thm list -> thm list -> thm list -> tactic + val mk_eq_corecUU_tac: Proof.context -> int -> term list -> term -> term -> thm -> thm -> + thm list -> thm list -> thm list -> thm list -> thm -> thm -> thm list -> thm list -> + thm list -> thm list -> thm list -> thm list -> thm list -> thm -> thm -> tactic + val mk_last_disc_tac: Proof.context -> term -> thm -> thm list -> tactic + val mk_rho_transfer_tac: Proof.context -> bool -> thm -> thm list -> tactic + val mk_unique_tac: Proof.context -> int -> term list -> term -> term -> thm -> thm -> thm list -> + thm list -> thm list -> thm list -> thm -> thm -> thm list -> thm list -> thm list -> + thm list -> thm list -> thm list -> thm list -> thm -> thm -> tactic +end; + +structure BNF_GFP_Grec_Sugar_Tactics : BNF_GFP_GREC_SUGAR_TACTICS = +struct + +open Ctr_Sugar +open BNF_Util +open BNF_Tactics +open BNF_FP_Def_Sugar_Tactics +open BNF_GFP_Grec_Tactics +open BNF_GFP_Grec_Sugar_Util + +fun apply_func f = + let + val arg_Ts = binder_fun_types (fastype_of f); + val args = map_index (fn (j, T) => Var (("a" ^ string_of_int j, 0), T)) arg_Ts; + in + list_comb (f, args) + end; + +fun instantiate_distrib thm ctxt t = + Drule.infer_instantiate' ctxt [SOME (Thm.incr_indexes_cterm 1 (Thm.cterm_of ctxt t))] thm; + +(* TODO (here and elsewhere): Use metaequality in goal instead and keep uninstianted version of + theorem? *) +val mk_if_distrib_of = instantiate_distrib @{thm if_distrib}; +val mk_case_sum_distrib_of = instantiate_distrib @{thm sum.case_distrib}; + +fun mk_case_dtor_tac ctxt u abs_inverse dtor_ctor ctr_defs exhaust cases = + let val exhaust' = Drule.infer_instantiate' ctxt [SOME (Thm.cterm_of ctxt u)] exhaust in + HEADGOAL (rtac ctxt exhaust') THEN + REPEAT_DETERM (HEADGOAL (etac ctxt ssubst THEN' + SELECT_GOAL (unfold_thms_tac ctxt cases THEN + unfold_thms_tac ctxt (abs_inverse :: dtor_ctor :: ctr_defs @ + @{thms prod.case sum.case})) THEN' + rtac ctxt refl)) + end; + +fun mk_cong_intro_ctr_or_friend_tac ctxt ctr_or_friend_def extra_simps cong_alg_intro = + HEADGOAL (REPEAT_DETERM_N 2 o subst_tac ctxt NONE [ctr_or_friend_def] THEN' + rtac ctxt cong_alg_intro) THEN + unfold_thms_tac ctxt (extra_simps @ sumprod_thms_rel @ + @{thms vimage2p_def prod.rel_eq sum.rel_eq}) THEN + REPEAT_DETERM (HEADGOAL (rtac ctxt conjI ORELSE' assume_tac ctxt ORELSE' rtac ctxt refl)); + +val shared_simps = + @{thms map_prod_simp map_sum.simps sum.case prod.case_eq_if split_beta' prod.sel + sum.disc(1)[THEN eq_True[THEN iffD2]] sum.disc(2)[THEN eq_False[THEN iffD2]] sum.sel + isl_map_sum map_sum_sel if_True if_False if_True_False Let_def[abs_def] o_def[abs_def] id_def + BNF_Composition.id_bnf_def}; + +fun 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 all_sig_maps + ssig_map_thms all_algLam_alg_pointfuls all_algrho_eqs eval_simps inner_fp_simps fun_def = + let + val fun_def' = + if null inner_fp_simps andalso num_args > 0 then + fun_def RS meta_eq_to_obj_eq RS (mk_curry_uncurryN_balanced ctxt num_args RS iffD2) RS sym + else + fun_def; + val case_trivial' = unfold_thms ctxt (case_eq_ifs @ ctr_defs @ shared_simps) case_trivial; + val case_eq_ifs' = map (Drule.abs_def o (fn thm => thm RS eq_reflection)) case_eq_ifs; + val if_distribs = @{thm if_distrib_fun} :: map (mk_if_distrib_of ctxt) + (eval :: map apply_func (ssig_map :: fpsig_nesting_maps)); + + val unfold_tac = unfold_thms_tac ctxt (case_trivial' :: pre_map_def :: abs_inverse :: + fp_map_ident :: (if null inner_fp_simps then [] else [corecUU]) @ fpsig_nesting_map_ident0s @ + fpsig_nesting_map_comps @ fpsig_nesting_map_thms @ live_nesting_map_ident0s @ ctr_defs @ + case_eq_ifs' @ all_sig_maps @ ssig_map_thms @ all_algLam_alg_pointfuls @ all_algrho_eqs @ + eval_simps @ if_distribs @ shared_simps); + in + HEADGOAL (subst_tac ctxt NONE [fun_def] THEN' subst_tac ctxt NONE [corecUU] THEN' + (if null inner_fp_simps then K all_tac else subst_tac ctxt NONE inner_fp_simps)) THEN + unfold_thms_tac ctxt [fun_def'] THEN + unfold_tac THEN HEADGOAL (CONVERSION Thm.eta_long_conversion) THEN unfold_tac THEN + HEADGOAL (rtac ctxt refl) + end; + +fun 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_maps ssig_map_thms + all_algLam_alg_pointfuls all_algrho_eqs eval_simps = + let + fun mk_abs_def thm = Drule.abs_def (thm RS eq_reflection); + + val nullary_disc_defs' = nullary_disc_defs + |> map (fn thm => thm RS sym) + |> maps (fn thm => [thm, thm RS @{thm subst[OF eq_commute, of "%e. e = z" for z]}]); + + val case_dtor' = unfold_thms ctxt shared_simps case_dtor; + val disc_sel_eq_cases' = map (mk_abs_def + o unfold_thms ctxt (case_dtor' :: ctr_defs @ shared_simps)) disc_sel_eq_cases; + val const_pointful_naturals' = map (unfold_thms ctxt shared_simps) const_pointful_naturals; + val const_pointful_naturals_sym' = map (fn thm => thm RS sym) const_pointful_naturals'; + val case_eq_ifs' = map mk_abs_def (@{thm sum.case_eq_if} :: case_eq_ifs); + + val distrib_consts = + abs :: rep :: ctor :: eval :: map apply_func (ssig_map :: fpsig_nesting_maps); + val if_distribs = @{thm if_distrib_fun} :: map (mk_if_distrib_of ctxt) distrib_consts; + val case_sum_distribs = map (mk_case_sum_distrib_of ctxt) distrib_consts; + + val simp_ctxt = (ctxt + |> Context_Position.set_visible false + |> put_simpset (simpset_of (Proof_Context.init_global @{theory Main})) + |> Raw_Simplifier.add_cong @{thm if_cong}) + addsimps pre_map_def :: abs_inverse :: fp_map_ident :: dtor_ctor :: rho_def :: + @{thm convol_def} :: fpsig_nesting_map_ident0s @ fpsig_nesting_map_comps @ + fpsig_nesting_map_thms @ live_nesting_map_ident0s @ ctr_defs @ disc_sel_eq_cases' @ + fp_nesting_k_map_disc_sels' @ case_eq_ifs' @ all_sig_maps @ ssig_map_thms @ + all_algLam_alg_pointfuls @ all_algrho_eqs @ eval_simps @ if_distribs @ case_sum_distribs @ + shared_simps; + + fun mk_main_simp const_pointful_naturals_maybe_sym' = + simp_tac (simp_ctxt addsimps const_pointful_naturals_maybe_sym'); + in + unfold_thms_tac ctxt [eq_corecUU] THEN + HEADGOAL (REPEAT_DETERM o rtac ctxt ext THEN' + rtac ctxt (corecUU_unique RS sym RS fun_cong) THEN' + subst_tac ctxt NONE [dtor_algrho RS (ctor_iff_dtor RS iffD2)] THEN' rtac ctxt ext) THEN + unfold_thms_tac ctxt (nullary_disc_defs' @ shared_simps) THEN + HEADGOAL (rtac ctxt meta_eq_to_obj_eq) THEN + REPEAT_DETERM_N (length const_pointful_naturals' + 1) + (ALLGOALS (mk_main_simp const_pointful_naturals_sym') THEN + ALLGOALS (mk_main_simp const_pointful_naturals')) + end; + +fun 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_maps + ssig_map_thms all_algLam_alg_pointfuls all_algrho_eqs eval_simps corecUU_unique fun_code = + let + val case_trivial' = unfold_thms ctxt (case_eq_ifs @ ctr_defs @ shared_simps) case_trivial; + val case_eq_ifs' = map (Drule.abs_def o (fn thm => thm RS eq_reflection)) case_eq_ifs; + val if_distribs = @{thm if_distrib_fun} :: map (mk_if_distrib_of ctxt) + (eval :: map apply_func (ssig_map :: fpsig_nesting_maps)); + + val unfold_tac = unfold_thms_tac ctxt (case_trivial' :: pre_map_def :: abs_inverse :: + fp_map_ident :: fpsig_nesting_map_ident0s @ fpsig_nesting_map_comps @ fpsig_nesting_map_thms @ + live_nesting_map_ident0s @ ctr_defs @ case_eq_ifs' @ all_sig_maps @ ssig_map_thms @ + all_algLam_alg_pointfuls @ all_algrho_eqs @ eval_simps @ if_distribs @ shared_simps); + in + HEADGOAL (rtac ctxt (mk_curry_uncurryN_balanced ctxt num_args RS iffD1) THEN' + rtac ctxt corecUU_unique THEN' rtac ctxt ext) THEN + unfold_thms_tac ctxt @{thms prod.case_eq_if} THEN + HEADGOAL (rtac ctxt (fun_code RS trans)) THEN + unfold_tac THEN HEADGOAL (CONVERSION Thm.eta_long_conversion) THEN unfold_tac THEN + HEADGOAL (rtac ctxt refl) + end; + +fun mk_last_disc_tac ctxt u exhaust discs' = + let val exhaust' = Drule.infer_instantiate' ctxt [SOME (Thm.cterm_of ctxt u)] exhaust in + HEADGOAL (rtac ctxt exhaust') THEN + REPEAT_DETERM (HEADGOAL (etac ctxt ssubst THEN' + simp_tac (ss_only (distinct Thm.eq_thm discs' @ @{thms simp_thms}) ctxt))) + end; + +val rho_transfer_simps = @{thms BNF_Def.vimage2p_def[abs_def] BNF_Composition.id_bnf_def}; + +fun mk_rho_transfer_tac ctxt unfold rel_def const_transfers = + (if unfold then unfold_thms_tac ctxt (rel_def :: rho_transfer_simps) else all_tac) THEN + HEADGOAL (transfer_prover_add_tac ctxt [] const_transfers); + +fun 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_maps + ssig_map_thms all_algLam_alg_pointfuls all_algrho_eqs eval_simps corecUU_unique eq_corecUU = + let + val case_trivial' = unfold_thms ctxt (case_eq_ifs @ ctr_defs @ shared_simps) case_trivial; + val case_eq_ifs' = map (Drule.abs_def o (fn thm => thm RS eq_reflection)) case_eq_ifs; + val if_distribs = @{thm if_distrib_fun} :: map (mk_if_distrib_of ctxt) + (eval :: map apply_func (ssig_map :: fpsig_nesting_maps)); + + val unfold_tac = unfold_thms_tac ctxt (case_trivial' :: pre_map_def :: abs_inverse :: + fp_map_ident :: fpsig_nesting_map_ident0s @ fpsig_nesting_map_comps @ fpsig_nesting_map_thms @ + live_nesting_map_ident0s @ ctr_defs @ case_eq_ifs' @ all_sig_maps @ ssig_map_thms @ + all_algLam_alg_pointfuls @ all_algrho_eqs @ eval_simps @ if_distribs @ shared_simps); + in + HEADGOAL (subst_tac ctxt NONE [eq_corecUU] THEN' + rtac ctxt (mk_curry_uncurryN_balanced ctxt num_args RS iffD1) THEN' + rtac ctxt corecUU_unique THEN' rtac ctxt ext THEN' + etac ctxt @{thm ssubst[of _ _ "\x. f x = u" for f u]}) THEN + unfold_tac THEN HEADGOAL (CONVERSION Thm.eta_long_conversion) THEN unfold_tac THEN + HEADGOAL (rtac ctxt refl) + end; + +end; diff -r 9bfcbab7cd99 -r 0701f25fac39 src/HOL/Tools/BNF/bnf_gfp_grec_sugar_util.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/BNF/bnf_gfp_grec_sugar_util.ML Tue Mar 22 12:39:37 2016 +0100 @@ -0,0 +1,481 @@ +(* Title: HOL/Tools/BNF/bnf_gfp_grec_sugar_util.ML + Author: Aymeric Bouzy, Ecole polytechnique + Author: Jasmin Blanchette, Inria, LORIA, MPII + Copyright 2015, 2016 + +Library for generalized corecursor sugar. +*) + +signature BNF_GFP_GREC_SUGAR_UTIL = +sig + type s_parse_info = + {outer_buffer: BNF_GFP_Grec.buffer, + ctr_guards: term Symtab.table, + inner_buffer: BNF_GFP_Grec.buffer} + + type rho_parse_info = + {pattern_ctrs: (term * term list) Symtab.table, + discs: term Symtab.table, + sels: term Symtab.table, + it: term, + mk_case: typ -> term} + + exception UNNATURAL of unit + + val generalize_types: int -> typ -> typ -> typ + val mk_curry_uncurryN_balanced: Proof.context -> int -> thm + val mk_const_transfer_goal: Proof.context -> string * typ -> term + val mk_abs_transfer: Proof.context -> string -> thm + val mk_rep_transfer: Proof.context -> string -> thm + val mk_pointful_natural_from_transfer: Proof.context -> thm -> thm + + val corec_parse_info_of: Proof.context -> typ list -> typ -> BNF_GFP_Grec.buffer -> s_parse_info + val friend_parse_info_of: Proof.context -> typ list -> typ -> BNF_GFP_Grec.buffer -> + s_parse_info * rho_parse_info +end; + +structure BNF_GFP_Grec_Sugar_Util : BNF_GFP_GREC_SUGAR_UTIL = +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_GFP_Grec +open BNF_GFP_Grec_Tactics + +val mk_case_sumN_balanced = Balanced_Tree.make mk_case_sum; + +fun not_codatatype ctxt T = + error ("Not a codatatype: " ^ Syntax.string_of_typ ctxt T); + +fun generalize_types max_j T U = + let + val vars = Unsynchronized.ref []; + + fun var_of T U = + (case AList.lookup (op =) (!vars) (T, U) of + SOME V => V + | NONE => + let val V = TVar ((Name.aT, length (!vars) + max_j), @{sort type}) in + vars := ((T, U), V) :: !vars; V + end); + + fun gen (T as Type (s, Ts)) (U as Type (s', Us)) = + if s = s' then Type (s, map2 gen Ts Us) else var_of T U + | gen T U = if T = U then T else var_of T U; + in + gen T U + end; + +fun mk_curry_uncurryN_balanced_raw ctxt n = + let + val ((As, B), names_ctxt) = ctxt + |> mk_TFrees (n + 1) + |>> split_last; + + val tupled_As = mk_tupleT_balanced As; + + val f_T = As ---> B; + val g_T = tupled_As --> B; + + val (((f, g), xs), _) = names_ctxt + |> yield_singleton (mk_Frees "f") f_T + ||>> yield_singleton (mk_Frees "g") g_T + ||>> mk_Frees "x" As; + + val tupled_xs = mk_tuple1_balanced As xs; + + val uncurried_f = mk_tupled_fun f tupled_xs xs; + val curried_g = abs_curried_balanced As g; + + val lhs = HOLogic.mk_eq (uncurried_f, g); + val rhs = HOLogic.mk_eq (f, curried_g); + val goal = fold_rev Logic.all [f, g] (mk_Trueprop_eq (lhs, rhs)); + + fun mk_tac ctxt = + HEADGOAL (rtac ctxt iffI THEN' dtac ctxt sym THEN' hyp_subst_tac ctxt) THEN + unfold_thms_tac ctxt @{thms prod.case} THEN + HEADGOAL (rtac ctxt refl THEN' hyp_subst_tac ctxt THEN' + REPEAT_DETERM o subst_tac ctxt NONE @{thms unit_abs_eta_conv case_prod_eta} THEN' + rtac ctxt refl); + in + Goal.prove_sorry ctxt [] [] goal (fn {context = ctxt, ...} => mk_tac ctxt) + |> Thm.close_derivation + end; + +val num_curry_uncurryN_balanced_precomp = 8; +val curry_uncurryN_balanced_precomp = + map (mk_curry_uncurryN_balanced_raw @{context}) (0 upto num_curry_uncurryN_balanced_precomp); + +fun mk_curry_uncurryN_balanced ctxt n = + if n <= num_curry_uncurryN_balanced_precomp then nth curry_uncurryN_balanced_precomp n + else mk_curry_uncurryN_balanced_raw ctxt n; + +fun mk_const_transfer_goal ctxt (s, var_T) = + let + val var_As = Term.add_tvarsT var_T []; + + val ((As, Bs), names_ctxt) = ctxt + |> Variable.declare_typ var_T + |> mk_TFrees' (map snd var_As) + ||>> mk_TFrees' (map snd var_As); + + val (Rs, _) = names_ctxt + |> mk_Frees "R" (map2 mk_pred2T As Bs); + + val T = Term.typ_subst_TVars (map fst var_As ~~ As) var_T; + val U = Term.typ_subst_TVars (map fst var_As ~~ Bs) var_T; + in + mk_parametricity_goal ctxt Rs (Const (s, T)) (Const (s, U)) + |> tap (fn goal => can type_of goal orelse + error ("Cannot transfer constant " ^ quote (Syntax.string_of_term ctxt (Const (s, T))) ^ + " from type " ^ quote (Syntax.string_of_typ ctxt T) ^ " to " ^ + quote (Syntax.string_of_typ ctxt U))) + end; + +fun mk_abs_transfer ctxt fpT_name = + let + val SOME {pre_bnf, absT_info = {absT, repT, abs, type_definition, ...}, ...} = + fp_sugar_of ctxt fpT_name; + in + if absT = repT then + raise Fail "no abs/rep" + else + let + val rel_def = rel_def_of_bnf pre_bnf; + + val absT = T_of_bnf pre_bnf + |> singleton (freeze_types ctxt (map dest_TVar (lives_of_bnf pre_bnf))); + + val goal = mk_const_transfer_goal ctxt (dest_Const (mk_abs absT abs)) + in + Variable.add_free_names ctxt goal [] + |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => + unfold_thms_tac ctxt [rel_def] THEN + HEADGOAL (rtac ctxt (@{thm Abs_transfer} OF [type_definition, type_definition])))) + end + end; + +fun mk_rep_transfer ctxt fpT_name = + let + val SOME {pre_bnf, absT_info = {absT, repT, rep, ...}, ...} = fp_sugar_of ctxt fpT_name; + in + if absT = repT then + raise Fail "no abs/rep" + else + let + val rel_def = rel_def_of_bnf pre_bnf; + + val absT = T_of_bnf pre_bnf + |> singleton (freeze_types ctxt (map dest_TVar (lives_of_bnf pre_bnf))); + + val goal = mk_const_transfer_goal ctxt (dest_Const (mk_rep absT rep)) + in + Variable.add_free_names ctxt goal [] + |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => + unfold_thms_tac ctxt [rel_def] THEN + HEADGOAL (rtac ctxt @{thm vimage2p_rel_fun}))) + end + end; + +exception UNNATURAL of unit; + +fun mk_pointful_natural_from_transfer ctxt transfer = + let + val _ $ (_ $ Const (s, T0) $ Const (_, U0)) = Thm.prop_of transfer; + val [T, U] = freeze_types ctxt [] [T0, U0]; + val var_T = generalize_types 0 T U; + + val var_As = map TVar (rev (Term.add_tvarsT var_T [])); + + val ((As, Bs), names_ctxt) = ctxt + |> mk_TFrees' (map Type.sort_of_atyp var_As) + ||>> mk_TFrees' (map Type.sort_of_atyp var_As); + + val TA = typ_subst_atomic (var_As ~~ As) var_T; + + val ((xs, fs), _) = names_ctxt + |> mk_Frees "x" (binder_types TA) + ||>> mk_Frees "f" (map2 (curry (op -->)) As Bs); + + val AB_fs = (As ~~ Bs) ~~ fs; + + fun build_applied_map TU t = + if op = TU then + t + else + (case try (build_map ctxt [] (the o AList.lookup (op =) AB_fs)) TU of + SOME mapx => mapx $ t + | NONE => raise UNNATURAL ()); + + fun unextensionalize (f $ (x as Free _), rhs) = unextensionalize (f, lambda x rhs) + | unextensionalize tu = tu; + + val TB = typ_subst_atomic (var_As ~~ Bs) var_T; + + val (binder_TAs, body_TA) = strip_type TA; + val (binder_TBs, body_TB) = strip_type TB; + + val n = length var_As; + val m = length binder_TAs; + + val A_nesting_bnfs = nesting_bnfs ctxt [[body_TA :: binder_TAs]] As; + val A_nesting_map_ids = map map_id_of_bnf A_nesting_bnfs; + val A_nesting_rel_Grps = map rel_Grp_of_bnf A_nesting_bnfs; + + val ta = Const (s, TA); + val tb = Const (s, TB); + val xfs = @{map 3} (curry build_applied_map) binder_TAs binder_TBs xs; + + val goal = (list_comb (tb, xfs), build_applied_map (body_TA, body_TB) (list_comb (ta, xs))) + |> unextensionalize |> mk_Trueprop_eq; + + val _ = if can type_of goal then () else raise UNNATURAL (); + + val vars = map (fst o dest_Free) (xs @ fs); + in + Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => + mk_natural_from_transfer_tac ctxt m (replicate n true) transfer A_nesting_map_ids + A_nesting_rel_Grps []) + |> Thm.close_derivation + end; + +type s_parse_info = + {outer_buffer: BNF_GFP_Grec.buffer, + ctr_guards: term Symtab.table, + inner_buffer: BNF_GFP_Grec.buffer}; + +type rho_parse_info = + {pattern_ctrs: (term * term list) Symtab.table, + discs: term Symtab.table, + sels: term Symtab.table, + it: term, + mk_case: typ -> term}; + +fun curry_friend (T, t) = + let + val prod_T = domain_type (fastype_of t); + val Ts = dest_tupleT_balanced (num_binder_types T) prod_T; + val xs = map_index (fn (i, T) => Free ("x" ^ string_of_int i, T)) Ts; + val body = mk_tuple_balanced xs; + in + (T, fold_rev Term.lambda xs (t $ body)) + end; + +fun curry_friends ({Oper, VLeaf, CLeaf, ctr_wrapper, friends} : buffer) = + {Oper = Oper, VLeaf = VLeaf, CLeaf = CLeaf, ctr_wrapper = ctr_wrapper, + friends = Symtab.map (K curry_friend) friends}; + +fun checked_gfp_sugar_of lthy (T as Type (T_name, _)) = + (case fp_sugar_of lthy T_name of + SOME (sugar as {fp = Greatest_FP, ...}) => sugar + | _ => not_codatatype lthy T) + | checked_gfp_sugar_of lthy T = not_codatatype lthy T; + +fun generic_spec_of friend ctxt arg_Ts res_T (raw_buffer0 as {VLeaf = VLeaf0, ...}) = + let + val thy = Proof_Context.theory_of ctxt; + + val tupled_arg_T = mk_tupleT_balanced arg_Ts; + + val {T = fpT, X, fp_res_index, fp_res = {ctors = ctors0, ...}, + absT_info = {abs = abs0, rep = rep0, ...}, + fp_ctr_sugar = {ctrXs_Tss, ctr_sugar = {ctrs = ctrs0, casex = case0, discs = discs0, + selss = selss0, sel_defs, ...}, ...}, ...} = + checked_gfp_sugar_of ctxt res_T; + + val VLeaf0_T = fastype_of VLeaf0; + val Y = domain_type VLeaf0_T; + + val raw_buffer = specialize_buffer_types raw_buffer0; + + 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 substYT = Tsubst Y tupled_arg_T; + val substY = substT Y tupled_arg_T; + + val Ys_rho_inner = if friend then [] else [(Y, tupled_arg_T)]; + val substYT_inner = substAT o Term.typ_subst_atomic Ys_rho_inner; + val substY_inner = substA o Term.subst_atomic_types Ys_rho_inner; + + val mid_T = substYT_inner (range_type VLeaf0_T); + + val substXT_mid = Tsubst X mid_T; + + val XifyT = typ_subst_nonatomic [(res_T, X)]; + val YifyT = typ_subst_nonatomic [(res_T, Y)]; + + val substXYT = Tsubst X Y; + + val ctor0 = nth ctors0 fp_res_index; + val ctor = enforce_type ctxt range_type res_T ctor0; + val preT = YifyT (domain_type (fastype_of ctor)); + + val n = length ctrs0; + val ks = 1 upto n; + + fun mk_ctr_guards () = + let + val ctr_Tss = map (map (substXT_mid o substAT)) ctrXs_Tss; + val preT = XifyT (domain_type (fastype_of ctor)); + val mid_preT = substXT_mid preT; + val abs = enforce_type ctxt range_type mid_preT abs0; + val absT = range_type (fastype_of abs); + + fun mk_ctr_guard k ctr_Ts (Const (s, _)) = + let + val xs = map_index (fn (i, T) => Free ("x" ^ string_of_int i, T)) ctr_Ts; + val body = mk_absumprod absT abs n k xs; + in + (s, fold_rev Term.lambda xs body) + end; + in + Symtab.make (@{map 3} mk_ctr_guard ks ctr_Tss ctrs0) + end; + + val substYT_mid = substYT o Tsubst Y mid_T; + + val outer_T = substYT_mid preT; + + val substY_outer = substY o substT Y outer_T; + + val outer_buffer = curry_friends (map_buffer substY_outer raw_buffer); + val ctr_guards = mk_ctr_guards (); + val inner_buffer = curry_friends (map_buffer substY_inner raw_buffer); + + val s_parse_info = + {outer_buffer = outer_buffer, ctr_guards = ctr_guards, inner_buffer = inner_buffer}; + + fun mk_friend_spec () = + let + fun encapsulate_nested U T free = + betapply (build_map ctxt [] (fn (T, _) => + if T = domain_type VLeaf0_T then Abs (Name.uu, T, VLeaf0 $ Bound 0) + else Abs (Name.uu, T, Bound 0)) (T, U), + free); + + val preT = YifyT (domain_type (fastype_of ctor)); + val YpreT = HOLogic.mk_prodT (Y, preT); + + val rep = rep0 |> enforce_type ctxt domain_type (substXT_mid (XifyT preT)); + + fun mk_disc k = + ctrXs_Tss + |> map_index (fn (i, Ts) => + Abs (Name.uu, mk_tupleT_balanced Ts, + if i + 1 = k then @{const HOL.True} else @{const HOL.False})) + |> mk_case_sumN_balanced + |> map_types substXYT + |> (fn tm => Library.foldl1 HOLogic.mk_comp [tm, rep, snd_const YpreT]) + |> map_types substAT; + + val all_discs = map mk_disc ks; + + fun mk_pair (Const (disc_name, _)) disc = SOME (disc_name, disc) + | mk_pair _ _ = NONE; + + val discs = @{map 2} mk_pair discs0 all_discs + |> map_filter I |> Symtab.make; + + fun mk_sel sel_def = + let + val (sel_name, case_functions) = + sel_def + |> Object_Logic.rulify ctxt + |> Thm.concl_of + |> perhaps (try drop_all) + |> perhaps (try HOLogic.dest_Trueprop) + |> HOLogic.dest_eq + |>> fst o strip_comb + |>> fst o dest_Const + ||> fst o dest_comb + ||> snd o strip_comb + ||> map (map_types (XifyT o substAT)); + + fun encapsulate_case_function case_function = + let + fun encapsulate bound_Ts [] case_function = + let val T = fastype_of1 (bound_Ts, case_function) in + encapsulate_nested (substXT_mid T) (substXYT T) case_function + end + | encapsulate bound_Ts (T :: Ts) case_function = + Abs (Name.uu, T, + encapsulate (T :: bound_Ts) Ts + (betapply (incr_boundvars 1 case_function, Bound 0))); + in + encapsulate [] (binder_types (fastype_of case_function)) case_function + end; + in + (sel_name, ctrXs_Tss + |> map (map_index (fn (i, T) => Free ("x" ^ string_of_int (i + 1), T))) + |> `(map mk_tuple_balanced) + |> uncurry (@{map 3} mk_tupled_fun (map encapsulate_case_function case_functions)) + |> mk_case_sumN_balanced + |> map_types substXYT + |> (fn tm => Library.foldl1 HOLogic.mk_comp [tm, rep, snd_const YpreT]) + |> map_types substAT) + end; + + val sels = Symtab.make (map mk_sel sel_defs); + + fun mk_disc_sels_pair disc sels = + if forall is_some sels then SOME (disc, map the sels) else NONE; + + val pattern_ctrs = (ctrs0, selss0) + ||> map (map (try dest_Const #> Option.mapPartial (fst #> Symtab.lookup sels))) + ||> @{map 2} mk_disc_sels_pair all_discs + |>> map (dest_Const #> fst) + |> op ~~ + |> map_filter (fn (s, opt) => if is_some opt then SOME (s, the opt) else NONE) + |> Symtab.make; + + val it = HOLogic.mk_comp (VLeaf0, fst_const YpreT); + + val mk_case = + let + val abs_fun_tms = case0 + |> fastype_of + |> substAT + |> XifyT + |> binder_fun_types + |> map_index (fn (i, T) => Free ("f" ^ string_of_int (i + 1), T)); + val arg_Uss = abs_fun_tms + |> map fastype_of + |> map binder_types; + val arg_Tss = arg_Uss + |> map (map substXYT); + val case0 = + arg_Tss + |> map (map_index (fn (i, T) => Free ("x" ^ string_of_int (i + 1), T))) + |> `(map mk_tuple_balanced) + ||> @{map 3} (@{map 3} encapsulate_nested) arg_Uss arg_Tss + |> uncurry (@{map 3} mk_tupled_fun abs_fun_tms) + |> mk_case_sumN_balanced + |> (fn tm => Library.foldl1 HOLogic.mk_comp [tm, rep, snd_const YpreT]) + |> fold_rev lambda abs_fun_tms + |> map_types (substAT o substXT_mid); + in + fn U => case0 + |> substT (body_type (fastype_of case0)) U + |> Syntax.check_term ctxt + end; + in + {pattern_ctrs = pattern_ctrs, discs = discs, sels = sels, it = it, mk_case = mk_case} + end; + in + (s_parse_info, mk_friend_spec) + end; + +fun corec_parse_info_of ctxt = + fst ooo generic_spec_of false ctxt; + +fun friend_parse_info_of ctxt = + apsnd (fn f => f ()) ooo generic_spec_of true ctxt; + +end; diff -r 9bfcbab7cd99 -r 0701f25fac39 src/HOL/Tools/BNF/bnf_gfp_grec_tactics.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/BNF/bnf_gfp_grec_tactics.ML Tue Mar 22 12:39:37 2016 +0100 @@ -0,0 +1,420 @@ +(* Title: HOL/Tools/BNF/bnf_gfp_grec_tactics.ML + Author: Jasmin Blanchette, Inria, LORIA, MPII + Author: Dmitriy Traytel, ETH Zurich + Copyright 2015, 2016 + +Tactics for generalized corecursor construction. +*) + +signature BNF_GFP_GREC_TACTICS = +sig + val transfer_prover_add_tac: Proof.context -> thm list -> thm list -> int -> tactic + + val mk_algLam_algLam_tac: Proof.context -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> + tactic + val mk_algLam_algrho_tac: Proof.context -> thm -> thm -> tactic + val mk_algLam_base_tac: Proof.context -> term -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> + thm list -> thm -> thm list -> thm list -> thm -> thm -> tactic + val mk_algLam_step_tac: Proof.context -> thm -> thm -> thm -> tactic + val mk_cong_locale_tac: Proof.context -> thm -> thm list -> thm -> thm -> thm list -> thm -> + thm -> tactic + val mk_corecU_pointfree_tac: Proof.context -> thm -> thm -> thm list -> thm -> thm list -> thm -> + thm list -> thm -> thm -> thm -> tactic + val mk_corecUU_pointfree_tac: Proof.context -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> + thm -> thm -> thm -> thm -> thm -> thm -> tactic + val mk_corecUU_unique_tac: Proof.context -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> + thm -> thm -> thm -> thm -> thm -> thm -> tactic + val mk_corecUU_Inl_tac: Proof.context -> term -> thm -> thm -> thm -> thm -> thm list -> thm -> + thm list -> thm -> thm -> thm -> thm -> tactic + val mk_dtor_algLam_tac: Proof.context -> thm -> thm -> thm -> thm -> thm -> thm -> thm list -> + thm -> thm -> thm list -> thm -> thm -> thm -> thm -> tactic + val mk_dtor_algrho_tac: Proof.context -> thm -> thm -> thm -> thm -> tactic + val mk_dtor_transfer_tac: Proof.context -> thm -> tactic + val mk_equivp_Retr_tac: Proof.context -> thm -> thm -> thm -> thm -> tactic + val mk_eval_core_embL_tac: Proof.context -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> + thm -> thm -> thm -> thm list -> thm list -> thm list -> thm -> tactic + val mk_eval_core_flat_tac: Proof.context -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> + thm list -> thm -> thm list -> thm -> thm -> thm -> thm list -> tactic + val mk_eval_core_k_as_ssig_tac: Proof.context -> thm -> thm -> thm -> thm list -> thm -> thm -> + thm -> thm list -> tactic + val mk_eval_embL_tac: Proof.context -> thm -> thm -> thm -> thm -> thm -> thm -> tactic + val mk_eval_flat_tac: Proof.context -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> + tactic + val mk_eval_sctr_tac: Proof.context -> thm -> thm -> thm -> thm -> tactic + val mk_eval_Oper_tac: Proof.context -> int -> thm -> thm -> thm -> thm -> thm -> thm list -> + thm -> thm -> tactic + val mk_eval_V_or_CLeaf_tac: Proof.context -> thm -> thm -> thm -> thm -> thm -> thm list -> thm -> + tactic + val mk_extdd_mor_tac: Proof.context -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> + thm -> thm -> thm -> tactic + val mk_extdd_o_VLeaf_tac: Proof.context -> thm -> thm -> thm -> thm list -> thm list -> thm -> + thm -> thm -> tactic + val mk_flat_embL_tac: Proof.context -> thm -> thm -> thm -> thm -> thm -> thm list -> thm list -> + thm list -> thm list -> tactic + val mk_flat_VLeaf_or_flat_tac: Proof.context -> thm -> thm -> thm list -> tactic + val mk_Lam_Inl_Inr_tac: Proof.context -> thm -> thm -> tactic + val mk_mor_cutSsig_flat_tac: Proof.context -> term -> thm -> thm -> thm -> thm -> thm -> thm -> + thm list -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> tactic + val mk_natural_from_transfer_tac: Proof.context -> int -> bool list -> thm -> thm list -> + thm list -> thm list -> tactic + val mk_natural_by_unfolding_tac: Proof.context -> thm list -> tactic + val mk_Retr_coinduct_tac: Proof.context -> thm -> thm -> tactic + val mk_sig_transfer_tac: Proof.context -> thm -> thm list -> thm -> tactic + val mk_transfer_by_transfer_prover_tac: Proof.context -> thm list -> thm list -> thm list -> + tactic +end; + +structure BNF_GFP_Grec_Tactics : BNF_GFP_GREC_TACTICS = +struct + +open BNF_Util +open BNF_Tactics +open BNF_FP_Util + +val o_assoc = @{thm o_assoc}; +val o_def = @{thm o_def}; + +fun ss_only_silent thms ctxt = + ss_only thms (ctxt |> Context_Position.set_visible false); + +fun context_relator_eq_add rel_eqs ctxt = + fold (snd oo Thm.proof_attributes (map (Attrib.attribute ctxt) @{attributes [relator_eq]})) + rel_eqs ctxt; +val context_transfer_rule_add = fold (snd oo Thm.proof_attributes [Transfer.transfer_add]); + +fun transfer_prover_add_tac ctxt rel_eqs transfers = + Transfer.transfer_prover_tac (ctxt + |> context_relator_eq_add rel_eqs + |> context_transfer_rule_add transfers); + +fun instantiate_natural_rule_with_id ctxt live = + Rule_Insts.of_rule ctxt ([], NONE :: replicate live (SOME @{const_name id})) []; + +fun instantiate_transfer_rule_with_Grp_UNIV ctxt alives thm = + let + val n = length alives; + val fs = map (prefix "f" o string_of_int) (1 upto n); + val ss = map2 (fn live => fn f => SOME (@{const_name BNF_Def.Grp} ^ " " ^ @{const_name top} ^ + " " ^ (if live then f else @{const_name id}))) alives fs; + val bs = map_filter (fn (live, f) => if live then SOME (Binding.name f, NONE, NoSyn) else NONE) + (alives ~~ fs); + in + Rule_Insts.of_rule ctxt ([], ss) bs thm + end; + +fun mk_algLam_algLam_tac ctxt dead_pre_map_comp dtor_inject unsig sig_map Lam_def eval_embL + old_dtor_algLam dtor_algLam = + HEADGOAL (rtac ctxt ext THEN' rtac ctxt (dtor_inject RS iffD1)) THEN + unfold_thms_tac ctxt (dead_pre_map_comp :: unsig :: sig_map :: Lam_def :: eval_embL :: + old_dtor_algLam :: dtor_algLam :: @{thms o_apply id_o map_sum.simps sum.case}) THEN + HEADGOAL (rtac ctxt refl); + +fun mk_algLam_algrho_tac ctxt algLam_def algrho_def = + HEADGOAL (rtac ctxt ext) THEN unfold_thms_tac ctxt [algLam_def, algrho_def, o_apply] THEN + HEADGOAL (rtac ctxt refl); + +fun 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 Sig_pointful_natural ssig_maps Lam_def flat_simps eval_core_simps eval + algLam_def = + HEADGOAL (rtac ctxt (infer_instantiate' ctxt [NONE, SOME (Thm.cterm_of ctxt dead_pre_map_dtor)] + (trans OF [dtor_unfold_unique, dtor_unfold_unique RS sym]) OF [ext, ext])) THEN + ALLGOALS (asm_simp_tac (ss_only_silent (dead_pre_map_id :: ctor_dtor :: dtor_ctor :: unsig :: + Sig_pointful_natural :: Lam_def :: eval :: algLam_def :: + unfold_thms ctxt [o_def] dead_pre_map_comp :: ssig_maps @ flat_simps @ eval_core_simps @ + @{thms o_apply id_apply id_def[symmetric] snd_conv convol_apply}) ctxt)); + +fun mk_algLam_step_tac ctxt proto_sctr_def old_algLam_pointful algLam_algLam_pointful = + HEADGOAL (rtac ctxt ext) THEN + unfold_thms_tac ctxt [proto_sctr_def, old_algLam_pointful, algLam_algLam_pointful, o_apply] THEN + HEADGOAL (rtac ctxt refl); + +fun mk_cong_locale_tac ctxt dead_pre_rel_mono dead_pre_rel_maps equivp_Retr + ssig_rel_mono ssig_rel_maps eval eval_core_transfer = + HEADGOAL (resolve_tac ctxt (Locale.get_unfolds @{context}) THEN' + etac ctxt ssig_rel_mono THEN' etac ctxt equivp_Retr) THEN + unfold_thms_tac ctxt (eval :: dead_pre_rel_maps @ @{thms id_apply}) THEN + HEADGOAL (rtac ctxt (@{thm predicate2I} RS (dead_pre_rel_mono RS @{thm predicate2D})) THEN' + etac ctxt @{thm rel_funD} THEN' assume_tac ctxt THEN' + rtac ctxt (eval_core_transfer RS @{thm rel_funD})) THEN + unfold_thms_tac ctxt (ssig_rel_maps @ @{thms vimage2p_rel_prod vimage2p_id}) THEN + unfold_thms_tac ctxt @{thms vimage2p_def} THEN HEADGOAL (assume_tac ctxt); + +fun mk_corecU_pointfree_tac ctxt dead_pre_map_comp dtor_unfold ssig_maps dead_ssig_map_comp0 + flat_simps flat_VLeaf eval_core_simps cutSsig_def mor_cutSsig_flat corecU_def = + unfold_thms_tac ctxt [corecU_def, dead_ssig_map_comp0, o_assoc] THEN + HEADGOAL (subst_tac ctxt NONE [ext RS mor_cutSsig_flat] THEN' + asm_simp_tac (ss_only_silent [dtor_unfold, o_apply] ctxt) THEN' + asm_simp_tac (ss_only_silent (dtor_unfold :: flat_VLeaf :: cutSsig_def :: ssig_maps @ + flat_simps @ eval_core_simps @ unfold_thms ctxt [o_def] dead_pre_map_comp :: + @{thms o_def id_apply id_def[symmetric] snd_conv convol_apply}) ctxt)); + +fun mk_corecUU_tail_tac ctxt dead_pre_map_comp0 dead_pre_map_comp dtor_ctor ssig_map_comp + flat_pointful_natural eval_core_pointful_natural eval eval_flat sctr_pointful_natural + eval_sctr_pointful = + asm_simp_tac (ss_only_silent (dtor_ctor :: flat_pointful_natural :: eval :: eval_flat :: + map (unfold_thms ctxt [o_def]) [dead_pre_map_comp, ssig_map_comp] @ + @{thms o_apply id_apply id_def[symmetric] convol_apply}) ctxt) THEN' + asm_simp_tac (ss_only_silent (eval_core_pointful_natural :: sctr_pointful_natural :: + eval_sctr_pointful :: map (unfold_thms ctxt [o_def]) [dead_pre_map_comp0, ssig_map_comp] @ + @{thms id_apply id_def[symmetric] convol_apply map_prod_simp}) ctxt); + +fun 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 eval_flat corecU_ctor + sctr_pointful_natural eval_sctr_pointful corecUU_def = + unfold_thms_tac ctxt [corecUU_def] THEN + HEADGOAL (rtac ctxt ext THEN' subst_tac ctxt NONE [corecU_ctor RS sym]) THEN + unfold_thms_tac ctxt [corecUU_def RS Drule.symmetric_thm] THEN + HEADGOAL (rtac ctxt (dtor_inject RS iffD1) THEN' + mk_corecUU_tail_tac ctxt dead_pre_map_comp0 dead_pre_map_comp dtor_ctor ssig_map_comp + flat_pointful_natural eval_core_pointful_natural eval eval_flat sctr_pointful_natural + eval_sctr_pointful); + +fun 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 eval_flat corecU_unique + sctr_pointful_natural eval_sctr_pointful corecUU_def prem = + unfold_thms_tac ctxt [corecUU_def] THEN + HEADGOAL (rtac ctxt corecU_unique THEN' rtac ctxt sym THEN' subst_tac ctxt NONE [prem] THEN' + rtac ctxt ext THEN' + mk_corecUU_tail_tac ctxt dead_pre_map_comp0 dead_pre_map_comp dtor_ctor ssig_map_comp + flat_pointful_natural eval_core_pointful_natural eval eval_flat sctr_pointful_natural + eval_sctr_pointful); + +fun mk_corecUU_Inl_tac ctxt inl_case' pre_map_comp dead_pre_map_ident dead_pre_map_comp0 ctor_dtor + ssig_maps ssig_map_id0 eval_core_simps eval_core_pointful_natural eval_VLeaf corecUU_pointfree + corecUU_unique = + HEADGOAL (rtac ctxt (infer_instantiate' ctxt [NONE, SOME (Thm.cterm_of ctxt inl_case')] + (trans OF [corecUU_unique, corecUU_unique RS sym]) OF [ext, ext]) THEN' + subst_tac ctxt NONE [corecUU_pointfree] THEN' + asm_simp_tac (ss_only_silent (dead_pre_map_comp0 :: eval_core_pointful_natural :: ssig_maps @ + @{thms o_apply sum.case convol_apply id_apply map_prod_simp}) ctxt) THEN' + asm_simp_tac (ss_only_silent (dead_pre_map_ident :: ctor_dtor :: ssig_map_id0 :: + eval_core_pointful_natural :: eval_VLeaf :: unfold_thms ctxt [o_def] pre_map_comp :: + ssig_maps @ eval_core_simps @ @{thms o_apply prod.map_id convol_apply snd_conv id_apply}) + ctxt)); + +fun 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_maps dead_ssig_map_comp0 Lam_pointful_natural + eval_core_simps eval eval_flat eval_VLeaf algLam_def = + unfold_thms_tac ctxt [dead_ssig_map_comp0, o_assoc] THEN + HEADGOAL (asm_simp_tac (ss_only_silent (sig_map_comp :: Oper_pointful_natural :: eval :: + eval_flat :: algLam_def :: unfold_thms ctxt [o_def] dead_pre_map_comp :: eval_core_simps @ + @{thms o_apply id_apply id_def[symmetric]}) ctxt) THEN' + asm_simp_tac (ss_only_silent (Lam_pointful_natural :: eval_VLeaf :: + map (unfold_thms ctxt [o_def]) [dead_pre_map_comp0, sig_map_comp] @ ssig_maps @ + eval_core_simps @ + @{thms o_apply convol_apply snd_conv fst_conv id_apply map_prod_simp}) ctxt) THEN' + asm_simp_tac (ss_only_silent (dead_pre_map_id :: eval_VLeaf :: + unfold_thms ctxt [o_def] pre_map_comp :: + @{thms id_apply id_def[symmetric] convol_def}) ctxt)); + +fun mk_dtor_algrho_tac ctxt eval k_as_ssig_natural_pointful eval_core_k_as_ssig algrho_def = + HEADGOAL (asm_simp_tac (ss_only_silent [eval, k_as_ssig_natural_pointful, algrho_def, + eval_core_k_as_ssig RS sym, o_apply] ctxt)); + +fun mk_dtor_transfer_tac ctxt dtor_rel = + HEADGOAL (rtac ctxt refl ORELSE' + rtac ctxt @{thm rel_funI} THEN' rtac ctxt (dtor_rel RS iffD1) THEN' assume_tac ctxt); + +fun mk_equivp_Retr_tac ctxt dead_pre_rel_refl dead_pre_rel_flip dead_pre_rel_mono + dead_pre_rel_compp = + HEADGOAL (EVERY' [etac ctxt @{thm equivpE}, rtac ctxt @{thm equivpI}, + rtac ctxt @{thm reflpI}, rtac ctxt dead_pre_rel_refl, etac ctxt @{thm reflpD}, + SELECT_GOAL (unfold_thms_tac ctxt @{thms symp_iff}), + REPEAT_DETERM o rtac ctxt ext, rtac ctxt (dead_pre_rel_flip RS sym RS trans), + rtac ctxt ((@{thm conversep_iff} RS sym) RSN (2, trans)), + asm_simp_tac (ss_only_silent @{thms conversep_eq} ctxt), + SELECT_GOAL (unfold_thms_tac ctxt @{thms transp_relcompp}), + rtac ctxt @{thm predicate2I}, etac ctxt @{thm relcomppE}, + etac ctxt (dead_pre_rel_mono RS @{thm rev_predicate2D[rotated -1]}), + SELECT_GOAL (unfold_thms_tac ctxt + (unfold_thms ctxt [@{thm eq_OO}] dead_pre_rel_compp :: @{thms relcompp_apply})), + REPEAT_DETERM o resolve_tac ctxt [exI, conjI], assume_tac ctxt, assume_tac ctxt]); + +fun mk_eval_core_k_as_ssig_tac ctxt pre_map_comp dead_pre_map_id sig_map_comp ssig_maps + Lam_natural_pointful Lam_Inr flat_VLeaf eval_core_simps = + HEADGOAL (asm_simp_tac (ss_only_silent (dead_pre_map_id :: flat_VLeaf :: (Lam_Inr RS sym) :: + o_apply :: id_apply :: @{thm id_def[symmetric]} :: + unfold_thms ctxt @{thms map_prod_def split_def} Lam_natural_pointful :: ssig_maps @ + eval_core_simps @ map (unfold_thms ctxt [o_def]) [pre_map_comp, sig_map_comp]) ctxt)); + +fun mk_eval_embL_tac ctxt dead_pre_map_comp0 dtor_unfold_unique embL_pointful_natural eval_core_embL + old_eval eval = + HEADGOAL (rtac ctxt (unfold_thms ctxt [o_apply] + (trans OF [dtor_unfold_unique, dtor_unfold_unique RS sym] OF [ext, ext]) + OF [Drule.asm_rl, old_eval RS sym])) THEN + unfold_thms_tac ctxt [dead_pre_map_comp0, embL_pointful_natural, eval_core_embL, eval, + o_apply] THEN + HEADGOAL (rtac ctxt refl); + +fun 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 cond_eval_o_flat = + HEADGOAL (rtac ctxt (unfold_thms ctxt [o_apply] cond_eval_o_flat)) THEN + unfold_thms_tac ctxt [dead_pre_map_comp0, flat_pointful_natural, eval_core_flat, eval, + o_apply] THEN + HEADGOAL (rtac ctxt refl THEN' + asm_simp_tac (ss_only_silent (ssig_map_id :: eval_core_pointful_natural :: eval :: + map (unfold_thms ctxt [o_def]) [dead_pre_map_comp0, ssig_map_comp] @ + @{thms id_apply id_def[symmetric] fst_conv map_prod_simp convol_apply}) + ctxt)); + +fun instantiate_map_comp_with_f_g ctxt = + Rule_Insts.of_rule ctxt ([], [NONE, SOME ("%x. f (g x)")]) + [(Binding.name "f", NONE, NoSyn), (Binding.name "g", NONE, NoSyn)]; + +fun 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 = + HEADGOAL (rtac ctxt old_ssig_induct) THEN + ALLGOALS (asm_simp_tac (ss_only_silent (Sig_pointful_natural :: unsig_thm :: Lam_def :: + (flat_embL RS sym) :: unfold_thms ctxt [o_def] dead_pre_map_comp :: embL_simps @ + old_eval_core_simps @ eval_core_simps @ + @{thms id_apply id_def[symmetric] o_apply map_sum.simps sum.case}) ctxt)) THEN + HEADGOAL (asm_simp_tac (Simplifier.add_cong old_sig_map_cong (ss_only_silent + (old_Lam_pointful_natural :: embL_pointful_natural :: + map (unfold_thms ctxt [o_def]) [dead_pre_map_comp0, instantiate_map_comp_with_f_g ctxt + dead_pre_map_comp0, old_sig_map_comp] @ @{thms map_prod_simp}) ctxt))); + +fun 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_maps ssig_map_comp flat_simps flat_natural flat_flat + Lam_natural_sym eval_core_simps = + HEADGOAL (rtac ctxt ssig_induct) THEN + ALLGOALS (full_simp_tac (ss_only_silent ((flat_flat RS sym) :: dead_pre_map_id :: + dead_pre_map_comp :: fp_map_id :: sig_map_comp :: ssig_map_comp :: ssig_maps @ flat_simps @ + eval_core_simps @ @{thms o_def id_def[symmetric] convol_apply id_apply snd_conv}) ctxt)) THEN + HEADGOAL (asm_simp_tac (Simplifier.add_cong sig_map_cong (ss_only_silent + (map (unfold_thms ctxt [o_def]) [dead_pre_map_comp0, sig_map_comp] @ + flat_natural :: Lam_natural_sym :: @{thms id_apply fst_conv map_prod_simp}) + ctxt))); + +fun mk_eval_sctr_tac ctxt proto_sctr_pointful_natural eval_Oper algLam sctr_def = + HEADGOAL (rtac ctxt ext) THEN + unfold_thms_tac ctxt [proto_sctr_pointful_natural, eval_Oper, algLam RS sym, sctr_def, + o_apply] THEN + HEADGOAL (rtac ctxt refl); + +fun 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 eval_core_simps eval = + HEADGOAL (rtac ctxt (trans OF [dtor_unfold_unique, dtor_unfold_unique RS sym] RS fun_cong + OF [ext, ext])) THEN + ALLGOALS (asm_simp_tac (ss_only_silent (dead_pre_map_id :: fp_map_id :: + unfold_thms ctxt @{thms o_def} dead_pre_map_comp :: V_or_CLeaf_map :: eval :: eval_core_simps @ + @{thms o_apply id_def[symmetric] id_apply snd_conv convol_apply}) ctxt)); + +fun 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 = + let val VLeaf_natural' = instantiate_natural_rule_with_id ctxt live VLeaf_natural in + unfold_thms_tac ctxt [sig_map_comp, VLeaf_natural', algLam_def, o_apply] THEN + unfold_thms_tac ctxt (sig_map_comp0 :: Oper_natural_pointful :: (eval_flat RS sym) :: o_apply :: + flat_simps) THEN + unfold_thms_tac ctxt (@{thm id_apply} :: sig_map_ident :: unfold_thms ctxt [o_def] sig_map_comp :: + flat_simps) THEN + HEADGOAL (rtac ctxt refl) + end; + +fun mk_extdd_mor_tac ctxt dead_pre_map_comp0 dead_pre_map_comp VLeaf_map ssig_map_comp + flat_pointful_natural eval_core_pointful_natural eval eval_flat eval_VLeaf cutSsig_def prem = + HEADGOAL (rtac ctxt ext) THEN + unfold_thms_tac ctxt (ssig_map_comp :: unfold_thms ctxt [o_def] dead_pre_map_comp :: + flat_pointful_natural :: eval :: eval_flat :: cutSsig_def :: + @{thms o_apply convol_o id_o id_apply id_def[symmetric]}) THEN + unfold_thms_tac ctxt (unfold_thms ctxt [dead_pre_map_comp0] prem :: dead_pre_map_comp0 :: + ssig_map_comp :: eval_core_pointful_natural :: + @{thms o_def[symmetric] o_apply map_prod_o_convol}) THEN + unfold_thms_tac ctxt (VLeaf_map :: eval_VLeaf :: @{thms o_def id_apply id_def[symmetric]}) THEN + HEADGOAL (rtac ctxt refl); + +fun mk_extdd_o_VLeaf_tac ctxt dead_pre_map_comp0 dead_pre_map_comp dtor_inject ssig_maps + eval_core_simps eval eval_VLeaf prem = + HEADGOAL (rtac ctxt ext THEN' rtac ctxt (dtor_inject RS iffD1) THEN' + asm_simp_tac (ss_only_silent (dead_pre_map_comp0 :: ssig_maps @ eval_core_simps @ eval :: + eval_VLeaf :: (mk_pointful ctxt prem RS sym) :: unfold_thms ctxt [o_def] dead_pre_map_comp :: + @{thms o_apply convol_apply snd_conv id_apply}) ctxt)); + +fun mk_flat_embL_tac ctxt old_ssig_induct fp_map_id Sig_pointful_natural old_sig_map_comp + old_sig_map_cong old_ssig_maps old_flat_simps flat_simps embL_simps = + HEADGOAL (rtac ctxt old_ssig_induct) THEN + ALLGOALS (asm_simp_tac (Simplifier.add_cong old_sig_map_cong (ss_only_silent + (fp_map_id :: Sig_pointful_natural :: unfold_thms ctxt [o_def] old_sig_map_comp :: + old_ssig_maps @ old_flat_simps @ flat_simps @ embL_simps @ + @{thms id_apply id_def[symmetric] map_sum.simps}) ctxt))); + +fun mk_flat_VLeaf_or_flat_tac ctxt ssig_induct cong simps = + HEADGOAL (rtac ctxt ssig_induct) THEN + ALLGOALS (asm_simp_tac (Simplifier.add_cong cong (ss_only_silent simps ctxt))); + +fun mk_Lam_Inl_Inr_tac ctxt unsig Lam_def = + TRYALL Goal.conjunction_tac THEN ALLGOALS (rtac ctxt ext) THEN + unfold_thms_tac ctxt (o_apply :: Lam_def :: unsig :: @{thms sum.case}) THEN + ALLGOALS (rtac ctxt refl); + +fun 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 = + HEADGOAL (rtac ctxt (infer_instantiate' ctxt [NONE, SOME (Thm.cterm_of ctxt eval_core_o_map)] + (trans OF [dtor_unfold_unique, dtor_unfold_unique RS sym]) OF [ext, ext]) THEN' + asm_simp_tac (ss_only_silent ((prem RS sym) :: dead_pre_map_comp0 :: ssig_map_comp :: + eval_core_pointful_natural :: eval_thm :: + @{thms o_apply map_prod_o_convol o_id convol_o id_o}) ctxt) THEN' + asm_simp_tac (ss_only_silent ((mk_pointful ctxt prem RS sym) :: dead_pre_map_comp0 :: + cutSsig_def_pointful_natural :: flat_simps @ + @{thms o_apply convol_apply map_prod_simp id_apply}) ctxt) THEN' + rtac ctxt (dead_pre_map_cong OF [Drule.asm_rl, refl]) THEN' + asm_simp_tac (ss_only_silent (ssig_map_comp :: cutSsig_def :: flat_pointful_natural :: + eval_core_flat :: unfold_thms ctxt [o_def] dead_pre_map_comp :: (dead_ssig_map_comp0 RS sym) :: + (flat_flat RS sym) :: + @{thms o_apply convol_o fst_convol id_apply id_def[symmetric]}) ctxt) THEN' + asm_simp_tac (ss_only_silent (eval_core_pointful_natural :: flat_VLeaf :: + map (unfold_thms ctxt [o_def]) [dead_pre_map_comp0, ssig_map_comp] @ + @{thms o_apply id_apply id_def[symmetric] map_prod_simp convol_def}) ctxt)); + +fun mk_natural_from_transfer_tac ctxt m alives transfer map_ids rel_Grps subst_rel_Grps = + let + val unfold_eq = unfold_thms ctxt @{thms Grp_UNIV_id[symmetric]}; + val (rel_Grps', subst_rel_Grps') = + apply2 (map (fn thm => unfold_eq (thm RS eq_reflection))) (rel_Grps, subst_rel_Grps); + val transfer' = instantiate_transfer_rule_with_Grp_UNIV ctxt alives (unfold_eq transfer) + |> unfold_thms ctxt rel_Grps'; + in + HEADGOAL (Method.insert_tac ctxt [transfer'] THEN' + EVERY' (map (subst_asm_tac ctxt NONE o single) subst_rel_Grps')) THEN + unfold_thms_tac ctxt (map_ids @ @{thms Grp_def rel_fun_def}) THEN + HEADGOAL (REPEAT_DETERM_N m o rtac ctxt ext THEN' + asm_full_simp_tac (ss_only_silent @{thms simp_thms id_apply o_apply mem_Collect_eq + top_greatest UNIV_I subset_UNIV[unfolded UNIV_def]} ctxt)) THEN + ALLGOALS (REPEAT_DETERM o etac ctxt @{thm meta_allE} THEN' REPEAT_DETERM o etac ctxt allE THEN' + forward_tac ctxt [sym] THEN' assume_tac ctxt) + end; + +fun mk_natural_by_unfolding_tac ctxt maps = + HEADGOAL (rtac ctxt ext) THEN + unfold_thms_tac ctxt (@{thms o_def[abs_def] id_apply id_def[symmetric]} @ maps) THEN + HEADGOAL (rtac ctxt refl); + +fun mk_Retr_coinduct_tac ctxt dtor_rel_coinduct rel_eq = + HEADGOAL (EVERY' [rtac ctxt allI, rtac ctxt impI, + rtac ctxt (@{thm ord_le_eq_trans} OF [dtor_rel_coinduct, rel_eq]), + etac ctxt @{thm predicate2D}, assume_tac ctxt]); + +fun mk_sig_transfer_tac ctxt pre_rel_def rel_eqs0 transfer = + let + val rel_eqs = no_refl rel_eqs0; + val rel_eq_syms = map (fn thm => thm RS sym) rel_eqs; + val transfer' = unfold_thms ctxt rel_eq_syms transfer + in + HEADGOAL (rtac ctxt transfer') ORELSE + unfold_thms_tac ctxt (pre_rel_def :: rel_eq_syms @ + @{thms BNF_Def.vimage2p_def BNF_Composition.id_bnf_def}) THEN + HEADGOAL (rtac ctxt transfer') + end; + +fun mk_transfer_by_transfer_prover_tac ctxt defs rel_eqs0 transfers = + let + val rel_eqs = no_refl rel_eqs0; + val rel_eq_syms = map (fn thm => thm RS sym) rel_eqs; + in + unfold_thms_tac ctxt (defs @ rel_eq_syms) THEN + HEADGOAL (transfer_prover_add_tac ctxt rel_eqs transfers) + end; + +end; diff -r 9bfcbab7cd99 -r 0701f25fac39 src/HOL/Tools/BNF/bnf_gfp_grec_unique_sugar.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/BNF/bnf_gfp_grec_unique_sugar.ML Tue Mar 22 12:39:37 2016 +0100 @@ -0,0 +1,73 @@ +(* Title: HOL/Tools/BNF/bnf_gfp_grec_unique_sugar.ML + Author: Jasmin Blanchette, Inria, LORIA, MPII + Copyright 2016 + +Proof method for proving uniqueness of corecursive equations ("corec_unique"). +*) + +signature BNF_GFP_GREC_UNIQUE_SUGAR = +sig + val corec_unique_tac: Proof.context -> int -> tactic +end; + +structure BNF_GFP_Grec_Unique_Sugar : BNF_GFP_GREC_UNIQUE_SUGAR = +struct + +open BNF_Util +open BNF_GFP_Grec +open BNF_GFP_Grec_Sugar_Util +open BNF_GFP_Grec_Sugar + +fun corec_unique_tac ctxt = + Subgoal.FOCUS (fn {context = ctxt, prems, concl, ...} => + let + (* Workaround for odd name clash for goals with "x" in their context *) + val (_, ctxt) = ctxt + |> yield_singleton (mk_Frees "x") @{typ unit}; + + val code_thm = (if null prems then error "No premise" else hd prems) + |> Object_Logic.rulify ctxt; + val code_goal = Thm.prop_of code_thm; + + val (fun_t, args) = strip_comb (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop code_goal))) + handle TERM _ => error "Wrong format for first premise"; + + val _ = is_Free fun_t orelse + error ("Expected free variable as function in premise, found " ^ + Syntax.string_of_term ctxt fun_t); + val _ = + (case filter_out is_Var args of + [] => () + | arg :: _ => + error ("Expected universal variable as argument to function in premise, found " ^ + Syntax.string_of_term ctxt arg)); + + val fun_T = fastype_of fun_t; + val (arg_Ts, res_T) = strip_type fun_T; + + val num_args_in_concl = length (snd (strip_comb (fst (HOLogic.dest_eq + (HOLogic.dest_Trueprop (Thm.term_of concl)))))) + handle TERM _ => error "Wrong format for conclusion"; + + val (corec_info, corec_parse_info) = + (case maybe_corec_info_of ctxt res_T of + SOME (info as {buffer, ...}) => (info, corec_parse_info_of ctxt arg_Ts res_T buffer) + | NONE => error ("No corecursor for " ^ quote (Syntax.string_of_typ ctxt res_T) ^ + " (use " ^ quote (#1 @{command_keyword coinduction_upto}) ^ " to derive it)")); + + val parsed_eq = parse_corec_equation ctxt [fun_t] code_goal; + val explored_eq = + explore_corec_equation ctxt false false "" fun_t corec_parse_info res_T parsed_eq; + + val ((_, corecUU_arg), _) = build_corecUU_arg_and_goals false fun_t explored_eq ctxt; + val eq_corecUU = derive_eq_corecUU ctxt corec_info fun_t corecUU_arg code_thm; + + val unique' = derive_unique ctxt Morphism.identity code_goal corec_info res_T eq_corecUU + |> funpow num_args_in_concl (fn thm => thm RS fun_cong); + in + HEADGOAL ((K all_tac APPEND' rtac ctxt sym) THEN' rtac ctxt unique' THEN' + REPEAT_DETERM_N num_args_in_concl o rtac ctxt ext) + end) ctxt THEN' + etac ctxt thin_rl; + +end;