diff --git a/src/HOL/Number_Theory/Residues.thy b/src/HOL/Number_Theory/Residues.thy --- a/src/HOL/Number_Theory/Residues.thy +++ b/src/HOL/Number_Theory/Residues.thy @@ -1,506 +1,505 @@ (* Title: HOL/Number_Theory/Residues.thy Author: Jeremy Avigad An algebraic treatment of residue rings, and resulting proofs of Euler's theorem and Wilson's theorem. *) section \Residue rings\ theory Residues imports Cong "HOL-Algebra.Multiplicative_Group" Totient begin definition QuadRes :: "int \ int \ bool" where "QuadRes p a = (\y. ([y^2 = a] (mod p)))" definition Legendre :: "int \ int \ int" where "Legendre a p = (if ([a = 0] (mod p)) then 0 else if QuadRes p a then 1 else -1)" subsection \A locale for residue rings\ definition residue_ring :: "int \ int ring" where "residue_ring m = \carrier = {0..m - 1}, monoid.mult = \x y. (x * y) mod m, one = 1, zero = 0, add = \x y. (x + y) mod m\" locale residues = fixes m :: int and R (structure) assumes m_gt_one: "m > 1" defines "R \ residue_ring m" begin lemma abelian_group: "abelian_group R" proof - have "\y\{0..m - 1}. (x + y) mod m = 0" if "0 \ x" "x < m" for x proof (cases "x = 0") case True with m_gt_one show ?thesis by simp next case False then have "(x + (m - x)) mod m = 0" by simp with m_gt_one that show ?thesis by (metis False atLeastAtMost_iff diff_ge_0_iff_ge diff_left_mono int_one_le_iff_zero_less less_le) qed with m_gt_one show ?thesis by (fastforce simp add: R_def residue_ring_def mod_add_right_eq ac_simps intro!: abelian_groupI) qed lemma comm_monoid: "comm_monoid R" unfolding R_def residue_ring_def apply (rule comm_monoidI) using m_gt_one apply auto apply (metis mod_mult_right_eq mult.assoc mult.commute) apply (metis mult.commute) done lemma cring: "cring R" apply (intro cringI abelian_group comm_monoid) unfolding R_def residue_ring_def apply (auto simp add: comm_semiring_class.distrib mod_add_eq mod_mult_left_eq) done end sublocale residues < cring by (rule cring) context residues begin text \ These lemmas translate back and forth between internal and external concepts. \ lemma res_carrier_eq: "carrier R = {0..m - 1}" by (auto simp: R_def residue_ring_def) lemma res_add_eq: "x \ y = (x + y) mod m" by (auto simp: R_def residue_ring_def) lemma res_mult_eq: "x \ y = (x * y) mod m" by (auto simp: R_def residue_ring_def) lemma res_zero_eq: "\ = 0" by (auto simp: R_def residue_ring_def) lemma res_one_eq: "\ = 1" by (auto simp: R_def residue_ring_def units_of_def) lemma res_units_eq: "Units R = {x. 0 < x \ x < m \ coprime x m}" using m_gt_one apply (auto simp add: Units_def R_def residue_ring_def ac_simps invertible_coprime intro: ccontr) apply (subst (asm) coprime_iff_invertible'_int) apply (auto simp add: cong_def) done lemma res_neg_eq: "\ x = (- x) mod m" using m_gt_one unfolding R_def a_inv_def m_inv_def residue_ring_def apply simp apply (rule the_equality) apply (simp add: mod_add_right_eq) apply (simp add: add.commute mod_add_right_eq) apply (metis add.right_neutral minus_add_cancel mod_add_right_eq mod_pos_pos_trivial) done lemma finite [iff]: "finite (carrier R)" by (simp add: res_carrier_eq) lemma finite_Units [iff]: "finite (Units R)" by (simp add: finite_ring_finite_units) text \ The function \a \ a mod m\ maps the integers to the residue classes. The following lemmas show that this mapping respects addition and multiplication on the integers. \ lemma mod_in_carrier [iff]: "a mod m \ carrier R" unfolding res_carrier_eq using insert m_gt_one by auto lemma add_cong: "(x mod m) \ (y mod m) = (x + y) mod m" by (auto simp: R_def residue_ring_def mod_simps) lemma mult_cong: "(x mod m) \ (y mod m) = (x * y) mod m" by (auto simp: R_def residue_ring_def mod_simps) lemma zero_cong: "\ = 0" by (auto simp: R_def residue_ring_def) lemma one_cong: "\ = 1 mod m" using m_gt_one by (auto simp: R_def residue_ring_def) (* FIXME revise algebra library to use 1? *) lemma pow_cong: "(x mod m) [^] n = x^n mod m" using m_gt_one apply (induct n) apply (auto simp add: nat_pow_def one_cong) apply (metis mult.commute mult_cong) done lemma neg_cong: "\ (x mod m) = (- x) mod m" by (metis mod_minus_eq res_neg_eq) lemma (in residues) prod_cong: "finite A \ (\i\A. (f i) mod m) = (\i\A. f i) mod m" by (induct set: finite) (auto simp: one_cong mult_cong) lemma (in residues) sum_cong: "finite A \ (\i\A. (f i) mod m) = (\i\A. f i) mod m" by (induct set: finite) (auto simp: zero_cong add_cong) lemma mod_in_res_units [simp]: assumes "1 < m" and "coprime a m" shows "a mod m \ Units R" proof (cases "a mod m = 0") case True with assms show ?thesis by (auto simp add: res_units_eq gcd_red_int [symmetric]) next case False from assms have "0 < m" by simp then have "0 \ a mod m" by (rule pos_mod_sign [of m a]) with False have "0 < a mod m" by simp with assms show ?thesis by (auto simp add: res_units_eq gcd_red_int [symmetric] ac_simps) qed lemma res_eq_to_cong: "(a mod m) = (b mod m) \ [a = b] (mod m)" by (auto simp: cong_def) text \Simplifying with these will translate a ring equation in R to a congruence.\ lemmas res_to_cong_simps = add_cong mult_cong pow_cong one_cong prod_cong sum_cong neg_cong res_eq_to_cong text \Other useful facts about the residue ring.\ lemma one_eq_neg_one: "\ = \ \ \ m = 2" apply (simp add: res_one_eq res_neg_eq) apply (metis add.commute add_diff_cancel mod_mod_trivial one_add_one uminus_add_conv_diff zero_neq_one zmod_zminus1_eq_if) done end subsection \Prime residues\ locale residues_prime = fixes p :: nat and R (structure) assumes p_prime [intro]: "prime p" defines "R \ residue_ring (int p)" sublocale residues_prime < residues p unfolding R_def residues_def using p_prime apply auto apply (metis (full_types) of_nat_1 of_nat_less_iff prime_gt_1_nat) done context residues_prime begin lemma p_coprime_left: "coprime p a \ \ p dvd a" using p_prime by (auto intro: prime_imp_coprime dest: coprime_common_divisor) lemma p_coprime_right: "coprime a p \ \ p dvd a" using p_coprime_left [of a] by (simp add: ac_simps) lemma p_coprime_left_int: "coprime (int p) a \ \ int p dvd a" using p_prime by (auto intro: prime_imp_coprime dest: coprime_common_divisor) lemma p_coprime_right_int: "coprime a (int p) \ \ int p dvd a" using p_coprime_left_int [of a] by (simp add: ac_simps) lemma is_field: "field R" proof - have "0 < x \ x < int p \ coprime (int p) x" for x by (rule prime_imp_coprime) (auto simp add: zdvd_not_zless) then show ?thesis by (intro cring.field_intro2 cring) (auto simp add: res_carrier_eq res_one_eq res_zero_eq res_units_eq ac_simps) qed lemma res_prime_units_eq: "Units R = {1..p - 1}" apply (subst res_units_eq) apply (auto simp add: p_coprime_right_int zdvd_not_zless) done end sublocale residues_prime < field by (rule is_field) section \Test cases: Euler's theorem and Wilson's theorem\ subsection \Euler's theorem\ lemma (in residues) totatives_eq: "totatives (nat m) = nat ` Units R" proof - from m_gt_one have "\m\ > 1" by simp then have "totatives (nat \m\) = nat ` abs ` Units R" by (auto simp add: totatives_def res_units_eq image_iff le_less) (use m_gt_one zless_nat_eq_int_zless in force) moreover have "\m\ = m" "abs ` Units R = Units R" using m_gt_one by (auto simp add: res_units_eq image_iff) ultimately show ?thesis by simp qed lemma (in residues) totient_eq: "totient (nat m) = card (Units R)" proof - have *: "inj_on nat (Units R)" by (rule inj_onI) (auto simp add: res_units_eq) then show ?thesis by (simp add: totient_def totatives_eq card_image) qed lemma (in residues_prime) totient_eq: "totient p = p - 1" using totient_eq by (simp add: res_prime_units_eq) lemma (in residues) euler_theorem: assumes "coprime a m" shows "[a ^ totient (nat m) = 1] (mod m)" proof - have "a ^ totient (nat m) mod m = 1 mod m" by (metis assms finite_Units m_gt_one mod_in_res_units one_cong totient_eq pow_cong units_power_order_eq_one) then show ?thesis using res_eq_to_cong by blast qed lemma euler_theorem: fixes a m :: nat assumes "coprime a m" shows "[a ^ totient m = 1] (mod m)" proof (cases "m = 0 \ m = 1") case True then show ?thesis by auto next case False with assms show ?thesis using residues.euler_theorem [of "int m" "int a"] cong_int_iff by (auto simp add: residues_def gcd_int_def) fastforce qed lemma fermat_theorem: fixes p a :: nat assumes "prime p" and "\ p dvd a" shows "[a ^ (p - 1) = 1] (mod p)" proof - from assms prime_imp_coprime [of p a] have "coprime a p" by (auto simp add: ac_simps) then have "[a ^ totient p = 1] (mod p)" by (rule euler_theorem) also have "totient p = p - 1" by (rule totient_prime) (rule assms) finally show ?thesis . qed subsection \Wilson's theorem\ lemma (in field) inv_pair_lemma: "x \ Units R \ y \ Units R \ {x, inv x} \ {y, inv y} \ {x, inv x} \ {y, inv y} = {}" apply auto apply (metis Units_inv_inv)+ done lemma (in residues_prime) wilson_theorem1: assumes a: "p > 2" shows "[fact (p - 1) = (-1::int)] (mod p)" proof - let ?Inverse_Pairs = "{{x, inv x}| x. x \ Units R - {\, \ \}}" have UR: "Units R = {\, \ \} \ \?Inverse_Pairs" by auto have "(\i\Units R. i) = (\i\{\, \ \}. i) \ (\i\\?Inverse_Pairs. i)" apply (subst UR) apply (subst finprod_Un_disjoint) apply (auto intro: funcsetI) using inv_one apply auto[1] using inv_eq_neg_one_eq apply auto done also have "(\i\{\, \ \}. i) = \ \" apply (subst finprod_insert) apply auto apply (frule one_eq_neg_one) using a apply force done also have "(\i\(\?Inverse_Pairs). i) = (\A\?Inverse_Pairs. (\y\A. y))" apply (subst finprod_Union_disjoint) apply (auto simp: pairwise_def disjnt_def) apply (metis Units_inv_inv)+ done also have "\ = \" apply (rule finprod_one_eqI) apply auto apply (subst finprod_insert) apply auto apply (metis inv_eq_self) done finally have "(\i\Units R. i) = \ \" by simp also have "(\i\Units R. i) = (\i\Units R. i mod p)" by (rule finprod_cong') (auto simp: res_units_eq) also have "\ = (\i\Units R. i) mod p" by (rule prod_cong) auto also have "\ = fact (p - 1) mod p" apply (simp add: fact_prod) using assms apply (subst res_prime_units_eq) apply (simp add: int_prod zmod_int prod_int_eq) done finally have "fact (p - 1) mod p = \ \" . then show ?thesis by (simp add: cong_def res_neg_eq res_one_eq zmod_int) qed lemma wilson_theorem: assumes "prime p" shows "[fact (p - 1) = - 1] (mod p)" proof (cases "p = 2") case True then show ?thesis by (simp add: cong_def fact_prod) next case False then show ?thesis using assms prime_ge_2_nat by (metis residues_prime.wilson_theorem1 residues_prime.intro le_eq_less_or_eq) qed text \ This result can be transferred to the multiplicative group of \\/p\\ for \p\ prime.\ lemma mod_nat_int_pow_eq: fixes n :: nat and p a :: int shows "a \ 0 \ p \ 0 \ (nat a ^ n) mod (nat p) = nat ((a ^ n) mod p)" by (simp add: int_one_le_iff_zero_less nat_mod_distrib order_less_imp_le nat_power_eq[symmetric]) theorem residue_prime_mult_group_has_gen: fixes p :: nat assumes prime_p : "prime p" shows "\a \ {1 .. p - 1}. {1 .. p - 1} = {a^i mod p|i . i \ UNIV}" proof - have "p \ 2" using prime_gt_1_nat[OF prime_p] by simp interpret R: residues_prime p "residue_ring p" by (simp add: residues_prime_def prime_p) have car: "carrier (residue_ring (int p)) - {\\<^bsub>residue_ring (int p)\<^esub>} = {1 .. int p - 1}" by (auto simp add: R.zero_cong R.res_carrier_eq) have "x [^]\<^bsub>residue_ring (int p)\<^esub> i = x ^ i mod (int p)" if "x \ {1 .. int p - 1}" for x and i :: nat using that R.pow_cong[of x i] by auto moreover obtain a where a: "a \ {1 .. int p - 1}" and a_gen: "{1 .. int p - 1} = {a[^]\<^bsub>residue_ring (int p)\<^esub>i|i::nat . i \ UNIV}" using field.finite_field_mult_group_has_gen[OF R.is_field] by (auto simp add: car[symmetric] carrier_mult_of) moreover have "nat ` {1 .. int p - 1} = {1 .. p - 1}" (is "?L = ?R") proof have "n \ ?R" if "n \ ?L" for n using that \p\2\ by force then show "?L \ ?R" by blast have "n \ ?L" if "n \ ?R" for n using that \p\2\ by (auto intro: rev_image_eqI [of "int n"]) then show "?R \ ?L" by blast qed moreover have "nat ` {a^i mod (int p) | i::nat. i \ UNIV} = {nat a^i mod p | i . i \ UNIV}" (is "?L = ?R") proof have "x \ ?R" if "x \ ?L" for x proof - from that obtain i where i: "x = nat (a^i mod (int p))" by blast then have "x = nat a ^ i mod p" using mod_nat_int_pow_eq[of a "int p" i] a \p\2\ by auto with i show ?thesis by blast qed then show "?L \ ?R" by blast have "x \ ?L" if "x \ ?R" for x proof - from that obtain i where i: "x = nat a^i mod p" by blast with mod_nat_int_pow_eq[of a "int p" i] a \p\2\ show ?thesis by auto qed then show "?R \ ?L" by blast qed ultimately have "{1 .. p - 1} = {nat a^i mod p | i. i \ UNIV}" by presburger moreover from a have "nat a \ {1 .. p - 1}" by force ultimately show ?thesis .. qed subsection \Upper bound for the number of $n$-th roots\ lemma roots_mod_prime_bound: fixes n c p :: nat assumes "prime p" "n > 0" defines "A \ {x\{.. n" proof - define R where "R = residue_ring (int p)" - term monom from assms(1) interpret residues_prime p R by unfold_locales (simp_all add: R_def) interpret R: UP_domain R "UP R" by (unfold_locales) let ?f = "UnivPoly.monom (UP R) \\<^bsub>R\<^esub> n \\<^bsub>(UP R)\<^esub> UnivPoly.monom (UP R) (int (c mod p)) 0" have in_carrier: "int (c mod p) \ carrier R" using prime_gt_1_nat[OF assms(1)] by (simp add: R_def residue_ring_def) have "deg R ?f = n" using assms in_carrier by (simp add: R.deg_minus_eq) hence f_not_zero: "?f \ \\<^bsub>UP R\<^esub>" using assms by (auto simp add : R.deg_nzero_nzero) have roots_bound: "finite {a \ carrier R. UnivPoly.eval R R id a ?f = \\<^bsub>R\<^esub>} \ card {a \ carrier R. UnivPoly.eval R R id a ?f = \\<^bsub>R\<^esub>} \ deg R ?f" using finite in_carrier by (intro R.roots_bound[OF _ f_not_zero]) simp have subs: "{x \ carrier R. x [^]\<^bsub>R\<^esub> n = int (c mod p)} \ {a \ carrier R. UnivPoly.eval R R id a ?f = \\<^bsub>R\<^esub>}" using in_carrier by (auto simp: R.evalRR_simps) then have "card {x \ carrier R. x [^]\<^bsub>R\<^esub> n = int (c mod p)} \ card {a \ carrier R. UnivPoly.eval R R id a ?f = \\<^bsub>R\<^esub>}" using finite by (intro card_mono) auto also have "\ \ n" using \deg R ?f = n\ roots_bound by linarith also { fix x assume "x \ carrier R" hence "x [^]\<^bsub>R\<^esub> n = (x ^ n) mod (int p)" by (subst pow_cong [symmetric]) (auto simp: R_def residue_ring_def) } hence "{x \ carrier R. x [^]\<^bsub>R\<^esub> n = int (c mod p)} = {x \ carrier R. [x ^ n = int c] (mod p)}" by (fastforce simp: cong_def zmod_int) also have "bij_betw int A {x \ carrier R. [x ^ n = int c] (mod p)}" by (rule bij_betwI[of int _ _ nat]) (use cong_int_iff in \force simp: R_def residue_ring_def A_def\)+ from bij_betw_same_card[OF this] have "card {x \ carrier R. [x ^ n = int c] (mod p)} = card A" .. finally show ?thesis . qed end diff --git a/src/HOL/Tools/BNF/bnf_fp_util.ML b/src/HOL/Tools/BNF/bnf_fp_util.ML --- a/src/HOL/Tools/BNF/bnf_fp_util.ML +++ b/src/HOL/Tools/BNF/bnf_fp_util.ML @@ -1,1008 +1,984 @@ (* Title: HOL/Tools/BNF/bnf_fp_util.ML Author: Dmitriy Traytel, TU Muenchen Author: Jasmin Blanchette, TU Muenchen Author: Martin Desharnais, TU Muenchen Author: Stefan Berghofer, TU Muenchen Copyright 2012, 2013, 2014 Shared library for the datatype and codatatype constructions. *) signature BNF_FP_UTIL = sig exception EMPTY_DATATYPE of string type fp_result = {Ts: typ list, bnfs: BNF_Def.bnf list, pre_bnfs: BNF_Def.bnf list, absT_infos: BNF_Comp.absT_info list, ctors: term list, dtors: term list, xtor_un_folds: term list, xtor_co_recs: term list, xtor_co_induct: thm, dtor_ctors: thm list, ctor_dtors: thm list, ctor_injects: thm list, dtor_injects: thm list, xtor_maps: thm list, xtor_map_unique: thm, xtor_setss: thm list list, xtor_rels: thm list, xtor_un_fold_thms: thm list, xtor_co_rec_thms: thm list, xtor_un_fold_unique: thm, xtor_co_rec_unique: thm, xtor_un_fold_o_maps: thm list, xtor_co_rec_o_maps: thm list, xtor_un_fold_transfers: thm list, xtor_co_rec_transfers: thm list, xtor_rel_co_induct: thm, dtor_set_inducts: thm list} val morph_fp_result: morphism -> fp_result -> fp_result val time: Proof.context -> Timer.real_timer -> string -> Timer.real_timer val fixpoint: ('a * 'a -> bool) -> ('a list -> 'a list) -> 'a list -> 'a list val IITN: string val LevN: string val algN: string val behN: string val bisN: string val carTN: string val caseN: string val coN: string val coinductN: string val coinduct_strongN: string val corecN: string val corec_discN: string val corec_disc_iffN: string val ctorN: string val ctor_dtorN: string val ctor_exhaustN: string val ctor_induct2N: string val ctor_inductN: string val ctor_injectN: string val ctor_foldN: string val ctor_fold_o_mapN: string val ctor_fold_transferN: string val ctor_fold_uniqueN: string val ctor_mapN: string val ctor_map_uniqueN: string val ctor_recN: string val ctor_rec_o_mapN: string val ctor_rec_transferN: string val ctor_rec_uniqueN: string val ctor_relN: string val ctor_rel_inductN: string val ctor_set_inclN: string val ctor_set_set_inclN: string val dtorN: string val dtor_coinductN: string val dtor_corecN: string val dtor_corec_o_mapN: string val dtor_corec_transferN: string val dtor_corec_uniqueN: string val dtor_ctorN: string val dtor_exhaustN: string val dtor_injectN: string val dtor_mapN: string val dtor_map_coinductN: string val dtor_map_coinduct_strongN: string val dtor_map_uniqueN: string val dtor_relN: string val dtor_rel_coinductN: string val dtor_set_inclN: string val dtor_set_set_inclN: string val dtor_coinduct_strongN: string val dtor_unfoldN: string val dtor_unfold_o_mapN: string val dtor_unfold_transferN: string val dtor_unfold_uniqueN: string val exhaustN: string val colN: string val inductN: string val injectN: string val isNodeN: string val lsbisN: string val mapN: string val map_uniqueN: string val min_algN: string val morN: string val nchotomyN: string val recN: string val rel_casesN: string val rel_coinductN: string val rel_inductN: string val rel_injectN: string val rel_introsN: string val rel_distinctN: string val rel_selN: string val rvN: string val corec_selN: string val set_inclN: string val set_set_inclN: string val setN: string val simpsN: string val strTN: string val str_initN: string val sum_bdN: string val sum_bdTN: string val uniqueN: string (* TODO: Don't index set facts. Isabelle packages traditionally generate uniform names. *) val mk_ctor_setN: int -> string val mk_dtor_setN: int -> string val mk_dtor_set_inductN: int -> string val mk_set_inductN: int -> string val co_prefix: BNF_Util.fp_kind -> string val split_conj_thm: thm -> thm list val split_conj_prems: int -> thm -> thm val mk_sumTN: typ list -> typ val mk_sumTN_balanced: typ list -> typ val mk_tupleT_balanced: typ list -> typ val mk_sumprodT_balanced: typ list list -> typ val mk_proj: typ -> int -> int -> term val mk_convol: term * term -> term val mk_rel_prod: term -> term -> term val mk_rel_sum: term -> term -> term val Inl_const: typ -> typ -> term val Inr_const: typ -> typ -> term val mk_tuple_balanced: term list -> term val mk_tuple1_balanced: typ list -> term list -> term val abs_curried_balanced: typ list -> term -> term val mk_tupled_fun: term -> term -> term list -> term val mk_case_sum: term * term -> term val mk_case_sumN: term list -> term val mk_case_absumprod: typ -> term -> term list -> term list list -> term list list -> term val mk_Inl: typ -> term -> term val mk_Inr: typ -> term -> term val mk_sumprod_balanced: typ -> int -> int -> term list -> term val mk_absumprod: typ -> term -> int -> int -> term list -> term val dest_sumT: typ -> typ * typ val dest_sumTN_balanced: int -> typ -> typ list val dest_tupleT_balanced: int -> typ -> typ list val dest_absumprodT: typ -> typ -> int -> int list -> typ -> typ list list val If_const: typ -> term val mk_Field: term -> term val mk_If: term -> term -> term -> term val mk_absumprodE: thm -> int list -> thm val mk_sum_caseN: int -> int -> thm val mk_sum_caseN_balanced: int -> int -> thm val mk_sum_Cinfinite: thm list -> thm val mk_sum_card_order: thm list -> thm val force_typ: Proof.context -> typ -> term -> term val mk_xtor_rel_co_induct_thm: BNF_Util.fp_kind -> term list -> term list -> term list -> term list -> term list -> term list -> term list -> term list -> ({prems: thm list, context: Proof.context} -> tactic) -> Proof.context -> thm val mk_xtor_co_iter_transfer_thms: BNF_Util.fp_kind -> term list -> term list -> term list -> term list -> term list -> term list -> term list -> ({prems: thm list, context: Proof.context} -> tactic) -> Proof.context -> thm list val mk_xtor_co_iter_o_map_thms: BNF_Util.fp_kind -> bool -> int -> thm -> thm list -> thm list -> thm list -> thm list -> thm list val derive_xtor_co_recs: BNF_Util.fp_kind -> binding list -> (typ list -> typ list) -> (typ list list * typ list) -> BNF_Def.bnf list -> term list -> term list -> thm -> thm list -> thm list -> thm list -> thm list -> (BNF_Comp.absT_info * BNF_Comp.absT_info) option list -> local_theory -> (term list * (thm list * thm * thm list * thm list)) * local_theory val raw_qualify: (binding -> binding) -> binding -> binding -> binding val fixpoint_bnf: bool -> (binding -> binding) -> (binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list -> BNF_Comp.absT_info list -> local_theory -> 'a) -> binding list -> (string * sort) list -> (string * sort) list -> (string * sort) list -> typ list -> BNF_Comp.comp_cache -> local_theory -> ((BNF_Def.bnf list * BNF_Comp.absT_info list) * BNF_Comp.comp_cache) * 'a - val fp_antiquote_setup: binding -> theory -> theory end; structure BNF_FP_Util : BNF_FP_UTIL = struct open Ctr_Sugar open BNF_Comp open BNF_Def open BNF_Util open BNF_FP_Util_Tactics exception EMPTY_DATATYPE of string; type fp_result = {Ts: typ list, bnfs: bnf list, pre_bnfs: BNF_Def.bnf list, absT_infos: BNF_Comp.absT_info list, ctors: term list, dtors: term list, xtor_un_folds: term list, xtor_co_recs: term list, xtor_co_induct: thm, dtor_ctors: thm list, ctor_dtors: thm list, ctor_injects: thm list, dtor_injects: thm list, xtor_maps: thm list, xtor_map_unique: thm, xtor_setss: thm list list, xtor_rels: thm list, xtor_un_fold_thms: thm list, xtor_co_rec_thms: thm list, xtor_un_fold_unique: thm, xtor_co_rec_unique: thm, xtor_un_fold_o_maps: thm list, xtor_co_rec_o_maps: thm list, xtor_un_fold_transfers: thm list, xtor_co_rec_transfers: thm list, xtor_rel_co_induct: thm, dtor_set_inducts: thm list}; fun morph_fp_result phi {Ts, bnfs, pre_bnfs, absT_infos, ctors, dtors, xtor_un_folds, xtor_co_recs, xtor_co_induct, dtor_ctors, ctor_dtors, ctor_injects, dtor_injects, xtor_maps, xtor_map_unique, xtor_setss, xtor_rels, xtor_un_fold_thms, xtor_co_rec_thms, xtor_un_fold_unique, xtor_co_rec_unique, xtor_un_fold_o_maps, xtor_co_rec_o_maps, xtor_un_fold_transfers, xtor_co_rec_transfers, xtor_rel_co_induct, dtor_set_inducts} = {Ts = map (Morphism.typ phi) Ts, bnfs = map (morph_bnf phi) bnfs, pre_bnfs = map (morph_bnf phi) pre_bnfs, absT_infos = map (morph_absT_info phi) absT_infos, ctors = map (Morphism.term phi) ctors, dtors = map (Morphism.term phi) dtors, xtor_un_folds = map (Morphism.term phi) xtor_un_folds, xtor_co_recs = map (Morphism.term phi) xtor_co_recs, xtor_co_induct = Morphism.thm phi xtor_co_induct, dtor_ctors = map (Morphism.thm phi) dtor_ctors, ctor_dtors = map (Morphism.thm phi) ctor_dtors, ctor_injects = map (Morphism.thm phi) ctor_injects, dtor_injects = map (Morphism.thm phi) dtor_injects, xtor_maps = map (Morphism.thm phi) xtor_maps, xtor_map_unique = Morphism.thm phi xtor_map_unique, xtor_setss = map (map (Morphism.thm phi)) xtor_setss, xtor_rels = map (Morphism.thm phi) xtor_rels, xtor_un_fold_thms = map (Morphism.thm phi) xtor_un_fold_thms, xtor_co_rec_thms = map (Morphism.thm phi) xtor_co_rec_thms, xtor_un_fold_unique = Morphism.thm phi xtor_un_fold_unique, xtor_co_rec_unique = Morphism.thm phi xtor_co_rec_unique, xtor_un_fold_o_maps = map (Morphism.thm phi) xtor_un_fold_o_maps, xtor_co_rec_o_maps = map (Morphism.thm phi) xtor_co_rec_o_maps, xtor_un_fold_transfers = map (Morphism.thm phi) xtor_un_fold_transfers, xtor_co_rec_transfers = map (Morphism.thm phi) xtor_co_rec_transfers, xtor_rel_co_induct = Morphism.thm phi xtor_rel_co_induct, dtor_set_inducts = map (Morphism.thm phi) dtor_set_inducts}; fun time ctxt timer msg = (if Config.get ctxt bnf_timing then warning (msg ^ ": " ^ string_of_int (Time.toMilliseconds (Timer.checkRealTimer timer)) ^ " ms") else (); Timer.startRealTimer ()); val preN = "pre_" val rawN = "raw_" val coN = "co" val unN = "un" val algN = "alg" val IITN = "IITN" val foldN = "fold" val unfoldN = unN ^ foldN val uniqueN = "unique" val transferN = "transfer" val simpsN = "simps" val ctorN = "ctor" val dtorN = "dtor" val ctor_foldN = ctorN ^ "_" ^ foldN val dtor_unfoldN = dtorN ^ "_" ^ unfoldN val ctor_fold_uniqueN = ctor_foldN ^ "_" ^ uniqueN val ctor_fold_o_mapN = ctor_foldN ^ "_o_" ^ mapN val dtor_unfold_uniqueN = dtor_unfoldN ^ "_" ^ uniqueN val dtor_unfold_o_mapN = dtor_unfoldN ^ "_o_" ^ mapN val ctor_fold_transferN = ctor_foldN ^ "_" ^ transferN val dtor_unfold_transferN = dtor_unfoldN ^ "_" ^ transferN val ctor_mapN = ctorN ^ "_" ^ mapN val dtor_mapN = dtorN ^ "_" ^ mapN val map_uniqueN = mapN ^ "_" ^ uniqueN val ctor_map_uniqueN = ctorN ^ "_" ^ map_uniqueN val dtor_map_uniqueN = dtorN ^ "_" ^ map_uniqueN val min_algN = "min_alg" val morN = "mor" val bisN = "bis" val lsbisN = "lsbis" val sum_bdTN = "sbdT" val sum_bdN = "sbd" val carTN = "carT" val strTN = "strT" val isNodeN = "isNode" val LevN = "Lev" val rvN = "recover" val behN = "beh" val setN = "set" val mk_ctor_setN = prefix (ctorN ^ "_") o mk_setN val mk_dtor_setN = prefix (dtorN ^ "_") o mk_setN fun mk_set_inductN i = mk_setN i ^ "_induct" val mk_dtor_set_inductN = prefix (dtorN ^ "_") o mk_set_inductN val str_initN = "str_init" val recN = "rec" val corecN = coN ^ recN val ctor_recN = ctorN ^ "_" ^ recN val ctor_rec_o_mapN = ctor_recN ^ "_o_" ^ mapN val ctor_rec_transferN = ctor_recN ^ "_" ^ transferN val ctor_rec_uniqueN = ctor_recN ^ "_" ^ uniqueN val dtor_corecN = dtorN ^ "_" ^ corecN val dtor_corec_o_mapN = dtor_corecN ^ "_o_" ^ mapN val dtor_corec_transferN = dtor_corecN ^ "_" ^ transferN val dtor_corec_uniqueN = dtor_corecN ^ "_" ^ uniqueN val ctor_dtorN = ctorN ^ "_" ^ dtorN val dtor_ctorN = dtorN ^ "_" ^ ctorN val nchotomyN = "nchotomy" val injectN = "inject" val exhaustN = "exhaust" val ctor_injectN = ctorN ^ "_" ^ injectN val ctor_exhaustN = ctorN ^ "_" ^ exhaustN val dtor_injectN = dtorN ^ "_" ^ injectN val dtor_exhaustN = dtorN ^ "_" ^ exhaustN val ctor_relN = ctorN ^ "_" ^ relN val dtor_relN = dtorN ^ "_" ^ relN val inductN = "induct" val coinductN = coN ^ inductN val ctor_inductN = ctorN ^ "_" ^ inductN val ctor_induct2N = ctor_inductN ^ "2" val dtor_map_coinductN = dtor_mapN ^ "_" ^ coinductN val dtor_coinductN = dtorN ^ "_" ^ coinductN val coinduct_strongN = coinductN ^ "_strong" val dtor_map_coinduct_strongN = dtor_mapN ^ "_" ^ coinduct_strongN val dtor_coinduct_strongN = dtorN ^ "_" ^ coinduct_strongN val colN = "col" val set_inclN = "set_incl" val ctor_set_inclN = ctorN ^ "_" ^ set_inclN val dtor_set_inclN = dtorN ^ "_" ^ set_inclN val set_set_inclN = "set_set_incl" val ctor_set_set_inclN = ctorN ^ "_" ^ set_set_inclN val dtor_set_set_inclN = dtorN ^ "_" ^ set_set_inclN val caseN = "case" val discN = "disc" val corec_discN = corecN ^ "_" ^ discN val iffN = "_iff" val corec_disc_iffN = corec_discN ^ iffN val distinctN = "distinct" val rel_distinctN = relN ^ "_" ^ distinctN val injectN = "inject" val rel_casesN = relN ^ "_cases" val rel_injectN = relN ^ "_" ^ injectN val rel_introsN = relN ^ "_intros" val rel_coinductN = relN ^ "_" ^ coinductN val rel_selN = relN ^ "_sel" val dtor_rel_coinductN = dtorN ^ "_" ^ rel_coinductN val rel_inductN = relN ^ "_" ^ inductN val ctor_rel_inductN = ctorN ^ "_" ^ rel_inductN val selN = "sel" val corec_selN = corecN ^ "_" ^ selN fun co_prefix fp = case_fp fp "" "co"; fun dest_sumT (Type (\<^type_name>\sum\, [T, T'])) = (T, T'); val dest_sumTN_balanced = Balanced_Tree.dest dest_sumT; fun dest_tupleT_balanced 0 \<^typ>\unit\ = [] | dest_tupleT_balanced n T = Balanced_Tree.dest HOLogic.dest_prodT n T; fun dest_absumprodT absT repT n ms = map2 dest_tupleT_balanced ms o dest_sumTN_balanced n o mk_repT absT repT; val mk_sumTN = Library.foldr1 mk_sumT; val mk_sumTN_balanced = Balanced_Tree.make mk_sumT; fun mk_tupleT_balanced [] = HOLogic.unitT | mk_tupleT_balanced Ts = Balanced_Tree.make HOLogic.mk_prodT Ts; val mk_sumprodT_balanced = mk_sumTN_balanced o map mk_tupleT_balanced; fun mk_proj T n k = let val (binders, _) = strip_typeN n T in fold_rev (fn T => fn t => Abs (Name.uu, T, t)) binders (Bound (n - k - 1)) end; fun mk_convol (f, g) = let val (fU, fTU) = `range_type (fastype_of f); val ((gT, gU), gTU) = `dest_funT (fastype_of g); val convolT = fTU --> gTU --> gT --> HOLogic.mk_prodT (fU, gU); in Const (\<^const_name>\convol\, convolT) $ f $ g end; fun mk_rel_prod R S = let val ((A1, A2), RT) = `dest_pred2T (fastype_of R); val ((B1, B2), ST) = `dest_pred2T (fastype_of S); val rel_prodT = RT --> ST --> mk_pred2T (HOLogic.mk_prodT (A1, B1)) (HOLogic.mk_prodT (A2, B2)); in Const (\<^const_name>\rel_prod\, rel_prodT) $ R $ S end; fun mk_rel_sum R S = let val ((A1, A2), RT) = `dest_pred2T (fastype_of R); val ((B1, B2), ST) = `dest_pred2T (fastype_of S); val rel_sumT = RT --> ST --> mk_pred2T (mk_sumT (A1, B1)) (mk_sumT (A2, B2)); in Const (\<^const_name>\rel_sum\, rel_sumT) $ R $ S end; fun Inl_const LT RT = Const (\<^const_name>\Inl\, LT --> mk_sumT (LT, RT)); fun mk_Inl RT t = Inl_const (fastype_of t) RT $ t; fun Inr_const LT RT = Const (\<^const_name>\Inr\, RT --> mk_sumT (LT, RT)); fun mk_Inr LT t = Inr_const LT (fastype_of t) $ t; fun mk_prod1 bound_Ts (t, u) = HOLogic.pair_const (fastype_of1 (bound_Ts, t)) (fastype_of1 (bound_Ts, u)) $ t $ u; fun mk_tuple1_balanced _ [] = HOLogic.unit | mk_tuple1_balanced bound_Ts ts = Balanced_Tree.make (mk_prod1 bound_Ts) ts; val mk_tuple_balanced = mk_tuple1_balanced []; fun abs_curried_balanced Ts t = t $ mk_tuple1_balanced (List.rev Ts) (map Bound (length Ts - 1 downto 0)) |> fold_rev (Term.abs o pair Name.uu) Ts; fun mk_sumprod_balanced T n k ts = Sum_Tree.mk_inj T n k (mk_tuple_balanced ts); fun mk_absumprod absT abs0 n k ts = let val abs = mk_abs absT abs0; in abs $ mk_sumprod_balanced (domain_type (fastype_of abs)) n k ts end; fun mk_case_sum (f, g) = let val (fT, T') = dest_funT (fastype_of f); val (gT, _) = dest_funT (fastype_of g); in Sum_Tree.mk_sumcase fT gT T' f g end; val mk_case_sumN = Library.foldr1 mk_case_sum; val mk_case_sumN_balanced = Balanced_Tree.make mk_case_sum; fun mk_tupled_fun f x xs = if xs = [x] then f else HOLogic.tupled_lambda x (Term.list_comb (f, xs)); fun mk_case_absumprod absT rep fs xss xss' = HOLogic.mk_comp (mk_case_sumN_balanced (@{map 3} mk_tupled_fun fs (map mk_tuple_balanced xss) xss'), mk_rep absT rep); fun If_const T = Const (\<^const_name>\If\, HOLogic.boolT --> T --> T --> T); fun mk_If p t f = let val T = fastype_of t in If_const T $ p $ t $ f end; fun mk_Field r = let val T = fst (dest_relT (fastype_of r)); in Const (\<^const_name>\Field\, mk_relT (T, T) --> HOLogic.mk_setT T) $ r end; (*dangerous; use with monotonic, converging functions only!*) fun fixpoint eq f X = if subset eq (f X, X) then X else fixpoint eq f (f X); (* stolen from "~~/src/HOL/Tools/Datatype/datatype_aux.ML" *) fun split_conj_thm th = ((th RS conjunct1) :: split_conj_thm (th RS conjunct2)) handle THM _ => [th]; fun split_conj_prems limit th = let fun split n i th = if i = n then th else split n (i + 1) (conjI RSN (i, th)) handle THM _ => th; in split limit 1 th end; fun mk_obj_sumEN_balanced n = Balanced_Tree.make (fn (thm1, thm2) => thm1 RSN (1, thm2 RSN (2, @{thm obj_sumE_f}))) (replicate n asm_rl); fun mk_tupled_allIN_balanced 0 = @{thm unit_all_impI} | mk_tupled_allIN_balanced n = let val (tfrees, _) = BNF_Util.mk_TFrees n \<^context>; val T = mk_tupleT_balanced tfrees; in @{thm asm_rl[of "\x. P x \ Q x" for P Q]} |> Thm.instantiate' [SOME (Thm.ctyp_of \<^context> T)] [] |> Raw_Simplifier.rewrite_goals_rule \<^context> @{thms split_paired_All[THEN eq_reflection]} |> (fn thm => impI RS funpow n (fn th => allI RS th) thm) |> Thm.varifyT_global end; fun mk_absumprodE type_definition ms = let val n = length ms in mk_obj_sumEN_balanced n OF map mk_tupled_allIN_balanced ms RS (type_definition RS @{thm type_copy_obj_one_point_absE}) end; fun mk_sum_caseN 1 1 = refl | mk_sum_caseN _ 1 = @{thm sum.case(1)} | mk_sum_caseN 2 2 = @{thm sum.case(2)} | mk_sum_caseN n k = trans OF [@{thm case_sum_step(2)}, mk_sum_caseN (n - 1) (k - 1)]; fun mk_sum_step base step thm = if Thm.eq_thm_prop (thm, refl) then base else trans OF [step, thm]; fun mk_sum_caseN_balanced 1 1 = refl | mk_sum_caseN_balanced n k = Balanced_Tree.access {left = mk_sum_step @{thm sum.case(1)} @{thm case_sum_step(1)}, right = mk_sum_step @{thm sum.case(2)} @{thm case_sum_step(2)}, init = refl} n k; fun mk_sum_Cinfinite [thm] = thm | mk_sum_Cinfinite (thm :: thms) = @{thm Cinfinite_csum_weak} OF [thm, mk_sum_Cinfinite thms]; fun mk_sum_card_order [thm] = thm | mk_sum_card_order (thm :: thms) = @{thm card_order_csum} OF [thm, mk_sum_card_order thms]; fun mk_xtor_rel_co_induct_thm fp pre_rels pre_phis rels phis xs ys xtors xtor's tac lthy = let val pre_relphis = map (fn rel => Term.list_comb (rel, phis @ pre_phis)) pre_rels; val relphis = map (fn rel => Term.list_comb (rel, phis)) rels; fun mk_xtor fp' xtor x = if fp = fp' then xtor $ x else x; val dtor = mk_xtor Greatest_FP; val ctor = mk_xtor Least_FP; fun flip f x y = if fp = Greatest_FP then f y x else f x y; fun mk_prem pre_relphi phi x y xtor xtor' = HOLogic.mk_Trueprop (list_all_free [x, y] (flip (curry HOLogic.mk_imp) (pre_relphi $ (dtor xtor x) $ (dtor xtor' y)) (phi $ (ctor xtor x) $ (ctor xtor' y)))); val prems = @{map 6} mk_prem pre_relphis pre_phis xs ys xtors xtor's; val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (flip mk_leq) relphis pre_phis)); in Goal.prove_sorry lthy (map (fst o dest_Free) (phis @ pre_phis)) prems concl tac |> Thm.close_derivation \<^here> |> (fn thm => thm OF (replicate (length pre_rels) @{thm allI[OF allI[OF impI]]})) end; fun mk_xtor_co_iter_transfer_thms fp pre_rels pre_iphis pre_ophis rels phis un_folds un_folds' tac lthy = let val pre_relphis = map (fn rel => Term.list_comb (rel, phis @ pre_iphis)) pre_rels; val relphis = map (fn rel => Term.list_comb (rel, phis)) rels; fun flip f x y = if fp = Greatest_FP then f y x else f x y; val arg_rels = map2 (flip mk_rel_fun) pre_relphis pre_ophis; fun mk_transfer relphi pre_phi un_fold un_fold' = fold_rev mk_rel_fun arg_rels (flip mk_rel_fun relphi pre_phi) $ un_fold $ un_fold'; val transfers = @{map 4} mk_transfer relphis pre_ophis un_folds un_folds'; val goal = fold_rev Logic.all (phis @ pre_ophis) (HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj transfers)); in Goal.prove_sorry lthy [] [] goal tac |> Thm.close_derivation \<^here> |> split_conj_thm end; fun mk_xtor_co_iter_o_map_thms fp is_rec m un_fold_unique xtor_maps xtor_un_folds sym_map_comps map_cong0s = let val n = length sym_map_comps; val rewrite_comp_comp2 = case_fp fp @{thm rewriteR_comp_comp2} @{thm rewriteL_comp_comp2}; val rewrite_comp_comp = case_fp fp @{thm rewriteR_comp_comp} @{thm rewriteL_comp_comp}; val map_cong_passive_args1 = replicate m (case_fp fp @{thm id_comp} @{thm comp_id} RS fun_cong); val map_cong_active_args1 = replicate n (if is_rec then case_fp fp @{thm convol_o} @{thm o_case_sum} RS fun_cong else refl); val map_cong_passive_args2 = replicate m (case_fp fp @{thm comp_id} @{thm id_comp} RS fun_cong); val map_cong_active_args2 = replicate n (if is_rec then case_fp fp @{thm map_prod_o_convol_id} @{thm case_sum_o_map_sum_id} else case_fp fp @{thm id_comp} @{thm comp_id} RS fun_cong); fun mk_map_congs passive active = map (fn thm => thm OF (passive @ active) RS @{thm ext}) map_cong0s; val map_cong1s = mk_map_congs map_cong_passive_args1 map_cong_active_args1; val map_cong2s = mk_map_congs map_cong_passive_args2 map_cong_active_args2; fun mk_rewrites map_congs = map2 (fn sym_map_comp => fn map_cong => mk_trans sym_map_comp map_cong RS rewrite_comp_comp) sym_map_comps map_congs; val rewrite1s = mk_rewrites map_cong1s; val rewrite2s = mk_rewrites map_cong2s; val unique_prems = @{map 4} (fn xtor_map => fn un_fold => fn rewrite1 => fn rewrite2 => mk_trans (rewrite_comp_comp2 OF [xtor_map, un_fold]) (mk_trans rewrite1 (mk_sym rewrite2))) xtor_maps xtor_un_folds rewrite1s rewrite2s; in split_conj_thm (un_fold_unique OF map (case_fp fp I mk_sym) unique_prems) end; fun force_typ ctxt T = Term.map_types Type_Infer.paramify_vars #> Type.constraint T #> Syntax.check_term ctxt #> singleton (Variable.polymorphic ctxt); fun absT_info_encodeT thy (SOME (src : absT_info, dst : absT_info)) src_absT = let val src_repT = mk_repT (#absT src) (#repT src) src_absT; val dst_absT = mk_absT thy (#repT dst) (#absT dst) src_repT; in dst_absT end | absT_info_encodeT _ NONE T = T; fun absT_info_decodeT thy = absT_info_encodeT thy o Option.map swap; fun absT_info_encode thy fp (opt as SOME (src : absT_info, dst : absT_info)) t = let val co_alg_funT = case_fp fp domain_type range_type; fun co_swap pair = case_fp fp I swap pair; val mk_co_comp = curry (HOLogic.mk_comp o co_swap); val mk_co_abs = case_fp fp mk_abs mk_rep; val mk_co_rep = case_fp fp mk_rep mk_abs; val co_abs = case_fp fp #abs #rep; val co_rep = case_fp fp #rep #abs; val src_absT = co_alg_funT (fastype_of t); val dst_absT = absT_info_encodeT thy opt src_absT; val co_src_abs = mk_co_abs src_absT (co_abs src); val co_dst_rep = mk_co_rep dst_absT (co_rep dst); in mk_co_comp (mk_co_comp t co_src_abs) co_dst_rep end | absT_info_encode _ _ NONE t = t; fun absT_info_decode thy fp = absT_info_encode thy fp o Option.map swap; fun mk_xtor_un_fold_xtor_thms ctxt fp un_folds xtors xtor_un_fold_unique map_id0s absT_info_opts = let val thy = Proof_Context.theory_of ctxt; fun mk_goal un_fold = let val rhs = list_comb (un_fold, @{map 2} (absT_info_encode thy fp) absT_info_opts xtors); val T = range_type (fastype_of rhs); in HOLogic.mk_eq (HOLogic.id_const T, rhs) end; val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map mk_goal un_folds)); fun mk_inverses NONE = [] | mk_inverses (SOME (src, dst)) = [#type_definition dst RS @{thm type_definition.Abs_inverse[OF _ UNIV_I]}, #type_definition src RS @{thm type_definition.Rep_inverse}]; val inverses = maps mk_inverses absT_info_opts; in Goal.prove_sorry ctxt [] [] goal (fn {context = ctxt, prems = _} => mk_xtor_un_fold_xtor_tac ctxt xtor_un_fold_unique map_id0s inverses) |> split_conj_thm |> map mk_sym end; fun derive_xtor_co_recs fp bs mk_Ts (Dss, resDs) pre_bnfs xtors0 un_folds0 xtor_un_fold_unique xtor_un_folds xtor_un_fold_transfers xtor_maps xtor_rels absT_info_opts lthy = let val thy = Proof_Context.theory_of lthy; fun co_swap pair = case_fp fp I swap pair; val mk_co_comp = curry (HOLogic.mk_comp o co_swap); fun mk_co_algT T U = case_fp fp (T --> U) (U --> T); val co_alg_funT = case_fp fp domain_type range_type; val mk_co_product = curry (case_fp fp mk_convol mk_case_sum); val co_proj1_const = case_fp fp fst_const (uncurry Inl_const o dest_sumT) o co_alg_funT; val co_proj2_const = case_fp fp snd_const (uncurry Inr_const o dest_sumT) o co_alg_funT; val mk_co_productT = curry (case_fp fp HOLogic.mk_prodT mk_sumT); val rewrite_comp_comp = case_fp fp @{thm rewriteL_comp_comp} @{thm rewriteR_comp_comp}; val n = length pre_bnfs; val live = live_of_bnf (hd pre_bnfs); val m = live - n; val ks = 1 upto n; val map_id0s = map map_id0_of_bnf pre_bnfs; val map_comps = map map_comp_of_bnf pre_bnfs; val map_cong0s = map map_cong0_of_bnf pre_bnfs; val map_transfers = map map_transfer_of_bnf pre_bnfs; val sym_map_comp0s = map (mk_sym o map_comp0_of_bnf) pre_bnfs; val deads = fold (union (op =)) Dss resDs; val ((((As, Bs), Xs), Ys), names_lthy) = lthy |> fold Variable.declare_typ deads |> mk_TFrees m ||>> mk_TFrees m ||>> mk_TFrees n ||>> mk_TFrees n; val XFTs = @{map 2} (fn Ds => mk_T_of_bnf Ds (As @ Xs)) Dss pre_bnfs; val co_algXFTs = @{map 2} mk_co_algT XFTs Xs; val Ts = mk_Ts As; val un_foldTs = @{map 2} (fn T => fn X => co_algXFTs ---> mk_co_algT T X) Ts Xs; val un_folds = @{map 2} (force_typ names_lthy) un_foldTs un_folds0; val ABs = As ~~ Bs; val XYs = Xs ~~ Ys; val Us = map (typ_subst_atomic ABs) Ts; val TFTs = @{map 2} (fn Ds => mk_T_of_bnf Ds (As @ Ts)) Dss pre_bnfs; val TFTs' = @{map 2} (absT_info_decodeT thy) absT_info_opts TFTs; val xtors = @{map 3} (force_typ names_lthy oo mk_co_algT) TFTs' Ts xtors0; val ids = map HOLogic.id_const As; val co_rec_Xs = @{map 2} mk_co_productT Ts Xs; val co_rec_Ys = @{map 2} mk_co_productT Us Ys; val co_rec_algXs = @{map 2} mk_co_algT co_rec_Xs Xs; val co_proj1s = map co_proj1_const co_rec_algXs; val co_rec_maps = @{map 2} (fn Ds => mk_map_of_bnf Ds (As @ case_fp fp co_rec_Xs Ts) (As @ case_fp fp Ts co_rec_Xs)) Dss pre_bnfs; val co_rec_Ts = @{map 2} (fn Ds => mk_T_of_bnf Ds (As @ co_rec_Xs)) Dss pre_bnfs val co_rec_argTs = @{map 2} mk_co_algT co_rec_Ts Xs; val co_rec_resTs = @{map 2} mk_co_algT Ts Xs; val (((co_rec_ss, fs), xs), names_lthy) = names_lthy |> mk_Frees "s" co_rec_argTs ||>> mk_Frees "f" co_rec_resTs ||>> mk_Frees "x" (case_fp fp TFTs' Xs); val co_rec_strs = @{map 4} (fn xtor => fn s => fn mapx => fn absT_info_opt => mk_co_product (mk_co_comp (absT_info_encode thy fp absT_info_opt xtor) (list_comb (mapx, ids @ co_proj1s))) s) xtors co_rec_ss co_rec_maps absT_info_opts; val theta = Xs ~~ co_rec_Xs; val co_rec_un_folds = map (subst_atomic_types theta) un_folds; val co_rec_spec0s = map (fn un_fold => list_comb (un_fold, co_rec_strs)) co_rec_un_folds; val co_rec_ids = @{map 2} (mk_co_comp o co_proj1_const) co_rec_algXs co_rec_spec0s; val co_rec_specs = @{map 2} (mk_co_comp o co_proj2_const) co_rec_algXs co_rec_spec0s; val co_recN = case_fp fp ctor_recN dtor_corecN; fun co_rec_bind i = nth bs (i - 1) |> Binding.prefix_name (co_recN ^ "_"); val co_rec_def_bind = rpair [] o Binding.concealed o Thm.def_binding o co_rec_bind; fun co_rec_spec i = fold_rev (Term.absfree o Term.dest_Free) co_rec_ss (nth co_rec_specs (i - 1)); val ((co_rec_frees, (_, co_rec_def_frees)), (lthy, lthy_old)) = lthy |> Local_Theory.open_target |> snd |> fold_map (fn i => Local_Theory.define ((co_rec_bind i, NoSyn), (co_rec_def_bind i, co_rec_spec i))) ks |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val co_rec_names = map (fst o dest_Const o Morphism.term phi) co_rec_frees; val co_recs = @{map 2} (fn name => fn resT => Const (name, co_rec_argTs ---> resT)) co_rec_names co_rec_resTs; val co_rec_defs = map (fn def => mk_unabs_def n (HOLogic.mk_obj_eq (Morphism.thm phi def))) co_rec_def_frees; val xtor_un_fold_xtor_thms = mk_xtor_un_fold_xtor_thms lthy fp (map (Term.subst_atomic_types (Xs ~~ Ts)) un_folds) xtors xtor_un_fold_unique map_id0s absT_info_opts; val co_rec_id_thms = let val goal = @{map 2} (fn T => fn t => HOLogic.mk_eq (t, HOLogic.id_const T)) Ts co_rec_ids |> Library.foldr1 HOLogic.mk_conj |> HOLogic.mk_Trueprop; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_xtor_co_rec_id_tac ctxt xtor_un_fold_xtor_thms xtor_un_fold_unique xtor_un_folds map_comps) |> Thm.close_derivation \<^here> |> split_conj_thm end; val co_rec_app_ss = map (fn co_rec => list_comb (co_rec, co_rec_ss)) co_recs; val co_products = @{map 2} (fn T => mk_co_product (HOLogic.id_const T)) Ts co_rec_app_ss; val co_rec_maps_rev = @{map 2} (fn Ds => mk_map_of_bnf Ds (As @ case_fp fp Ts co_rec_Xs) (As @ case_fp fp co_rec_Xs Ts)) Dss pre_bnfs; fun mk_co_app f g x = case_fp fp (f $ (g $ x)) (g $ (f $ x)); val co_rec_expand_thms = map (fn thm => thm RS case_fp fp @{thm convol_expand_snd} @{thm case_sum_expand_Inr_pointfree}) co_rec_id_thms; val xtor_co_rec_thms = let fun mk_goal co_rec s mapx xtor x absT_info_opt = let val lhs = mk_co_app co_rec xtor x; val rhs = mk_co_app s (list_comb (mapx, ids @ co_products) |> absT_info_decode thy fp absT_info_opt) x; in mk_Trueprop_eq (lhs, rhs) end; val goals = @{map 6} mk_goal co_rec_app_ss co_rec_ss co_rec_maps_rev xtors xs absT_info_opts; in map2 (fn goal => fn un_fold => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_xtor_co_rec_tac ctxt un_fold co_rec_defs co_rec_expand_thms)) |> Thm.close_derivation \<^here>) goals xtor_un_folds end; val co_product_fs = @{map 2} (fn T => mk_co_product (HOLogic.id_const T)) Ts fs; val co_rec_expand'_thms = map (fn thm => thm RS case_fp fp @{thm convol_expand_snd'} @{thm case_sum_expand_Inr'}) co_rec_id_thms; val xtor_co_rec_unique_thm = let fun mk_prem f s mapx xtor absT_info_opt = let val lhs = mk_co_comp f xtor; val rhs = mk_co_comp s (list_comb (mapx, ids @ co_product_fs)) |> absT_info_decode thy fp absT_info_opt; in mk_Trueprop_eq (co_swap (lhs, rhs)) end; val prems = @{map 5} mk_prem fs co_rec_ss co_rec_maps_rev xtors absT_info_opts; val concl = @{map 2} (curry HOLogic.mk_eq) fs co_rec_app_ss |> Library.foldr1 HOLogic.mk_conj |> HOLogic.mk_Trueprop; val goal = Logic.list_implies (prems, concl); val vars = Variable.add_free_names lthy goal []; fun mk_inverses NONE = [] | mk_inverses (SOME (src, dst)) = [#type_definition dst RS @{thm type_copy_Rep_o_Abs} RS rewrite_comp_comp, #type_definition src RS @{thm type_copy_Abs_o_Rep}]; val inverses = maps mk_inverses absT_info_opts; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_xtor_co_rec_unique_tac ctxt fp co_rec_defs co_rec_expand'_thms xtor_un_fold_unique map_id0s sym_map_comp0s inverses) |> Thm.close_derivation \<^here> end; val xtor_co_rec_o_map_thms = if forall is_none absT_info_opts then mk_xtor_co_iter_o_map_thms fp true m xtor_co_rec_unique_thm (map (mk_pointfree2 lthy) xtor_maps) (map (mk_pointfree2 lthy) xtor_co_rec_thms) sym_map_comp0s map_cong0s else replicate n refl (* FIXME *); val ABphiTs = @{map 2} mk_pred2T As Bs; val XYphiTs = @{map 2} mk_pred2T Xs Ys; val ((ABphis, XYphis), names_lthy) = names_lthy |> mk_Frees "R" ABphiTs ||>> mk_Frees "S" XYphiTs; val xtor_co_rec_transfer_thms = if forall is_none absT_info_opts then let val pre_rels = @{map 2} (fn Ds => mk_rel_of_bnf Ds (As @ co_rec_Xs) (Bs @ co_rec_Ys)) Dss pre_bnfs; val rels = @{map 3} (fn T => fn T' => Thm.prop_of #> HOLogic.dest_Trueprop #> fst o dest_comb #> fst o dest_comb #> funpow n (snd o dest_comb) #> case_fp fp (fst o dest_comb #> snd o dest_comb) (snd o dest_comb) #> head_of #> force_typ names_lthy (ABphiTs ---> mk_pred2T T T')) Ts Us xtor_un_fold_transfers; fun tac {context = ctxt, prems = _} = mk_xtor_co_rec_transfer_tac ctxt fp n m co_rec_defs xtor_un_fold_transfers map_transfers xtor_rels; val mk_rel_co_product = case_fp fp mk_rel_prod mk_rel_sum; val rec_phis = map2 (fn rel => mk_rel_co_product (Term.list_comb (rel, ABphis))) rels XYphis; in mk_xtor_co_iter_transfer_thms fp pre_rels rec_phis XYphis rels ABphis co_recs (map (subst_atomic_types (ABs @ XYs)) co_recs) tac lthy end else replicate n TrueI (* FIXME *); val notes = [(case_fp fp ctor_recN dtor_corecN, xtor_co_rec_thms), (case_fp fp ctor_rec_uniqueN dtor_corec_uniqueN, split_conj_thm xtor_co_rec_unique_thm), (case_fp fp ctor_rec_o_mapN dtor_corec_o_mapN, xtor_co_rec_o_map_thms), (case_fp fp ctor_rec_transferN dtor_corec_transferN, xtor_co_rec_transfer_thms)] |> map (apsnd (map single)) |> maps (fn (thmN, thmss) => map2 (fn b => fn thms => ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])])) bs thmss); val lthy = lthy |> Config.get lthy bnf_internals ? snd o Local_Theory.notes notes; in ((co_recs, (xtor_co_rec_thms, xtor_co_rec_unique_thm, xtor_co_rec_o_map_thms, xtor_co_rec_transfer_thms)), lthy) end; fun raw_qualify extra_qualify base_b = let val qs = Binding.path_of base_b; val n = Binding.name_of base_b; in Binding.prefix_name rawN #> fold_rev (fn (s, mand) => Binding.qualify mand s) (qs @ [(n, true)]) #> extra_qualify #> Binding.concealed end; fun fixpoint_bnf force_out_of_line extra_qualify construct_fp bs resBs Ds0 Xs rhsXs comp_cache0 lthy = let val time = time lthy; val timer = time (Timer.startRealTimer ()); fun flatten_tyargs Ass = subtract (op =) Xs (filter (fn T => exists (fn Ts => member (op =) Ts T) Ass) resBs) @ Xs; val ((bnfs, (deadss, livess)), (comp_cache_unfold_set, lthy')) = apfst (apsnd split_list o split_list) (@{fold_map 2} (fn b => bnf_of_typ true Smart_Inline (raw_qualify extra_qualify b) flatten_tyargs Xs Ds0) bs rhsXs ((comp_cache0, empty_unfolds), lthy)); fun norm_qualify i = Binding.qualify true (Binding.name_of (nth bs (Int.max (0, i - 1)))) #> extra_qualify #> Binding.concealed; val Ass = map (map dest_TFree) livess; val Ds' = fold (fold Term.add_tfreesT) deadss []; val Ds = union (op =) Ds' Ds0; val missing = resBs |> fold (subtract (op =)) (Ds' :: Ass); val (dead_phantoms, live_phantoms) = List.partition (member (op =) Ds0) missing; val resBs' = resBs |> fold (subtract (op =)) [dead_phantoms, Ds]; val timer = time (timer "Construction of BNFs"); val ((kill_posss, _), (bnfs', ((comp_cache', unfold_set'), lthy''))) = normalize_bnfs norm_qualify Ass Ds (K (resBs' @ Xs)) bnfs (comp_cache_unfold_set, lthy'); val Dss = @{map 3} (fn lives => fn kill_posss => fn deads => deads @ map (nth lives) kill_posss) livess kill_posss deadss; val all_Dss = Dss |> force_out_of_line ? map (fn Ds' => union (op =) Ds' (map TFree Ds0)); fun pre_qualify b = Binding.qualify false (Binding.name_of b) #> extra_qualify #> not (Config.get lthy'' bnf_internals) ? Binding.concealed; val ((pre_bnfs, (deadss, absT_infos)), lthy''') = lthy'' |> @{fold_map 5} (fn b => seal_bnf (pre_qualify b) unfold_set' (Binding.prefix_name preN b)) bs (replicate (length rhsXs) (force_out_of_line orelse not (null live_phantoms))) Dss all_Dss bnfs' |>> split_list |>> apsnd split_list; val timer = time (timer "Normalization & sealing of BNFs"); val res = construct_fp bs resBs (map TFree dead_phantoms, deadss) pre_bnfs absT_infos lthy'''; val timer = time (timer "FP construction in total"); in (((pre_bnfs, absT_infos), comp_cache'), res) end; -fun fp_antiquote_setup binding = - Thy_Output.antiquotation_pretty_source_embedded binding - (Args.type_name {proper = true, strict = true}) - (fn ctxt => fn fcT_name => - (case Ctr_Sugar.ctr_sugar_of ctxt fcT_name of - NONE => error ("Not a known freely generated type name: " ^ quote fcT_name) - | SOME {T = T0, ctrs = ctrs0, ...} => - let - val freezeT = Term.map_atyps (fn TVar ((s, _), S) => TFree (s, S) | T => T); - - val T = freezeT T0; - val ctrs = map (Term.map_types freezeT) ctrs0; - - val pretty_typ_bracket = Syntax.pretty_typ (Config.put pretty_priority 1001 ctxt); - fun pretty_ctr ctr = - Pretty.block (Pretty.breaks (Syntax.pretty_term ctxt ctr :: - map pretty_typ_bracket (binder_types (fastype_of ctr)))); - in - Pretty.block (Pretty.keyword1 (Binding.name_of binding) :: Pretty.brk 1 :: - Syntax.pretty_typ ctxt T :: Pretty.str " =" :: Pretty.brk 1 :: - flat (separate [Pretty.brk 1, Pretty.str "| "] (map (single o pretty_ctr) ctrs))) - end)); - end; diff --git a/src/HOL/Tools/BNF/bnf_gfp.ML b/src/HOL/Tools/BNF/bnf_gfp.ML --- a/src/HOL/Tools/BNF/bnf_gfp.ML +++ b/src/HOL/Tools/BNF/bnf_gfp.ML @@ -1,2611 +1,2609 @@ (* Title: HOL/Tools/BNF/bnf_gfp.ML Author: Dmitriy Traytel, TU Muenchen Author: Andrei Popescu, TU Muenchen Author: Jasmin Blanchette, TU Muenchen Copyright 2012 Codatatype construction. *) signature BNF_GFP = sig val construct_gfp: mixfix list -> binding list -> binding list -> binding list -> binding list list -> binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list -> BNF_Comp.absT_info list -> local_theory -> BNF_FP_Util.fp_result * local_theory end; structure BNF_GFP : BNF_GFP = struct open BNF_Def open BNF_Util open BNF_Tactics open BNF_Comp open BNF_FP_Util open BNF_FP_Def_Sugar open BNF_GFP_Util open BNF_GFP_Tactics datatype wit_tree = Wit_Leaf of int | Wit_Node of (int * int * int list) * wit_tree list; fun mk_tree_args (I, T) (I', Ts) = (sort_distinct int_ord (I @ I'), T :: Ts); fun finish Iss m seen i (nwit, I) = let val treess = map (fn j => if j < m orelse member (op =) seen j then [([j], Wit_Leaf j)] else map_index (finish Iss m (insert (op =) j seen) j) (nth Iss (j - m)) |> flat |> minimize_wits) I; in map (fn (I, t) => (I, Wit_Node ((i - m, nwit, filter (fn i => i < m) I), t))) (fold_rev (map_product mk_tree_args) treess [([], [])]) |> minimize_wits end; fun tree_to_ctor_wit vars _ _ (Wit_Leaf j) = ([j], nth vars j) | tree_to_ctor_wit vars ctors witss (Wit_Node ((i, nwit, I), subtrees)) = (I, nth ctors i $ (Term.list_comb (snd (nth (nth witss i) nwit), map (snd o tree_to_ctor_wit vars ctors witss) subtrees))); fun tree_to_coind_wits _ (Wit_Leaf _) = [] | tree_to_coind_wits lwitss (Wit_Node ((i, nwit, I), subtrees)) = ((i, I), nth (nth lwitss i) nwit) :: maps (tree_to_coind_wits lwitss) subtrees; (*all BNFs have the same lives*) fun construct_gfp mixfixes map_bs rel_bs pred_bs set_bss0 bs resBs (resDs, Dss) bnfs absT_infos lthy = let val time = time lthy; val timer = time (Timer.startRealTimer ()); val live = live_of_bnf (hd bnfs); val n = length bnfs; (*active*) val ks = 1 upto n; val m = live - n; (*passive, if 0 don't generate a new BNF*) val ls = 1 upto m; val internals = Config.get lthy bnf_internals; val b_names = map Binding.name_of bs; val b_name = mk_common_name b_names; val b = Binding.name b_name; fun mk_internal_of_b name = Binding.prefix_name (name ^ "_") #> Binding.prefix true b_name #> Binding.concealed; fun mk_internal_b name = mk_internal_of_b name b; fun mk_internal_bs name = map (mk_internal_of_b name) bs; val external_bs = map2 (Binding.prefix false) b_names bs |> not internals ? map Binding.concealed; val deads = fold (union (op =)) Dss resDs; val names_lthy = fold Variable.declare_typ deads lthy; val passives = map fst (subtract (op = o apsnd TFree) deads resBs); (* tvars *) val ((((((passiveAs, activeAs), passiveBs), activeBs), passiveCs), activeCs), idxT) = names_lthy |> variant_tfrees passives ||>> mk_TFrees n ||>> variant_tfrees passives ||>> mk_TFrees n ||>> mk_TFrees m ||>> mk_TFrees n ||> fst o mk_TFrees 1 ||> the_single; val allAs = passiveAs @ activeAs; val allBs' = passiveBs @ activeBs; val Ass = replicate n allAs; val allBs = passiveAs @ activeBs; val Bss = replicate n allBs; val allCs = passiveAs @ activeCs; val allCs' = passiveBs @ activeCs; val Css' = replicate n allCs'; (* types *) val dead_poss = map (fn x => if member (op =) deads (TFree x) then SOME (TFree x) else NONE) resBs; fun mk_param NONE passive = (hd passive, tl passive) | mk_param (SOME a) passive = (a, passive); val mk_params = fold_map mk_param dead_poss #> fst; fun mk_FTs Ts = map2 (fn Ds => mk_T_of_bnf Ds Ts) Dss bnfs; val (params, params') = `(map Term.dest_TFree) (mk_params passiveAs); val FTsAs = mk_FTs allAs; val FTsBs = mk_FTs allBs; val FTsCs = mk_FTs allCs; val ATs = map HOLogic.mk_setT passiveAs; val BTs = map HOLogic.mk_setT activeAs; val B'Ts = map HOLogic.mk_setT activeBs; val B''Ts = map HOLogic.mk_setT activeCs; val sTs = map2 (fn T => fn U => T --> U) activeAs FTsAs; val s'Ts = map2 (fn T => fn U => T --> U) activeBs FTsBs; val s''Ts = map2 (fn T => fn U => T --> U) activeCs FTsCs; val fTs = map2 (fn T => fn U => T --> U) activeAs activeBs; val self_fTs = map (fn T => T --> T) activeAs; val gTs = map2 (fn T => fn U => T --> U) activeBs activeCs; val all_gTs = map2 (fn T => fn U => T --> U) allBs allCs'; val RTs = map2 (fn T => fn U => HOLogic.mk_prodT (T, U)) activeAs activeBs; val sRTs = map2 (fn T => fn U => HOLogic.mk_prodT (T, U)) activeAs activeAs; val R'Ts = map2 (fn T => fn U => HOLogic.mk_prodT (T, U)) activeBs activeCs; val setsRTs = map HOLogic.mk_setT sRTs; val setRTs = map HOLogic.mk_setT RTs; val all_sbisT = HOLogic.mk_tupleT setsRTs; val setR'Ts = map HOLogic.mk_setT R'Ts; val FRTs = mk_FTs (passiveAs @ RTs); (* terms *) val mapsAsAs = @{map 4} mk_map_of_bnf Dss Ass Ass bnfs; val mapsAsBs = @{map 4} mk_map_of_bnf Dss Ass Bss bnfs; val mapsBsCs' = @{map 4} mk_map_of_bnf Dss Bss Css' bnfs; val mapsAsCs' = @{map 4} mk_map_of_bnf Dss Ass Css' bnfs; val map_fsts = @{map 4} mk_map_of_bnf Dss (replicate n (passiveAs @ RTs)) Ass bnfs; val map_snds = @{map 4} mk_map_of_bnf Dss (replicate n (passiveAs @ RTs)) Bss bnfs; fun mk_setss Ts = @{map 3} mk_sets_of_bnf (map (replicate live) Dss) (map (replicate live) (replicate n Ts)) bnfs; val setssAs = mk_setss allAs; val setssAs' = transpose setssAs; val bis_setss = mk_setss (passiveAs @ RTs); val relsAsBs = @{map 4} mk_rel_of_bnf Dss Ass Bss bnfs; val bds = @{map 3} mk_bd_of_bnf Dss Ass bnfs; val sum_bd = Library.foldr1 (uncurry mk_csum) bds; val sum_bdT = fst (dest_relT (fastype_of sum_bd)); val (sum_bdT_params, sum_bdT_params') = `(map TFree) (Term.add_tfreesT sum_bdT []); val ((((((((((zs, zs'), Bs), ss), fs), self_fs), all_gs), xFs), yFs), yFs_copy), _) = lthy |> mk_Frees' "b" activeAs ||>> mk_Frees "B" BTs ||>> mk_Frees "s" sTs ||>> mk_Frees "f" fTs ||>> mk_Frees "f" self_fTs ||>> mk_Frees "g" all_gTs ||>> mk_Frees "x" FTsAs ||>> mk_Frees "y" FTsBs ||>> mk_Frees "y" FTsBs; val passive_UNIVs = map HOLogic.mk_UNIV passiveAs; val passive_eqs = map HOLogic.eq_const passiveAs; val active_UNIVs = map HOLogic.mk_UNIV activeAs; val passive_ids = map HOLogic.id_const passiveAs; val active_ids = map HOLogic.id_const activeAs; val fsts = map fst_const RTs; val snds = map snd_const RTs; (* thms *) val bd_card_orders = map bd_card_order_of_bnf bnfs; val bd_card_order = hd bd_card_orders val bd_Card_orders = map bd_Card_order_of_bnf bnfs; val bd_Card_order = hd bd_Card_orders; val bd_Cinfinites = map bd_Cinfinite_of_bnf bnfs; val bd_Cinfinite = hd bd_Cinfinites; val in_monos = map in_mono_of_bnf bnfs; val map_comp0s = map map_comp0_of_bnf bnfs; val sym_map_comps = map mk_sym map_comp0s; val map_comps = map map_comp_of_bnf bnfs; val map_cong0s = map map_cong0_of_bnf bnfs; val map_id0s = map map_id0_of_bnf bnfs; val map_ids = map map_id_of_bnf bnfs; val set_bdss = map set_bd_of_bnf bnfs; val set_mapss = map set_map_of_bnf bnfs; val rel_congs = map rel_cong0_of_bnf bnfs; val rel_converseps = map rel_conversep_of_bnf bnfs; val rel_Grps = map rel_Grp_of_bnf bnfs; val le_rel_OOs = map le_rel_OO_of_bnf bnfs; val in_rels = map in_rel_of_bnf bnfs; val timer = time (timer "Extracted terms & thms"); (* derived thms *) (*map g1 ... gm g(m+1) ... g(m+n) (map id ... id f(m+1) ... f(m+n) x) = map g1 ... gm (g(m+1) o f(m+1)) ... (g(m+n) o f(m+n)) x*) fun mk_map_comp_id x mapAsBs mapBsCs mapAsCs map_comp0 = let val lhs = Term.list_comb (mapBsCs, all_gs) $ (Term.list_comb (mapAsBs, passive_ids @ fs) $ x); val rhs = Term.list_comb (mapAsCs, take m all_gs @ map HOLogic.mk_comp (drop m all_gs ~~ fs)) $ x; val goal = mk_Trueprop_eq (lhs, rhs); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_map_comp_id_tac ctxt map_comp0) |> Thm.close_derivation \<^here> end; val map_comp_id_thms = @{map 5} mk_map_comp_id xFs mapsAsBs mapsBsCs' mapsAsCs' map_comps; (*forall a : set(m+1) x. f(m+1) a = a; ...; forall a : set(m+n) x. f(m+n) a = a ==> map id ... id f(m+1) ... f(m+n) x = x*) fun mk_map_cong0L x mapAsAs sets map_cong0 map_id = let fun mk_prem set f z z' = HOLogic.mk_Trueprop (mk_Ball (set $ x) (Term.absfree z' (HOLogic.mk_eq (f $ z, z)))); val prems = @{map 4} mk_prem (drop m sets) self_fs zs zs'; val goal = mk_Trueprop_eq (Term.list_comb (mapAsAs, passive_ids @ self_fs) $ x, x); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, goal)) (fn {context = ctxt, prems = _} => mk_map_cong0L_tac ctxt m map_cong0 map_id) |> Thm.close_derivation \<^here> end; val map_cong0L_thms = @{map 5} mk_map_cong0L xFs mapsAsAs setssAs map_cong0s map_ids; val in_mono'_thms = map (fn thm => (thm OF (replicate m subset_refl)) RS @{thm set_mp}) in_monos; val map_arg_cong_thms = let val prems = map2 (curry mk_Trueprop_eq) yFs yFs_copy; val maps = map (fn mapx => Term.list_comb (mapx, all_gs)) mapsBsCs'; val concls = @{map 3} (fn x => fn y => fn mapx => mk_Trueprop_eq (mapx $ x, mapx $ y)) yFs yFs_copy maps; val goals = map2 (fn prem => fn concl => Logic.mk_implies (prem, concl)) prems concls; in map (fn goal => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => (hyp_subst_tac ctxt THEN' rtac ctxt refl) 1)) |> Thm.close_derivation \<^here>) goals end; val timer = time (timer "Derived simple theorems"); (* coalgebra *) val coalg_bind = mk_internal_b (coN ^ algN) ; val coalg_def_bind = (Thm.def_binding coalg_bind, []); (*forall i = 1 ... n: (\x \ Bi. si \ Fi_in UNIV .. UNIV B1 ... Bn)*) val coalg_spec = let val ins = @{map 3} mk_in (replicate n (passive_UNIVs @ Bs)) setssAs FTsAs; fun mk_coalg_conjunct B s X z z' = mk_Ball B (Term.absfree z' (HOLogic.mk_mem (s $ z, X))); val rhs = Library.foldr1 HOLogic.mk_conj (@{map 5} mk_coalg_conjunct Bs ss ins zs zs') in fold_rev (Term.absfree o Term.dest_Free) (Bs @ ss) rhs end; val ((coalg_free, (_, coalg_def_free)), (lthy, lthy_old)) = lthy |> Local_Theory.open_target |> snd |> Local_Theory.define ((coalg_bind, NoSyn), (coalg_def_bind, coalg_spec)) ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val coalg = fst (Term.dest_Const (Morphism.term phi coalg_free)); val coalg_def = mk_unabs_def (2 * n) (HOLogic.mk_obj_eq (Morphism.thm phi coalg_def_free)); fun mk_coalg Bs ss = let val args = Bs @ ss; val Ts = map fastype_of args; val coalgT = Library.foldr (op -->) (Ts, HOLogic.boolT); in Term.list_comb (Const (coalg, coalgT), args) end; val((((((zs, zs'), Bs), B's), ss), s's), _) = lthy |> mk_Frees' "b" activeAs ||>> mk_Frees "B" BTs ||>> mk_Frees "B'" B'Ts ||>> mk_Frees "s" sTs ||>> mk_Frees "s'" s'Ts; val coalg_prem = HOLogic.mk_Trueprop (mk_coalg Bs ss); val coalg_in_thms = map (fn i => coalg_def RS iffD1 RS mk_conjunctN n i RS bspec) ks val coalg_set_thmss = let val coalg_prem = HOLogic.mk_Trueprop (mk_coalg Bs ss); fun mk_prem x B = mk_Trueprop_mem (x, B); fun mk_concl s x B set = HOLogic.mk_Trueprop (mk_leq (set $ (s $ x)) B); val prems = map2 mk_prem zs Bs; val conclss = @{map 3} (fn s => fn x => fn sets => map2 (mk_concl s x) Bs (drop m sets)) ss zs setssAs; val goalss = map2 (fn prem => fn concls => map (fn concl => Logic.list_implies (coalg_prem :: [prem], concl)) concls) prems conclss; in map (fn goals => map (fn goal => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_coalg_set_tac ctxt coalg_def)) |> Thm.close_derivation \<^here>) goals) goalss end; fun mk_tcoalg BTs = mk_coalg (map HOLogic.mk_UNIV BTs); val tcoalg_thm = let val goal = HOLogic.mk_Trueprop (mk_tcoalg activeAs ss); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => (rtac ctxt (coalg_def RS iffD2) 1 THEN CONJ_WRAP (K (EVERY' [rtac ctxt ballI, rtac ctxt CollectI, CONJ_WRAP' (K (EVERY' [rtac ctxt @{thm subset_UNIV}])) allAs] 1)) ss)) |> Thm.close_derivation \<^here> end; val timer = time (timer "Coalgebra definition & thms"); (* morphism *) val mor_bind = mk_internal_b morN; val mor_def_bind = (Thm.def_binding mor_bind, []); (*fbetw) forall i = 1 ... n: (\x \ Bi. fi x \ B'i)*) (*mor) forall i = 1 ... n: (\x \ Bi. Fi_map id ... id f1 ... fn (si x) = si' (fi x)*) val mor_spec = let fun mk_fbetw f B1 B2 z z' = mk_Ball B1 (Term.absfree z' (HOLogic.mk_mem (f $ z, B2))); fun mk_mor B mapAsBs f s s' z z' = mk_Ball B (Term.absfree z' (HOLogic.mk_eq (Term.list_comb (mapAsBs, passive_ids @ fs @ [s $ z]), s' $ (f $ z)))); val rhs = HOLogic.mk_conj (Library.foldr1 HOLogic.mk_conj (@{map 5} mk_fbetw fs Bs B's zs zs'), Library.foldr1 HOLogic.mk_conj (@{map 7} mk_mor Bs mapsAsBs fs ss s's zs zs')) in fold_rev (Term.absfree o Term.dest_Free) (Bs @ ss @ B's @ s's @ fs) rhs end; val ((mor_free, (_, mor_def_free)), (lthy, lthy_old)) = lthy |> Local_Theory.open_target |> snd |> Local_Theory.define ((mor_bind, NoSyn), (mor_def_bind, mor_spec)) ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val mor = fst (Term.dest_Const (Morphism.term phi mor_free)); val mor_def = mk_unabs_def (5 * n) (HOLogic.mk_obj_eq (Morphism.thm phi mor_def_free)); fun mk_mor Bs1 ss1 Bs2 ss2 fs = let val args = Bs1 @ ss1 @ Bs2 @ ss2 @ fs; val Ts = map fastype_of (Bs1 @ ss1 @ Bs2 @ ss2 @ fs); val morT = Library.foldr (op -->) (Ts, HOLogic.boolT); in Term.list_comb (Const (mor, morT), args) end; val ((((((((((((((zs, z's), Bs), Bs_copy), B's), B''s), ss), s's), s''s), fs), fs_copy), gs), RFs), Rs), _) = lthy |> mk_Frees "b" activeAs ||>> mk_Frees "b" activeBs ||>> mk_Frees "B" BTs ||>> mk_Frees "B" BTs ||>> mk_Frees "B'" B'Ts ||>> mk_Frees "B''" B''Ts ||>> mk_Frees "s" sTs ||>> mk_Frees "s'" s'Ts ||>> mk_Frees "s''" s''Ts ||>> mk_Frees "f" fTs ||>> mk_Frees "f" fTs ||>> mk_Frees "g" gTs ||>> mk_Frees "x" FRTs ||>> mk_Frees "R" setRTs; val mor_prem = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs); val (mor_image_thms, morE_thms) = let val prem = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs); fun mk_image_goal f B1 B2 = Logic.mk_implies (prem, HOLogic.mk_Trueprop (mk_leq (mk_image f $ B1) B2)); val image_goals = @{map 3} mk_image_goal fs Bs B's; fun mk_elim_goal B mapAsBs f s s' x = Logic.list_implies ([prem, mk_Trueprop_mem (x, B)], mk_Trueprop_eq (Term.list_comb (mapAsBs, passive_ids @ fs @ [s $ x]), s' $ (f $ x))); val elim_goals = @{map 6} mk_elim_goal Bs mapsAsBs fs ss s's zs; fun prove goal = Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_mor_elim_tac ctxt mor_def)) |> Thm.close_derivation \<^here>; in (map prove image_goals, map prove elim_goals) end; val mor_image'_thms = map (fn thm => @{thm set_mp} OF [thm, imageI]) mor_image_thms; val mor_incl_thm = let val prems = map2 (HOLogic.mk_Trueprop oo mk_leq) Bs Bs_copy; val concl = HOLogic.mk_Trueprop (mk_mor Bs ss Bs_copy ss active_ids); val vars = fold (Variable.add_free_names lthy) (concl :: prems) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, concl)) (fn {context = ctxt, prems = _} => mk_mor_incl_tac ctxt mor_def map_ids) |> Thm.close_derivation \<^here> end; val mor_id_thm = mor_incl_thm OF (replicate n subset_refl); val mor_comp_thm = let val prems = [HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs), HOLogic.mk_Trueprop (mk_mor B's s's B''s s''s gs)]; val concl = HOLogic.mk_Trueprop (mk_mor Bs ss B''s s''s (map2 (curry HOLogic.mk_comp) gs fs)); val vars = fold (Variable.add_free_names lthy) (concl :: prems) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, concl)) (fn {context = ctxt, prems = _} => mk_mor_comp_tac ctxt mor_def mor_image'_thms morE_thms map_comp_id_thms) |> Thm.close_derivation \<^here> end; val mor_cong_thm = let val prems = map HOLogic.mk_Trueprop (map2 (curry HOLogic.mk_eq) fs_copy fs @ [mk_mor Bs ss B's s's fs]) val concl = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs_copy); val vars = fold (Variable.add_free_names lthy) (concl :: prems) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, concl)) (fn {context = ctxt, prems = _} => (hyp_subst_tac ctxt THEN' assume_tac ctxt) 1) |> Thm.close_derivation \<^here> end; val mor_UNIV_thm = let fun mk_conjunct mapAsBs f s s' = HOLogic.mk_eq (HOLogic.mk_comp (Term.list_comb (mapAsBs, passive_ids @ fs), s), HOLogic.mk_comp (s', f)); val lhs = mk_mor active_UNIVs ss (map HOLogic.mk_UNIV activeBs) s's fs; val rhs = Library.foldr1 HOLogic.mk_conj (@{map 4} mk_conjunct mapsAsBs fs ss s's); val vars = fold (Variable.add_free_names lthy) [lhs, rhs] []; in Goal.prove_sorry lthy vars [] (mk_Trueprop_eq (lhs, rhs)) (fn {context = ctxt, prems = _} => mk_mor_UNIV_tac ctxt morE_thms mor_def) |> Thm.close_derivation \<^here> end; val mor_str_thm = let val maps = map2 (fn Ds => fn bnf => Term.list_comb (mk_map_of_bnf Ds allAs (passiveAs @ FTsAs) bnf, passive_ids @ ss)) Dss bnfs; val goal = HOLogic.mk_Trueprop (mk_mor active_UNIVs ss (map HOLogic.mk_UNIV FTsAs) maps ss); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_mor_str_tac ctxt ks mor_UNIV_thm) |> Thm.close_derivation \<^here> end; val timer = time (timer "Morphism definition & thms"); (* bisimulation *) val bis_bind = mk_internal_b bisN; val bis_def_bind = (Thm.def_binding bis_bind, []); fun mk_bis_le_conjunct R B1 B2 = mk_leq R (mk_Times (B1, B2)); val bis_le = Library.foldr1 HOLogic.mk_conj (@{map 3} mk_bis_le_conjunct Rs Bs B's) val bis_spec = let val fst_args = passive_ids @ fsts; val snd_args = passive_ids @ snds; fun mk_bis R s s' b1 b2 RF map1 map2 sets = list_all_free [b1, b2] (HOLogic.mk_imp (HOLogic.mk_mem (HOLogic.mk_prod (b1, b2), R), mk_Bex (mk_in (passive_UNIVs @ Rs) sets (snd (dest_Free RF))) (Term.absfree (dest_Free RF) (HOLogic.mk_conj (HOLogic.mk_eq (Term.list_comb (map1, fst_args) $ RF, s $ b1), HOLogic.mk_eq (Term.list_comb (map2, snd_args) $ RF, s' $ b2)))))); val rhs = HOLogic.mk_conj (bis_le, Library.foldr1 HOLogic.mk_conj (@{map 9} mk_bis Rs ss s's zs z's RFs map_fsts map_snds bis_setss)) in fold_rev (Term.absfree o Term.dest_Free) (Bs @ ss @ B's @ s's @ Rs) rhs end; val ((bis_free, (_, bis_def_free)), (lthy, lthy_old)) = lthy |> Local_Theory.open_target |> snd |> Local_Theory.define ((bis_bind, NoSyn), (bis_def_bind, bis_spec)) ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val bis = fst (Term.dest_Const (Morphism.term phi bis_free)); val bis_def = mk_unabs_def (5 * n) (HOLogic.mk_obj_eq (Morphism.thm phi bis_def_free)); fun mk_bis Bs1 ss1 Bs2 ss2 Rs = let val args = Bs1 @ ss1 @ Bs2 @ ss2 @ Rs; val Ts = map fastype_of args; val bisT = Library.foldr (op -->) (Ts, HOLogic.boolT); in Term.list_comb (Const (bis, bisT), args) end; val (((((((((((((((((zs, z's), Bs), B's), B''s), ss), s's), s''s), fs), (Rtuple, Rtuple')), Rs), Rs_copy), R's), sRs), (idx, idx')), Idx), Ris), _) = lthy |> mk_Frees "b" activeAs ||>> mk_Frees "b" activeBs ||>> mk_Frees "B" BTs ||>> mk_Frees "B'" B'Ts ||>> mk_Frees "B''" B''Ts ||>> mk_Frees "s" sTs ||>> mk_Frees "s'" s'Ts ||>> mk_Frees "s''" s''Ts ||>> mk_Frees "f" fTs ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "Rtuple") all_sbisT ||>> mk_Frees "R" setRTs ||>> mk_Frees "R" setRTs ||>> mk_Frees "R'" setR'Ts ||>> mk_Frees "R" setsRTs ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "i") idxT ||>> yield_singleton (mk_Frees "I") (HOLogic.mk_setT idxT) ||>> mk_Frees "Ri" (map (fn T => idxT --> T) setRTs); val bis_cong_thm = let val prems = map HOLogic.mk_Trueprop (mk_bis Bs ss B's s's Rs :: map2 (curry HOLogic.mk_eq) Rs_copy Rs) val concl = HOLogic.mk_Trueprop (mk_bis Bs ss B's s's Rs_copy); val vars = fold (Variable.add_free_names lthy) (concl :: prems) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, concl)) (fn {context = ctxt, prems = _} => (hyp_subst_tac ctxt THEN' assume_tac ctxt) 1) |> Thm.close_derivation \<^here> end; val bis_rel_thm = let fun mk_conjunct R s s' b1 b2 rel = list_all_free [b1, b2] (HOLogic.mk_imp (HOLogic.mk_mem (HOLogic.mk_prod (b1, b2), R), Term.list_comb (rel, passive_eqs @ map mk_in_rel Rs) $ (s $ b1) $ (s' $ b2))); val rhs = HOLogic.mk_conj (bis_le, Library.foldr1 HOLogic.mk_conj (@{map 6} mk_conjunct Rs ss s's zs z's relsAsBs)) val goal = mk_Trueprop_eq (mk_bis Bs ss B's s's Rs, rhs); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_bis_rel_tac ctxt m bis_def in_rels map_comps map_cong0s set_mapss) |> Thm.close_derivation \<^here> end; val bis_converse_thm = let val goal = Logic.mk_implies (HOLogic.mk_Trueprop (mk_bis Bs ss B's s's Rs), HOLogic.mk_Trueprop (mk_bis B's s's Bs ss (map mk_converse Rs))); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_bis_converse_tac ctxt m bis_rel_thm rel_congs rel_converseps) |> Thm.close_derivation \<^here> end; val bis_O_thm = let val prems = [HOLogic.mk_Trueprop (mk_bis Bs ss B's s's Rs), HOLogic.mk_Trueprop (mk_bis B's s's B''s s''s R's)]; val concl = HOLogic.mk_Trueprop (mk_bis Bs ss B''s s''s (map2 (curry mk_rel_comp) Rs R's)); val vars = fold (Variable.add_free_names lthy) (concl :: prems) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, concl)) (fn {context = ctxt, prems = _} => mk_bis_O_tac ctxt m bis_rel_thm rel_congs le_rel_OOs) |> Thm.close_derivation \<^here> end; val bis_Gr_thm = let val concl = HOLogic.mk_Trueprop (mk_bis Bs ss B's s's (map2 mk_Gr Bs fs)); val vars = fold (Variable.add_free_names lthy) ([coalg_prem, mor_prem, concl]) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies ([coalg_prem, mor_prem], concl)) (fn {context = ctxt, prems = _} => mk_bis_Gr_tac ctxt bis_rel_thm rel_Grps mor_image_thms morE_thms coalg_in_thms) |> Thm.close_derivation \<^here> end; val bis_image2_thm = bis_cong_thm OF ((bis_O_thm OF [bis_Gr_thm RS bis_converse_thm, bis_Gr_thm]) :: replicate n @{thm image2_Gr}); val bis_Id_on_thm = bis_cong_thm OF ((mor_id_thm RSN (2, bis_Gr_thm)) :: replicate n @{thm Id_on_Gr}); val bis_Union_thm = let val prem = HOLogic.mk_Trueprop (mk_Ball Idx (Term.absfree idx' (mk_bis Bs ss B's s's (map (fn R => R $ idx) Ris)))); val concl = HOLogic.mk_Trueprop (mk_bis Bs ss B's s's (map (mk_UNION Idx) Ris)); val vars = fold (Variable.add_free_names lthy) [prem, concl] []; in Goal.prove_sorry lthy vars [] (Logic.mk_implies (prem, concl)) (fn {context = ctxt, prems = _} => mk_bis_Union_tac ctxt bis_def in_mono'_thms) |> Thm.close_derivation \<^here> end; (* self-bisimulation *) fun mk_sbis Bs ss Rs = mk_bis Bs ss Bs ss Rs; (* largest self-bisimulation *) val lsbis_binds = mk_internal_bs lsbisN; fun lsbis_bind i = nth lsbis_binds (i - 1); val lsbis_def_bind = rpair [] o Thm.def_binding o lsbis_bind; val all_sbis = HOLogic.mk_Collect (fst Rtuple', snd Rtuple', list_exists_free sRs (HOLogic.mk_conj (HOLogic.mk_eq (Rtuple, HOLogic.mk_tuple sRs), mk_sbis Bs ss sRs))); fun lsbis_spec i = fold_rev (Term.absfree o Term.dest_Free) (Bs @ ss) (mk_UNION all_sbis (Term.absfree Rtuple' (mk_nthN n Rtuple i))); val ((lsbis_frees, (_, lsbis_def_frees)), (lthy, lthy_old)) = lthy |> Local_Theory.open_target |> snd |> fold_map (fn i => Local_Theory.define ((lsbis_bind i, NoSyn), (lsbis_def_bind i, lsbis_spec i))) ks |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val lsbis_defs = map (fn def => mk_unabs_def (2 * n) (HOLogic.mk_obj_eq (Morphism.thm phi def))) lsbis_def_frees; val lsbiss = map (fst o Term.dest_Const o Morphism.term phi) lsbis_frees; fun mk_lsbis Bs ss i = let val args = Bs @ ss; val Ts = map fastype_of args; val RT = mk_relT (`I (HOLogic.dest_setT (fastype_of (nth Bs (i - 1))))); val lsbisT = Library.foldr (op -->) (Ts, RT); in Term.list_comb (Const (nth lsbiss (i - 1), lsbisT), args) end; val (((((zs, zs'), Bs), ss), sRs), _) = lthy |> mk_Frees' "b" activeAs ||>> mk_Frees "B" BTs ||>> mk_Frees "s" sTs ||>> mk_Frees "R" setsRTs; val sbis_prem = HOLogic.mk_Trueprop (mk_sbis Bs ss sRs); val coalg_prem = HOLogic.mk_Trueprop (mk_coalg Bs ss); val sbis_lsbis_thm = let val goal = HOLogic.mk_Trueprop (mk_sbis Bs ss (map (mk_lsbis Bs ss) ks)); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_sbis_lsbis_tac ctxt lsbis_defs bis_Union_thm bis_cong_thm) |> Thm.close_derivation \<^here> end; val lsbis_incl_thms = map (fn i => sbis_lsbis_thm RS (bis_def RS iffD1 RS conjunct1 RS mk_conjunctN n i)) ks; val lsbisE_thms = map (fn i => (mk_specN 2 (sbis_lsbis_thm RS (bis_def RS iffD1 RS conjunct2 RS mk_conjunctN n i))) RS mp) ks; val incl_lsbis_thms = let fun mk_concl i R = HOLogic.mk_Trueprop (mk_leq R (mk_lsbis Bs ss i)); val goals = map2 (fn i => fn R => Logic.mk_implies (sbis_prem, mk_concl i R)) ks sRs; in @{map 3} (fn goal => fn i => fn def => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_incl_lsbis_tac ctxt n i def)) |> Thm.close_derivation \<^here>) goals ks lsbis_defs end; val equiv_lsbis_thms = let fun mk_concl i B = HOLogic.mk_Trueprop (mk_equiv B (mk_lsbis Bs ss i)); val goals = map2 (fn i => fn B => Logic.mk_implies (coalg_prem, mk_concl i B)) ks Bs; in @{map 3} (fn goal => fn l_incl => fn incl_l => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_equiv_lsbis_tac ctxt sbis_lsbis_thm l_incl incl_l bis_Id_on_thm bis_converse_thm bis_O_thm) |> Thm.close_derivation \<^here>)) goals lsbis_incl_thms incl_lsbis_thms end; val timer = time (timer "Bisimulations"); (* bounds *) val (lthy, sbd, sbdT, sbd_card_order, sbd_Cinfinite, sbd_Card_order, set_sbdss) = if n = 1 then (lthy, sum_bd, sum_bdT, bd_card_order, bd_Cinfinite, bd_Card_order, set_bdss) else let val sbdT_bind = mk_internal_b sum_bdTN; val ((sbdT_name, (sbdT_glob_info, sbdT_loc_info)), lthy) = typedef (sbdT_bind, sum_bdT_params', NoSyn) (HOLogic.mk_UNIV sum_bdT) NONE (fn ctxt => EVERY' [rtac ctxt exI, rtac ctxt UNIV_I] 1) lthy; val sbdT = Type (sbdT_name, sum_bdT_params); val Abs_sbdT = Const (#Abs_name sbdT_glob_info, sum_bdT --> sbdT); val sbd_bind = mk_internal_b sum_bdN; val sbd_def_bind = (Thm.def_binding sbd_bind, []); val sbd_spec = mk_dir_image sum_bd Abs_sbdT; val ((sbd_free, (_, sbd_def_free)), (lthy, lthy_old)) = lthy |> Local_Theory.open_target |> snd |> Local_Theory.define ((sbd_bind, NoSyn), (sbd_def_bind, sbd_spec)) ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val sbd_def = HOLogic.mk_obj_eq (Morphism.thm phi sbd_def_free); val sbd = Const (fst (Term.dest_Const (Morphism.term phi sbd_free)), mk_relT (`I sbdT)); val Abs_sbdT_inj = mk_Abs_inj_thm (#Abs_inject sbdT_loc_info); val Abs_sbdT_bij = mk_Abs_bij_thm lthy Abs_sbdT_inj (#Abs_cases sbdT_loc_info); val sum_Cinfinite = mk_sum_Cinfinite bd_Cinfinites; val sum_Card_order = sum_Cinfinite RS conjunct2; val sum_card_order = mk_sum_card_order bd_card_orders; val sbd_ordIso = @{thm ssubst_Pair_rhs} OF [@{thm dir_image} OF [Abs_sbdT_inj, sum_Card_order], sbd_def]; val sbd_card_order = @{thm iffD2[OF arg_cong[of _ _ card_order]]} OF [sbd_def, @{thm card_order_dir_image} OF [Abs_sbdT_bij, sum_card_order]]; val sbd_Cinfinite = @{thm Cinfinite_cong} OF [sbd_ordIso, sum_Cinfinite]; val sbd_Card_order = sbd_Cinfinite RS conjunct2; fun mk_set_sbd i bd_Card_order bds = map (fn thm => @{thm ordLeq_ordIso_trans} OF [bd_Card_order RS mk_ordLeq_csum n i thm, sbd_ordIso]) bds; val set_sbdss = @{map 3} mk_set_sbd ks bd_Card_orders set_bdss; in (lthy, sbd, sbdT, sbd_card_order, sbd_Cinfinite, sbd_Card_order, set_sbdss) end; val sbdTs = replicate n sbdT; val sum_sbdT = mk_sumTN sbdTs; val sum_sbd_listT = HOLogic.listT sum_sbdT; val sum_sbd_list_setT = HOLogic.mk_setT sum_sbd_listT; val bdTs = passiveAs @ replicate n sbdT; val to_sbd_maps = @{map 4} mk_map_of_bnf Dss Ass (replicate n bdTs) bnfs; val bdFTs = mk_FTs bdTs; val sbdFT = mk_sumTN bdFTs; val treeT = HOLogic.mk_prodT (sum_sbd_list_setT, sum_sbd_listT --> sbdFT); val treeQT = HOLogic.mk_setT treeT; val treeTs = passiveAs @ replicate n treeT; val treeQTs = passiveAs @ replicate n treeQT; val treeFTs = mk_FTs treeTs; val tree_maps = @{map 4} mk_map_of_bnf Dss (replicate n bdTs) (replicate n treeTs) bnfs; val final_maps = @{map 4} mk_map_of_bnf Dss (replicate n treeTs) (replicate n treeQTs) bnfs; val isNode_setss = mk_setss (passiveAs @ replicate n sbdT); val root = HOLogic.mk_set sum_sbd_listT [HOLogic.mk_list sum_sbdT []]; val Zero = HOLogic.mk_tuple (map (fn U => absdummy U root) activeAs); val Lev_recT = fastype_of Zero; val Nil = HOLogic.mk_tuple (@{map 3} (fn i => fn z => fn z'=> Term.absfree z' (mk_InN activeAs z i)) ks zs zs'); val rv_recT = fastype_of Nil; val (((((((((((((((zs, zs'), zs_copy), ss), (nat, nat')), (sumx, sumx')), (kks, kks')), (kl, kl')), (kl_copy, kl'_copy)), (Kl, Kl')), (lab, lab')), (Kl_lab, Kl_lab')), xs), (Lev_rec, Lev_rec')), (rv_rec, rv_rec')), _) = lthy |> mk_Frees' "b" activeAs ||>> mk_Frees "b" activeAs ||>> mk_Frees "s" sTs ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "n") HOLogic.natT ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "sumx") sum_sbdT ||>> mk_Frees' "k" sbdTs ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "kl") sum_sbd_listT ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "kl") sum_sbd_listT ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "Kl") sum_sbd_list_setT ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "lab") (sum_sbd_listT --> sbdFT) ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "Kl_lab") treeT ||>> mk_Frees "x" bdFTs ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "rec") Lev_recT ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "rec") rv_recT; val (k, k') = (hd kks, hd kks') val timer = time (timer "Bounds"); (* tree coalgebra *) val isNode_binds = mk_internal_bs isNodeN; fun isNode_bind i = nth isNode_binds (i - 1); val isNode_def_bind = rpair [] o Thm.def_binding o isNode_bind; val isNodeT = Library.foldr (op -->) (map fastype_of [Kl, lab, kl], HOLogic.boolT); val Succs = @{map 3} (fn i => fn k => fn k' => HOLogic.mk_Collect (fst k', snd k', HOLogic.mk_mem (mk_InN sbdTs k i, mk_Succ Kl kl))) ks kks kks'; fun isNode_spec sets x i = let val active_sets = drop m (map (fn set => set $ x) sets); val rhs = list_exists_free [x] (Library.foldr1 HOLogic.mk_conj (HOLogic.mk_eq (lab $ kl, mk_InN bdFTs x i) :: map2 (curry HOLogic.mk_eq) active_sets Succs)); in fold_rev (Term.absfree o Term.dest_Free) [Kl, lab, kl] rhs end; val ((isNode_frees, (_, isNode_def_frees)), (lthy, lthy_old)) = lthy |> Local_Theory.open_target |> snd |> @{fold_map 3} (fn i => fn x => fn sets => Local_Theory.define ((isNode_bind i, NoSyn), (isNode_def_bind i, isNode_spec sets x i))) ks xs isNode_setss |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val isNode_defs = map (fn def => mk_unabs_def 3 (HOLogic.mk_obj_eq (Morphism.thm phi def))) isNode_def_frees; val isNodes = map (fst o Term.dest_Const o Morphism.term phi) isNode_frees; fun mk_isNode kl i = Term.list_comb (Const (nth isNodes (i - 1), isNodeT), [Kl, lab, kl]); val isTree = let val empty = HOLogic.mk_mem (HOLogic.mk_list sum_sbdT [], Kl); val tree = mk_Ball Kl (Term.absfree kl' (Library.foldr1 HOLogic.mk_conj (@{map 4} (fn Succ => fn i => fn k => fn k' => mk_Ball Succ (Term.absfree k' (mk_isNode (mk_append (kl, HOLogic.mk_list sum_sbdT [mk_InN sbdTs k i])) i))) Succs ks kks kks'))); in HOLogic.mk_conj (empty, tree) end; val carT_binds = mk_internal_bs carTN; fun carT_bind i = nth carT_binds (i - 1); val carT_def_bind = rpair [] o Thm.def_binding o carT_bind; fun carT_spec i = HOLogic.mk_Collect (fst Kl_lab', snd Kl_lab', list_exists_free [Kl, lab] (HOLogic.mk_conj (HOLogic.mk_eq (Kl_lab, HOLogic.mk_prod (Kl, lab)), HOLogic.mk_conj (isTree, mk_isNode (HOLogic.mk_list sum_sbdT []) i)))); val ((carT_frees, (_, carT_def_frees)), (lthy, lthy_old)) = lthy |> Local_Theory.open_target |> snd |> fold_map (fn i => Local_Theory.define ((carT_bind i, NoSyn), (carT_def_bind i, carT_spec i))) ks |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val carT_defs = map (fn def => HOLogic.mk_obj_eq (Morphism.thm phi def)) carT_def_frees; val carTs = map (fst o Term.dest_Const o Morphism.term phi) carT_frees; fun mk_carT i = Const (nth carTs (i - 1), HOLogic.mk_setT treeT); val strT_binds = mk_internal_bs strTN; fun strT_bind i = nth strT_binds (i - 1); val strT_def_bind = rpair [] o Thm.def_binding o strT_bind; fun strT_spec mapFT FT i = let fun mk_f i k k' = let val in_k = mk_InN sbdTs k i; in Term.absfree k' (HOLogic.mk_prod (mk_Shift Kl in_k, mk_shift lab in_k)) end; val f = Term.list_comb (mapFT, passive_ids @ @{map 3} mk_f ks kks kks'); val (fTs1, fTs2) = apsnd tl (chop (i - 1) (map (fn T => T --> FT) bdFTs)); val fs = map mk_undefined fTs1 @ (f :: map mk_undefined fTs2); in HOLogic.mk_case_prod (Term.absfree Kl' (Term.absfree lab' (mk_case_sumN fs $ (lab $ HOLogic.mk_list sum_sbdT [])))) end; val ((strT_frees, (_, strT_def_frees)), (lthy, lthy_old)) = lthy |> Local_Theory.open_target |> snd |> @{fold_map 3} (fn i => fn mapFT => fn FT => Local_Theory.define ((strT_bind i, NoSyn), (strT_def_bind i, strT_spec mapFT FT i))) ks tree_maps treeFTs |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val strT_defs = map (fn def => trans OF [HOLogic.mk_obj_eq (Morphism.thm phi def) RS fun_cong, @{thm prod.case}]) strT_def_frees; val strTs = map (fst o Term.dest_Const o Morphism.term phi) strT_frees; fun mk_strT FT i = Const (nth strTs (i - 1), treeT --> FT); val carTAs = map mk_carT ks; val strTAs = map2 mk_strT treeFTs ks; val coalgT_thm = Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (mk_coalg carTAs strTAs)) (fn {context = ctxt, prems = _} => mk_coalgT_tac ctxt m (coalg_def :: isNode_defs @ carT_defs) strT_defs set_mapss) |> Thm.close_derivation \<^here>; val timer = time (timer "Tree coalgebra"); fun mk_to_sbd s x i i' = mk_toCard (nth (nth setssAs (i - 1)) (m + i' - 1) $ (s $ x)) sbd; fun mk_from_sbd s x i i' = mk_fromCard (nth (nth setssAs (i - 1)) (m + i' - 1) $ (s $ x)) sbd; fun mk_to_sbd_thmss thm = map (map (fn set_sbd => thm OF [set_sbd, sbd_Card_order]) o drop m) set_sbdss; val to_sbd_inj_thmss = mk_to_sbd_thmss @{thm toCard_inj}; val from_to_sbd_thmss = mk_to_sbd_thmss @{thm fromCard_toCard}; val Lev_bind = mk_internal_b LevN; val Lev_def_bind = rpair [] (Thm.def_binding Lev_bind); val Lev_spec = let fun mk_Suc i s setsAs a a' = let val sets = drop m setsAs; fun mk_set i' set b = let val Cons = HOLogic.mk_eq (kl_copy, mk_Cons (mk_InN sbdTs (mk_to_sbd s a i i' $ b) i') kl) val b_set = HOLogic.mk_mem (b, set $ (s $ a)); val kl_rec = HOLogic.mk_mem (kl, mk_nthN n Lev_rec i' $ b); in HOLogic.mk_Collect (fst kl'_copy, snd kl'_copy, list_exists_free [b, kl] (HOLogic.mk_conj (Cons, HOLogic.mk_conj (b_set, kl_rec)))) end; in Term.absfree a' (Library.foldl1 mk_union (@{map 3} mk_set ks sets zs_copy)) end; val Suc = Term.absdummy HOLogic.natT (Term.absfree Lev_rec' (HOLogic.mk_tuple (@{map 5} mk_Suc ks ss setssAs zs zs'))); val rhs = mk_rec_nat Zero Suc; in fold_rev (Term.absfree o Term.dest_Free) ss rhs end; val ((Lev_free, (_, Lev_def_free)), (lthy, lthy_old)) = lthy |> Local_Theory.open_target |> snd |> Local_Theory.define ((Lev_bind, NoSyn), (Lev_def_bind, Lev_spec)) ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val Lev_def = mk_unabs_def n (HOLogic.mk_obj_eq (Morphism.thm phi Lev_def_free)); val Lev = fst (Term.dest_Const (Morphism.term phi Lev_free)); fun mk_Lev ss nat i = let val Ts = map fastype_of ss; val LevT = Library.foldr (op -->) (Ts, HOLogic.natT --> HOLogic.mk_tupleT (map (fn U => domain_type U --> sum_sbd_list_setT) Ts)); in mk_nthN n (Term.list_comb (Const (Lev, LevT), ss) $ nat) i end; val Lev_0s = flat (mk_rec_simps n @{thm rec_nat_0_imp} [Lev_def]); val Lev_Sucs = flat (mk_rec_simps n @{thm rec_nat_Suc_imp} [Lev_def]); val rv_bind = mk_internal_b rvN; val rv_def_bind = rpair [] (Thm.def_binding rv_bind); val rv_spec = let fun mk_Cons i s b b' = let fun mk_case i' = Term.absfree k' (mk_nthN n rv_rec i' $ (mk_from_sbd s b i i' $ k)); in Term.absfree b' (mk_case_sumN (map mk_case ks) $ sumx) end; val Cons = Term.absfree sumx' (Term.absdummy sum_sbd_listT (Term.absfree rv_rec' (HOLogic.mk_tuple (@{map 4} mk_Cons ks ss zs zs')))); val rhs = mk_rec_list Nil Cons; in fold_rev (Term.absfree o Term.dest_Free) ss rhs end; val ((rv_free, (_, rv_def_free)), (lthy, lthy_old)) = lthy |> Local_Theory.open_target |> snd |> Local_Theory.define ((rv_bind, NoSyn), (rv_def_bind, rv_spec)) ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val rv_def = mk_unabs_def n (HOLogic.mk_obj_eq (Morphism.thm phi rv_def_free)); val rv = fst (Term.dest_Const (Morphism.term phi rv_free)); fun mk_rv ss kl i = let val Ts = map fastype_of ss; val As = map domain_type Ts; val rvT = Library.foldr (op -->) (Ts, fastype_of kl --> HOLogic.mk_tupleT (map (fn U => U --> mk_sumTN As) As)); in mk_nthN n (Term.list_comb (Const (rv, rvT), ss) $ kl) i end; val rv_Nils = flat (mk_rec_simps n @{thm rec_list_Nil_imp} [rv_def]); val rv_Conss = flat (mk_rec_simps n @{thm rec_list_Cons_imp} [rv_def]); val beh_binds = mk_internal_bs behN; fun beh_bind i = nth beh_binds (i - 1); val beh_def_bind = rpair [] o Thm.def_binding o beh_bind; fun beh_spec i z = let fun mk_case i to_sbd_map s k k' = Term.absfree k' (mk_InN bdFTs (Term.list_comb (to_sbd_map, passive_ids @ map (mk_to_sbd s k i) ks) $ (s $ k)) i); val Lab = Term.absfree kl' (mk_case_sumN (@{map 5} mk_case ks to_sbd_maps ss zs zs') $ (mk_rv ss kl i $ z)); val rhs = HOLogic.mk_prod (mk_UNION (HOLogic.mk_UNIV HOLogic.natT) (Term.absfree nat' (mk_Lev ss nat i $ z)), Lab); in fold_rev (Term.absfree o Term.dest_Free) (ss @ [z]) rhs end; val ((beh_frees, (_, beh_def_frees)), (lthy, lthy_old)) = lthy |> Local_Theory.open_target |> snd |> @{fold_map 2} (fn i => fn z => Local_Theory.define ((beh_bind i, NoSyn), (beh_def_bind i, beh_spec i z))) ks zs |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val beh_defs = map (fn def => mk_unabs_def (n + 1) (HOLogic.mk_obj_eq (Morphism.thm phi def))) beh_def_frees; val behs = map (fst o Term.dest_Const o Morphism.term phi) beh_frees; fun mk_beh ss i = let val Ts = map fastype_of ss; val behT = Library.foldr (op -->) (Ts, nth activeAs (i - 1) --> treeT); in Term.list_comb (Const (nth behs (i - 1), behT), ss) end; val ((((((zs, zs_copy), zs_copy2), ss), (nat, nat')), (kl, kl')), _) = lthy |> mk_Frees "b" activeAs ||>> mk_Frees "b" activeAs ||>> mk_Frees "b" activeAs ||>> mk_Frees "s" sTs ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "n") HOLogic.natT ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "kl") sum_sbd_listT; val (length_Lev_thms, length_Lev'_thms) = let fun mk_conjunct i z = HOLogic.mk_imp (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z), HOLogic.mk_eq (mk_size kl, nat)); val goal = list_all_free (kl :: zs) (Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct ks zs)); val vars = Variable.add_free_names lthy goal []; val cts = map (SOME o Thm.cterm_of lthy) [Term.absfree nat' goal, nat]; val length_Lev = Goal.prove_sorry lthy vars [] (HOLogic.mk_Trueprop goal) (fn {context = ctxt, prems = _} => mk_length_Lev_tac ctxt cts Lev_0s Lev_Sucs) |> Thm.close_derivation \<^here>; val length_Lev' = mk_specN (n + 1) length_Lev; val length_Levs = map (fn i => length_Lev' RS mk_conjunctN n i RS mp) ks; fun mk_goal i z = Logic.mk_implies (mk_Trueprop_mem (kl, mk_Lev ss nat i $ z), mk_Trueprop_mem (kl, mk_Lev ss (mk_size kl) i $ z)); val goals = map2 mk_goal ks zs; val length_Levs' = map2 (fn goal => fn length_Lev => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_length_Lev'_tac ctxt length_Lev)) |> Thm.close_derivation \<^here>) goals length_Levs; in (length_Levs, length_Levs') end; val rv_last_thmss = let fun mk_conjunct i z i' z_copy = list_exists_free [z_copy] (HOLogic.mk_eq (mk_rv ss (mk_append (kl, HOLogic.mk_list sum_sbdT [mk_InN sbdTs k i'])) i $ z, mk_InN activeAs z_copy i')); val goal = list_all_free (k :: zs) (Library.foldr1 HOLogic.mk_conj (map2 (fn i => fn z => Library.foldr1 HOLogic.mk_conj (map2 (mk_conjunct i z) ks zs_copy)) ks zs)); val vars = Variable.add_free_names lthy goal []; val cTs = [SOME (Thm.ctyp_of lthy sum_sbdT)]; val cts = map (SOME o Thm.cterm_of lthy) [Term.absfree kl' goal, kl]; val rv_last = Goal.prove_sorry lthy vars [] (HOLogic.mk_Trueprop goal) (fn {context = ctxt, prems = _} => mk_rv_last_tac ctxt cTs cts rv_Nils rv_Conss) |> Thm.close_derivation \<^here>; val rv_last' = mk_specN (n + 1) rv_last; in map (fn i => map (fn i' => rv_last' RS mk_conjunctN n i RS mk_conjunctN n i') ks) ks end; val set_Lev_thmsss = let fun mk_conjunct i z = let fun mk_conjunct' i' sets s z' = let fun mk_conjunct'' i'' set z'' = HOLogic.mk_imp (HOLogic.mk_mem (z'', set $ (s $ z')), HOLogic.mk_mem (mk_append (kl, HOLogic.mk_list sum_sbdT [mk_InN sbdTs (mk_to_sbd s z' i' i'' $ z'') i'']), mk_Lev ss (HOLogic.mk_Suc nat) i $ z)); in HOLogic.mk_imp (HOLogic.mk_eq (mk_rv ss kl i $ z, mk_InN activeAs z' i'), (Library.foldr1 HOLogic.mk_conj (@{map 3} mk_conjunct'' ks (drop m sets) zs_copy2))) end; in HOLogic.mk_imp (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z), Library.foldr1 HOLogic.mk_conj (@{map 4} mk_conjunct' ks setssAs ss zs_copy)) end; val goal = list_all_free (kl :: zs @ zs_copy @ zs_copy2) (Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct ks zs)); val vars = Variable.add_free_names lthy goal []; val cts = map (SOME o Thm.cterm_of lthy) [Term.absfree nat' goal, nat]; val set_Lev = Goal.prove_sorry lthy vars [] (HOLogic.mk_Trueprop goal) (fn {context = ctxt, prems = _} => mk_set_Lev_tac ctxt cts Lev_0s Lev_Sucs rv_Nils rv_Conss from_to_sbd_thmss) |> Thm.close_derivation \<^here>; val set_Lev' = mk_specN (3 * n + 1) set_Lev; in map (fn i => map (fn i' => map (fn i'' => set_Lev' RS mk_conjunctN n i RS mp RS mk_conjunctN n i' RS mp RS mk_conjunctN n i'' RS mp) ks) ks) ks end; val set_image_Lev_thmsss = let fun mk_conjunct i z = let fun mk_conjunct' i' sets = let fun mk_conjunct'' i'' set s z'' = HOLogic.mk_imp (HOLogic.mk_eq (mk_rv ss kl i $ z, mk_InN activeAs z'' i''), HOLogic.mk_mem (k, mk_image (mk_to_sbd s z'' i'' i') $ (set $ (s $ z'')))); in HOLogic.mk_imp (HOLogic.mk_mem (mk_append (kl, HOLogic.mk_list sum_sbdT [mk_InN sbdTs k i']), mk_Lev ss (HOLogic.mk_Suc nat) i $ z), (Library.foldr1 HOLogic.mk_conj (@{map 4} mk_conjunct'' ks sets ss zs_copy))) end; in HOLogic.mk_imp (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z), Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct' ks (drop m setssAs'))) end; val goal = list_all_free (kl :: k :: zs @ zs_copy) (Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct ks zs)); val vars = Variable.add_free_names lthy goal []; val cts = map (SOME o Thm.cterm_of lthy) [Term.absfree nat' goal, nat]; val set_image_Lev = Goal.prove_sorry lthy vars [] (HOLogic.mk_Trueprop goal) (fn {context = ctxt, prems = _} => mk_set_image_Lev_tac ctxt cts Lev_0s Lev_Sucs rv_Nils rv_Conss from_to_sbd_thmss to_sbd_inj_thmss) |> Thm.close_derivation \<^here>; val set_image_Lev' = mk_specN (2 * n + 2) set_image_Lev; in map (fn i => map (fn i' => map (fn i'' => set_image_Lev' RS mk_conjunctN n i RS mp RS mk_conjunctN n i'' RS mp RS mk_conjunctN n i' RS mp) ks) ks) ks end; val mor_beh_thm = let val goal = HOLogic.mk_Trueprop (mk_mor active_UNIVs ss carTAs strTAs (map (mk_beh ss) ks)); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_mor_beh_tac ctxt m mor_def mor_cong_thm beh_defs carT_defs strT_defs isNode_defs to_sbd_inj_thmss from_to_sbd_thmss Lev_0s Lev_Sucs rv_Nils rv_Conss length_Lev_thms length_Lev'_thms rv_last_thmss set_Lev_thmsss set_image_Lev_thmsss set_mapss map_comp_id_thms map_cong0s) |> Thm.close_derivation \<^here> end; val timer = time (timer "Behavioral morphism"); val lsbisAs = map (mk_lsbis carTAs strTAs) ks; fun mk_str_final i = mk_univ (HOLogic.mk_comp (Term.list_comb (nth final_maps (i - 1), passive_ids @ map mk_proj lsbisAs), nth strTAs (i - 1))); val car_finals = map2 mk_quotient carTAs lsbisAs; val str_finals = map mk_str_final ks; val coalgT_set_thmss = map (map (fn thm => coalgT_thm RS thm)) coalg_set_thmss; val equiv_LSBIS_thms = map (fn thm => coalgT_thm RS thm) equiv_lsbis_thms; val congruent_str_final_thms = let fun mk_goal R final_map strT = HOLogic.mk_Trueprop (mk_congruent R (HOLogic.mk_comp (Term.list_comb (final_map, passive_ids @ map mk_proj lsbisAs), strT))); val goals = @{map 3} mk_goal lsbisAs final_maps strTAs; in @{map 4} (fn goal => fn lsbisE => fn map_comp_id => fn map_cong0 => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_congruent_str_final_tac ctxt m lsbisE map_comp_id map_cong0 equiv_LSBIS_thms) |> Thm.close_derivation \<^here>) goals lsbisE_thms map_comp_id_thms map_cong0s end; val coalg_final_thm = Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (mk_coalg car_finals str_finals)) (fn {context = ctxt, prems = _} => mk_coalg_final_tac ctxt m coalg_def congruent_str_final_thms equiv_LSBIS_thms set_mapss coalgT_set_thmss) |> Thm.close_derivation \<^here>; val mor_T_final_thm = Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (mk_mor carTAs strTAs car_finals str_finals (map mk_proj lsbisAs))) (fn {context = ctxt, prems = _} => mk_mor_T_final_tac ctxt mor_def congruent_str_final_thms equiv_LSBIS_thms) |> Thm.close_derivation \<^here>; val mor_final_thm = mor_comp_thm OF [mor_beh_thm, mor_T_final_thm]; val in_car_final_thms = map (fn thm => thm OF [mor_final_thm, UNIV_I]) mor_image'_thms; val timer = time (timer "Final coalgebra"); val ((T_names, (T_glob_infos, T_loc_infos)), lthy) = lthy |> @{fold_map 4} (fn b => fn mx => fn car_final => fn in_car_final => typedef (b, params, mx) car_final NONE (fn ctxt => EVERY' [rtac ctxt exI, rtac ctxt in_car_final] 1)) bs mixfixes car_finals in_car_final_thms |>> apsnd split_list o split_list; val Ts = map (fn name => Type (name, params')) T_names; fun mk_Ts passive = map (Term.typ_subst_atomic (passiveAs ~~ passive)) Ts; val Ts' = mk_Ts passiveBs; val Rep_Ts = map2 (fn info => fn T => Const (#Rep_name info, T --> treeQT)) T_glob_infos Ts; val Abs_Ts = map2 (fn info => fn T => Const (#Abs_name info, treeQT --> T)) T_glob_infos Ts; val Reps = map #Rep T_loc_infos; val Rep_injects = map #Rep_inject T_loc_infos; val Abs_inverses = map #Abs_inverse T_loc_infos; val timer = time (timer "THE TYPEDEFs & Rep/Abs thms"); val UNIVs = map HOLogic.mk_UNIV Ts; val FTs = mk_FTs (passiveAs @ Ts); val FTs_setss = mk_setss (passiveAs @ Ts); val map_FTs = map2 (fn Ds => mk_map_of_bnf Ds treeQTs (passiveAs @ Ts)) Dss bnfs; val unfold_fTs = map2 (curry op -->) activeAs Ts; val emptys = map (fn T => HOLogic.mk_set T []) passiveAs; val Zeros = map (fn empty => HOLogic.mk_tuple (map (fn U => absdummy U empty) Ts)) emptys; val hrecTs = map fastype_of Zeros; val (((zs, ss), (Jzs, Jzs')), _) = lthy |> mk_Frees "b" activeAs ||>> mk_Frees "s" sTs ||>> mk_Frees' "z" Ts; fun dtor_bind i = nth external_bs (i - 1) |> Binding.prefix_name (dtorN ^ "_"); val dtor_def_bind = rpair [] o Binding.concealed o Thm.def_binding o dtor_bind; fun dtor_spec rep str map_FT Jz Jz' = Term.absfree Jz' (Term.list_comb (map_FT, map HOLogic.id_const passiveAs @ Abs_Ts) $ (str $ (rep $ Jz))); val ((dtor_frees, (_, dtor_def_frees)), (lthy, lthy_old)) = lthy |> Local_Theory.open_target |> snd |> @{fold_map 6} (fn i => fn rep => fn str => fn mapx => fn Jz => fn Jz' => Local_Theory.define ((dtor_bind i, NoSyn), (dtor_def_bind i, dtor_spec rep str mapx Jz Jz'))) ks Rep_Ts str_finals map_FTs Jzs Jzs' |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; fun mk_dtors passive = map (Term.subst_atomic_types (map (Morphism.typ phi) params' ~~ (mk_params passive)) o Morphism.term phi) dtor_frees; val dtors = mk_dtors passiveAs; val dtor's = mk_dtors passiveBs; val dtor_defs = map (fn def => HOLogic.mk_obj_eq (Morphism.thm phi def) RS fun_cong) dtor_def_frees; val coalg_final_set_thmss = map (map (fn thm => coalg_final_thm RS thm)) coalg_set_thmss; val (mor_Rep_thm, mor_Abs_thm) = let val mor_Rep = Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (mk_mor UNIVs dtors car_finals str_finals Rep_Ts)) (fn {context = ctxt, prems = _} => mk_mor_Rep_tac ctxt (mor_def :: dtor_defs) Reps Abs_inverses coalg_final_set_thmss map_comp_id_thms map_cong0L_thms) |> Thm.close_derivation \<^here>; val mor_Abs = Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (mk_mor car_finals str_finals UNIVs dtors Abs_Ts)) (fn {context = ctxt, prems = _} => mk_mor_Abs_tac ctxt (mor_def :: dtor_defs) Abs_inverses) |> Thm.close_derivation \<^here>; in (mor_Rep, mor_Abs) end; val timer = time (timer "dtor definitions & thms"); fun unfold_bind i = nth external_bs (i - 1) |> Binding.prefix_name (dtor_unfoldN ^ "_"); val unfold_def_bind = rpair [] o Binding.concealed o Thm.def_binding o unfold_bind; fun unfold_spec abs f z = fold_rev (Term.absfree o Term.dest_Free) (ss @ [z]) (abs $ (f $ z)); val ((unfold_frees, (_, unfold_def_frees)), (lthy, lthy_old)) = lthy |> Local_Theory.open_target |> snd |> @{fold_map 4} (fn i => fn abs => fn f => fn z => Local_Theory.define ((unfold_bind i, NoSyn), (unfold_def_bind i, unfold_spec abs f z))) ks Abs_Ts (map (fn i => HOLogic.mk_comp (mk_proj (nth lsbisAs (i - 1)), mk_beh ss i)) ks) zs |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val unfolds = map (Morphism.term phi) unfold_frees; val unfold_names = map (fst o dest_Const) unfolds; fun mk_unfolds passives actives = @{map 3} (fn name => fn T => fn active => Const (name, Library.foldr (op -->) (map2 (curry op -->) actives (mk_FTs (passives @ actives)), active --> T))) unfold_names (mk_Ts passives) actives; fun mk_unfold Ts ss i = Term.list_comb (Const (nth unfold_names (i - 1), Library.foldr (op -->) (map fastype_of ss, domain_type (fastype_of (nth ss (i - 1))) --> nth Ts (i - 1))), ss); val unfold_defs = map (fn def => mk_unabs_def (n + 1) (HOLogic.mk_obj_eq (Morphism.thm phi def))) unfold_def_frees; val (((ss, TRs), unfold_fs), _) = lthy |> mk_Frees "s" sTs ||>> mk_Frees "r" (map (mk_relT o `I) Ts) ||>> mk_Frees "f" unfold_fTs; val mor_unfold_thm = let val Abs_inverses' = map2 (curry op RS) in_car_final_thms Abs_inverses; val morEs' = map (fn thm => (thm OF [mor_final_thm, UNIV_I]) RS sym) morE_thms; val goal = HOLogic.mk_Trueprop (mk_mor active_UNIVs ss UNIVs dtors (map (mk_unfold Ts ss) ks)); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_mor_unfold_tac ctxt m mor_UNIV_thm dtor_defs unfold_defs Abs_inverses' morEs' map_comp_id_thms map_cong0s) |> Thm.close_derivation \<^here> end; val dtor_unfold_thms = map (fn thm => (thm OF [mor_unfold_thm, UNIV_I]) RS sym) morE_thms; val (raw_coind_thms, raw_coind_thm) = let val prem = HOLogic.mk_Trueprop (mk_sbis UNIVs dtors TRs); val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (fn R => fn T => mk_leq R (Id_const T)) TRs Ts)); val vars = fold (Variable.add_free_names lthy) [prem, concl] []; in `split_conj_thm (Goal.prove_sorry lthy vars [] (Logic.mk_implies (prem, concl)) (fn {context = ctxt, prems = _} => mk_raw_coind_tac ctxt bis_def bis_cong_thm bis_O_thm bis_converse_thm bis_Gr_thm tcoalg_thm coalgT_thm mor_T_final_thm sbis_lsbis_thm lsbis_incl_thms incl_lsbis_thms equiv_LSBIS_thms mor_Rep_thm Rep_injects) |> Thm.close_derivation \<^here>) end; val (unfold_unique_mor_thms, unfold_unique_mor_thm) = let val prem = HOLogic.mk_Trueprop (mk_mor active_UNIVs ss UNIVs dtors unfold_fs); fun mk_fun_eq f i = HOLogic.mk_eq (f, mk_unfold Ts ss i); val unique = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_fun_eq unfold_fs ks)); val vars = fold (Variable.add_free_names lthy) [prem, unique] []; val bis_thm = tcoalg_thm RSN (2, tcoalg_thm RS bis_image2_thm); val mor_thm = mor_comp_thm OF [mor_final_thm, mor_Abs_thm]; val unique_mor = Goal.prove_sorry lthy vars [] (Logic.mk_implies (prem, unique)) (fn {context = ctxt, prems = _} => mk_unfold_unique_mor_tac ctxt raw_coind_thms bis_thm mor_thm unfold_defs) |> Thm.close_derivation \<^here>; in `split_conj_thm unique_mor end; val (dtor_unfold_unique_thms, dtor_unfold_unique_thm) = `split_conj_thm (split_conj_prems n (mor_UNIV_thm RS iffD2 RS unfold_unique_mor_thm)); val unfold_dtor_thms = map (fn thm => mor_id_thm RS thm RS sym) unfold_unique_mor_thms; val unfold_o_dtor_thms = let val mor = mor_comp_thm OF [mor_str_thm, mor_unfold_thm]; in map2 (fn unique => fn unfold_ctor => trans OF [mor RS unique, unfold_ctor]) unfold_unique_mor_thms unfold_dtor_thms end; val timer = time (timer "unfold definitions & thms"); val map_dtors = map2 (fn Ds => fn bnf => Term.list_comb (mk_map_of_bnf Ds (passiveAs @ Ts) (passiveAs @ FTs) bnf, map HOLogic.id_const passiveAs @ dtors)) Dss bnfs; fun ctor_bind i = nth external_bs (i - 1) |> Binding.prefix_name (ctorN ^ "_"); val ctor_def_bind = rpair [] o Binding.concealed o Thm.def_binding o ctor_bind; fun ctor_spec i = mk_unfold Ts map_dtors i; val ((ctor_frees, (_, ctor_def_frees)), (lthy, lthy_old)) = lthy |> Local_Theory.open_target |> snd |> fold_map (fn i => Local_Theory.define ((ctor_bind i, NoSyn), (ctor_def_bind i, ctor_spec i))) ks |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; fun mk_ctors params = map (Term.subst_atomic_types (map (Morphism.typ phi) params' ~~ params) o Morphism.term phi) ctor_frees; val ctors = mk_ctors params'; val ctor_defs = map (fn def => HOLogic.mk_obj_eq (Morphism.thm phi def)) ctor_def_frees; val ctor_o_dtor_thms = map2 (Local_Defs.fold lthy o single) ctor_defs unfold_o_dtor_thms; val dtor_o_ctor_thms = let fun mk_goal dtor ctor FT = mk_Trueprop_eq (HOLogic.mk_comp (dtor, ctor), HOLogic.id_const FT); val goals = @{map 3} mk_goal dtors ctors FTs; in @{map 5} (fn goal => fn ctor_def => fn unfold => fn map_comp_id => fn map_cong0L => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_dtor_o_ctor_tac ctxt ctor_def unfold map_comp_id map_cong0L unfold_o_dtor_thms) |> Thm.close_derivation \<^here>) goals ctor_defs dtor_unfold_thms map_comp_id_thms map_cong0L_thms end; val dtor_ctor_thms = map (fn thm => thm RS @{thm pointfree_idE}) dtor_o_ctor_thms; val ctor_dtor_thms = map (fn thm => thm RS @{thm pointfree_idE}) ctor_o_dtor_thms; val bij_dtor_thms = map2 (fn thm1 => fn thm2 => @{thm o_bij} OF [thm1, thm2]) ctor_o_dtor_thms dtor_o_ctor_thms; val inj_dtor_thms = map (fn thm => thm RS @{thm bij_is_inj}) bij_dtor_thms; val surj_dtor_thms = map (fn thm => thm RS @{thm bij_is_surj}) bij_dtor_thms; val dtor_nchotomy_thms = map (fn thm => thm RS @{thm surjD}) surj_dtor_thms; val dtor_inject_thms = map (fn thm => thm RS @{thm inj_eq}) inj_dtor_thms; val dtor_exhaust_thms = map (fn thm => thm RS exE) dtor_nchotomy_thms; val bij_ctor_thms = map2 (fn thm1 => fn thm2 => @{thm o_bij} OF [thm1, thm2]) dtor_o_ctor_thms ctor_o_dtor_thms; val inj_ctor_thms = map (fn thm => thm RS @{thm bij_is_inj}) bij_ctor_thms; val surj_ctor_thms = map (fn thm => thm RS @{thm bij_is_surj}) bij_ctor_thms; val ctor_nchotomy_thms = map (fn thm => thm RS @{thm surjD}) surj_ctor_thms; val ctor_inject_thms = map (fn thm => thm RS @{thm inj_eq}) inj_ctor_thms; val ctor_exhaust_thms = map (fn thm => thm RS exE) ctor_nchotomy_thms; val timer = time (timer "ctor definitions & thms"); val (((((Jzs, Jzs_copy), Jzs1), Jzs2), phis), _) = lthy |> mk_Frees "z" Ts ||>> mk_Frees "z'" Ts ||>> mk_Frees "z1" Ts ||>> mk_Frees "z2" Ts ||>> mk_Frees "P" (map2 mk_pred2T Ts Ts); val (coinduct_params, dtor_coinduct_thm) = let val rels = map (Term.subst_atomic_types ((activeAs ~~ Ts) @ (activeBs ~~ Ts))) relsAsBs; fun mk_concl phi z1 z2 = HOLogic.mk_imp (phi $ z1 $ z2, HOLogic.mk_eq (z1, z2)); val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (@{map 3} mk_concl phis Jzs1 Jzs2)); fun mk_rel_prem phi dtor rel Jz Jz_copy = let val concl = Term.list_comb (rel, passive_eqs @ phis) $ (dtor $ Jz) $ (dtor $ Jz_copy); in HOLogic.mk_Trueprop (list_all_free [Jz, Jz_copy] (HOLogic.mk_imp (phi $ Jz $ Jz_copy, concl))) end; val rel_prems = @{map 5} mk_rel_prem phis dtors rels Jzs Jzs_copy; val dtor_coinduct_goal = Logic.list_implies (rel_prems, concl); val dtor_coinduct = Variable.add_free_names lthy dtor_coinduct_goal [] |> (fn vars => Goal.prove_sorry lthy vars [] dtor_coinduct_goal (fn {context = ctxt, prems = _} => mk_dtor_coinduct_tac ctxt m raw_coind_thm bis_rel_thm rel_congs)) |> Thm.close_derivation \<^here>; in (rev (Term.add_tfrees dtor_coinduct_goal []), dtor_coinduct) end; val timer = time (timer "coinduction"); fun mk_dtor_map_DEADID_thm dtor_inject map_id0 = trans OF [iffD2 OF [dtor_inject, id_apply], map_id0 RS sym]; fun mk_dtor_map_unique_DEADID_thm () = let val (funs, algs) = HOLogic.conjuncts (HOLogic.dest_Trueprop (Thm.concl_of dtor_unfold_unique_thm)) |> map_split HOLogic.dest_eq ||> snd o strip_comb o hd |> @{apply 2} (map (fst o dest_Var)); fun mk_fun_insts T ix = Thm.cterm_of lthy (Var (ix, T --> T)); val theta = (funs ~~ @{map 2} mk_fun_insts Ts funs) @ (algs ~~ map (Thm.cterm_of lthy) dtors); val dtor_unfold_dtors = (dtor_unfold_unique_thm OF map (fn thm => mk_trans (thm RS @{thm arg_cong2[of _ _ _ _ "(\)", OF _ refl]}) @{thm trans[OF id_o o_id[symmetric]]}) map_id0s) |> split_conj_thm |> map mk_sym; in infer_instantiate lthy theta dtor_unfold_unique_thm |> Morphism.thm (Local_Theory.target_morphism lthy) |> unfold_thms lthy dtor_unfold_dtors |> (fn thm => thm OF replicate n sym) end; (* thm trans[OF x.dtor_unfold_unique x.dtor_unfold_unique[symmetric, OF trans[OF arg_cong2[of _ _ _ _ "(o)", OF pre_x.map_id0 refl] trans[OF id_o o_id[symmetric]]]], OF sym] *) fun mk_dtor_Jrel_DEADID_thm dtor_inject bnf = trans OF [rel_eq_of_bnf bnf RS @{thm predicate2_eqD}, dtor_inject] RS sym; val JphiTs = map2 mk_pred2T passiveAs passiveBs; val Jpsi1Ts = map2 mk_pred2T passiveAs passiveCs; val Jpsi2Ts = map2 mk_pred2T passiveCs passiveBs; val prodTsTs' = map2 (curry HOLogic.mk_prodT) Ts Ts'; val fstsTsTs' = map fst_const prodTsTs'; val sndsTsTs' = map snd_const prodTsTs'; val activephiTs = map2 mk_pred2T activeAs activeBs; val activeJphiTs = map2 mk_pred2T Ts Ts'; val rels = map2 (fn Ds => mk_rel_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs; val ((((Jzs, Jz's), Jphis), activeJphis), _) = lthy |> mk_Frees "z" Ts ||>> mk_Frees "y" Ts' ||>> mk_Frees "R" JphiTs ||>> mk_Frees "JR" activeJphiTs; fun mk_Jrel_DEADID_coinduct_thm () = mk_xtor_rel_co_induct_thm Greatest_FP rels activeJphis (map HOLogic.eq_const Ts) Jphis Jzs Jz's dtors dtor's (fn {context = ctxt, prems} => (unfold_thms_tac ctxt @{thms le_fun_def le_bool_def all_simps(1,2)[symmetric]} THEN REPEAT_DETERM (rtac ctxt allI 1) THEN rtac ctxt (dtor_coinduct_thm OF prems) 1)) lthy; (*register new codatatypes as BNFs*) val (timer, Jbnfs, (dtor_Jmap_o_thms, dtor_Jmap_thms), dtor_Jmap_unique_thm, dtor_Jset_thmss', dtor_Jrel_thms, Jrel_coinduct_thm, Jbnf_notes, dtor_Jset_induct_thms, lthy) = if m = 0 then (timer, replicate n DEADID_bnf, map_split (`(mk_pointfree2 lthy)) (map2 mk_dtor_map_DEADID_thm dtor_inject_thms map_ids), mk_dtor_map_unique_DEADID_thm (), replicate n [], map2 mk_dtor_Jrel_DEADID_thm dtor_inject_thms bnfs, mk_Jrel_DEADID_coinduct_thm (), [], [], lthy) else let val fTs = map2 (curry op -->) passiveAs passiveBs; val gTs = map2 (curry op -->) passiveBs passiveCs; val uTs = map2 (curry op -->) Ts Ts'; val (((((nat, nat'), (Jzs, Jzs')), (hrecs, hrecs')), (fs, fs')), _) = lthy |> yield_singleton (apfst (op ~~) oo mk_Frees' "n") HOLogic.natT ||>> mk_Frees' "z" Ts ||>> mk_Frees' "rec" hrecTs ||>> mk_Frees' "f" fTs; val map_FTFT's = map2 (fn Ds => mk_map_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs; fun mk_maps ATs BTs Ts mk_T = map2 (fn Ds => mk_map_of_bnf Ds (ATs @ Ts) (BTs @ map mk_T Ts)) Dss bnfs; fun mk_Fmap mk_const fs Ts Fmap = Term.list_comb (Fmap, fs @ map mk_const Ts); fun mk_map mk_const mk_T Ts fs Ts' dtors mk_maps = mk_unfold Ts' (map2 (fn dtor => fn Fmap => HOLogic.mk_comp (mk_Fmap mk_const fs Ts Fmap, dtor)) dtors (mk_maps Ts mk_T)); val mk_map_id = mk_map HOLogic.id_const I; val mk_mapsAB = mk_maps passiveAs passiveBs; val fs_maps = map (mk_map_id Ts fs Ts' dtors mk_mapsAB) ks; val set_bss = map (flat o map2 (fn B => fn b => if member (op =) resDs (TFree B) then [] else [b]) resBs) set_bss0; fun col_bind j = mk_internal_b (colN ^ (if m = 1 then "" else string_of_int j)); val col_def_bind = rpair [] o Thm.def_binding o col_bind; fun col_spec j Zero hrec hrec' = let fun mk_Suc dtor sets z z' = let val (set, sets) = apfst (fn xs => nth xs (j - 1)) (chop m sets); fun mk_UN set k = mk_UNION (set $ (dtor $ z)) (mk_nthN n hrec k); in Term.absfree z' (mk_union (set $ (dtor $ z), Library.foldl1 mk_union (map2 mk_UN sets ks))) end; val Suc = Term.absdummy HOLogic.natT (Term.absfree hrec' (HOLogic.mk_tuple (@{map 4} mk_Suc dtors FTs_setss Jzs Jzs'))); in mk_rec_nat Zero Suc end; val ((col_frees, (_, col_def_frees)), (lthy, lthy_old)) = lthy |> Local_Theory.open_target |> snd |> @{fold_map 4} (fn j => fn Zero => fn hrec => fn hrec' => Local_Theory.define ((col_bind j, NoSyn), (col_def_bind j, col_spec j Zero hrec hrec'))) ls Zeros hrecs hrecs' |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val col_defs = map (fn def => HOLogic.mk_obj_eq (Morphism.thm phi def)) col_def_frees; val cols = map (fst o Term.dest_Const o Morphism.term phi) col_frees; fun mk_col Ts nat i j T = let val hrecT = HOLogic.mk_tupleT (map (fn U => U --> HOLogic.mk_setT T) Ts) val colT = HOLogic.natT --> hrecT; in mk_nthN n (Term.list_comb (Const (nth cols (j - 1), colT), [nat])) i end; val col_0ss = mk_rec_simps n @{thm rec_nat_0_imp} col_defs; val col_Sucss = mk_rec_simps n @{thm rec_nat_Suc_imp} col_defs; val col_0ss' = transpose col_0ss; val col_Sucss' = transpose col_Sucss; fun mk_set Ts i j T = Abs (Name.uu, nth Ts (i - 1), mk_UNION (HOLogic.mk_UNIV HOLogic.natT) (Term.absfree nat' (mk_col Ts nat i j T $ Bound 1))); val setss = map (fn i => map2 (mk_set Ts i) ls passiveAs) ks; val (Jbnf_consts, lthy) = @{fold_map 8} (fn b => fn map_b => fn rel_b => fn pred_b => fn set_bs => fn mapx => fn sets => fn T => fn lthy => define_bnf_consts Hardly_Inline (user_policy Note_Some lthy) false (SOME deads) map_b rel_b pred_b set_bs (((((((b, T), fold_rev Term.absfree fs' mapx), sets), sbd), [Const (\<^const_name>\undefined\, T)]), NONE), NONE) lthy) bs map_bs rel_bs pred_bs set_bss fs_maps setss Ts lthy; val (_, Jconsts, Jconst_defs, mk_Jconsts) = @{split_list 4} Jbnf_consts; val (_, Jsetss, Jbds_Ds, _, _, _) = @{split_list 6} Jconsts; val (Jmap_defs, Jset_defss, Jbd_defs, _, Jrel_defs, Jpred_defs) = @{split_list 6} Jconst_defs; val (mk_Jmaps_Ds, mk_Jt_Ds, _, mk_Jrels_Ds, mk_Jpreds_Ds, _, _) = @{split_list 7} mk_Jconsts; val Jrel_unabs_defs = map (fn def => mk_unabs_def m (HOLogic.mk_obj_eq def)) Jrel_defs; val Jpred_unabs_defs = map (fn def => mk_unabs_def m (HOLogic.mk_obj_eq def)) Jpred_defs; val Jset_defs = flat Jset_defss; fun mk_Jmaps As Bs = map (fn mk => mk deads As Bs) mk_Jmaps_Ds; fun mk_Jsetss As = map2 (fn mk => fn Jsets => map (mk deads As) Jsets) mk_Jt_Ds Jsetss; val Jbds = map2 (fn mk => mk deads passiveAs) mk_Jt_Ds Jbds_Ds; fun mk_Jrels As Bs = map (fn mk => mk deads As Bs) mk_Jrels_Ds; fun mk_Jpreds As = map (fn mk => mk deads As) mk_Jpreds_Ds; val Jmaps = mk_Jmaps passiveAs passiveBs; val (Jsetss_by_range, Jsetss_by_bnf) = `transpose (mk_Jsetss passiveAs); val timer = time (timer "bnf constants for the new datatypes"); val ((((((((((((((((((((ys, ys'), (nat, nat')), (Jzs, Jzs')), Jz's), Jzs_copy), Jz's_copy), dtor_set_induct_phiss), Jphis), Jpsi1s), Jpsi2s), activeJphis), fs), fs_copy), gs), us), (Jys, Jys')), (Jys_copy, Jys'_copy)), (ys_copy, ys'_copy)), Kss), names_lthy) = lthy |> mk_Frees' "y" passiveAs ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "n") HOLogic.natT ||>> mk_Frees' "z" Ts ||>> mk_Frees "y" Ts' ||>> mk_Frees "z'" Ts ||>> mk_Frees "y'" Ts' ||>> mk_Freess "P" (map (fn A => map (mk_pred2T A) Ts) passiveAs) ||>> mk_Frees "R" JphiTs ||>> mk_Frees "R" Jpsi1Ts ||>> mk_Frees "Q" Jpsi2Ts ||>> mk_Frees "JR" activeJphiTs ||>> mk_Frees "f" fTs ||>> mk_Frees "f" fTs ||>> mk_Frees "g" gTs ||>> mk_Frees "u" uTs ||>> mk_Frees' "b" Ts' ||>> mk_Frees' "b" Ts' ||>> mk_Frees' "y" passiveAs ||>> mk_Freess "K" (map (fn AT => map (fn T => T --> AT) Ts) ATs); val fs_Jmaps = map (fn m => Term.list_comb (m, fs)) Jmaps; val fs_copy_Jmaps = map (fn m => Term.list_comb (m, fs_copy)) Jmaps; val gs_Jmaps = map (fn m => Term.list_comb (m, gs)) (mk_Jmaps passiveBs passiveCs); val fgs_Jmaps = map (fn m => Term.list_comb (m, map2 (curry HOLogic.mk_comp) gs fs)) (mk_Jmaps passiveAs passiveCs); val (dtor_Jmap_thms, Jmap_thms) = let fun mk_goal fs_Jmap map dtor dtor' = mk_Trueprop_eq (HOLogic.mk_comp (dtor', fs_Jmap), HOLogic.mk_comp (Term.list_comb (map, fs @ fs_Jmaps), dtor)); val goals = @{map 4} mk_goal fs_Jmaps map_FTFT's dtors dtor's; val maps = @{map 5} (fn goal => fn unfold => fn map_comp => fn map_cong0 => fn map_arg_cong => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jmap_defs THEN mk_map_tac ctxt m n map_arg_cong unfold map_comp map_cong0)) |> Thm.close_derivation \<^here>) goals dtor_unfold_thms map_comps map_cong0s map_arg_cong_thms; in map_split (fn thm => (thm RS @{thm comp_eq_dest}, thm)) maps end; val (dtor_Jmap_unique_thms, dtor_Jmap_unique_thm) = let fun mk_prem u map dtor dtor' = mk_Trueprop_eq (HOLogic.mk_comp (dtor', u), HOLogic.mk_comp (Term.list_comb (map, fs @ us), dtor)); val prems = @{map 4} mk_prem us map_FTFT's dtors dtor's; val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_eq) us fs_Jmaps)); val vars = fold (Variable.add_free_names lthy) (goal :: prems) []; in `split_conj_thm (Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, goal)) (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jmap_defs THEN mk_dtor_map_unique_tac ctxt dtor_unfold_unique_thm sym_map_comps) |> Thm.close_derivation \<^here>) end; val Jmap_comp0_thms = let val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (@{map 3} (fn fmap => fn gmap => fn fgmap => HOLogic.mk_eq (HOLogic.mk_comp (gmap, fmap), fgmap)) fs_Jmaps gs_Jmaps fgs_Jmaps)) val vars = Variable.add_free_names lthy goal []; in split_conj_thm (Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_map_comp0_tac ctxt Jmap_thms map_comp0s dtor_Jmap_unique_thm) |> Thm.close_derivation \<^here>) end; val timer = time (timer "map functions for the new codatatypes"); val Jset_minimal_thms = let fun mk_passive_prem set dtor x K = Logic.all x (HOLogic.mk_Trueprop (mk_leq (set $ (dtor $ x)) (K $ x))); fun mk_active_prem dtor x1 K1 set x2 K2 = fold_rev Logic.all [x1, x2] (Logic.mk_implies (mk_Trueprop_mem (x2, set $ (dtor $ x1)), HOLogic.mk_Trueprop (mk_leq (K2 $ x2) (K1 $ x1)))); val premss = map2 (fn j => fn Ks => @{map 4} mk_passive_prem (map (fn xs => nth xs (j - 1)) FTs_setss) dtors Jzs Ks @ flat (@{map 4} (fn sets => fn s => fn x1 => fn K1 => @{map 3} (mk_active_prem s x1 K1) (drop m sets) Jzs_copy Ks) FTs_setss dtors Jzs Ks)) ls Kss; val col_minimal_thms = let fun mk_conjunct j T i K x = mk_leq (mk_col Ts nat i j T $ x) (K $ x); fun mk_concl j T Ks = list_all_free Jzs (Library.foldr1 HOLogic.mk_conj (@{map 3} (mk_conjunct j T) ks Ks Jzs)); val concls = @{map 3} mk_concl ls passiveAs Kss; val goals = map2 (fn prems => fn concl => Logic.list_implies (prems, HOLogic.mk_Trueprop concl)) premss concls val ctss = map (fn phi => map (SOME o Thm.cterm_of lthy) [Term.absfree nat' phi, nat]) concls; in @{map 4} (fn goal => fn cts => fn col_0s => fn col_Sucs => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_col_minimal_tac ctxt m cts col_0s col_Sucs)) |> Thm.close_derivation \<^here>) goals ctss col_0ss' col_Sucss' end; fun mk_conjunct set K x = mk_leq (set $ x) (K $ x); fun mk_concl sets Ks = Library.foldr1 HOLogic.mk_conj (@{map 3} mk_conjunct sets Ks Jzs); val concls = map2 mk_concl Jsetss_by_range Kss; val goals = map2 (fn prems => fn concl => Logic.list_implies (prems, HOLogic.mk_Trueprop concl)) premss concls; in map2 (fn goal => fn col_minimal => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jset_defs THEN mk_Jset_minimal_tac ctxt n col_minimal)) |> Thm.close_derivation \<^here>) goals col_minimal_thms end; val (dtor_Jset_incl_thmss, dtor_set_Jset_incl_thmsss) = let fun mk_set_incl_Jset dtor x set Jset = HOLogic.mk_Trueprop (mk_leq (set $ (dtor $ x)) (Jset $ x)); fun mk_set_Jset_incl_Jset dtor x y set Jset1 Jset2 = Logic.mk_implies (mk_Trueprop_mem (x, set $ (dtor $ y)), HOLogic.mk_Trueprop (mk_leq (Jset1 $ x) (Jset2 $ y))); val set_incl_Jset_goalss = @{map 4} (fn dtor => fn x => fn sets => fn Jsets => map2 (mk_set_incl_Jset dtor x) (take m sets) Jsets) dtors Jzs FTs_setss Jsetss_by_bnf; (*x(k) : F(i)set(m+k) (dtor(i) y(i)) ==> J(k)set(j) x(k) <= J(i)set(j) y(i)*) val set_Jset_incl_Jset_goalsss = @{map 4} (fn dtori => fn yi => fn sets => fn Jsetsi => @{map 3} (fn xk => fn set => fn Jsetsk => map2 (mk_set_Jset_incl_Jset dtori xk yi set) Jsetsk Jsetsi) Jzs_copy (drop m sets) Jsetss_by_bnf) dtors Jzs FTs_setss Jsetss_by_bnf; in (map2 (fn goals => fn rec_Sucs => map2 (fn goal => fn rec_Suc => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jset_defs THEN mk_set_incl_Jset_tac ctxt rec_Suc)) |> Thm.close_derivation \<^here>) goals rec_Sucs) set_incl_Jset_goalss col_Sucss, map2 (fn goalss => fn rec_Sucs => map2 (fn k => fn goals => map2 (fn goal => fn rec_Suc => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jset_defs THEN mk_set_Jset_incl_Jset_tac ctxt n rec_Suc k)) |> Thm.close_derivation \<^here>) goals rec_Sucs) ks goalss) set_Jset_incl_Jset_goalsss col_Sucss) end; val set_incl_Jset_thmss' = transpose dtor_Jset_incl_thmss; val set_Jset_incl_Jset_thmsss' = transpose (map transpose dtor_set_Jset_incl_thmsss); val set_Jset_thmss = map (map (fn thm => thm RS @{thm set_mp})) dtor_Jset_incl_thmss; val set_Jset_Jset_thmsss = map (map (map (fn thm => thm RS @{thm set_mp}))) dtor_set_Jset_incl_thmsss; val set_Jset_thmss' = transpose set_Jset_thmss; val set_Jset_Jset_thmsss' = transpose (map transpose set_Jset_Jset_thmsss); val dtor_Jset_induct_thms = let val incls = maps (map (fn thm => thm RS @{thm subset_Collect_iff})) dtor_Jset_incl_thmss @ @{thms subset_Collect_iff[OF subset_refl]}; val cTs = map (SOME o Thm.ctyp_of lthy) params'; fun mk_induct_tinst phis jsets y y' = @{map 4} (fn phi => fn jset => fn Jz => fn Jz' => SOME (Thm.cterm_of lthy (Term.absfree Jz' (HOLogic.mk_Collect (fst y', snd y', HOLogic.mk_conj (HOLogic.mk_mem (y, jset $ Jz), phi $ y $ Jz)))))) phis jsets Jzs Jzs'; in @{map 6} (fn set_minimal => fn set_set_inclss => fn jsets => fn y => fn y' => fn phis => ((set_minimal |> Thm.instantiate' cTs (mk_induct_tinst phis jsets y y') |> unfold_thms lthy incls) OF (replicate n ballI @ maps (map (fn thm => thm RS @{thm subset_CollectI})) set_set_inclss)) |> singleton (Proof_Context.export names_lthy lthy) |> rule_by_tactic lthy (ALLGOALS (TRY o etac lthy asm_rl))) Jset_minimal_thms set_Jset_incl_Jset_thmsss' Jsetss_by_range ys ys' dtor_set_induct_phiss end; val (dtor_Jset_thmss', dtor_Jset_thmss) = let fun mk_simp_goal relate pas_set act_sets sets dtor z set = relate (set $ z, mk_union (pas_set $ (dtor $ z), Library.foldl1 mk_union (map2 (fn X => mk_UNION (X $ (dtor $ z))) act_sets sets))); fun mk_goals eq = map2 (fn i => fn sets => @{map 4} (fn Fsets => mk_simp_goal eq (nth Fsets (i - 1)) (drop m Fsets) sets) FTs_setss dtors Jzs sets) ls Jsetss_by_range; val le_goals = map (HOLogic.mk_Trueprop o Library.foldr1 HOLogic.mk_conj) (mk_goals (uncurry mk_leq)); val set_le_thmss = map split_conj_thm (@{map 4} (fn goal => fn Jset_minimal => fn set_Jsets => fn set_Jset_Jsetss => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_set_le_tac ctxt n Jset_minimal set_Jsets set_Jset_Jsetss)) |> Thm.close_derivation \<^here>) le_goals Jset_minimal_thms set_Jset_thmss' set_Jset_Jset_thmsss'); val ge_goalss = map (map HOLogic.mk_Trueprop) (mk_goals (uncurry mk_leq o swap)); val set_ge_thmss = @{map 3} (@{map 3} (fn goal => fn set_incl_Jset => fn set_Jset_incl_Jsets => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_set_ge_tac ctxt n set_incl_Jset set_Jset_incl_Jsets)) |> Thm.close_derivation \<^here>)) ge_goalss set_incl_Jset_thmss' set_Jset_incl_Jset_thmsss' in map2 (map2 (fn le => fn ge => equalityI OF [le, ge])) set_le_thmss set_ge_thmss |> `transpose end; val timer = time (timer "set functions for the new codatatypes"); val colss = map2 (fn j => fn T => map (fn i => mk_col Ts nat i j T) ks) ls passiveAs; val colss' = map2 (fn j => fn T => map (fn i => mk_col Ts' nat i j T) ks) ls passiveBs; val col_natural_thmss = let fun mk_col_natural f map z col col' = HOLogic.mk_eq (mk_image f $ (col $ z), col' $ (map $ z)); fun mk_goal f cols cols' = list_all_free Jzs (Library.foldr1 HOLogic.mk_conj (@{map 4} (mk_col_natural f) fs_Jmaps Jzs cols cols')); val goals = @{map 3} mk_goal fs colss colss'; val ctss = map (fn phi => map (SOME o Thm.cterm_of lthy) [Term.absfree nat' phi, nat]) goals; val thms = @{map 4} (fn goal => fn cts => fn rec_0s => fn rec_Sucs => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] (HOLogic.mk_Trueprop goal) (fn {context = ctxt, prems = _} => mk_col_natural_tac ctxt cts rec_0s rec_Sucs dtor_Jmap_thms set_mapss)) |> Thm.close_derivation \<^here>) goals ctss col_0ss' col_Sucss'; in map (split_conj_thm o mk_specN n) thms end; val col_bd_thmss = let fun mk_col_bd z col bd = mk_ordLeq (mk_card_of (col $ z)) bd; fun mk_goal bds cols = list_all_free Jzs (Library.foldr1 HOLogic.mk_conj (@{map 3} mk_col_bd Jzs cols bds)); val goals = map (mk_goal Jbds) colss; val ctss = map (fn phi => map (SOME o Thm.cterm_of lthy) [Term.absfree nat' phi, nat]) (map (mk_goal (replicate n sbd)) colss); val thms = @{map 5} (fn j => fn goal => fn cts => fn rec_0s => fn rec_Sucs => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] (HOLogic.mk_Trueprop goal) (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jbd_defs THEN mk_col_bd_tac ctxt m j cts rec_0s rec_Sucs sbd_Card_order sbd_Cinfinite set_sbdss)) |> Thm.close_derivation \<^here>) ls goals ctss col_0ss' col_Sucss'; in map (split_conj_thm o mk_specN n) thms end; val map_cong0_thms = let val cTs = map (SOME o Thm.ctyp_of lthy o Term.typ_subst_atomic (passiveAs ~~ passiveBs) o TFree) coinduct_params; fun mk_prem z set f g y y' = mk_Ball (set $ z) (Term.absfree y' (HOLogic.mk_eq (f $ y, g $ y))); fun mk_prems sets z = Library.foldr1 HOLogic.mk_conj (@{map 5} (mk_prem z) sets fs fs_copy ys ys') fun mk_map_cong0 sets z fmap gmap = HOLogic.mk_imp (mk_prems sets z, HOLogic.mk_eq (fmap $ z, gmap $ z)); fun mk_coind_body sets (x, T) z fmap gmap y y_copy = HOLogic.mk_conj (HOLogic.mk_mem (z, HOLogic.mk_Collect (x, T, mk_prems sets z)), HOLogic.mk_conj (HOLogic.mk_eq (y, fmap $ z), HOLogic.mk_eq (y_copy, gmap $ z))) fun mk_cphi sets (z' as (x, T)) z fmap gmap y' y y'_copy y_copy = HOLogic.mk_exists (x, T, mk_coind_body sets z' z fmap gmap y y_copy) |> Term.absfree y'_copy |> Term.absfree y' |> Thm.cterm_of lthy; val cphis = @{map 9} mk_cphi Jsetss_by_bnf Jzs' Jzs fs_Jmaps fs_copy_Jmaps Jys' Jys Jys'_copy Jys_copy; val coinduct = Thm.instantiate' cTs (map SOME cphis) dtor_coinduct_thm; val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (@{map 4} mk_map_cong0 Jsetss_by_bnf Jzs fs_Jmaps fs_copy_Jmaps)); val vars = Variable.add_free_names lthy goal []; val thm = Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_mcong_tac ctxt m (rtac ctxt coinduct) map_comps dtor_Jmap_thms map_cong0s set_mapss set_Jset_thmss set_Jset_Jset_thmsss in_rels) |> Thm.close_derivation \<^here>; in split_conj_thm thm end; val in_Jrels = map (fn def => trans OF [def, @{thm OO_Grp_alt}] RS @{thm predicate2_eqD}) Jrel_unabs_defs; val Jrels = mk_Jrels passiveAs passiveBs; val Jpreds = mk_Jpreds passiveAs; val Jrelphis = map (fn rel => Term.list_comb (rel, Jphis)) Jrels; val relphis = map (fn rel => Term.list_comb (rel, Jphis @ Jrelphis)) rels; val Jrelpsi1s = map (fn rel => Term.list_comb (rel, Jpsi1s)) (mk_Jrels passiveAs passiveCs); val Jrelpsi2s = map (fn rel => Term.list_comb (rel, Jpsi2s)) (mk_Jrels passiveCs passiveBs); val Jrelpsi12s = map (fn rel => Term.list_comb (rel, map2 (curry mk_rel_compp) Jpsi1s Jpsi2s)) Jrels; val dtor_Jrel_thms = let fun mk_goal Jz Jz' dtor dtor' Jrelphi relphi = mk_Trueprop_eq (Jrelphi $ Jz $ Jz', relphi $ (dtor $ Jz) $ (dtor' $ Jz')); val goals = @{map 6} mk_goal Jzs Jz's dtors dtor's Jrelphis relphis; in @{map 12} (fn i => fn goal => fn in_rel => fn map_comp0 => fn map_cong0 => fn dtor_map => fn dtor_sets => fn dtor_inject => fn dtor_ctor => fn set_map0s => fn dtor_set_incls => fn dtor_set_set_inclss => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_dtor_rel_tac ctxt in_Jrels i in_rel map_comp0 map_cong0 dtor_map dtor_sets dtor_inject dtor_ctor set_map0s dtor_set_incls dtor_set_set_inclss)) |> Thm.close_derivation \<^here>) ks goals in_rels map_comps map_cong0s dtor_Jmap_thms dtor_Jset_thmss' dtor_inject_thms dtor_ctor_thms set_mapss dtor_Jset_incl_thmss dtor_set_Jset_incl_thmsss end; val passiveABs = map2 (curry HOLogic.mk_prodT) passiveAs passiveBs; val zip_ranTs = passiveABs @ prodTsTs'; val allJphis = Jphis @ activeJphis; val zipFTs = mk_FTs zip_ranTs; val zipTs = @{map 3} (fn T => fn T' => fn FT => T --> T' --> FT) Ts Ts' zipFTs; val zip_zTs = mk_Ts passiveABs; val (((zips, (abs, abs')), (zip_zs, zip_zs')), _) = names_lthy |> mk_Frees "zip" zipTs ||>> mk_Frees' "ab" passiveABs ||>> mk_Frees' "z" zip_zTs; val Iphi_sets = map2 (fn phi => fn T => HOLogic.Collect_const T $ HOLogic.mk_case_prod phi) allJphis zip_ranTs; val in_phis = map2 (mk_in Iphi_sets) (mk_setss zip_ranTs) zipFTs; val fstABs = map fst_const passiveABs; val all_fsts = fstABs @ fstsTsTs'; val map_all_fsts = map2 (fn Ds => fn bnf => Term.list_comb (mk_map_of_bnf Ds zip_ranTs (passiveAs @ Ts) bnf, all_fsts)) Dss bnfs; val Jmap_fsts = map2 (fn map => fn T => if m = 0 then HOLogic.id_const T else Term.list_comb (map, fstABs)) (mk_Jmaps passiveABs passiveAs) Ts; val sndABs = map snd_const passiveABs; val all_snds = sndABs @ sndsTsTs'; val map_all_snds = map2 (fn Ds => fn bnf => Term.list_comb (mk_map_of_bnf Ds zip_ranTs (passiveBs @ Ts') bnf, all_snds)) Dss bnfs; val Jmap_snds = map2 (fn map => fn T => if m = 0 then HOLogic.id_const T else Term.list_comb (map, sndABs)) (mk_Jmaps passiveABs passiveBs) Ts; val zip_unfolds = map (mk_unfold zip_zTs (map HOLogic.mk_case_prod zips)) ks; val zip_setss = mk_Jsetss passiveABs |> transpose; fun Jrel_coinduct_tac {context = ctxt, prems = CIHs} = let fun mk_helper_prem phi in_phi zip x y map map' dtor dtor' = let val zipxy = zip $ x $ y; in HOLogic.mk_Trueprop (list_all_free [x, y] (HOLogic.mk_imp (phi $ x $ y, HOLogic.mk_conj (HOLogic.mk_mem (zipxy, in_phi), HOLogic.mk_conj (HOLogic.mk_eq (map $ zipxy, dtor $ x), HOLogic.mk_eq (map' $ zipxy, dtor' $ y)))))) end; val helper_prems = @{map 9} mk_helper_prem activeJphis in_phis zips Jzs Jz's map_all_fsts map_all_snds dtors dtor's; fun mk_helper_coind_phi fst phi x alt y map zip_unfold = list_exists_free [if fst then y else x] (HOLogic.mk_conj (phi $ x $ y, HOLogic.mk_eq (alt, map $ (zip_unfold $ HOLogic.mk_prod (x, y))))) val coind1_phis = @{map 6} (mk_helper_coind_phi true) activeJphis Jzs Jzs_copy Jz's Jmap_fsts zip_unfolds; val coind2_phis = @{map 6} (mk_helper_coind_phi false) activeJphis Jzs Jz's_copy Jz's Jmap_snds zip_unfolds; fun mk_cts zs z's phis = @{map 3} (fn z => fn z' => fn phi => SOME (Thm.cterm_of lthy (fold_rev (Term.absfree o Term.dest_Free) [z', z] phi))) zs z's phis @ map (SOME o Thm.cterm_of lthy) (splice z's zs); val cts1 = mk_cts Jzs Jzs_copy coind1_phis; val cts2 = mk_cts Jz's Jz's_copy coind2_phis; fun mk_helper_coind_concl z alt coind_phi = HOLogic.mk_imp (coind_phi, HOLogic.mk_eq (alt, z)); val helper_coind1_concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (@{map 3} mk_helper_coind_concl Jzs Jzs_copy coind1_phis)); val helper_coind2_concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (@{map 3} mk_helper_coind_concl Jz's Jz's_copy coind2_phis)); fun mk_helper_coind_thms fst concl cts = let val vars = fold (Variable.add_free_names lthy) (concl :: helper_prems) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (helper_prems, concl)) (fn {context = ctxt, prems = _} => mk_rel_coinduct_coind_tac ctxt fst m (infer_instantiate' ctxt cts dtor_coinduct_thm) ks map_comps map_cong0s map_arg_cong_thms set_mapss dtor_unfold_thms dtor_Jmap_thms in_rels) |> Thm.close_derivation \<^here> |> split_conj_thm end; val helper_coind1_thms = mk_helper_coind_thms true helper_coind1_concl cts1; val helper_coind2_thms = mk_helper_coind_thms false helper_coind2_concl cts2; fun mk_helper_ind_phi phi ab fst snd z active_phi x y zip_unfold = list_all_free [x, y] (HOLogic.mk_imp (HOLogic.mk_conj (active_phi $ x $ y, HOLogic.mk_eq (z, zip_unfold $ HOLogic.mk_prod (x, y))), phi $ (fst $ ab) $ (snd $ ab))); val helper_ind_phiss = @{map 4} (fn Jphi => fn ab => fn fst => fn snd => @{map 5} (mk_helper_ind_phi Jphi ab fst snd) zip_zs activeJphis Jzs Jz's zip_unfolds) Jphis abs fstABs sndABs; val ctss = map2 (fn ab' => fn phis => map2 (fn z' => fn phi => SOME (Thm.cterm_of lthy (Term.absfree ab' (Term.absfree z' phi)))) zip_zs' phis @ map (SOME o Thm.cterm_of lthy) zip_zs) abs' helper_ind_phiss; fun mk_helper_ind_concl ab' z ind_phi set = mk_Ball (set $ z) (Term.absfree ab' ind_phi); val mk_helper_ind_concls = @{map 3} (fn ab' => fn ind_phis => fn zip_sets => @{map 3} (mk_helper_ind_concl ab') zip_zs ind_phis zip_sets) abs' helper_ind_phiss zip_setss |> map (HOLogic.mk_Trueprop o Library.foldr1 HOLogic.mk_conj); val helper_ind_thmss = if m = 0 then replicate n [] else @{map 4} (fn concl => fn j => fn set_induct => fn cts => fold (Variable.add_free_names lthy) (concl :: helper_prems) [] |> (fn vars => Goal.prove_sorry lthy vars [] (Logic.list_implies (helper_prems, concl)) (fn {context = ctxt, prems = _} => mk_rel_coinduct_ind_tac ctxt m ks dtor_unfold_thms set_mapss j (infer_instantiate' ctxt cts set_induct))) |> Thm.close_derivation \<^here> |> split_conj_thm) mk_helper_ind_concls ls dtor_Jset_induct_thms ctss |> transpose; in mk_rel_coinduct_tac ctxt CIHs in_rels in_Jrels helper_ind_thmss helper_coind1_thms helper_coind2_thms end; val Jrel_coinduct_thm = mk_xtor_rel_co_induct_thm Greatest_FP rels activeJphis Jrels Jphis Jzs Jz's dtors dtor's Jrel_coinduct_tac lthy; val le_Jrel_OO_thm = let fun mk_le_Jrel_OO Jrelpsi1 Jrelpsi2 Jrelpsi12 = mk_leq (mk_rel_compp (Jrelpsi1, Jrelpsi2)) Jrelpsi12; val goals = @{map 3} mk_le_Jrel_OO Jrelpsi1s Jrelpsi2s Jrelpsi12s; val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj goals); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_le_rel_OO_tac ctxt Jrel_coinduct_thm dtor_Jrel_thms le_rel_OOs) |> Thm.close_derivation \<^here> end; val timer = time (timer "helpers for BNF properties"); fun close_wit I wit = (I, fold_rev Term.absfree (map (nth ys') I) wit); val all_unitTs = replicate live HOLogic.unitT; val unitTs = replicate n HOLogic.unitT; val unit_funs = replicate n (Term.absdummy HOLogic.unitT HOLogic.unit); fun mk_map_args I = map (fn i => if member (op =) I i then Term.absdummy HOLogic.unitT (nth ys i) else mk_undefined (HOLogic.unitT --> nth passiveAs i)) (0 upto (m - 1)); fun mk_nat_wit Ds bnf (I, wit) () = let val passiveI = filter (fn i => i < m) I; val map_args = mk_map_args passiveI; in Term.absdummy HOLogic.unitT (Term.list_comb (mk_map_of_bnf Ds all_unitTs (passiveAs @ unitTs) bnf, map_args @ unit_funs) $ wit) end; fun mk_dummy_wit Ds bnf I = let val map_args = mk_map_args I; in Term.absdummy HOLogic.unitT (Term.list_comb (mk_map_of_bnf Ds all_unitTs (passiveAs @ unitTs) bnf, map_args @ unit_funs) $ mk_undefined (mk_T_of_bnf Ds all_unitTs bnf)) end; val nat_witss = map2 (fn Ds => fn bnf => mk_wits_of_bnf (replicate (nwits_of_bnf bnf) Ds) (replicate (nwits_of_bnf bnf) (replicate live HOLogic.unitT)) bnf |> map (fn (I, wit) => (I, Lazy.lazy (mk_nat_wit Ds bnf (I, Term.list_comb (wit, map (K HOLogic.unit) I)))))) Dss bnfs; val nat_wit_thmss = map2 (curry op ~~) nat_witss (map wit_thmss_of_bnf bnfs) val Iss = map (map fst) nat_witss; fun filter_wits (I, wit) = let val J = filter (fn i => i < m) I; in (J, (length J < length I, wit)) end; val wit_treess = map_index (fn (i, Is) => map_index (finish Iss m [i+m] (i+m)) Is) Iss |> map (minimize_wits o map filter_wits o minimize_wits o flat); val coind_wit_argsss = map (map (tree_to_coind_wits nat_wit_thmss o snd o snd) o filter (fst o snd)) wit_treess; val nonredundant_coind_wit_argsss = fold (fn i => fn argsss => nth_map (i - 1) (filter_out (fn xs => exists (fn ys => let val xs' = (map (fst o fst) xs, snd (fst (hd xs))); val ys' = (map (fst o fst) ys, snd (fst (hd ys))); in eq_pair (subset (op =)) (eq_set (op =)) (xs', ys') andalso not (fst xs' = fst ys') end) (flat argsss))) argsss) ks coind_wit_argsss; fun prepare_args args = let val I = snd (fst (hd args)); val (dummys, args') = map_split (fn i => (case find_first (fn arg => fst (fst arg) = i - 1) args of SOME (_, ((_, wit), thms)) => (NONE, (Lazy.force wit, thms)) | NONE => (SOME (i - 1), (mk_dummy_wit (nth Dss (i - 1)) (nth bnfs (i - 1)) I, [])))) ks; in ((I, dummys), apsnd flat (split_list args')) end; fun mk_coind_wits ((I, dummys), (args, thms)) = ((I, dummys), (map (fn i => mk_unfold Ts args i $ HOLogic.unit) ks, thms)); val coind_witss = maps (map (mk_coind_wits o prepare_args)) nonredundant_coind_wit_argsss; val witss = map2 (fn Ds => fn bnf => mk_wits_of_bnf (replicate (nwits_of_bnf bnf) Ds) (replicate (nwits_of_bnf bnf) (passiveAs @ Ts)) bnf) Dss bnfs; val ctor_witss = map (map (uncurry close_wit o tree_to_ctor_wit ys ctors witss o snd o snd) o filter_out (fst o snd)) wit_treess; fun mk_coind_wit_thms ((I, dummys), (wits, wit_thms)) = let fun mk_goal sets y y_copy y'_copy j = let fun mk_conjunct set z dummy wit = mk_Ball (set $ z) (Term.absfree y'_copy (if dummy = NONE orelse member (op =) I (j - 1) then HOLogic.mk_imp (HOLogic.mk_eq (z, wit), if member (op =) I (j - 1) then HOLogic.mk_eq (y_copy, y) else \<^term>\False\) else \<^term>\True\)); in HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (@{map 4} mk_conjunct sets Jzs dummys wits)) end; val goals = @{map 5} mk_goal Jsetss_by_range ys ys_copy ys'_copy ls; in map2 (fn goal => fn induct => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_coind_wit_tac ctxt induct dtor_unfold_thms (flat set_mapss) wit_thms)) |> Thm.close_derivation \<^here>) goals dtor_Jset_induct_thms |> map split_conj_thm |> transpose |> map (map_filter (try (fn thm => thm RS bspec RS mp))) |> curry op ~~ (map_index Library.I (map (close_wit I) wits)) |> filter (fn (_, thms) => length thms = m) end; val coind_wit_thms = maps mk_coind_wit_thms coind_witss; val (wit_thmss, all_witss) = fold (fn ((i, wit), thms) => fn witss => nth_map i (fn (thms', wits) => (thms @ thms', wit :: wits)) witss) coind_wit_thms (map (pair []) ctor_witss) |> map (apsnd (map snd o minimize_wits)) |> split_list; val timer = time (timer "witnesses"); val map_id0_tacs = map2 (fn thm => fn thm' => fn ctxt => mk_map_id0_tac ctxt Jmap_thms thm thm') dtor_unfold_unique_thms unfold_dtor_thms; val map_comp0_tacs = map (fn thm => fn ctxt => rtac ctxt (thm RS sym) 1) Jmap_comp0_thms; val map_cong0_tacs = map (fn thm => fn ctxt => mk_map_cong0_tac ctxt m thm) map_cong0_thms; val set_map0_tacss = map (map (fn col => fn ctxt => unfold_thms_tac ctxt Jset_defs THEN mk_set_map0_tac ctxt col)) (transpose col_natural_thmss); val Jbd_card_orders = map (fn def => Local_Defs.fold lthy [def] sbd_card_order) Jbd_defs; val Jbd_Cinfinites = map (fn def => Local_Defs.fold lthy [def] sbd_Cinfinite) Jbd_defs; val bd_co_tacs = map (fn thm => fn ctxt => rtac ctxt thm 1) Jbd_card_orders; val bd_cinf_tacs = map (fn thm => fn ctxt => rtac ctxt (thm RS conjunct1) 1) Jbd_Cinfinites; val set_bd_tacss = map2 (fn Cinf => map (fn col => fn ctxt => unfold_thms_tac ctxt Jset_defs THEN mk_set_bd_tac ctxt Cinf col)) Jbd_Cinfinites (transpose col_bd_thmss); val le_rel_OO_tacs = map (fn i => fn ctxt => rtac ctxt (le_Jrel_OO_thm RS mk_conjunctN n i) 1) ks; val rel_OO_Grp_tacs = map (fn def => fn ctxt => rtac ctxt def 1) Jrel_unabs_defs; val pred_set_tacs = map (fn def => fn ctxt => rtac ctxt def 1) Jpred_unabs_defs; val tacss = @{map 10} zip_axioms map_id0_tacs map_comp0_tacs map_cong0_tacs set_map0_tacss bd_co_tacs bd_cinf_tacs set_bd_tacss le_rel_OO_tacs rel_OO_Grp_tacs pred_set_tacs; fun wit_tac thms ctxt = mk_wit_tac ctxt n dtor_ctor_thms (flat dtor_Jset_thmss) (maps wit_thms_of_bnf bnfs) thms; val (Jbnfs, lthy) = @{fold_map 7} (fn tacs => fn map_b => fn rel_b => fn pred_b => fn set_bs => fn wit_thms => fn consts => bnf_def Hardly_Inline (user_policy Note_Some) false I tacs (wit_tac wit_thms) (SOME deads) map_b rel_b pred_b set_bs consts) tacss map_bs rel_bs pred_bs set_bss wit_thmss (((((((replicate n Binding.empty ~~ Ts) ~~ Jmaps) ~~ Jsetss_by_bnf) ~~ Jbds) ~~ all_witss) ~~ map SOME Jrels) ~~ map SOME Jpreds) lthy; val timer = time (timer "registered new codatatypes as BNFs"); val ls' = if m = 1 then [0] else ls; val Jbnf_common_notes = map2 (fn i => fn thm => (mk_dtor_set_inductN i, [thm])) ls' dtor_Jset_induct_thms |> map (fn (thmN, thms) => ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])])); val Jbnf_notes = [(dtor_mapN, map single dtor_Jmap_thms), (dtor_map_uniqueN, map single dtor_Jmap_unique_thms), (dtor_relN, map single dtor_Jrel_thms), (dtor_set_inclN, dtor_Jset_incl_thmss), (dtor_set_set_inclN, map flat dtor_set_Jset_incl_thmsss)] @ map2 (fn i => fn thms => (mk_dtor_setN i, map single thms)) ls' dtor_Jset_thmss |> maps (fn (thmN, thmss) => map2 (fn b => fn thms => ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])])) bs thmss) in (timer, Jbnfs, (Jmap_thms, dtor_Jmap_thms), dtor_Jmap_unique_thm, dtor_Jset_thmss', dtor_Jrel_thms, Jrel_coinduct_thm, Jbnf_common_notes @ Jbnf_notes, dtor_Jset_induct_thms, lthy) end; val ((Jphis, activephis), _) = lthy |> mk_Frees "R" JphiTs ||>> mk_Frees "S" activephiTs; val dtor_unfold_o_Jmap_thms = mk_xtor_co_iter_o_map_thms Greatest_FP false m dtor_unfold_unique_thm dtor_Jmap_o_thms (map (mk_pointfree2 lthy) dtor_unfold_thms) sym_map_comps map_cong0s; val rels = map2 (fn Ds => mk_rel_of_bnf Ds allAs allBs') Dss bnfs; val Jrels = if m = 0 then map HOLogic.eq_const Ts else map (mk_rel_of_bnf deads passiveAs passiveBs) Jbnfs; val dtor_unfold_transfer_thms = mk_xtor_co_iter_transfer_thms Greatest_FP rels activephis activephis Jrels Jphis (mk_unfolds passiveAs activeAs) (mk_unfolds passiveBs activeBs) (fn {context = ctxt, prems = _} => mk_unfold_transfer_tac ctxt m Jrel_coinduct_thm (map map_transfer_of_bnf bnfs) dtor_unfold_thms) lthy; val timer = time (timer "relator coinduction"); fun mk_Ts As = map (typ_subst_atomic (passiveAs ~~ As)) Ts; val export = map (Morphism.term (Local_Theory.target_morphism lthy)) val ((corecs, (dtor_corec_thms, dtor_corec_unique_thm, dtor_corec_o_Jmap_thms, dtor_corec_transfer_thms)), lthy) = lthy |> derive_xtor_co_recs Greatest_FP external_bs mk_Ts (Dss, resDs) bnfs (export dtors) (export unfolds) dtor_unfold_unique_thm dtor_unfold_thms dtor_unfold_transfer_thms dtor_Jmap_thms dtor_Jrel_thms (replicate n NONE); val timer = time (timer "recursor"); val common_notes = [(dtor_coinductN, [dtor_coinduct_thm]), (dtor_rel_coinductN, [Jrel_coinduct_thm])] |> map (fn (thmN, thms) => ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])])); val notes = [(ctor_dtorN, ctor_dtor_thms), (ctor_exhaustN, ctor_exhaust_thms), (ctor_injectN, ctor_inject_thms), (dtor_ctorN, dtor_ctor_thms), (dtor_exhaustN, dtor_exhaust_thms), (dtor_injectN, dtor_inject_thms), (dtor_unfoldN, dtor_unfold_thms), (dtor_unfold_o_mapN, dtor_unfold_o_Jmap_thms), (dtor_unfold_transferN, dtor_unfold_transfer_thms), (dtor_unfold_uniqueN, dtor_unfold_unique_thms)] |> map (apsnd (map single)) |> maps (fn (thmN, thmss) => map2 (fn b => fn thms => ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])])) bs thmss); val lthy' = lthy |> internals ? snd o Local_Theory.notes (common_notes @ notes @ Jbnf_notes); val fp_res = {Ts = Ts, bnfs = Jbnfs, pre_bnfs = bnfs, absT_infos = absT_infos, ctors = ctors, dtors = dtors, xtor_un_folds = unfolds, xtor_co_recs = export corecs, xtor_co_induct = dtor_coinduct_thm, dtor_ctors = dtor_ctor_thms, ctor_dtors = ctor_dtor_thms, ctor_injects = ctor_inject_thms, dtor_injects = dtor_inject_thms, xtor_maps = dtor_Jmap_thms, xtor_map_unique = dtor_Jmap_unique_thm, xtor_setss = dtor_Jset_thmss', xtor_rels = dtor_Jrel_thms, xtor_un_fold_thms = dtor_unfold_thms, xtor_co_rec_thms = dtor_corec_thms, xtor_un_fold_unique = dtor_unfold_unique_thm, xtor_co_rec_unique = dtor_corec_unique_thm, xtor_un_fold_o_maps = dtor_unfold_o_Jmap_thms, xtor_co_rec_o_maps = dtor_corec_o_Jmap_thms, xtor_un_fold_transfers = dtor_unfold_transfer_thms, xtor_co_rec_transfers = dtor_corec_transfer_thms, xtor_rel_co_induct = Jrel_coinduct_thm, dtor_set_inducts = dtor_Jset_induct_thms}; in timer; (fp_res, lthy') end; val _ = Outer_Syntax.local_theory \<^command_keyword>\codatatype\ "define coinductive datatypes" (parse_co_datatype_cmd Greatest_FP construct_gfp); -val _ = Theory.setup (fp_antiquote_setup \<^binding>\codatatype\); - end; diff --git a/src/HOL/Tools/BNF/bnf_lfp.ML b/src/HOL/Tools/BNF/bnf_lfp.ML --- a/src/HOL/Tools/BNF/bnf_lfp.ML +++ b/src/HOL/Tools/BNF/bnf_lfp.ML @@ -1,1875 +1,1873 @@ (* Title: HOL/Tools/BNF/bnf_lfp.ML Author: Dmitriy Traytel, TU Muenchen Author: Andrei Popescu, TU Muenchen Copyright 2012 Datatype construction. *) signature BNF_LFP = sig val construct_lfp: mixfix list -> binding list -> binding list -> binding list -> binding list list -> binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list -> BNF_Comp.absT_info list -> local_theory -> BNF_FP_Util.fp_result * local_theory end; structure BNF_LFP : BNF_LFP = struct open BNF_Def open BNF_Util open BNF_Tactics open BNF_Comp open BNF_FP_Util open BNF_FP_Def_Sugar open BNF_LFP_Util open BNF_LFP_Tactics (*all BNFs have the same lives*) fun construct_lfp mixfixes map_bs rel_bs pred_bs set_bss0 bs resBs (resDs, Dss) bnfs absT_infos lthy = let val time = time lthy; val timer = time (Timer.startRealTimer ()); val live = live_of_bnf (hd bnfs); val n = length bnfs; (*active*) val ks = 1 upto n; val m = live - n; (*passive, if 0 don't generate a new BNF*) val internals = Config.get lthy bnf_internals; val b_names = map Binding.name_of bs; val b_name = mk_common_name b_names; val b = Binding.name b_name; fun mk_internal_of_b name = Binding.prefix_name (name ^ "_") #> Binding.prefix true b_name #> Binding.concealed; fun mk_internal_b name = mk_internal_of_b name b; fun mk_internal_bs name = map (mk_internal_of_b name) bs; val external_bs = map2 (Binding.prefix false) b_names bs |> not internals ? map Binding.concealed; val deads = fold (union (op =)) Dss resDs; val names_lthy = fold Variable.declare_typ deads lthy; val passives = map fst (subtract (op = o apsnd TFree) deads resBs); (* tvars *) val (((((passiveAs, activeAs), passiveBs), activeBs), passiveCs), activeCs) = names_lthy |> variant_tfrees passives ||>> mk_TFrees n ||>> variant_tfrees passives ||>> mk_TFrees n ||>> variant_tfrees passives ||>> mk_TFrees n |> fst; val allAs = passiveAs @ activeAs; val allBs' = passiveBs @ activeBs; val Ass = replicate n allAs; val allBs = passiveAs @ activeBs; val Bss = replicate n allBs; val allCs = passiveAs @ activeCs; val allCs' = passiveBs @ activeCs; val Css' = replicate n allCs'; (* types *) val dead_poss = map (fn x => if member (op =) deads (TFree x) then SOME (TFree x) else NONE) resBs; fun mk_param NONE passive = (hd passive, tl passive) | mk_param (SOME a) passive = (a, passive); val mk_params = fold_map mk_param dead_poss #> fst; fun mk_FTs Ts = map2 (fn Ds => mk_T_of_bnf Ds Ts) Dss bnfs; val (params, params') = `(map Term.dest_TFree) (mk_params passiveAs); val FTsAs = mk_FTs allAs; val FTsBs = mk_FTs allBs; val FTsCs = mk_FTs allCs; val BTs = map HOLogic.mk_setT activeAs; val B'Ts = map HOLogic.mk_setT activeBs; val B''Ts = map HOLogic.mk_setT activeCs; val sTs = map2 (curry op -->) FTsAs activeAs; val s'Ts = map2 (curry op -->) FTsBs activeBs; val s''Ts = map2 (curry op -->) FTsCs activeCs; val fTs = map2 (curry op -->) activeAs activeBs; val inv_fTs = map2 (curry op -->) activeBs activeAs; val self_fTs = map2 (curry op -->) activeAs activeAs; val gTs = map2 (curry op -->) activeBs activeCs; val all_gTs = map2 (curry op -->) allBs allCs'; (* terms *) val mapsAsAs = @{map 4} mk_map_of_bnf Dss Ass Ass bnfs; val mapsAsBs = @{map 4} mk_map_of_bnf Dss Ass Bss bnfs; val mapsBsCs' = @{map 4} mk_map_of_bnf Dss Bss Css' bnfs; val mapsAsCs' = @{map 4} mk_map_of_bnf Dss Ass Css' bnfs; fun mk_setss Ts = @{map 3} mk_sets_of_bnf (map (replicate live) Dss) (map (replicate live) (replicate n Ts)) bnfs; val setssAs = mk_setss allAs; val bd0s = @{map 3} mk_bd_of_bnf Dss Ass bnfs; val bds = @{map 3} (fn bd0 => fn Ds => fn bnf => mk_csum bd0 (mk_card_of (HOLogic.mk_UNIV (mk_T_of_bnf Ds (replicate live (fst (dest_relT (fastype_of bd0)))) bnf)))) bd0s Dss bnfs; val witss = map wits_of_bnf bnfs; val ((((((((zs, zs'), Bs), ss), fs), self_fs), all_gs), (xFs, xFs')), _) = lthy |> mk_Frees' "z" activeAs ||>> mk_Frees "B" BTs ||>> mk_Frees "s" sTs ||>> mk_Frees "f" fTs ||>> mk_Frees "f" self_fTs ||>> mk_Frees "g" all_gTs ||>> mk_Frees' "x" FTsAs; val passive_UNIVs = map HOLogic.mk_UNIV passiveAs; val active_UNIVs = map HOLogic.mk_UNIV activeAs; val passive_ids = map HOLogic.id_const passiveAs; val active_ids = map HOLogic.id_const activeAs; (* thms *) val bd0_card_orders = map bd_card_order_of_bnf bnfs; val bd0_Card_orders = map bd_Card_order_of_bnf bnfs; val bd0_Cinfinites = map bd_Cinfinite_of_bnf bnfs; val set_bd0ss = map set_bd_of_bnf bnfs; val bd_Card_order = @{thm Card_order_csum}; val bd_Card_orders = replicate n bd_Card_order; val bd_Cinfinites = map (fn thm => thm RS @{thm Cinfinite_csum1}) bd0_Cinfinites; val bd_Cnotzeros = map (fn thm => thm RS @{thm Cinfinite_Cnotzero}) bd_Cinfinites; val bd_Cinfinite = hd bd_Cinfinites; val set_bdss = map2 (fn set_bd0s => fn bd0_Card_order => map (fn thm => ctrans OF [thm, bd0_Card_order RS @{thm ordLeq_csum1}]) set_bd0s) set_bd0ss bd0_Card_orders; val in_bds = map in_bd_of_bnf bnfs; val sym_map_comps = map (fn bnf => map_comp0_of_bnf bnf RS sym) bnfs; val map_comps = map map_comp_of_bnf bnfs; val map_cong0s = map map_cong0_of_bnf bnfs; val map_id0s = map map_id0_of_bnf bnfs; val map_ids = map map_id_of_bnf bnfs; val set_mapss = map set_map_of_bnf bnfs; val rel_mono_strong0s = map rel_mono_strong0_of_bnf bnfs; val le_rel_OOs = map le_rel_OO_of_bnf bnfs; val timer = time (timer "Extracted terms & thms"); (* nonemptiness check *) fun new_wit X (wit: nonemptiness_witness) = subset (op =) (#I wit, (0 upto m - 1) @ map snd X); val all = m upto m + n - 1; fun enrich X = map_filter (fn i => (case find_first (fn (_, i') => i = i') X of NONE => (case find_index (new_wit X) (nth witss (i - m)) of ~1 => NONE | j => SOME (j, i)) | SOME ji => SOME ji)) all; val reachable = fixpoint (op =) enrich []; val _ = (case subtract (op =) (map snd reachable) all of [] => () | i :: _ => raise EMPTY_DATATYPE (Binding.name_of (nth bs (i - m)))); val wit_thms = flat (map2 (fn bnf => fn (j, _) => nth (wit_thmss_of_bnf bnf) j) bnfs reachable); val timer = time (timer "Checked nonemptiness"); (* derived thms *) (*map g1 ... gm g(m+1) ... g(m+n) (map id ... id f(m+1) ... f(m+n) x) = map g1 ... gm (g(m+1) o f(m+1)) ... (g(m+n) o f(m+n)) x*) fun mk_map_comp_id x mapAsBs mapBsCs mapAsCs map_comp0 = let val lhs = Term.list_comb (mapBsCs, all_gs) $ (Term.list_comb (mapAsBs, passive_ids @ fs) $ x); val rhs = Term.list_comb (mapAsCs, take m all_gs @ map HOLogic.mk_comp (drop m all_gs ~~ fs)) $ x; val vars = fold (Variable.add_free_names lthy) [lhs, rhs] []; in Goal.prove_sorry lthy vars [] (mk_Trueprop_eq (lhs, rhs)) (fn {context = ctxt, prems = _} => mk_map_comp_id_tac ctxt map_comp0) |> Thm.close_derivation \<^here> end; val map_comp_id_thms = @{map 5} mk_map_comp_id xFs mapsAsBs mapsBsCs' mapsAsCs' map_comps; (*forall a : set(m+1) x. f(m+1) a = a; ...; forall a : set(m+n) x. f(m+n) a = a ==> map id ... id f(m+1) ... f(m+n) x = x*) fun mk_map_cong0L x mapAsAs sets map_cong0 map_id = let fun mk_prem set f z z' = HOLogic.mk_Trueprop (mk_Ball (set $ x) (Term.absfree z' (HOLogic.mk_eq (f $ z, z)))); val prems = @{map 4} mk_prem (drop m sets) self_fs zs zs'; val goal = mk_Trueprop_eq (Term.list_comb (mapAsAs, passive_ids @ self_fs) $ x, x); val vars = fold (Variable.add_free_names lthy) (goal :: prems) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, goal)) (fn {context = ctxt, prems = _} => mk_map_cong0L_tac ctxt m map_cong0 map_id) |> Thm.close_derivation \<^here> end; val map_cong0L_thms = @{map 5} mk_map_cong0L xFs mapsAsAs setssAs map_cong0s map_ids; val in_mono'_thms = map (fn bnf => in_mono_of_bnf bnf OF (replicate m subset_refl)) bnfs; val in_cong'_thms = map (fn bnf => in_cong_of_bnf bnf OF (replicate m refl)) bnfs; val timer = time (timer "Derived simple theorems"); (* algebra *) val alg_bind = mk_internal_b algN; val alg_def_bind = (Thm.def_binding alg_bind, []); (*forall i = 1 ... n: (\x \ Fi_in UNIV .. UNIV B1 ... Bn. si x \ Bi)*) val alg_spec = let val ins = @{map 3} mk_in (replicate n (passive_UNIVs @ Bs)) setssAs FTsAs; fun mk_alg_conjunct B s X x x' = mk_Ball X (Term.absfree x' (HOLogic.mk_mem (s $ x, B))); val rhs = Library.foldr1 HOLogic.mk_conj (@{map 5} mk_alg_conjunct Bs ss ins xFs xFs') in fold_rev (Term.absfree o Term.dest_Free) (Bs @ ss) rhs end; val ((alg_free, (_, alg_def_free)), (lthy, lthy_old)) = lthy |> Local_Theory.open_target |> snd |> Local_Theory.define ((alg_bind, NoSyn), (alg_def_bind, alg_spec)) ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val alg = fst (Term.dest_Const (Morphism.term phi alg_free)); val alg_def = mk_unabs_def (2 * n) (HOLogic.mk_obj_eq (Morphism.thm phi alg_def_free)); fun mk_alg Bs ss = let val args = Bs @ ss; val Ts = map fastype_of args; val algT = Library.foldr (op -->) (Ts, HOLogic.boolT); in Term.list_comb (Const (alg, algT), args) end; val ((((((((zs, zs'), Bs), B's), ss), s's), fs), (xFs, xFs')), _) = lthy |> mk_Frees' "z" activeAs ||>> mk_Frees "B" BTs ||>> mk_Frees "B'" B'Ts ||>> mk_Frees "s" sTs ||>> mk_Frees "s'" s'Ts ||>> mk_Frees "f" fTs ||>> mk_Frees' "x" FTsAs; val alg_set_thms = let val alg_prem = HOLogic.mk_Trueprop (mk_alg Bs ss); fun mk_prem x set B = HOLogic.mk_Trueprop (mk_leq (set $ x) B); fun mk_concl s x B = mk_Trueprop_mem (s $ x, B); val premss = map2 ((fn x => fn sets => map2 (mk_prem x) (drop m sets) Bs)) xFs setssAs; val concls = @{map 3} mk_concl ss xFs Bs; val goals = map2 (fn prems => fn concl => Logic.list_implies (alg_prem :: prems, concl)) premss concls; in map (fn goal => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_alg_set_tac ctxt alg_def)) |> Thm.close_derivation \<^here>) goals end; val timer = time (timer "Algebra definition & thms"); val alg_not_empty_thms = let val alg_prem = HOLogic.mk_Trueprop (mk_alg Bs ss); val concls = map (HOLogic.mk_Trueprop o mk_not_empty) Bs; val goals = map (fn concl => Logic.mk_implies (alg_prem, concl)) concls; in map2 (fn goal => fn alg_set => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_alg_not_empty_tac ctxt alg_set alg_set_thms wit_thms)) |> Thm.close_derivation \<^here>) goals alg_set_thms end; val timer = time (timer "Proved nonemptiness"); (* morphism *) val mor_bind = mk_internal_b morN; val mor_def_bind = (Thm.def_binding mor_bind, []); (*fbetw) forall i = 1 ... n: (\x \ Bi. f x \ B'i)*) (*mor) forall i = 1 ... n: (\x \ Fi_in UNIV ... UNIV B1 ... Bn. f (s1 x) = s1' (Fi_map id ... id f1 ... fn x))*) val mor_spec = let fun mk_fbetw f B1 B2 z z' = mk_Ball B1 (Term.absfree z' (HOLogic.mk_mem (f $ z, B2))); fun mk_mor sets mapAsBs f s s' T x x' = mk_Ball (mk_in (passive_UNIVs @ Bs) sets T) (Term.absfree x' (HOLogic.mk_eq (f $ (s $ x), s' $ (Term.list_comb (mapAsBs, passive_ids @ fs) $ x)))); val rhs = HOLogic.mk_conj (Library.foldr1 HOLogic.mk_conj (@{map 5} mk_fbetw fs Bs B's zs zs'), Library.foldr1 HOLogic.mk_conj (@{map 8} mk_mor setssAs mapsAsBs fs ss s's FTsAs xFs xFs')) in fold_rev (Term.absfree o Term.dest_Free) (Bs @ ss @ B's @ s's @ fs) rhs end; val ((mor_free, (_, mor_def_free)), (lthy, lthy_old)) = lthy |> Local_Theory.open_target |> snd |> Local_Theory.define ((mor_bind, NoSyn), (mor_def_bind, mor_spec)) ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val mor = fst (Term.dest_Const (Morphism.term phi mor_free)); val mor_def = mk_unabs_def (5 * n) (HOLogic.mk_obj_eq (Morphism.thm phi mor_def_free)); fun mk_mor Bs1 ss1 Bs2 ss2 fs = let val args = Bs1 @ ss1 @ Bs2 @ ss2 @ fs; val Ts = map fastype_of (Bs1 @ ss1 @ Bs2 @ ss2 @ fs); val morT = Library.foldr (op -->) (Ts, HOLogic.boolT); in Term.list_comb (Const (mor, morT), args) end; val (((((((((((Bs, Bs_copy), B's), B''s), ss), s's), s''s), fs), fs_copy), gs), xFs), _) = lthy |> mk_Frees "B" BTs ||>> mk_Frees "B" BTs ||>> mk_Frees "B'" B'Ts ||>> mk_Frees "B''" B''Ts ||>> mk_Frees "s" sTs ||>> mk_Frees "s'" s'Ts ||>> mk_Frees "s''" s''Ts ||>> mk_Frees "f" fTs ||>> mk_Frees "f" fTs ||>> mk_Frees "g" gTs ||>> mk_Frees "x" FTsAs; val morE_thms = let val prem = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs); fun mk_elim_prem sets x T = HOLogic.mk_Trueprop (HOLogic.mk_mem (x, mk_in (passive_UNIVs @ Bs) sets T)); fun mk_elim_goal sets mapAsBs f s s' x T = Logic.list_implies ([prem, mk_elim_prem sets x T], mk_Trueprop_eq (f $ (s $ x), s' $ Term.list_comb (mapAsBs, passive_ids @ fs @ [x]))); val elim_goals = @{map 7} mk_elim_goal setssAs mapsAsBs fs ss s's xFs FTsAs; fun prove goal = Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_mor_elim_tac ctxt mor_def)) |> Thm.close_derivation \<^here>; in map prove elim_goals end; val mor_incl_thm = let val prems = map2 (HOLogic.mk_Trueprop oo mk_leq) Bs Bs_copy; val concl = HOLogic.mk_Trueprop (mk_mor Bs ss Bs_copy ss active_ids); val vars = fold (Variable.add_free_names lthy) (concl :: prems) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, concl)) (fn {context = ctxt, prems = _} => mk_mor_incl_tac ctxt mor_def map_ids) |> Thm.close_derivation \<^here> end; val mor_comp_thm = let val prems = [HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs), HOLogic.mk_Trueprop (mk_mor B's s's B''s s''s gs)]; val concl = HOLogic.mk_Trueprop (mk_mor Bs ss B''s s''s (map2 (curry HOLogic.mk_comp) gs fs)); val vars = fold (Variable.add_free_names lthy) (concl :: prems) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, concl)) (fn {context = ctxt, prems = _} => mk_mor_comp_tac ctxt mor_def set_mapss map_comp_id_thms) |> Thm.close_derivation \<^here> end; val mor_cong_thm = let val prems = map HOLogic.mk_Trueprop (map2 (curry HOLogic.mk_eq) fs_copy fs @ [mk_mor Bs ss B's s's fs]) val concl = HOLogic.mk_Trueprop (mk_mor Bs ss B's s's fs_copy); val vars = fold (Variable.add_free_names lthy) (concl :: prems) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, concl)) (fn {context = ctxt, prems = _} => (hyp_subst_tac ctxt THEN' assume_tac ctxt) 1) |> Thm.close_derivation \<^here> end; val mor_str_thm = let val maps = map2 (fn Ds => fn bnf => Term.list_comb (mk_map_of_bnf Ds (passiveAs @ FTsAs) allAs bnf, passive_ids @ ss)) Dss bnfs; val goal = HOLogic.mk_Trueprop (mk_mor (map HOLogic.mk_UNIV FTsAs) maps active_UNIVs ss ss); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_mor_str_tac ctxt ks mor_def) |> Thm.close_derivation \<^here> end; val mor_UNIV_thm = let fun mk_conjunct mapAsBs f s s' = HOLogic.mk_eq (HOLogic.mk_comp (f, s), HOLogic.mk_comp (s', Term.list_comb (mapAsBs, passive_ids @ fs))); val lhs = mk_mor active_UNIVs ss (map HOLogic.mk_UNIV activeBs) s's fs; val rhs = Library.foldr1 HOLogic.mk_conj (@{map 4} mk_conjunct mapsAsBs fs ss s's); val vars = fold (Variable.add_free_names lthy) [lhs, rhs] []; in Goal.prove_sorry lthy vars [] (mk_Trueprop_eq (lhs, rhs)) (fn {context = ctxt, prems = _} => mk_mor_UNIV_tac ctxt m morE_thms mor_def) |> Thm.close_derivation \<^here> end; val timer = time (timer "Morphism definition & thms"); (* bounds *) val sum_bd = Library.foldr1 (uncurry mk_csum) bds; val sum_bdT = fst (dest_relT (fastype_of sum_bd)); val (sum_bdT_params, sum_bdT_params') = `(map TFree) (Term.add_tfreesT sum_bdT []); val (lthy, sbd, sbd_Cinfinite, sbd_Card_order, set_sbdss, in_sbds) = if n = 1 then (lthy, sum_bd, bd_Cinfinite, bd_Card_order, set_bdss, in_bds) else let val sbdT_bind = mk_internal_b sum_bdTN; val ((sbdT_name, (sbdT_glob_info, sbdT_loc_info)), lthy) = typedef (sbdT_bind, sum_bdT_params', NoSyn) (HOLogic.mk_UNIV sum_bdT) NONE (fn ctxt => EVERY' [rtac ctxt exI, rtac ctxt UNIV_I] 1) lthy; val sbdT = Type (sbdT_name, sum_bdT_params); val Abs_sbdT = Const (#Abs_name sbdT_glob_info, sum_bdT --> sbdT); val sbd_bind = mk_internal_b sum_bdN; val sbd_def_bind = (Thm.def_binding sbd_bind, []); val sbd_spec = mk_dir_image sum_bd Abs_sbdT; val ((sbd_free, (_, sbd_def_free)), (lthy, lthy_old)) = lthy |> Local_Theory.open_target |> snd |> Local_Theory.define ((sbd_bind, NoSyn), (sbd_def_bind, sbd_spec)) ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val sbd_def = HOLogic.mk_obj_eq (Morphism.thm phi sbd_def_free); val sbd = Const (fst (Term.dest_Const (Morphism.term phi sbd_free)), mk_relT (`I sbdT)); val Abs_sbdT_inj = mk_Abs_inj_thm (#Abs_inject sbdT_loc_info); val sum_Cinfinite = mk_sum_Cinfinite bd_Cinfinites; val sum_Card_order = sum_Cinfinite RS conjunct2; val sbd_ordIso = @{thm ssubst_Pair_rhs} OF [@{thm dir_image} OF [Abs_sbdT_inj, sum_Card_order], sbd_def]; val sbd_Cinfinite = @{thm Cinfinite_cong} OF [sbd_ordIso, sum_Cinfinite]; val sbd_Card_order = sbd_Cinfinite RS conjunct2; fun mk_set_sbd i bd_Card_order bds = map (fn thm => @{thm ordLeq_ordIso_trans} OF [bd_Card_order RS mk_ordLeq_csum n i thm, sbd_ordIso]) bds; val set_sbdss = @{map 3} mk_set_sbd ks bd_Card_orders set_bdss; fun mk_in_bd_sum i Co Cnz bd = Cnz RS ((@{thm ordLeq_ordIso_trans} OF [Co RS mk_ordLeq_csum n i (Co RS @{thm ordLeq_refl}), sbd_ordIso]) RS (bd RS @{thm ordLeq_transitive[OF _ cexp_mono2_Cnotzero[OF _ Card_order_csum]]})); val in_sbds = @{map 4} mk_in_bd_sum ks bd_Card_orders bd_Cnotzeros in_bds; in (lthy, sbd, sbd_Cinfinite, sbd_Card_order, set_sbdss, in_sbds) end; val sbd_Cnotzero = sbd_Cinfinite RS @{thm Cinfinite_Cnotzero}; val suc_bd = mk_cardSuc sbd; val field_suc_bd = mk_Field suc_bd; val suc_bdT = fst (dest_relT (fastype_of suc_bd)); fun mk_Asuc_bd [] = mk_cexp ctwo suc_bd | mk_Asuc_bd As = mk_cexp (mk_csum (Library.foldr1 (uncurry mk_csum) (map mk_card_of As)) ctwo) suc_bd; val suc_bd_Card_order = sbd_Card_order RS @{thm cardSuc_Card_order}; val suc_bd_Cinfinite = sbd_Cinfinite RS @{thm Cinfinite_cardSuc}; val suc_bd_Cnotzero = suc_bd_Cinfinite RS @{thm Cinfinite_Cnotzero}; val suc_bd_worel = suc_bd_Card_order RS @{thm Card_order_wo_rel} val basis_Asuc = if m = 0 then @{thm ordLeq_refl[OF Card_order_ctwo]} else @{thm ordLeq_csum2[OF Card_order_ctwo]}; val Asuc_bd_Cinfinite = suc_bd_Cinfinite RS (basis_Asuc RS @{thm Cinfinite_cexp}); val suc_bd_Asuc_bd = @{thm ordLess_ordLeq_trans[OF ordLess_ctwo_cexp cexp_mono1]} OF [suc_bd_Card_order, basis_Asuc, suc_bd_Card_order]; val Asuc_bd = mk_Asuc_bd passive_UNIVs; val Asuc_bdT = fst (dest_relT (fastype_of Asuc_bd)); val II_BTs = replicate n (HOLogic.mk_setT Asuc_bdT); val II_sTs = map2 (fn Ds => fn bnf => mk_T_of_bnf Ds (passiveAs @ replicate n Asuc_bdT) bnf --> Asuc_bdT) Dss bnfs; val ((((((Bs, ss), idxs), Asi_name), (idx, idx')), (jdx, jdx')), _) = lthy |> mk_Frees "B" BTs ||>> mk_Frees "s" sTs ||>> mk_Frees "i" (replicate n suc_bdT) ||>> (fn ctxt => apfst the_single (mk_fresh_names ctxt 1 "Asi")) ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "i") suc_bdT ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "j") suc_bdT; val suc_bd_limit_thm = let val prem = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map (fn idx => HOLogic.mk_mem (idx, field_suc_bd)) idxs)); fun mk_conjunct idx = HOLogic.mk_conj (mk_not_eq idx jdx, HOLogic.mk_mem (HOLogic.mk_prod (idx, jdx), suc_bd)); val concl = HOLogic.mk_Trueprop (mk_Bex field_suc_bd (Term.absfree jdx' (Library.foldr1 HOLogic.mk_conj (map mk_conjunct idxs)))); val vars = fold (Variable.add_free_names lthy) [prem, concl] []; in Goal.prove_sorry lthy vars [] (Logic.list_implies ([prem], concl)) (fn {context = ctxt, prems = _} => mk_bd_limit_tac ctxt n suc_bd_Cinfinite) |> Thm.close_derivation \<^here> end; val timer = time (timer "Bounds"); (* minimal algebra *) fun mk_minG Asi i k = mk_UNION (mk_underS suc_bd $ i) (Term.absfree jdx' (mk_nthN n (Asi $ jdx) k)); fun mk_minH_component Asi i sets Ts s k = HOLogic.mk_binop \<^const_name>\sup\ (mk_minG Asi i k, mk_image s $ mk_in (passive_UNIVs @ map (mk_minG Asi i) ks) sets Ts); fun mk_min_algs ss = let val BTs = map (range_type o fastype_of) ss; val Ts = passiveAs @ BTs; val (Asi, Asi') = `Free (Asi_name, suc_bdT --> Library.foldr1 HOLogic.mk_prodT (map HOLogic.mk_setT BTs)); in mk_worec suc_bd (Term.absfree Asi' (Term.absfree idx' (HOLogic.mk_tuple (@{map 4} (mk_minH_component Asi idx) (mk_setss Ts) (mk_FTs Ts) ss ks)))) end; val (min_algs_thms, min_algs_mono_thms, card_of_min_algs_thm, least_min_algs_thm) = let val i_field = HOLogic.mk_mem (idx, field_suc_bd); val min_algs = mk_min_algs ss; val min_algss = map (fn k => mk_nthN n (min_algs $ idx) k) ks; val concl = HOLogic.mk_Trueprop (HOLogic.mk_eq (min_algs $ idx, HOLogic.mk_tuple (@{map 4} (mk_minH_component min_algs idx) setssAs FTsAs ss ks))); val goal = Logic.mk_implies (HOLogic.mk_Trueprop i_field, concl); val vars = Variable.add_free_names lthy goal []; val min_algs_thm = Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_min_algs_tac ctxt suc_bd_worel in_cong'_thms) |> Thm.close_derivation \<^here>; val min_algs_thms = map (fn k => min_algs_thm RS mk_nthI n k) ks; fun mk_mono_goal min_alg = HOLogic.mk_Trueprop (mk_relChain suc_bd (Term.absfree idx' min_alg)); val monos = map2 (fn goal => fn min_algs => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_min_algs_mono_tac ctxt min_algs)) |> Thm.close_derivation \<^here>) (map mk_mono_goal min_algss) min_algs_thms; fun mk_card_conjunct min_alg = mk_ordLeq (mk_card_of min_alg) Asuc_bd; val card_conjunction = Library.foldr1 HOLogic.mk_conj (map mk_card_conjunct min_algss); val card_cT = Thm.ctyp_of lthy suc_bdT; val card_ct = Thm.cterm_of lthy (Term.absfree idx' card_conjunction); val card_of = let val goal = HOLogic.mk_Trueprop (HOLogic.mk_imp (i_field, card_conjunction)); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_min_algs_card_of_tac ctxt card_cT card_ct m suc_bd_worel min_algs_thms in_sbds sbd_Card_order sbd_Cnotzero suc_bd_Card_order suc_bd_Cinfinite suc_bd_Cnotzero suc_bd_Asuc_bd Asuc_bd_Cinfinite) |> Thm.close_derivation \<^here> end; val least_prem = HOLogic.mk_Trueprop (mk_alg Bs ss); val least_conjunction = Library.foldr1 HOLogic.mk_conj (map2 mk_leq min_algss Bs); val least_cT = Thm.ctyp_of lthy suc_bdT; val least_ct = Thm.cterm_of lthy (Term.absfree idx' least_conjunction); val least = let val goal = Logic.mk_implies (least_prem, HOLogic.mk_Trueprop (HOLogic.mk_imp (i_field, least_conjunction))); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_min_algs_least_tac ctxt least_cT least_ct suc_bd_worel min_algs_thms alg_set_thms) |> Thm.close_derivation \<^here> end; in (min_algs_thms, monos, card_of, least) end; val timer = time (timer "min_algs definition & thms"); val min_alg_binds = mk_internal_bs min_algN; fun min_alg_bind i = nth min_alg_binds (i - 1); val min_alg_def_bind = rpair [] o Thm.def_binding o min_alg_bind; fun min_alg_spec i = let val rhs = mk_UNION (field_suc_bd) (Term.absfree idx' (mk_nthN n (mk_min_algs ss $ idx) i)); in fold_rev (Term.absfree o Term.dest_Free) ss rhs end; val ((min_alg_frees, (_, min_alg_def_frees)), (lthy, lthy_old)) = lthy |> Local_Theory.open_target |> snd |> fold_map (fn i => Local_Theory.define ((min_alg_bind i, NoSyn), (min_alg_def_bind i, min_alg_spec i))) ks |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val min_algs = map (fst o Term.dest_Const o Morphism.term phi) min_alg_frees; val min_alg_defs = map (fn def => mk_unabs_def n (HOLogic.mk_obj_eq (Morphism.thm phi def))) min_alg_def_frees; fun mk_min_alg ss i = let val T = HOLogic.mk_setT (range_type (fastype_of (nth ss (i - 1)))) val Ts = map fastype_of ss; val min_algT = Library.foldr (op -->) (Ts, T); in Term.list_comb (Const (nth min_algs (i - 1), min_algT), ss) end; val min_algs = map (mk_min_alg ss) ks; val ((Bs, ss), _) = lthy |> mk_Frees "B" BTs ||>> mk_Frees "s" sTs; val (alg_min_alg_thm, card_of_min_alg_thms, least_min_alg_thms, mor_incl_min_alg_thm) = let val alg_min_alg = let val goal = HOLogic.mk_Trueprop (mk_alg min_algs ss); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_alg_min_alg_tac ctxt m alg_def min_alg_defs suc_bd_limit_thm sbd_Cinfinite set_sbdss min_algs_thms min_algs_mono_thms) |> Thm.close_derivation \<^here> end; fun mk_card_of_thm min_alg def = let val goal = HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of min_alg) Asuc_bd); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_card_of_min_alg_tac ctxt def card_of_min_algs_thm suc_bd_Card_order suc_bd_Asuc_bd Asuc_bd_Cinfinite) |> Thm.close_derivation \<^here> end; fun mk_least_thm min_alg B def = let val prem = HOLogic.mk_Trueprop (mk_alg Bs ss); val goal = Logic.mk_implies (prem, HOLogic.mk_Trueprop (mk_leq min_alg B)); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_least_min_alg_tac ctxt def least_min_algs_thm) |> Thm.close_derivation \<^here> end; val leasts = @{map 3} mk_least_thm min_algs Bs min_alg_defs; val incl = let val prem = HOLogic.mk_Trueprop (mk_alg Bs ss); val goal = Logic.mk_implies (prem, HOLogic.mk_Trueprop (mk_mor min_algs ss Bs ss active_ids)); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => EVERY' (rtac ctxt mor_incl_thm :: map (etac ctxt) leasts) 1) |> Thm.close_derivation \<^here> end; in (alg_min_alg, map2 mk_card_of_thm min_algs min_alg_defs, leasts, incl) end; val timer = time (timer "Minimal algebra definition & thms"); val II_repT = HOLogic.mk_prodT (HOLogic.mk_tupleT II_BTs, HOLogic.mk_tupleT II_sTs); val IIT_bind = mk_internal_b IITN; val ((IIT_name, (IIT_glob_info, IIT_loc_info)), lthy) = typedef (IIT_bind, params, NoSyn) (HOLogic.mk_UNIV II_repT) NONE (fn ctxt => EVERY' [rtac ctxt exI, rtac ctxt UNIV_I] 1) lthy; val IIT = Type (IIT_name, params'); val Abs_IIT = Const (#Abs_name IIT_glob_info, II_repT --> IIT); val Rep_IIT = Const (#Rep_name IIT_glob_info, IIT --> II_repT); val Abs_IIT_inverse_thm = UNIV_I RS #Abs_inverse IIT_loc_info; val initT = IIT --> Asuc_bdT; val active_initTs = replicate n initT; val init_FTs = map2 (fn Ds => mk_T_of_bnf Ds (passiveAs @ active_initTs)) Dss bnfs; val init_fTs = map (fn T => initT --> T) activeAs; val ((((II_Bs, II_ss), (iidx, iidx')), init_xFs), _) = lthy |> mk_Frees "IIB" II_BTs ||>> mk_Frees "IIs" II_sTs ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "i") IIT ||>> mk_Frees "x" init_FTs; val II = HOLogic.mk_Collect (fst iidx', IIT, list_exists_free (II_Bs @ II_ss) (HOLogic.mk_conj (HOLogic.mk_eq (iidx, Abs_IIT $ (HOLogic.mk_prod (HOLogic.mk_tuple II_Bs, HOLogic.mk_tuple II_ss))), mk_alg II_Bs II_ss))); val select_Bs = map (mk_nthN n (HOLogic.mk_fst (Rep_IIT $ iidx))) ks; val select_ss = map (mk_nthN n (HOLogic.mk_snd (Rep_IIT $ iidx))) ks; val str_init_binds = mk_internal_bs str_initN; fun str_init_bind i = nth str_init_binds (i - 1); val str_init_def_bind = rpair [] o Thm.def_binding o str_init_bind; fun str_init_spec i = let val init_xF = nth init_xFs (i - 1) val select_s = nth select_ss (i - 1); val map = mk_map_of_bnf (nth Dss (i - 1)) (passiveAs @ active_initTs) (passiveAs @ replicate n Asuc_bdT) (nth bnfs (i - 1)); val map_args = passive_ids @ replicate n (mk_rapp iidx Asuc_bdT); val rhs = select_s $ (Term.list_comb (map, map_args) $ init_xF); in fold_rev (Term.absfree o Term.dest_Free) [init_xF, iidx] rhs end; val ((str_init_frees, (_, str_init_def_frees)), (lthy, lthy_old)) = lthy |> Local_Theory.open_target |> snd |> fold_map (fn i => Local_Theory.define ((str_init_bind i, NoSyn), (str_init_def_bind i, str_init_spec i))) ks |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val str_inits = map (Term.subst_atomic_types (map (`(Morphism.typ phi)) params') o Morphism.term phi) str_init_frees; val str_init_defs = map (fn def => mk_unabs_def 2 (HOLogic.mk_obj_eq (Morphism.thm phi def))) str_init_def_frees; val car_inits = map (mk_min_alg str_inits) ks; val (((((((((Bs, ss), Asuc_fs), (iidx, iidx')), init_xs), (init_xFs, init_xFs')), init_fs), init_fs_copy), init_phis), _) = lthy |> mk_Frees "B" BTs ||>> mk_Frees "s" sTs ||>> mk_Frees "f" (map (fn T => Asuc_bdT --> T) activeAs) ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "i") IIT ||>> mk_Frees "ix" active_initTs ||>> mk_Frees' "x" init_FTs ||>> mk_Frees "f" init_fTs ||>> mk_Frees "f" init_fTs ||>> mk_Frees "P" (replicate n (mk_pred1T initT)); val alg_init_thm = infer_instantiate' lthy (map (SOME o Thm.cterm_of lthy) str_inits) alg_min_alg_thm; val alg_select_thm = Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (mk_Ball II (Term.absfree iidx' (mk_alg select_Bs select_ss)))) (fn {context = ctxt, prems = _} => mk_alg_select_tac ctxt Abs_IIT_inverse_thm) |> Thm.close_derivation \<^here>; val mor_select_thm = let val i_prem = mk_Trueprop_mem (iidx, II); val mor_prem = HOLogic.mk_Trueprop (mk_mor select_Bs select_ss active_UNIVs ss Asuc_fs); val prems = [i_prem, mor_prem]; val concl = HOLogic.mk_Trueprop (mk_mor car_inits str_inits active_UNIVs ss (map (fn f => HOLogic.mk_comp (f, mk_rapp iidx Asuc_bdT)) Asuc_fs)); val vars = fold (Variable.add_free_names lthy) (concl :: prems) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, concl)) (fn {context = ctxt, prems = _} => mk_mor_select_tac ctxt mor_def mor_cong_thm mor_comp_thm mor_incl_min_alg_thm alg_def alg_select_thm alg_set_thms set_mapss str_init_defs) |> Thm.close_derivation \<^here> end; val init_unique_mor_thms = let val prems = map2 (HOLogic.mk_Trueprop oo curry HOLogic.mk_mem) init_xs car_inits val mor_prems = map HOLogic.mk_Trueprop [mk_mor car_inits str_inits Bs ss init_fs, mk_mor car_inits str_inits Bs ss init_fs_copy]; fun mk_fun_eq f g x = HOLogic.mk_eq (f $ x, g $ x); val unique = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (@{map 3} mk_fun_eq init_fs init_fs_copy init_xs)); val cts = map (Thm.cterm_of lthy) ss; val all_prems = prems @ mor_prems; val vars = fold (Variable.add_free_names lthy) (unique :: all_prems) []; val unique_mor = Goal.prove_sorry lthy vars [] (Logic.list_implies (all_prems, unique)) (fn {context = ctxt, prems = _} => mk_init_unique_mor_tac ctxt cts m alg_def alg_init_thm least_min_alg_thms in_mono'_thms alg_set_thms morE_thms map_cong0s) |> Thm.close_derivation \<^here>; in split_conj_thm unique_mor end; val init_setss = mk_setss (passiveAs @ active_initTs); val active_init_setss = map (drop m) init_setss; val init_ins = map2 (fn sets => mk_in (passive_UNIVs @ car_inits) sets) init_setss init_FTs; fun mk_closed phis = let fun mk_conjunct phi str_init init_sets init_in x x' = let val prem = Library.foldr1 HOLogic.mk_conj (map2 (fn set => mk_Ball (set $ x)) init_sets phis); val concl = phi $ (str_init $ x); in mk_Ball init_in (Term.absfree x' (HOLogic.mk_imp (prem, concl))) end; in Library.foldr1 HOLogic.mk_conj (@{map 6} mk_conjunct phis str_inits active_init_setss init_ins init_xFs init_xFs') end; val init_induct_thm = let val prem = HOLogic.mk_Trueprop (mk_closed init_phis); val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_Ball car_inits init_phis)); val vars = fold (Variable.add_free_names lthy) [concl, prem] []; in Goal.prove_sorry lthy vars [] (Logic.mk_implies (prem, concl)) (fn {context = ctxt, prems = _} => mk_init_induct_tac ctxt m alg_def alg_init_thm least_min_alg_thms alg_set_thms) |> Thm.close_derivation \<^here> end; val timer = time (timer "Initiality definition & thms"); val ((T_names, (T_glob_infos, T_loc_infos)), lthy) = lthy |> @{fold_map 3} (fn b => fn mx => fn car_init => typedef (b, params, mx) car_init NONE (fn ctxt => EVERY' [rtac ctxt iffD2, rtac ctxt @{thm ex_in_conv}, resolve_tac ctxt alg_not_empty_thms, rtac ctxt alg_init_thm] 1)) bs mixfixes car_inits |>> apsnd split_list o split_list; val Ts = map (fn name => Type (name, params')) T_names; fun mk_Ts passive = map (Term.typ_subst_atomic (passiveAs ~~ passive)) Ts; val Ts' = mk_Ts passiveBs; val Rep_Ts = map2 (fn info => fn T => Const (#Rep_name info, T --> initT)) T_glob_infos Ts; val Abs_Ts = map2 (fn info => fn T => Const (#Abs_name info, initT --> T)) T_glob_infos Ts; val type_defs = map #type_definition T_loc_infos; val Reps = map #Rep T_loc_infos; val Rep_inverses = map #Rep_inverse T_loc_infos; val Abs_inverses = map #Abs_inverse T_loc_infos; val timer = time (timer "THE TYPEDEFs & Rep/Abs thms"); val UNIVs = map HOLogic.mk_UNIV Ts; val FTs = mk_FTs (passiveAs @ Ts); val FTs' = mk_FTs (passiveBs @ Ts'); fun mk_set_Ts T = passiveAs @ replicate n (HOLogic.mk_setT T); val setFTss = map (mk_FTs o mk_set_Ts) passiveAs; val FTs_setss = mk_setss (passiveAs @ Ts); val FTs'_setss = mk_setss (passiveBs @ Ts'); val map_FT_inits = map2 (fn Ds => mk_map_of_bnf Ds (passiveAs @ Ts) (passiveAs @ active_initTs)) Dss bnfs; val fTs = map2 (curry op -->) Ts activeAs; val foldT = Library.foldr1 HOLogic.mk_prodT (map2 (curry op -->) Ts activeAs); val ((ss, (fold_f, fold_f')), _) = lthy |> mk_Frees "s" sTs ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "f") foldT; fun ctor_bind i = nth external_bs (i - 1) |> Binding.prefix_name (ctorN ^ "_"); val ctor_def_bind = rpair [] o Binding.concealed o Thm.def_binding o ctor_bind; fun ctor_spec abs str map_FT_init = Library.foldl1 HOLogic.mk_comp [abs, str, Term.list_comb (map_FT_init, map HOLogic.id_const passiveAs @ Rep_Ts)]; val ((ctor_frees, (_, ctor_def_frees)), (lthy, lthy_old)) = lthy |> Local_Theory.open_target |> snd |> @{fold_map 4} (fn i => fn abs => fn str => fn mapx => Local_Theory.define ((ctor_bind i, NoSyn), (ctor_def_bind i, ctor_spec abs str mapx))) ks Abs_Ts str_inits map_FT_inits |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; fun mk_ctors passive = map (Term.subst_atomic_types (map (Morphism.typ phi) params' ~~ (mk_params passive)) o Morphism.term phi) ctor_frees; val ctors = mk_ctors passiveAs; val ctor's = mk_ctors passiveBs; val ctor_defs = map (fn def => HOLogic.mk_obj_eq (Morphism.thm phi def)) ctor_def_frees; val (mor_Rep_thm, mor_Abs_thm) = let val defs = mor_def :: ctor_defs; val mor_Rep = Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (mk_mor UNIVs ctors car_inits str_inits Rep_Ts)) (fn {context = ctxt, prems = _} => mk_mor_Rep_tac ctxt m defs Reps Abs_inverses alg_min_alg_thm alg_set_thms set_mapss) |> Thm.close_derivation \<^here>; fun mk_ct initFT str abs = Term.absdummy initFT (abs $ (str $ Bound 0)) val cts = @{map 3} (Thm.cterm_of lthy ooo mk_ct) init_FTs str_inits Abs_Ts; val mor_Abs = Goal.prove_sorry lthy [] [] (HOLogic.mk_Trueprop (mk_mor car_inits str_inits UNIVs ctors Abs_Ts)) (fn {context = ctxt, prems = _} => mk_mor_Abs_tac ctxt cts defs Abs_inverses map_comp_id_thms map_cong0L_thms) |> Thm.close_derivation \<^here>; in (mor_Rep, mor_Abs) end; val timer = time (timer "ctor definitions & thms"); val fold_fun = Term.absfree fold_f' (mk_mor UNIVs ctors active_UNIVs ss (map (mk_nthN n fold_f) ks)); val foldx = HOLogic.choice_const foldT $ fold_fun; fun fold_bind i = nth external_bs (i - 1) |> Binding.prefix_name (ctor_foldN ^ "_"); val fold_def_bind = rpair [] o Binding.concealed o Thm.def_binding o fold_bind; fun fold_spec i = fold_rev (Term.absfree o Term.dest_Free) ss (mk_nthN n foldx i); val ((fold_frees, (_, fold_def_frees)), (lthy, lthy_old)) = lthy |> Local_Theory.open_target |> snd |> fold_map (fn i => Local_Theory.define ((fold_bind i, NoSyn), (fold_def_bind i, fold_spec i))) ks |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val folds = map (Morphism.term phi) fold_frees; val fold_names = map (fst o dest_Const) folds; fun mk_folds passives actives = @{map 3} (fn name => fn T => fn active => Const (name, Library.foldr (op -->) (map2 (curry op -->) (mk_FTs (passives @ actives)) actives, T --> active))) fold_names (mk_Ts passives) actives; fun mk_fold Ts ss i = Term.list_comb (Const (nth fold_names (i - 1), Library.foldr (op -->) (map fastype_of ss, nth Ts (i - 1) --> range_type (fastype_of (nth ss (i - 1))))), ss); val fold_defs = map (fn def => mk_unabs_def n (HOLogic.mk_obj_eq (Morphism.thm phi def))) fold_def_frees; (* algebra copies *) val ((((((Bs, B's), ss), s's), inv_fs), fs), _) = lthy |> mk_Frees "B" BTs ||>> mk_Frees "B'" B'Ts ||>> mk_Frees "s" sTs ||>> mk_Frees "s'" s'Ts ||>> mk_Frees "f" inv_fTs ||>> mk_Frees "f" fTs; val copy_thm = let val prems = HOLogic.mk_Trueprop (mk_alg Bs ss) :: @{map 3} (HOLogic.mk_Trueprop ooo mk_bij_betw) inv_fs B's Bs; val concl = HOLogic.mk_Trueprop (list_exists_free s's (HOLogic.mk_conj (mk_alg B's s's, mk_mor B's s's Bs ss inv_fs))); val vars = fold (Variable.add_free_names lthy) (concl :: prems) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, concl)) (fn {context = ctxt, prems = _} => mk_copy_tac ctxt m alg_def mor_def alg_set_thms set_mapss) |> Thm.close_derivation \<^here> end; val init_ex_mor_thm = let val goal = HOLogic.mk_Trueprop (list_exists_free fs (mk_mor UNIVs ctors active_UNIVs ss fs)); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_init_ex_mor_tac ctxt Abs_IIT_inverse_thm (alg_min_alg_thm RS copy_thm) card_of_min_alg_thms mor_Rep_thm mor_comp_thm mor_select_thm mor_incl_thm) |> Thm.close_derivation \<^here> end; val mor_fold_thm = let val mor_cong = mor_cong_thm OF (map (mk_nth_conv n) ks); val cT = Thm.ctyp_of lthy foldT; val ct = Thm.cterm_of lthy fold_fun val goal = HOLogic.mk_Trueprop (mk_mor UNIVs ctors active_UNIVs ss (map (mk_fold Ts ss) ks)); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_mor_fold_tac ctxt cT ct fold_defs init_ex_mor_thm mor_cong) |> Thm.close_derivation \<^here> end; val ctor_fold_thms = map (fn morE => rule_by_tactic lthy ((rtac lthy CollectI THEN' CONJ_WRAP' (K (rtac lthy @{thm subset_UNIV})) (1 upto m + n)) 1) (mor_fold_thm RS morE)) morE_thms; val (fold_unique_mor_thms, fold_unique_mor_thm) = let val prem = HOLogic.mk_Trueprop (mk_mor UNIVs ctors active_UNIVs ss fs); fun mk_fun_eq f i = HOLogic.mk_eq (f, mk_fold Ts ss i); val unique = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_fun_eq fs ks)); val vars = fold (Variable.add_free_names lthy) [prem, unique] []; val unique_mor = Goal.prove_sorry lthy vars [] (Logic.mk_implies (prem, unique)) (fn {context = ctxt, prems = _} => mk_fold_unique_mor_tac ctxt type_defs init_unique_mor_thms Reps mor_comp_thm mor_Abs_thm mor_fold_thm) |> Thm.close_derivation \<^here>; in `split_conj_thm unique_mor end; val (ctor_fold_unique_thms, ctor_fold_unique_thm) = `split_conj_thm (mk_conjIN n RS (mor_UNIV_thm RS iffD2 RS fold_unique_mor_thm)) val fold_ctor_thms = map (fn thm => (mor_incl_thm OF replicate n @{thm subset_UNIV}) RS thm RS sym) fold_unique_mor_thms; val ctor_o_fold_thms = let val mor = mor_comp_thm OF [mor_fold_thm, mor_str_thm]; in map2 (fn unique => fn fold_ctor => trans OF [mor RS unique, fold_ctor]) fold_unique_mor_thms fold_ctor_thms end; val timer = time (timer "fold definitions & thms"); val map_ctors = map2 (fn Ds => fn bnf => Term.list_comb (mk_map_of_bnf Ds (passiveAs @ FTs) (passiveAs @ Ts) bnf, map HOLogic.id_const passiveAs @ ctors)) Dss bnfs; fun dtor_bind i = nth external_bs (i - 1) |> Binding.prefix_name (dtorN ^ "_"); val dtor_def_bind = rpair [] o Binding.concealed o Thm.def_binding o dtor_bind; fun dtor_spec i = mk_fold Ts map_ctors i; val ((dtor_frees, (_, dtor_def_frees)), (lthy, lthy_old)) = lthy |> Local_Theory.open_target |> snd |> fold_map (fn i => Local_Theory.define ((dtor_bind i, NoSyn), (dtor_def_bind i, dtor_spec i))) ks |>> apsnd split_list o split_list ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; fun mk_dtors params = map (Term.subst_atomic_types (map (Morphism.typ phi) params' ~~ params) o Morphism.term phi) dtor_frees; val dtors = mk_dtors params'; val dtor_defs = map (fn def => HOLogic.mk_obj_eq (Morphism.thm phi def)) dtor_def_frees; val ctor_o_dtor_thms = map2 (Local_Defs.fold lthy o single) dtor_defs ctor_o_fold_thms; val dtor_o_ctor_thms = let fun mk_goal dtor ctor FT = mk_Trueprop_eq (HOLogic.mk_comp (dtor, ctor), HOLogic.id_const FT); val goals = @{map 3} mk_goal dtors ctors FTs; in @{map 5} (fn goal => fn dtor_def => fn foldx => fn map_comp_id => fn map_cong0L => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_dtor_o_ctor_tac ctxt dtor_def foldx map_comp_id map_cong0L ctor_o_fold_thms) |> Thm.close_derivation \<^here>) goals dtor_defs ctor_fold_thms map_comp_id_thms map_cong0L_thms end; val dtor_ctor_thms = map (fn thm => thm RS @{thm pointfree_idE}) dtor_o_ctor_thms; val ctor_dtor_thms = map (fn thm => thm RS @{thm pointfree_idE}) ctor_o_dtor_thms; val bij_dtor_thms = map2 (fn thm1 => fn thm2 => @{thm o_bij} OF [thm1, thm2]) ctor_o_dtor_thms dtor_o_ctor_thms; val inj_dtor_thms = map (fn thm => thm RS @{thm bij_is_inj}) bij_dtor_thms; val surj_dtor_thms = map (fn thm => thm RS @{thm bij_is_surj}) bij_dtor_thms; val dtor_nchotomy_thms = map (fn thm => thm RS @{thm surjD}) surj_dtor_thms; val dtor_inject_thms = map (fn thm => thm RS @{thm inj_eq}) inj_dtor_thms; val dtor_exhaust_thms = map (fn thm => thm RS exE) dtor_nchotomy_thms; val bij_ctor_thms = map2 (fn thm1 => fn thm2 => @{thm o_bij} OF [thm1, thm2]) dtor_o_ctor_thms ctor_o_dtor_thms; val inj_ctor_thms = map (fn thm => thm RS @{thm bij_is_inj}) bij_ctor_thms; val surj_ctor_thms = map (fn thm => thm RS @{thm bij_is_surj}) bij_ctor_thms; val ctor_nchotomy_thms = map (fn thm => thm RS @{thm surjD}) surj_ctor_thms; val ctor_inject_thms = map (fn thm => thm RS @{thm inj_eq}) inj_ctor_thms; val ctor_exhaust_thms = map (fn thm => thm RS exE) ctor_nchotomy_thms; val timer = time (timer "dtor definitions & thms"); val (((((((Izs, (Izs1, Izs1'))), (Izs2, Izs2')), xFs), yFs), init_phis), _) = lthy |> mk_Frees "z" Ts ||>> mk_Frees' "z1" Ts ||>> mk_Frees' "z2" Ts' ||>> mk_Frees "x" FTs ||>> mk_Frees "y" FTs' ||>> mk_Frees "P" (replicate n (mk_pred1T initT)); val phis = map2 retype_const_or_free (map mk_pred1T Ts) init_phis; val phi2s = map2 retype_const_or_free (map2 mk_pred2T Ts Ts') init_phis; val (ctor_induct_thm, induct_params) = let fun mk_prem phi ctor sets x = let fun mk_IH phi set z = let val prem = mk_Trueprop_mem (z, set $ x); val concl = HOLogic.mk_Trueprop (phi $ z); in Logic.all z (Logic.mk_implies (prem, concl)) end; val IHs = @{map 3} mk_IH phis (drop m sets) Izs; val concl = HOLogic.mk_Trueprop (phi $ (ctor $ x)); in Logic.all x (Logic.list_implies (IHs, concl)) end; val prems = @{map 4} mk_prem phis ctors FTs_setss xFs; fun mk_concl phi z = phi $ z; val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_concl phis Izs)); val goal = Logic.list_implies (prems, concl); val vars = Variable.add_free_names lthy goal []; in (Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_ctor_induct_tac ctxt m set_mapss init_induct_thm morE_thms mor_Abs_thm Rep_inverses Abs_inverses Reps) |> Thm.close_derivation \<^here>, rev (Term.add_tfrees goal [])) end; val cTs = map (SOME o Thm.ctyp_of lthy o TFree) induct_params; val weak_ctor_induct_thms = let fun insts i = (replicate (i - 1) TrueI) @ (asm_rl :: replicate (n - i) TrueI); in map (fn i => (ctor_induct_thm OF insts i) RS mk_conjunctN n i) ks end; val (ctor_induct2_thm, induct2_params) = let fun mk_prem phi ctor ctor' sets sets' x y = let fun mk_IH phi set set' z1 z2 = let val prem1 = mk_Trueprop_mem (z1, (set $ x)); val prem2 = mk_Trueprop_mem (z2, (set' $ y)); val concl = HOLogic.mk_Trueprop (phi $ z1 $ z2); in fold_rev Logic.all [z1, z2] (Logic.list_implies ([prem1, prem2], concl)) end; val IHs = @{map 5} mk_IH phi2s (drop m sets) (drop m sets') Izs1 Izs2; val concl = HOLogic.mk_Trueprop (phi $ (ctor $ x) $ (ctor' $ y)); in fold_rev Logic.all [x, y] (Logic.list_implies (IHs, concl)) end; val prems = @{map 7} mk_prem phi2s ctors ctor's FTs_setss FTs'_setss xFs yFs; fun mk_concl phi z1 z2 = phi $ z1 $ z2; val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (@{map 3} mk_concl phi2s Izs1 Izs2)); fun mk_t phi (z1, z1') (z2, z2') = Term.absfree z1' (HOLogic.mk_all (fst z2', snd z2', phi $ z1 $ z2)); val cts = @{map 3} (SOME o Thm.cterm_of lthy ooo mk_t) phi2s (Izs1 ~~ Izs1') (Izs2 ~~ Izs2'); val goal = Logic.list_implies (prems, concl); val vars = Variable.add_free_names lthy goal []; in (Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_ctor_induct2_tac ctxt cTs cts ctor_induct_thm weak_ctor_induct_thms) |> Thm.close_derivation \<^here>, rev (Term.add_tfrees goal [])) end; val timer = time (timer "induction"); fun mk_ctor_map_DEADID_thm ctor_inject map_id0 = trans OF [id_apply, iffD2 OF [ctor_inject, map_id0 RS sym]]; fun mk_ctor_map_unique_DEADID_thm () = let val (funs, algs) = HOLogic.conjuncts (HOLogic.dest_Trueprop (Thm.concl_of ctor_fold_unique_thm)) |> map_split HOLogic.dest_eq ||> snd o strip_comb o hd |> @{apply 2} (map (fst o dest_Var)); fun mk_fun_insts T ix = Thm.cterm_of lthy (Var (ix, T --> T)); val theta = (funs ~~ @{map 2} mk_fun_insts Ts funs) @ (algs ~~ map (Thm.cterm_of lthy) ctors); val ctor_fold_ctors = (ctor_fold_unique_thm OF map (fn thm => mk_trans @{thm id_o} (mk_sym (thm RS @{thm trans[OF arg_cong2[of _ _ _ _ "(\)", OF refl] o_id]}))) map_id0s) |> split_conj_thm |> map mk_sym; in infer_instantiate lthy theta ctor_fold_unique_thm |> unfold_thms lthy ctor_fold_ctors |> Morphism.thm (Local_Theory.target_morphism lthy) end; fun mk_ctor_Irel_DEADID_thm ctor_inject bnf = trans OF [ctor_inject, rel_eq_of_bnf bnf RS @{thm predicate2_eqD} RS sym]; val IphiTs = map2 mk_pred2T passiveAs passiveBs; val Ipsi1Ts = map2 mk_pred2T passiveAs passiveCs; val Ipsi2Ts = map2 mk_pred2T passiveCs passiveBs; val activephiTs = map2 mk_pred2T activeAs activeBs; val activeIphiTs = map2 mk_pred2T Ts Ts'; val rels = map2 (fn Ds => mk_rel_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs; (*register new datatypes as BNFs*) val (timer, Ibnfs, (ctor_Imap_o_thms, ctor_Imap_thms), ctor_Imap_unique_thm, ctor_Iset_thmss', ctor_Irel_thms, Ibnf_notes, lthy) = if m = 0 then (timer, replicate n DEADID_bnf, map_split (`(mk_pointfree2 lthy)) (map2 mk_ctor_map_DEADID_thm ctor_inject_thms map_ids), mk_ctor_map_unique_DEADID_thm (), replicate n [], map2 mk_ctor_Irel_DEADID_thm ctor_inject_thms bnfs, [], lthy) else let val fTs = map2 (curry op -->) passiveAs passiveBs; val uTs = map2 (curry op -->) Ts Ts'; val ((((fs, fs'), (AFss, AFss')), (ys, ys')), _) = lthy |> mk_Frees' "f" fTs ||>> mk_Freess' "z" setFTss ||>> mk_Frees' "y" passiveAs; val map_FTFT's = map2 (fn Ds => mk_map_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs; fun mk_passive_maps ATs BTs Ts = map2 (fn Ds => mk_map_of_bnf Ds (ATs @ Ts) (BTs @ Ts)) Dss bnfs; fun mk_map_fold_arg fs Ts ctor fmap = HOLogic.mk_comp (ctor, Term.list_comb (fmap, fs @ map HOLogic.id_const Ts)); fun mk_map Ts fs Ts' ctors mk_maps = mk_fold Ts (map2 (mk_map_fold_arg fs Ts') ctors (mk_maps Ts')); val pmapsABT' = mk_passive_maps passiveAs passiveBs; val fs_maps = map (mk_map Ts fs Ts' ctor's pmapsABT') ks; val ls = 1 upto m; val setsss = map (mk_setss o mk_set_Ts) passiveAs; fun mk_col l T z z' sets = let fun mk_UN set = mk_Union T $ (set $ z); in Term.absfree z' (mk_union (nth sets (l - 1) $ z, Library.foldl1 mk_union (map mk_UN (drop m sets)))) end; val colss = @{map 5} (fn l => fn T => @{map 3} (mk_col l T)) ls passiveAs AFss AFss' setsss; val setss_by_range = map (fn cols => map (mk_fold Ts cols) ks) colss; val setss_by_bnf = transpose setss_by_range; val set_bss = map (flat o map2 (fn B => fn b => if member (op =) deads (TFree B) then [] else [b]) resBs) set_bss0; val ctor_witss = let val witss = map2 (fn Ds => fn bnf => mk_wits_of_bnf (replicate (nwits_of_bnf bnf) Ds) (replicate (nwits_of_bnf bnf) (passiveAs @ Ts)) bnf) Dss bnfs; fun close_wit (I, wit) = fold_rev Term.absfree (map (nth ys') I) wit; fun wit_apply (arg_I, arg_wit) (fun_I, fun_wit) = (union (op =) arg_I fun_I, fun_wit $ arg_wit); fun gen_arg support i = if i < m then [([i], nth ys i)] else maps (mk_wit support (nth ctors (i - m)) (i - m)) (nth support (i - m)) and mk_wit support ctor i (I, wit) = let val args = map (gen_arg (nth_map i (remove (op =) (I, wit)) support)) I; in (args, [([], wit)]) |-> fold (map_product wit_apply) |> map (apsnd (fn t => ctor $ t)) |> minimize_wits end; in @{map 3} (fn ctor => fn i => map close_wit o minimize_wits o maps (mk_wit witss ctor i)) ctors (0 upto n - 1) witss end; val (lthy, sbd0, sbd0_card_order, sbd0_Cinfinite, set_sbd0ss) = if n = 1 then (lthy, hd bd0s, hd bd0_card_orders, hd bd0_Cinfinites, set_bd0ss) else let val sum_bd0 = Library.foldr1 (uncurry mk_csum) bd0s; val sum_bd0T = fst (dest_relT (fastype_of sum_bd0)); val (sum_bd0T_params, sum_bd0T_params') = `(map TFree) (Term.add_tfreesT sum_bd0T []); val sbd0T_bind = mk_internal_b (sum_bdTN ^ "0"); val ((sbd0T_name, (sbd0T_glob_info, sbd0T_loc_info)), lthy) = typedef (sbd0T_bind, sum_bd0T_params', NoSyn) (HOLogic.mk_UNIV sum_bd0T) NONE (fn ctxt => EVERY' [rtac ctxt exI, rtac ctxt UNIV_I] 1) lthy; val sbd0T = Type (sbd0T_name, sum_bd0T_params); val Abs_sbd0T = Const (#Abs_name sbd0T_glob_info, sum_bd0T --> sbd0T); val sbd0_bind = mk_internal_b (sum_bdN ^ "0"); val sbd0_def_bind = (Thm.def_binding sbd0_bind, []); val sbd0_spec = mk_dir_image sum_bd0 Abs_sbd0T; val ((sbd0_free, (_, sbd0_def_free)), (lthy, lthy_old)) = lthy |> Local_Theory.open_target |> snd |> Local_Theory.define ((sbd0_bind, NoSyn), (sbd0_def_bind, sbd0_spec)) ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val sbd0_def = HOLogic.mk_obj_eq (Morphism.thm phi sbd0_def_free); val sbd0 = Const (fst (Term.dest_Const (Morphism.term phi sbd0_free)), mk_relT (`I sbd0T)); val Abs_sbd0T_inj = mk_Abs_inj_thm (#Abs_inject sbd0T_loc_info); val Abs_sbd0T_bij = mk_Abs_bij_thm lthy Abs_sbd0T_inj (#Abs_cases sbd0T_loc_info); val sum_Cinfinite = mk_sum_Cinfinite bd0_Cinfinites; val sum_Card_order = sum_Cinfinite RS conjunct2; val sum_card_order = mk_sum_card_order bd0_card_orders; val sbd0_ordIso = @{thm ssubst_Pair_rhs} OF [@{thm dir_image} OF [Abs_sbd0T_inj, sum_Card_order], sbd0_def]; val sbd0_Cinfinite = @{thm Cinfinite_cong} OF [sbd0_ordIso, sum_Cinfinite]; val sbd0_card_order = @{thm iffD2[OF arg_cong[of _ _ card_order]]} OF [sbd0_def, @{thm card_order_dir_image} OF [Abs_sbd0T_bij, sum_card_order]]; fun mk_set_sbd0 i bd0_Card_order bd0s = map (fn thm => @{thm ordLeq_ordIso_trans} OF [bd0_Card_order RS mk_ordLeq_csum n i thm, sbd0_ordIso]) bd0s; val set_sbd0ss = @{map 3} mk_set_sbd0 ks bd0_Card_orders set_bd0ss; in (lthy, sbd0, sbd0_card_order, sbd0_Cinfinite, set_sbd0ss) end; val (Ibnf_consts, lthy) = @{fold_map 9} (fn b => fn map_b => fn rel_b => fn pred_b => fn set_bs => fn mapx => fn sets => fn wits => fn T => fn lthy => define_bnf_consts Hardly_Inline (user_policy Note_Some lthy) false (SOME deads) map_b rel_b pred_b set_bs (((((((b, T), fold_rev Term.absfree fs' mapx), sets), sbd0), wits), NONE), NONE) lthy) bs map_bs rel_bs pred_bs set_bss fs_maps setss_by_bnf ctor_witss Ts lthy; val ((((((((((((((Izs, (Izs1, Izs1')), (Izs2, Izs2')), xFs), yFs))), Iphis), Ipsi1s), Ipsi2s), fs), fs_copy), us), (ys, ys')), _) = lthy |> mk_Frees "z" Ts ||>> mk_Frees' "z1" Ts ||>> mk_Frees' "z2" Ts' ||>> mk_Frees "x" FTs ||>> mk_Frees "y" FTs' ||>> mk_Frees "R" IphiTs ||>> mk_Frees "R" Ipsi1Ts ||>> mk_Frees "Q" Ipsi2Ts ||>> mk_Frees "f" fTs ||>> mk_Frees "f" fTs ||>> mk_Frees "u" uTs ||>> mk_Frees' "y" passiveAs; val (_, Iconsts, Iconst_defs, mk_Iconsts) = @{split_list 4} Ibnf_consts; val (_, Isetss, Ibds_Ds, Iwitss_Ds, _, _) = @{split_list 6} Iconsts; val (Imap_defs, Iset_defss, Ibd_defs, Iwit_defss, Irel_defs, Ipred_defs) = @{split_list 6} Iconst_defs; val (mk_Imaps_Ds, mk_It_Ds, _, mk_Irels_Ds, mk_Ipreds_Ds, _, _) = @{split_list 7} mk_Iconsts; val Irel_unabs_defs = map (fn def => mk_unabs_def m (HOLogic.mk_obj_eq def)) Irel_defs; val Ipred_unabs_defs = map (fn def => mk_unabs_def m (HOLogic.mk_obj_eq def)) Ipred_defs; val Iset_defs = flat Iset_defss; fun mk_Imaps As Bs = map (fn mk => mk deads As Bs) mk_Imaps_Ds; fun mk_Isetss As = map2 (fn mk => fn Isets => map (mk deads As) Isets) mk_It_Ds Isetss; val Ibds = map2 (fn mk => mk deads passiveAs) mk_It_Ds Ibds_Ds; val Iwitss = map2 (fn mk => fn Iwits => map (mk deads passiveAs o snd) Iwits) mk_It_Ds Iwitss_Ds; fun mk_Irels As Bs = map (fn mk => mk deads As Bs) mk_Irels_Ds; fun mk_Ipreds As = map (fn mk => mk deads As) mk_Ipreds_Ds; val Imaps = mk_Imaps passiveAs passiveBs; val fs_Imaps = map (fn m => Term.list_comb (m, fs)) Imaps; val fs_copy_Imaps = map (fn m => Term.list_comb (m, fs_copy)) Imaps; val (Isetss_by_range, Isetss_by_bnf) = `transpose (mk_Isetss passiveAs); val map_setss = map (fn T => map2 (fn Ds => mk_map_of_bnf Ds (passiveAs @ Ts) (mk_set_Ts T)) Dss bnfs) passiveAs; val timer = time (timer "bnf constants for the new datatypes"); val (ctor_Imap_thms, ctor_Imap_o_thms) = let fun mk_goal fs_map map ctor ctor' = mk_Trueprop_eq (HOLogic.mk_comp (fs_map, ctor), HOLogic.mk_comp (ctor', Term.list_comb (map, fs @ fs_Imaps))); val goals = @{map 4} mk_goal fs_Imaps map_FTFT's ctors ctor's; val maps = @{map 4} (fn goal => fn foldx => fn map_comp_id => fn map_cong0 => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Imap_defs THEN mk_map_tac ctxt m n foldx map_comp_id map_cong0)) |> Thm.close_derivation \<^here>) goals ctor_fold_thms map_comp_id_thms map_cong0s; in `(map (fn thm => thm RS @{thm comp_eq_dest})) maps end; val (ctor_Imap_unique_thms, ctor_Imap_unique_thm) = let fun mk_prem u map ctor ctor' = mk_Trueprop_eq (HOLogic.mk_comp (u, ctor), HOLogic.mk_comp (ctor', Term.list_comb (map, fs @ us))); val prems = @{map 4} mk_prem us map_FTFT's ctors ctor's; val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_eq) us fs_Imaps)); val vars = fold (Variable.add_free_names lthy) (goal :: prems) []; val unique = Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, goal)) (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Imap_defs THEN mk_ctor_map_unique_tac ctxt ctor_fold_unique_thm sym_map_comps) |> Thm.close_derivation \<^here>; in `split_conj_thm unique end; val timer = time (timer "map functions for the new datatypes"); val ctor_Iset_thmss = let fun mk_goal sets ctor set col map = mk_Trueprop_eq (HOLogic.mk_comp (set, ctor), HOLogic.mk_comp (col, Term.list_comb (map, passive_ids @ sets))); val goalss = @{map 3} (fn sets => @{map 4} (mk_goal sets) ctors sets) Isetss_by_range colss map_setss; val setss = map (map2 (fn foldx => fn goal => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Iset_defs THEN mk_set_tac ctxt foldx) |> Thm.close_derivation \<^here>) ctor_fold_thms) goalss; fun mk_simp_goal pas_set act_sets sets ctor z set = mk_Trueprop_eq (set $ (ctor $ z), mk_union (pas_set $ z, Library.foldl1 mk_union (map2 (fn X => mk_UNION (X $ z)) act_sets sets))); val simp_goalss = map2 (fn i => fn sets => @{map 4} (fn Fsets => mk_simp_goal (nth Fsets (i - 1)) (drop m Fsets) sets) FTs_setss ctors xFs sets) ls Isetss_by_range; val ctor_setss = @{map 3} (fn i => @{map 3} (fn set_nats => fn goal => fn set => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_ctor_set_tac ctxt set (nth set_nats (i - 1)) (drop m set_nats))) |> Thm.close_derivation \<^here>) set_mapss) ls simp_goalss setss; in ctor_setss end; fun mk_set_thms ctor_set = (@{thm xt1(3)} OF [ctor_set, @{thm Un_upper1}]) :: map (fn i => (@{thm xt1(3)} OF [ctor_set, @{thm Un_upper2}]) RS (mk_Un_upper n i RS subset_trans) RSN (2, @{thm UN_upper} RS subset_trans)) (1 upto n); val set_Iset_thmsss = transpose (map (map mk_set_thms) ctor_Iset_thmss); val timer = time (timer "set functions for the new datatypes"); val cxs = map (SOME o Thm.cterm_of lthy) Izs; val Isetss_by_range' = map (map (Term.subst_atomic_types (passiveAs ~~ passiveBs))) Isetss_by_range; val Iset_Imap0_thmss = let fun mk_set_map0 f map z set set' = HOLogic.mk_eq (mk_image f $ (set $ z), set' $ (map $ z)); fun mk_cphi f map z set set' = Thm.cterm_of lthy (Term.absfree (dest_Free z) (mk_set_map0 f map z set set')); val csetss = map (map (Thm.cterm_of lthy)) Isetss_by_range'; val cphiss = @{map 3} (fn f => fn sets => fn sets' => (@{map 4} (mk_cphi f) fs_Imaps Izs sets sets')) fs Isetss_by_range Isetss_by_range'; val inducts = map (fn cphis => Thm.instantiate' cTs (map SOME cphis @ cxs) ctor_induct_thm) cphiss; val goals = @{map 3} (fn f => fn sets => fn sets' => HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (@{map 4} (mk_set_map0 f) fs_Imaps Izs sets sets'))) fs Isetss_by_range Isetss_by_range'; fun mk_tac ctxt induct = mk_set_nat_tac ctxt m (rtac ctxt induct) set_mapss ctor_Imap_thms; val thms = @{map 5} (fn goal => fn csets => fn ctor_sets => fn induct => fn i => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_tac ctxt induct csets ctor_sets i)) |> Thm.close_derivation \<^here>) goals csetss ctor_Iset_thmss inducts ls; in map split_conj_thm thms end; val Iset_bd_thmss = let fun mk_set_bd z bd set = mk_ordLeq (mk_card_of (set $ z)) bd; fun mk_cphi z set = Thm.cterm_of lthy (Term.absfree (dest_Free z) (mk_set_bd z sbd0 set)); val cphiss = map (map2 mk_cphi Izs) Isetss_by_range; val inducts = map (fn cphis => Thm.instantiate' cTs (map SOME cphis @ cxs) ctor_induct_thm) cphiss; val goals = map (fn sets => HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (@{map 3} mk_set_bd Izs Ibds sets))) Isetss_by_range; fun mk_tac ctxt induct = mk_set_bd_tac ctxt m (rtac ctxt induct) sbd0_Cinfinite set_sbd0ss; val thms = @{map 4} (fn goal => fn ctor_sets => fn induct => fn i => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Ibd_defs THEN mk_tac ctxt induct ctor_sets i)) |> Thm.close_derivation \<^here>) goals ctor_Iset_thmss inducts ls; in map split_conj_thm thms end; val Imap_cong0_thms = let fun mk_prem z set f g y y' = mk_Ball (set $ z) (Term.absfree y' (HOLogic.mk_eq (f $ y, g $ y))); fun mk_map_cong0 sets z fmap gmap = HOLogic.mk_imp (Library.foldr1 HOLogic.mk_conj (@{map 5} (mk_prem z) sets fs fs_copy ys ys'), HOLogic.mk_eq (fmap $ z, gmap $ z)); fun mk_cphi sets z fmap gmap = Thm.cterm_of lthy (Term.absfree (dest_Free z) (mk_map_cong0 sets z fmap gmap)); val cphis = @{map 4} mk_cphi Isetss_by_bnf Izs fs_Imaps fs_copy_Imaps; val induct = Thm.instantiate' cTs (map SOME cphis @ cxs) ctor_induct_thm; val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (@{map 4} mk_map_cong0 Isetss_by_bnf Izs fs_Imaps fs_copy_Imaps)); val vars = Variable.add_free_names lthy goal []; val thm = Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_mcong_tac ctxt (rtac ctxt induct) set_Iset_thmsss map_cong0s ctor_Imap_thms) |> Thm.close_derivation \<^here>; in split_conj_thm thm end; val in_rels = map in_rel_of_bnf bnfs; val in_Irels = map (fn def => trans OF [def, @{thm OO_Grp_alt}] RS @{thm predicate2_eqD}) Irel_unabs_defs; val ctor_Iset_incl_thmss = map (map hd) set_Iset_thmsss; val ctor_set_Iset_incl_thmsss = map (transpose o map tl) set_Iset_thmsss; val ctor_Iset_thmss' = transpose ctor_Iset_thmss; val Irels = mk_Irels passiveAs passiveBs; val Ipreds = mk_Ipreds passiveAs; val Irelphis = map (fn rel => Term.list_comb (rel, Iphis)) Irels; val relphis = map (fn rel => Term.list_comb (rel, Iphis @ Irelphis)) rels; val Irelpsi1s = map (fn rel => Term.list_comb (rel, Ipsi1s)) (mk_Irels passiveAs passiveCs); val Irelpsi2s = map (fn rel => Term.list_comb (rel, Ipsi2s)) (mk_Irels passiveCs passiveBs); val Irelpsi12s = map (fn rel => Term.list_comb (rel, map2 (curry mk_rel_compp) Ipsi1s Ipsi2s)) Irels; val ctor_Irel_thms = let fun mk_goal xF yF ctor ctor' Irelphi relphi = mk_Trueprop_eq (Irelphi $ (ctor $ xF) $ (ctor' $ yF), relphi $ xF $ yF); val goals = @{map 6} mk_goal xFs yFs ctors ctor's Irelphis relphis; in @{map 12} (fn i => fn goal => fn in_rel => fn map_comp0 => fn map_cong0 => fn ctor_map => fn ctor_sets => fn ctor_inject => fn ctor_dtor => fn set_map0s => fn ctor_set_incls => fn ctor_set_set_inclss => Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_ctor_rel_tac ctxt in_Irels i in_rel map_comp0 map_cong0 ctor_map ctor_sets ctor_inject ctor_dtor set_map0s ctor_set_incls ctor_set_set_inclss)) |> Thm.close_derivation \<^here>) ks goals in_rels map_comps map_cong0s ctor_Imap_thms ctor_Iset_thmss' ctor_inject_thms ctor_dtor_thms set_mapss ctor_Iset_incl_thmss ctor_set_Iset_incl_thmsss end; val le_Irel_OO_thm = let fun mk_le_Irel_OO Irelpsi1 Irelpsi2 Irelpsi12 Iz1 Iz2 = HOLogic.mk_imp (mk_rel_compp (Irelpsi1, Irelpsi2) $ Iz1 $ Iz2, Irelpsi12 $ Iz1 $ Iz2); val goals = @{map 5} mk_le_Irel_OO Irelpsi1s Irelpsi2s Irelpsi12s Izs1 Izs2; val cTs = map (SOME o Thm.ctyp_of lthy o TFree) induct2_params; val cxs = map (SOME o Thm.cterm_of lthy) (splice Izs1 Izs2); fun mk_cphi z1 z2 goal = SOME (Thm.cterm_of lthy (Term.absfree z1 (Term.absfree z2 goal))); val cphis = @{map 3} mk_cphi Izs1' Izs2' goals; val induct = Thm.instantiate' cTs (cphis @ cxs) ctor_induct2_thm; val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj goals); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_le_rel_OO_tac ctxt m induct ctor_nchotomy_thms ctor_Irel_thms rel_mono_strong0s le_rel_OOs) |> Thm.close_derivation \<^here> end; val timer = time (timer "helpers for BNF properties"); val map_id0_tacs = map (fn thm => fn ctxt => mk_map_id0_tac ctxt map_id0s thm) ctor_Imap_unique_thms; val map_comp0_tacs = map2 (fn thm => fn i => fn ctxt => mk_map_comp0_tac ctxt map_comps ctor_Imap_thms thm i) ctor_Imap_unique_thms ks; val map_cong0_tacs = map (fn thm => fn ctxt => mk_map_cong0_tac ctxt m thm) Imap_cong0_thms; val set_map0_tacss = map (map (fn thm => fn ctxt => mk_set_map0_tac ctxt thm)) (transpose Iset_Imap0_thmss); val bd_co_tacs = replicate n (fn ctxt => unfold_thms_tac ctxt Ibd_defs THEN rtac ctxt sbd0_card_order 1); val bd_cinf_tacs = replicate n (fn ctxt => unfold_thms_tac ctxt Ibd_defs THEN rtac ctxt (sbd0_Cinfinite RS conjunct1) 1); val set_bd_tacss = map (map (fn thm => fn ctxt => rtac ctxt thm 1)) (transpose Iset_bd_thmss); val le_rel_OO_tacs = map (fn i => fn ctxt => (rtac ctxt @{thm predicate2I} THEN' etac ctxt (le_Irel_OO_thm RS mk_conjunctN n i RS mp)) 1) ks; val rel_OO_Grp_tacs = map (fn def => fn ctxt => rtac ctxt def 1) Irel_unabs_defs; val pred_set_tacs = map (fn def => fn ctxt => rtac ctxt def 1) Ipred_unabs_defs; val tacss = @{map 10} zip_axioms map_id0_tacs map_comp0_tacs map_cong0_tacs set_map0_tacss bd_co_tacs bd_cinf_tacs set_bd_tacss le_rel_OO_tacs rel_OO_Grp_tacs pred_set_tacs; fun wit_tac ctxt = unfold_thms_tac ctxt (flat Iwit_defss) THEN mk_wit_tac ctxt n (flat ctor_Iset_thmss) (maps wit_thms_of_bnf bnfs); val (Ibnfs, lthy) = @{fold_map 6} (fn tacs => fn map_b => fn rel_b => fn pred_b => fn set_bs => fn consts => bnf_def Do_Inline (user_policy Note_Some) false I tacs wit_tac (SOME deads) map_b rel_b pred_b set_bs consts) tacss map_bs rel_bs pred_bs set_bss (((((((replicate n Binding.empty ~~ Ts) ~~ Imaps) ~~ Isetss_by_bnf) ~~ Ibds) ~~ Iwitss) ~~ map SOME Irels) ~~ map SOME Ipreds) lthy; val timer = time (timer "registered new datatypes as BNFs"); val ls' = if m = 1 then [0] else ls val Ibnf_common_notes = [(ctor_map_uniqueN, [ctor_Imap_unique_thm])] |> map (fn (thmN, thms) => ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])])); val Ibnf_notes = [(ctor_mapN, map single ctor_Imap_thms), (ctor_relN, map single ctor_Irel_thms), (ctor_set_inclN, ctor_Iset_incl_thmss), (ctor_set_set_inclN, map flat ctor_set_Iset_incl_thmsss)] @ map2 (fn i => fn thms => (mk_ctor_setN i, map single thms)) ls' ctor_Iset_thmss |> maps (fn (thmN, thmss) => map2 (fn b => fn thms => ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])])) bs thmss) in (timer, Ibnfs, (ctor_Imap_o_thms, ctor_Imap_thms), ctor_Imap_unique_thm, ctor_Iset_thmss', ctor_Irel_thms, Ibnf_common_notes @ Ibnf_notes, lthy) end; val ((((((xFs, yFs)), Iphis), activephis), activeIphis), _) = lthy |> mk_Frees "x" FTs ||>> mk_Frees "y" FTs' ||>> mk_Frees "R" IphiTs ||>> mk_Frees "S" activephiTs ||>> mk_Frees "IR" activeIphiTs; val ctor_fold_o_Imap_thms = mk_xtor_co_iter_o_map_thms Least_FP false m ctor_fold_unique_thm ctor_Imap_o_thms (map (mk_pointfree2 lthy) ctor_fold_thms) sym_map_comps map_cong0s; val Irels = if m = 0 then map HOLogic.eq_const Ts else map (mk_rel_of_bnf deads passiveAs passiveBs) Ibnfs; val Irel_induct_thm = mk_xtor_rel_co_induct_thm Least_FP rels activeIphis Irels Iphis xFs yFs ctors ctor's (fn {context = ctxt, prems = IHs} => mk_rel_induct_tac ctxt IHs m ctor_induct2_thm ks ctor_Irel_thms rel_mono_strong0s) lthy; val rels = map2 (fn Ds => mk_rel_of_bnf Ds allAs allBs') Dss bnfs; val ctor_fold_transfer_thms = mk_xtor_co_iter_transfer_thms Least_FP rels activephis activephis Irels Iphis (mk_folds passiveAs activeAs) (mk_folds passiveBs activeBs) (fn {context = ctxt, prems = _} => mk_fold_transfer_tac ctxt m Irel_induct_thm (map map_transfer_of_bnf bnfs) ctor_fold_thms) lthy; val timer = time (timer "relator induction"); fun mk_Ts As = map (typ_subst_atomic (passiveAs ~~ As)) Ts; val export = map (Morphism.term (Local_Theory.target_morphism lthy)) val ((recs, (ctor_rec_thms, ctor_rec_unique_thm, ctor_rec_o_Imap_thms, ctor_rec_transfer_thms)), lthy) = lthy |> derive_xtor_co_recs Least_FP external_bs mk_Ts (Dss, resDs) bnfs (export ctors) (export folds) ctor_fold_unique_thm ctor_fold_thms ctor_fold_transfer_thms ctor_Imap_thms ctor_Irel_thms (replicate n NONE); val timer = time (timer "recursor"); val common_notes = [(ctor_inductN, [ctor_induct_thm]), (ctor_induct2N, [ctor_induct2_thm]), (ctor_rel_inductN, [Irel_induct_thm])] |> map (fn (thmN, thms) => ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])])); val notes = [(ctor_dtorN, ctor_dtor_thms), (ctor_exhaustN, ctor_exhaust_thms), (ctor_foldN, ctor_fold_thms), (ctor_fold_o_mapN, ctor_fold_o_Imap_thms), (ctor_fold_transferN, ctor_fold_transfer_thms), (ctor_fold_uniqueN, ctor_fold_unique_thms), (ctor_injectN, ctor_inject_thms), (dtor_ctorN, dtor_ctor_thms), (dtor_exhaustN, dtor_exhaust_thms), (dtor_injectN, dtor_inject_thms)] |> map (apsnd (map single)) |> maps (fn (thmN, thmss) => map2 (fn b => fn thms => ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])])) bs thmss); val lthy' = lthy |> internals ? snd o Local_Theory.notes (common_notes @ notes @ Ibnf_notes); val fp_res = {Ts = Ts, bnfs = Ibnfs, pre_bnfs = bnfs, absT_infos = absT_infos, ctors = ctors, dtors = dtors, xtor_un_folds = folds, xtor_co_recs = export recs, xtor_co_induct = ctor_induct_thm, dtor_ctors = dtor_ctor_thms, ctor_dtors = ctor_dtor_thms, ctor_injects = ctor_inject_thms, dtor_injects = dtor_inject_thms, xtor_maps = ctor_Imap_thms, xtor_map_unique = ctor_Imap_unique_thm, xtor_setss = ctor_Iset_thmss', xtor_rels = ctor_Irel_thms, xtor_un_fold_thms = ctor_fold_thms, xtor_co_rec_thms = ctor_rec_thms, xtor_un_fold_unique = ctor_fold_unique_thm, xtor_co_rec_unique = ctor_rec_unique_thm, xtor_un_fold_o_maps = ctor_fold_o_Imap_thms, xtor_co_rec_o_maps = ctor_rec_o_Imap_thms, xtor_un_fold_transfers = ctor_fold_transfer_thms, xtor_co_rec_transfers = ctor_rec_transfer_thms, xtor_rel_co_induct = Irel_induct_thm, dtor_set_inducts = []}; in timer; (fp_res, lthy') end; val _ = Outer_Syntax.local_theory \<^command_keyword>\datatype\ "define inductive datatypes" (parse_co_datatype_cmd Least_FP construct_lfp); -val _ = Theory.setup (fp_antiquote_setup \<^binding>\datatype\); - end; diff --git a/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML b/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML --- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML +++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML @@ -1,1190 +1,1270 @@ (* Title: HOL/Tools/Ctr_Sugar/ctr_sugar.ML Author: Jasmin Blanchette, TU Muenchen Author: Martin Desharnais, TU Muenchen Copyright 2012, 2013 Wrapping existing freely generated type's constructors. *) signature CTR_SUGAR = sig datatype ctr_sugar_kind = Datatype | Codatatype | Record | Unknown type ctr_sugar = {kind: ctr_sugar_kind, T: typ, ctrs: term list, casex: term, discs: term list, selss: term list list, exhaust: thm, nchotomy: thm, injects: thm list, distincts: thm list, case_thms: thm list, case_cong: thm, case_cong_weak: thm, case_distribs: thm list, split: thm, split_asm: thm, disc_defs: thm list, disc_thmss: thm list list, discIs: thm list, disc_eq_cases: thm list, sel_defs: thm list, sel_thmss: thm list list, distinct_discsss: thm list list list, exhaust_discs: thm list, exhaust_sels: thm list, collapses: thm list, expands: thm list, split_sels: thm list, split_sel_asms: thm list, case_eq_ifs: thm list}; val morph_ctr_sugar: morphism -> ctr_sugar -> ctr_sugar val transfer_ctr_sugar: theory -> ctr_sugar -> ctr_sugar val ctr_sugar_of: Proof.context -> string -> ctr_sugar option val ctr_sugar_of_global: theory -> string -> ctr_sugar option val ctr_sugars_of: Proof.context -> ctr_sugar list val ctr_sugars_of_global: theory -> ctr_sugar list val ctr_sugar_of_case: Proof.context -> string -> ctr_sugar option val ctr_sugar_of_case_global: theory -> string -> ctr_sugar option val ctr_sugar_interpretation: string -> (ctr_sugar -> local_theory -> local_theory) -> theory -> theory val interpret_ctr_sugar: (string -> bool) -> ctr_sugar -> local_theory -> local_theory val register_ctr_sugar_raw: ctr_sugar -> local_theory -> local_theory val register_ctr_sugar: (string -> bool) -> ctr_sugar -> local_theory -> local_theory val default_register_ctr_sugar_global: (string -> bool) -> ctr_sugar -> theory -> theory val mk_half_pairss: 'a list * 'a list -> ('a * 'a) list list val join_halves: int -> 'a list list -> 'a list list -> 'a list * 'a list list list val mk_ctr: typ list -> term -> term val mk_case: typ list -> typ -> term -> term val mk_disc_or_sel: typ list -> term -> term val name_of_ctr: term -> string val name_of_disc: term -> string val dest_ctr: Proof.context -> string -> term -> term * term list val dest_case: Proof.context -> string -> typ list -> term -> (ctr_sugar * term list * term list) option type ('c, 'a) ctr_spec = (binding * 'c) * 'a list val disc_of_ctr_spec: ('c, 'a) ctr_spec -> binding val ctr_of_ctr_spec: ('c, 'a) ctr_spec -> 'c val args_of_ctr_spec: ('c, 'a) ctr_spec -> 'a list val code_plugin: string type ctr_options = (string -> bool) * bool type ctr_options_cmd = (Proof.context -> string -> bool) * bool val fake_local_theory_for_sel_defaults: (binding * typ) list -> Proof.context -> Proof.context val free_constructors: ctr_sugar_kind -> ({prems: thm list, context: Proof.context} -> tactic) list list -> ((ctr_options * binding) * (term, binding) ctr_spec list) * term list -> local_theory -> ctr_sugar * local_theory val free_constructors_cmd: ctr_sugar_kind -> ((((Proof.context -> Plugin_Name.filter) * bool) * binding) * ((binding * string) * binding list) list) * string list -> Proof.context -> Proof.state val default_ctr_options: ctr_options val default_ctr_options_cmd: ctr_options_cmd val parse_bound_term: (binding * string) parser val parse_ctr_options: ctr_options_cmd parser val parse_ctr_spec: 'c parser -> 'a parser -> ('c, 'a) ctr_spec parser val parse_sel_default_eqs: string list parser end; structure Ctr_Sugar : CTR_SUGAR = struct open Ctr_Sugar_Util open Ctr_Sugar_Tactics open Ctr_Sugar_Code datatype ctr_sugar_kind = Datatype | Codatatype | Record | Unknown; type ctr_sugar = {kind: ctr_sugar_kind, T: typ, ctrs: term list, casex: term, discs: term list, selss: term list list, exhaust: thm, nchotomy: thm, injects: thm list, distincts: thm list, case_thms: thm list, case_cong: thm, case_cong_weak: thm, case_distribs: thm list, split: thm, split_asm: thm, disc_defs: thm list, disc_thmss: thm list list, discIs: thm list, disc_eq_cases: thm list, sel_defs: thm list, sel_thmss: thm list list, distinct_discsss: thm list list list, exhaust_discs: thm list, exhaust_sels: thm list, collapses: thm list, expands: thm list, split_sels: thm list, split_sel_asms: thm list, case_eq_ifs: thm list}; fun morph_ctr_sugar phi ({kind, T, ctrs, casex, discs, selss, exhaust, nchotomy, injects, distincts, case_thms, case_cong, case_cong_weak, case_distribs, split, split_asm, disc_defs, disc_thmss, discIs, disc_eq_cases, sel_defs, sel_thmss, distinct_discsss, exhaust_discs, exhaust_sels, collapses, expands, split_sels, split_sel_asms, case_eq_ifs} : ctr_sugar) = {kind = kind, T = Morphism.typ phi T, ctrs = map (Morphism.term phi) ctrs, casex = Morphism.term phi casex, discs = map (Morphism.term phi) discs, selss = map (map (Morphism.term phi)) selss, exhaust = Morphism.thm phi exhaust, nchotomy = Morphism.thm phi nchotomy, injects = map (Morphism.thm phi) injects, distincts = map (Morphism.thm phi) distincts, case_thms = map (Morphism.thm phi) case_thms, case_cong = Morphism.thm phi case_cong, case_cong_weak = Morphism.thm phi case_cong_weak, case_distribs = map (Morphism.thm phi) case_distribs, split = Morphism.thm phi split, split_asm = Morphism.thm phi split_asm, disc_defs = map (Morphism.thm phi) disc_defs, disc_thmss = map (map (Morphism.thm phi)) disc_thmss, discIs = map (Morphism.thm phi) discIs, disc_eq_cases = map (Morphism.thm phi) disc_eq_cases, sel_defs = map (Morphism.thm phi) sel_defs, sel_thmss = map (map (Morphism.thm phi)) sel_thmss, distinct_discsss = map (map (map (Morphism.thm phi))) distinct_discsss, exhaust_discs = map (Morphism.thm phi) exhaust_discs, exhaust_sels = map (Morphism.thm phi) exhaust_sels, collapses = map (Morphism.thm phi) collapses, expands = map (Morphism.thm phi) expands, split_sels = map (Morphism.thm phi) split_sels, split_sel_asms = map (Morphism.thm phi) split_sel_asms, case_eq_ifs = map (Morphism.thm phi) case_eq_ifs}; val transfer_ctr_sugar = morph_ctr_sugar o Morphism.transfer_morphism; structure Data = Generic_Data ( - type T = ctr_sugar Symtab.table; + type T = (Position.T * ctr_sugar) Symtab.table; val empty = Symtab.empty; val extend = I; fun merge data : T = Symtab.merge (K true) data; ); fun ctr_sugar_of_generic context = - Option.map (transfer_ctr_sugar (Context.theory_of context)) o Symtab.lookup (Data.get context); + Option.map (transfer_ctr_sugar (Context.theory_of context) o #2) o Symtab.lookup (Data.get context); fun ctr_sugars_of_generic context = - Symtab.fold (cons o transfer_ctr_sugar (Context.theory_of context) o snd) (Data.get context) []; + Symtab.fold (cons o transfer_ctr_sugar (Context.theory_of context) o #2 o #2) (Data.get context) []; fun ctr_sugar_of_case_generic context s = find_first (fn {casex = Const (s', _), ...} => s' = s | _ => false) (ctr_sugars_of_generic context); val ctr_sugar_of = ctr_sugar_of_generic o Context.Proof; val ctr_sugar_of_global = ctr_sugar_of_generic o Context.Theory; val ctr_sugars_of = ctr_sugars_of_generic o Context.Proof; val ctr_sugars_of_global = ctr_sugars_of_generic o Context.Theory; val ctr_sugar_of_case = ctr_sugar_of_case_generic o Context.Proof; val ctr_sugar_of_case_global = ctr_sugar_of_case_generic o Context.Theory; structure Ctr_Sugar_Plugin = Plugin(type T = ctr_sugar); fun ctr_sugar_interpretation name f = Ctr_Sugar_Plugin.interpretation name (fn ctr_sugar => fn lthy => f (transfer_ctr_sugar (Proof_Context.theory_of lthy) ctr_sugar) lthy); val interpret_ctr_sugar = Ctr_Sugar_Plugin.data; -fun register_ctr_sugar_raw (ctr_sugar as {T = Type (s, _), ...}) = +fun register_ctr_sugar_raw (ctr_sugar as {T = Type (name, _), ...}) = Local_Theory.declaration {syntax = false, pervasive = true} - (fn phi => Data.map (Symtab.update (s, morph_ctr_sugar phi ctr_sugar))); + (fn phi => fn context => + let val pos = Position.thread_data () + in Data.map (Symtab.update (name, (pos, morph_ctr_sugar phi ctr_sugar))) context end); fun register_ctr_sugar plugins ctr_sugar = register_ctr_sugar_raw ctr_sugar #> interpret_ctr_sugar plugins ctr_sugar; -fun default_register_ctr_sugar_global plugins (ctr_sugar as {T = Type (s, _), ...}) thy = - let val tab = Data.get (Context.Theory thy) in - if Symtab.defined tab s then - thy +fun default_register_ctr_sugar_global plugins (ctr_sugar as {T = Type (name, _), ...}) thy = + let + val tab = Data.get (Context.Theory thy); + val pos = Position.thread_data (); + in + if Symtab.defined tab name then thy else thy - |> Context.theory_map (Data.put (Symtab.update_new (s, ctr_sugar) tab)) + |> Context.theory_map (Data.put (Symtab.update_new (name, (pos, ctr_sugar)) tab)) |> Named_Target.theory_map (Ctr_Sugar_Plugin.data plugins ctr_sugar) end; val is_prefix = "is_"; val un_prefix = "un_"; val not_prefix = "not_"; fun mk_unN 1 1 suf = un_prefix ^ suf | mk_unN _ l suf = un_prefix ^ suf ^ string_of_int l; val caseN = "case"; val case_congN = "case_cong"; val case_eq_ifN = "case_eq_if"; val collapseN = "collapse"; val discN = "disc"; val disc_eq_caseN = "disc_eq_case"; val discIN = "discI"; val distinctN = "distinct"; val distinct_discN = "distinct_disc"; val exhaustN = "exhaust"; val exhaust_discN = "exhaust_disc"; val expandN = "expand"; val injectN = "inject"; val nchotomyN = "nchotomy"; val selN = "sel"; val exhaust_selN = "exhaust_sel"; val splitN = "split"; val split_asmN = "split_asm"; val split_selN = "split_sel"; val split_sel_asmN = "split_sel_asm"; val splitsN = "splits"; val split_selsN = "split_sels"; val case_cong_weak_thmsN = "case_cong_weak"; val case_distribN = "case_distrib"; val cong_attrs = @{attributes [cong]}; val dest_attrs = @{attributes [dest]}; val safe_elim_attrs = @{attributes [elim!]}; val iff_attrs = @{attributes [iff]}; val inductsimp_attrs = @{attributes [induct_simp]}; val nitpicksimp_attrs = @{attributes [nitpick_simp]}; val simp_attrs = @{attributes [simp]}; fun unflat_lookup eq xs ys = map (fn xs' => permute_like_unique eq xs xs' ys); fun mk_half_pairss' _ ([], []) = [] | mk_half_pairss' indent (x :: xs, _ :: ys) = indent @ fold_rev (cons o single o pair x) ys (mk_half_pairss' ([] :: indent) (xs, ys)); fun mk_half_pairss p = mk_half_pairss' [[]] p; fun join_halves n half_xss other_half_xss = (splice (flat half_xss) (flat other_half_xss), map2 (map2 append) (Library.chop_groups n half_xss) (transpose (Library.chop_groups n other_half_xss))); fun mk_undefined T = Const (\<^const_name>\undefined\, T); fun mk_ctr Ts t = let val Type (_, Ts0) = body_type (fastype_of t) in subst_nonatomic_types (Ts0 ~~ Ts) t end; fun mk_case Ts T t = let val (Type (_, Ts0), body) = strip_type (fastype_of t) |>> List.last in subst_nonatomic_types ((body, T) :: (Ts0 ~~ Ts)) t end; fun mk_disc_or_sel Ts t = subst_nonatomic_types (snd (Term.dest_Type (domain_type (fastype_of t))) ~~ Ts) t; val name_of_ctr = name_of_const "constructor" body_type; fun name_of_disc t = (case head_of t of Abs (_, _, \<^const>\Not\ $ (t' $ Bound 0)) => Long_Name.map_base_name (prefix not_prefix) (name_of_disc t') | Abs (_, _, Const (\<^const_name>\HOL.eq\, _) $ Bound 0 $ t') => Long_Name.map_base_name (prefix is_prefix) (name_of_disc t') | Abs (_, _, \<^const>\Not\ $ (Const (\<^const_name>\HOL.eq\, _) $ Bound 0 $ t')) => Long_Name.map_base_name (prefix (not_prefix ^ is_prefix)) (name_of_disc t') | t' => name_of_const "discriminator" (perhaps (try domain_type)) t'); val base_name_of_ctr = Long_Name.base_name o name_of_ctr; fun dest_ctr ctxt s t = let val (f, args) = Term.strip_comb t in (case ctr_sugar_of ctxt s of SOME {ctrs, ...} => (case find_first (can (fo_match ctxt f)) ctrs of SOME f' => (f', args) | NONE => raise Fail "dest_ctr") | NONE => raise Fail "dest_ctr") end; fun dest_case ctxt s Ts t = (case Term.strip_comb t of (Const (c, _), args as _ :: _) => (case ctr_sugar_of ctxt s of SOME (ctr_sugar as {casex = Const (case_name, _), discs = discs0, selss = selss0, ...}) => if case_name = c then let val n = length discs0 in if n < length args then let val (branches, obj :: leftovers) = chop n args; val discs = map (mk_disc_or_sel Ts) discs0; val selss = map (map (mk_disc_or_sel Ts)) selss0; val conds = map (rapp obj) discs; val branch_argss = map (fn sels => map (rapp obj) sels @ leftovers) selss; val branches' = map2 (curry Term.betapplys) branches branch_argss; in SOME (ctr_sugar, conds, branches') end else NONE end else NONE | _ => NONE) | _ => NONE); fun const_or_free_name (Const (s, _)) = Long_Name.base_name s | const_or_free_name (Free (s, _)) = s | const_or_free_name t = raise TERM ("const_or_free_name", [t]) fun extract_sel_default ctxt t = let fun malformed () = error ("Malformed selector default value equation: " ^ Syntax.string_of_term ctxt t); val ((sel, (ctr, vars)), rhs) = fst (Term.replace_dummy_patterns (Syntax.check_term ctxt t) 0) |> HOLogic.dest_eq |>> (Term.dest_comb #>> const_or_free_name ##> (Term.strip_comb #>> (Term.dest_Const #> fst))) handle TERM _ => malformed (); in if forall (is_Free orf is_Var) vars andalso not (has_duplicates (op aconv) vars) then ((ctr, sel), fold_rev Term.lambda vars rhs) else malformed () end; (* Ideally, we would enrich the context with constants rather than free variables. *) fun fake_local_theory_for_sel_defaults sel_bTs = Proof_Context.allow_dummies #> Proof_Context.add_fixes (map (fn (b, T) => (b, SOME T, NoSyn)) sel_bTs) #> snd; type ('c, 'a) ctr_spec = (binding * 'c) * 'a list; fun disc_of_ctr_spec ((disc, _), _) = disc; fun ctr_of_ctr_spec ((_, ctr), _) = ctr; fun args_of_ctr_spec (_, args) = args; val code_plugin = Plugin_Name.declare_setup \<^binding>\code\; fun prepare_free_constructors kind prep_plugins prep_term ((((raw_plugins, discs_sels), raw_case_binding), ctr_specs), sel_default_eqs) no_defs_lthy = let val plugins = prep_plugins no_defs_lthy raw_plugins; (* TODO: sanity checks on arguments *) val raw_ctrs = map ctr_of_ctr_spec ctr_specs; val raw_disc_bindings = map disc_of_ctr_spec ctr_specs; val raw_sel_bindingss = map args_of_ctr_spec ctr_specs; val n = length raw_ctrs; val ks = 1 upto n; val _ = n > 0 orelse error "No constructors specified"; val ctrs0 = map (prep_term no_defs_lthy) raw_ctrs; val (fcT_name, As0) = (case body_type (fastype_of (hd ctrs0)) of Type T' => T' | _ => error "Expected type constructor in body type of constructor"); val _ = forall ((fn Type (T_name, _) => T_name = fcT_name | _ => false) o body_type o fastype_of) (tl ctrs0) orelse error "Constructors not constructing same type"; val fc_b_name = Long_Name.base_name fcT_name; val fc_b = Binding.name fc_b_name; fun qualify mandatory = Binding.qualify mandatory fc_b_name; val (unsorted_As, [B, C]) = no_defs_lthy |> variant_tfrees (map (fst o dest_TFree_or_TVar) As0) ||> fst o mk_TFrees 2; val As = map2 (resort_tfree_or_tvar o snd o dest_TFree_or_TVar) As0 unsorted_As; val fcT = Type (fcT_name, As); val ctrs = map (mk_ctr As) ctrs0; val ctr_Tss = map (binder_types o fastype_of) ctrs; val ms = map length ctr_Tss; fun can_definitely_rely_on_disc k = not (Binding.is_empty (nth raw_disc_bindings (k - 1))) orelse nth ms (k - 1) = 0; fun can_rely_on_disc k = can_definitely_rely_on_disc k orelse (k = 1 andalso not (can_definitely_rely_on_disc 2)); fun should_omit_disc_binding k = n = 1 orelse (n = 2 andalso can_rely_on_disc (3 - k)); val equal_binding = \<^binding>\=\; fun is_disc_binding_valid b = not (Binding.is_empty b orelse Binding.eq_name (b, equal_binding)); val standard_disc_binding = Binding.name o prefix is_prefix o base_name_of_ctr; val disc_bindings = raw_disc_bindings |> @{map 4} (fn k => fn m => fn ctr => fn disc => qualify false (if Binding.is_empty disc then if m = 0 then equal_binding else if should_omit_disc_binding k then disc else standard_disc_binding ctr else if Binding.eq_name (disc, standard_binding) then standard_disc_binding ctr else disc)) ks ms ctrs0; fun standard_sel_binding m l = Binding.name o mk_unN m l o base_name_of_ctr; val sel_bindingss = @{map 3} (fn ctr => fn m => map2 (fn l => fn sel => qualify false (if Binding.is_empty sel orelse Binding.eq_name (sel, standard_binding) then standard_sel_binding m l ctr else sel)) (1 upto m) o pad_list Binding.empty m) ctrs0 ms raw_sel_bindingss; val add_bindings = Variable.add_fixes (distinct (op =) (filter Symbol_Pos.is_identifier (map Binding.name_of (disc_bindings @ flat sel_bindingss)))) #> snd; val case_Ts = map (fn Ts => Ts ---> B) ctr_Tss; val (((((((((u, exh_y), xss), yss), fs), gs), w), (p, p'))), _) = no_defs_lthy |> add_bindings |> yield_singleton (mk_Frees fc_b_name) fcT ||>> yield_singleton (mk_Frees "y") fcT (* for compatibility with "datatype_realizer.ML" *) ||>> mk_Freess "x" ctr_Tss ||>> mk_Freess "y" ctr_Tss ||>> mk_Frees "f" case_Ts ||>> mk_Frees "g" case_Ts ||>> yield_singleton (mk_Frees "z") B ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "P") HOLogic.boolT; val q = Free (fst p', mk_pred1T B); val xctrs = map2 (curry Term.list_comb) ctrs xss; val yctrs = map2 (curry Term.list_comb) ctrs yss; val xfs = map2 (curry Term.list_comb) fs xss; val xgs = map2 (curry Term.list_comb) gs xss; (* TODO: Eta-expension is for compatibility with the old datatype package (but it also provides nicer names). Consider removing. *) val eta_fs = map2 (fold_rev Term.lambda) xss xfs; val eta_gs = map2 (fold_rev Term.lambda) xss xgs; val case_binding = qualify false (if Binding.is_empty raw_case_binding orelse Binding.eq_name (raw_case_binding, standard_binding) then Binding.prefix_name (caseN ^ "_") fc_b else raw_case_binding); fun mk_case_disj xctr xf xs = list_exists_free xs (HOLogic.mk_conj (HOLogic.mk_eq (u, xctr), HOLogic.mk_eq (w, xf))); val case_rhs = fold_rev (fold_rev Term.lambda) [fs, [u]] (Const (\<^const_name>\The\, (B --> HOLogic.boolT) --> B) $ Term.lambda w (Library.foldr1 HOLogic.mk_disj (@{map 3} mk_case_disj xctrs xfs xss))); val ((raw_case, (_, raw_case_def)), (lthy, lthy_old)) = no_defs_lthy |> Local_Theory.open_target |> snd |> Local_Theory.define ((case_binding, NoSyn), ((Binding.concealed (Thm.def_binding case_binding), []), case_rhs)) ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy_old lthy; val case_def = Morphism.thm phi raw_case_def; val case0 = Morphism.term phi raw_case; val casex = mk_case As B case0; val casexC = mk_case As C case0; val casexBool = mk_case As HOLogic.boolT case0; fun mk_uu_eq () = HOLogic.mk_eq (u, u); val exist_xs_u_eq_ctrs = map2 (fn xctr => fn xs => list_exists_free xs (HOLogic.mk_eq (u, xctr))) xctrs xss; val unique_disc_no_def = TrueI; (*arbitrary marker*) val alternate_disc_no_def = FalseE; (*arbitrary marker*) fun alternate_disc_lhs get_udisc k = HOLogic.mk_not (let val b = nth disc_bindings (k - 1) in if is_disc_binding_valid b then get_udisc b (k - 1) else nth exist_xs_u_eq_ctrs (k - 1) end); val no_discs_sels = not discs_sels andalso forall (forall Binding.is_empty) (raw_disc_bindings :: raw_sel_bindingss) andalso null sel_default_eqs; val (all_sels_distinct, discs, selss, disc_defs, sel_defs, sel_defss, lthy) = if no_discs_sels then (true, [], [], [], [], [], lthy) else let val all_sel_bindings = flat sel_bindingss; val num_all_sel_bindings = length all_sel_bindings; val uniq_sel_bindings = distinct Binding.eq_name all_sel_bindings; val all_sels_distinct = (length uniq_sel_bindings = num_all_sel_bindings); val sel_binding_index = if all_sels_distinct then 1 upto num_all_sel_bindings else map (fn b => find_index (curry Binding.eq_name b) uniq_sel_bindings) all_sel_bindings; val all_proto_sels = flat (@{map 3} (fn k => fn xs => map (pair k o pair xs)) ks xss xss); val sel_infos = AList.group (op =) (sel_binding_index ~~ all_proto_sels) |> sort (int_ord o apply2 fst) |> map snd |> curry (op ~~) uniq_sel_bindings; val sel_bindings = map fst sel_infos; val sel_defaults = if null sel_default_eqs then [] else let val sel_Ts = map (curry (op -->) fcT o fastype_of o snd o snd o hd o snd) sel_infos; val fake_lthy = fake_local_theory_for_sel_defaults (sel_bindings ~~ sel_Ts) no_defs_lthy; in map (extract_sel_default fake_lthy o prep_term fake_lthy) sel_default_eqs end; fun disc_free b = Free (Binding.name_of b, mk_pred1T fcT); fun disc_spec b exist_xs_u_eq_ctr = mk_Trueprop_eq (disc_free b $ u, exist_xs_u_eq_ctr); fun alternate_disc k = Term.lambda u (alternate_disc_lhs (K o rapp u o disc_free) (3 - k)); fun mk_sel_case_args b proto_sels T = @{map 3} (fn Const (c, _) => fn Ts => fn k => (case AList.lookup (op =) proto_sels k of NONE => (case filter (curry (op =) (c, Binding.name_of b) o fst) sel_defaults of [] => fold_rev (Term.lambda o curry Free Name.uu) Ts (mk_undefined T) | [(_, t)] => t | _ => error "Multiple default values for selector/constructor pair") | SOME (xs, x) => fold_rev Term.lambda xs x)) ctrs ctr_Tss ks; fun sel_spec b proto_sels = let val _ = (case duplicates (op =) (map fst proto_sels) of k :: _ => error ("Duplicate selector name " ^ quote (Binding.name_of b) ^ " for constructor " ^ quote (Syntax.string_of_term lthy (nth ctrs (k - 1)))) | [] => ()) val T = (case distinct (op =) (map (fastype_of o snd o snd) proto_sels) of [T] => T | T :: T' :: _ => error ("Inconsistent range type for selector " ^ quote (Binding.name_of b) ^ ": " ^ quote (Syntax.string_of_typ lthy T) ^ " vs. " ^ quote (Syntax.string_of_typ lthy T'))); in mk_Trueprop_eq (Free (Binding.name_of b, fcT --> T) $ u, Term.list_comb (mk_case As T case0, mk_sel_case_args b proto_sels T) $ u) end; fun unflat_selss xs = unflat_lookup Binding.eq_name sel_bindings xs sel_bindingss; val (((raw_discs, raw_disc_defs), (raw_sels, raw_sel_defs)), (lthy', lthy)) = lthy |> Local_Theory.open_target |> snd |> apfst split_list o @{fold_map 3} (fn k => fn exist_xs_u_eq_ctr => fn b => if Binding.is_empty b then if n = 1 then pair (Term.lambda u (mk_uu_eq ()), unique_disc_no_def) else pair (alternate_disc k, alternate_disc_no_def) else if Binding.eq_name (b, equal_binding) then pair (Term.lambda u exist_xs_u_eq_ctr, refl) else Specification.definition (SOME (b, NONE, NoSyn)) [] [] ((Thm.def_binding b, []), disc_spec b exist_xs_u_eq_ctr) #>> apsnd snd) ks exist_xs_u_eq_ctrs disc_bindings ||>> apfst split_list o fold_map (fn (b, proto_sels) => Specification.definition (SOME (b, NONE, NoSyn)) [] [] ((Thm.def_binding b, []), sel_spec b proto_sels) #>> apsnd snd) sel_infos ||> `Local_Theory.close_target; val phi = Proof_Context.export_morphism lthy lthy'; val disc_defs = map (Morphism.thm phi) raw_disc_defs; val sel_defs = map (Morphism.thm phi) raw_sel_defs; val sel_defss = unflat_selss sel_defs; val discs0 = map (Morphism.term phi) raw_discs; val selss0 = unflat_selss (map (Morphism.term phi) raw_sels); val discs = map (mk_disc_or_sel As) discs0; val selss = map (map (mk_disc_or_sel As)) selss0; in (all_sels_distinct, discs, selss, disc_defs, sel_defs, sel_defss, lthy') end; fun mk_imp_p Qs = Logic.list_implies (Qs, HOLogic.mk_Trueprop p); val exhaust_goal = let fun mk_prem xctr xs = fold_rev Logic.all xs (mk_imp_p [mk_Trueprop_eq (exh_y, xctr)]) in fold_rev Logic.all [p, exh_y] (mk_imp_p (map2 mk_prem xctrs xss)) end; val inject_goalss = let fun mk_goal _ _ [] [] = [] | mk_goal xctr yctr xs ys = [fold_rev Logic.all (xs @ ys) (mk_Trueprop_eq (HOLogic.mk_eq (xctr, yctr), Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_eq) xs ys)))]; in @{map 4} mk_goal xctrs yctrs xss yss end; val half_distinct_goalss = let fun mk_goal ((xs, xc), (xs', xc')) = fold_rev Logic.all (xs @ xs') (HOLogic.mk_Trueprop (HOLogic.mk_not (HOLogic.mk_eq (xc, xc')))); in map (map mk_goal) (mk_half_pairss (`I (xss ~~ xctrs))) end; val goalss = [exhaust_goal] :: inject_goalss @ half_distinct_goalss; fun after_qed ([exhaust_thm] :: thmss) lthy = let val ((((((((u, u'), (xss, xss')), fs), gs), h), v), p), _) = lthy |> add_bindings |> yield_singleton (apfst (op ~~) oo mk_Frees' fc_b_name) fcT ||>> mk_Freess' "x" ctr_Tss ||>> mk_Frees "f" case_Ts ||>> mk_Frees "g" case_Ts ||>> yield_singleton (mk_Frees "h") (B --> C) ||>> yield_singleton (mk_Frees (fc_b_name ^ "'")) fcT ||>> yield_singleton (mk_Frees "P") HOLogic.boolT; val xfs = map2 (curry Term.list_comb) fs xss; val xgs = map2 (curry Term.list_comb) gs xss; val fcase = Term.list_comb (casex, fs); val ufcase = fcase $ u; val vfcase = fcase $ v; val eta_fcase = Term.list_comb (casex, eta_fs); val eta_gcase = Term.list_comb (casex, eta_gs); val eta_ufcase = eta_fcase $ u; val eta_vgcase = eta_gcase $ v; fun mk_uu_eq () = HOLogic.mk_eq (u, u); val uv_eq = mk_Trueprop_eq (u, v); val ((inject_thms, inject_thmss), half_distinct_thmss) = chop n thmss |>> `flat; val rho_As = map (fn (T, U) => (dest_TVar T, Thm.ctyp_of lthy U)) (map Logic.varifyT_global As ~~ As); fun inst_thm t thm = Thm.instantiate' [] [SOME (Thm.cterm_of lthy t)] (Thm.instantiate (rho_As, []) (Drule.zero_var_indexes thm)); val uexhaust_thm = inst_thm u exhaust_thm; val exhaust_cases = map base_name_of_ctr ctrs; val other_half_distinct_thmss = map (map (fn thm => thm RS not_sym)) half_distinct_thmss; val (distinct_thms, (distinct_thmsss', distinct_thmsss)) = join_halves n half_distinct_thmss other_half_distinct_thmss ||> `transpose; val nchotomy_thm = let val goal = HOLogic.mk_Trueprop (HOLogic.mk_all (fst u', snd u', Library.foldr1 HOLogic.mk_disj exist_xs_u_eq_ctrs)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_nchotomy_tac ctxt n exhaust_thm) |> Thm.close_derivation \<^here> end; val case_thms = let val goals = @{map 3} (fn xctr => fn xf => fn xs => fold_rev Logic.all (fs @ xs) (mk_Trueprop_eq (fcase $ xctr, xf))) xctrs xfs xss; in @{map 4} (fn k => fn goal => fn injects => fn distinctss => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => mk_case_tac ctxt n k case_def injects distinctss) |> Thm.close_derivation \<^here>) ks goals inject_thmss distinct_thmsss end; val (case_cong_thm, case_cong_weak_thm) = let fun mk_prem xctr xs xf xg = fold_rev Logic.all xs (Logic.mk_implies (mk_Trueprop_eq (v, xctr), mk_Trueprop_eq (xf, xg))); val goal = Logic.list_implies (uv_eq :: @{map 4} mk_prem xctrs xss xfs xgs, mk_Trueprop_eq (eta_ufcase, eta_vgcase)); val weak_goal = Logic.mk_implies (uv_eq, mk_Trueprop_eq (ufcase, vfcase)); val vars = Variable.add_free_names lthy goal []; val weak_vars = Variable.add_free_names lthy weak_goal []; in (Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_case_cong_tac ctxt uexhaust_thm case_thms), Goal.prove_sorry lthy weak_vars [] weak_goal (fn {context = ctxt, prems = _} => etac ctxt arg_cong 1)) |> apply2 (Thm.close_derivation \<^here>) end; val split_lhs = q $ ufcase; fun mk_split_conjunct xctr xs f_xs = list_all_free xs (HOLogic.mk_imp (HOLogic.mk_eq (u, xctr), q $ f_xs)); fun mk_split_disjunct xctr xs f_xs = list_exists_free xs (HOLogic.mk_conj (HOLogic.mk_eq (u, xctr), HOLogic.mk_not (q $ f_xs))); fun mk_split_goal xctrs xss xfs = mk_Trueprop_eq (split_lhs, Library.foldr1 HOLogic.mk_conj (@{map 3} mk_split_conjunct xctrs xss xfs)); fun mk_split_asm_goal xctrs xss xfs = mk_Trueprop_eq (split_lhs, HOLogic.mk_not (Library.foldr1 HOLogic.mk_disj (@{map 3} mk_split_disjunct xctrs xss xfs))); fun prove_split selss goal = Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_split_tac ctxt uexhaust_thm case_thms selss inject_thmss distinct_thmsss)) |> Thm.close_derivation \<^here>; fun prove_split_asm asm_goal split_thm = Variable.add_free_names lthy asm_goal [] |> (fn vars => Goal.prove_sorry lthy vars [] asm_goal (fn {context = ctxt, ...} => mk_split_asm_tac ctxt split_thm)) |> Thm.close_derivation \<^here>; val (split_thm, split_asm_thm) = let val goal = mk_split_goal xctrs xss xfs; val asm_goal = mk_split_asm_goal xctrs xss xfs; val thm = prove_split (replicate n []) goal; val asm_thm = prove_split_asm asm_goal thm; in (thm, asm_thm) end; val (sel_defs, all_sel_thms, sel_thmss, nontriv_disc_defs, disc_thmss, nontriv_disc_thmss, discI_thms, nontriv_discI_thms, distinct_disc_thms, distinct_disc_thmsss, exhaust_disc_thms, exhaust_sel_thms, all_collapse_thms, safe_collapse_thms, expand_thms, split_sel_thms, split_sel_asm_thms, case_eq_if_thms, disc_eq_case_thms) = if no_discs_sels then ([], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], []) else let val udiscs = map (rapp u) discs; val uselss = map (map (rapp u)) selss; val usel_ctrs = map2 (curry Term.list_comb) ctrs uselss; val usel_fs = map2 (curry Term.list_comb) fs uselss; val vdiscs = map (rapp v) discs; val vselss = map (map (rapp v)) selss; fun make_sel_thm xs' case_thm sel_def = zero_var_indexes (Variable.gen_all lthy (Drule.rename_bvars' (map (SOME o fst) xs') (Drule.forall_intr_vars (case_thm RS (sel_def RS trans))))); val sel_thmss = @{map 3} (map oo make_sel_thm) xss' case_thms sel_defss; fun has_undefined_rhs thm = (case snd (HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of thm))) of Const (\<^const_name>\undefined\, _) => true | _ => false); val all_sel_thms = (if all_sels_distinct andalso null sel_default_eqs then flat sel_thmss else map_product (fn s => fn (xs', c) => make_sel_thm xs' c s) sel_defs (xss' ~~ case_thms)) |> filter_out has_undefined_rhs; fun mk_unique_disc_def () = let val m = the_single ms; val goal = mk_Trueprop_eq (mk_uu_eq (), the_single exist_xs_u_eq_ctrs); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_unique_disc_def_tac ctxt m uexhaust_thm) |> Thm.close_derivation \<^here> end; fun mk_alternate_disc_def k = let val goal = mk_Trueprop_eq (alternate_disc_lhs (K (nth udiscs)) (3 - k), nth exist_xs_u_eq_ctrs (k - 1)); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_alternate_disc_def_tac ctxt k (nth disc_defs (2 - k)) (nth distinct_thms (2 - k)) uexhaust_thm) |> Thm.close_derivation \<^here> end; val has_alternate_disc_def = exists (fn def => Thm.eq_thm_prop (def, alternate_disc_no_def)) disc_defs; val nontriv_disc_defs = disc_defs |> filter_out (member Thm.eq_thm_prop [unique_disc_no_def, alternate_disc_no_def, refl]); val disc_defs' = map2 (fn k => fn def => if Thm.eq_thm_prop (def, unique_disc_no_def) then mk_unique_disc_def () else if Thm.eq_thm_prop (def, alternate_disc_no_def) then mk_alternate_disc_def k else def) ks disc_defs; val discD_thms = map (fn def => def RS iffD1) disc_defs'; val discI_thms = map2 (fn m => fn def => funpow m (fn thm => exI RS thm) (def RS iffD2)) ms disc_defs'; val not_discI_thms = map2 (fn m => fn def => funpow m (fn thm => allI RS thm) (unfold_thms lthy @{thms not_ex} (def RS @{thm ssubst[of _ _ Not]}))) ms disc_defs'; val (disc_thmss', disc_thmss) = let fun mk_thm discI _ [] = refl RS discI | mk_thm _ not_discI [distinct] = distinct RS not_discI; fun mk_thms discI not_discI distinctss = map (mk_thm discI not_discI) distinctss; in @{map 3} mk_thms discI_thms not_discI_thms distinct_thmsss' |> `transpose end; val nontriv_disc_thmss = map2 (fn b => if is_disc_binding_valid b then I else K []) disc_bindings disc_thmss; fun is_discI_triv b = (n = 1 andalso Binding.is_empty b) orelse Binding.eq_name (b, equal_binding); val nontriv_discI_thms = flat (map2 (fn b => if is_discI_triv b then K [] else single) disc_bindings discI_thms); val (distinct_disc_thms, (distinct_disc_thmsss', distinct_disc_thmsss)) = let fun mk_goal [] = [] | mk_goal [((_, udisc), (_, udisc'))] = [Logic.all u (Logic.mk_implies (HOLogic.mk_Trueprop udisc, HOLogic.mk_Trueprop (HOLogic.mk_not udisc')))]; fun prove tac goal = Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => tac ctxt) |> Thm.close_derivation \<^here>; val half_pairss = mk_half_pairss (`I (ms ~~ discD_thms ~~ udiscs)); val half_goalss = map mk_goal half_pairss; val half_thmss = @{map 3} (fn [] => K (K []) | [goal] => fn [(((m, discD), _), _)] => fn disc_thm => [prove (fn ctxt => mk_half_distinct_disc_tac ctxt m discD disc_thm) goal]) half_goalss half_pairss (flat disc_thmss'); val other_half_goalss = map (mk_goal o map swap) half_pairss; val other_half_thmss = map2 (map2 (fn thm => prove (fn ctxt => mk_other_half_distinct_disc_tac ctxt thm))) half_thmss other_half_goalss; in join_halves n half_thmss other_half_thmss ||> `transpose |>> has_alternate_disc_def ? K [] end; val exhaust_disc_thm = let fun mk_prem udisc = mk_imp_p [HOLogic.mk_Trueprop udisc]; val goal = fold_rev Logic.all [p, u] (mk_imp_p (map mk_prem udiscs)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_exhaust_disc_tac ctxt n exhaust_thm discI_thms) |> Thm.close_derivation \<^here> end; val (safe_collapse_thms, all_collapse_thms) = let fun mk_goal m udisc usel_ctr = let val prem = HOLogic.mk_Trueprop udisc; val concl = mk_Trueprop_eq ((usel_ctr, u) |> m = 0 ? swap); in (prem aconv concl, Logic.all u (Logic.mk_implies (prem, concl))) end; val (trivs, goals) = @{map 3} mk_goal ms udiscs usel_ctrs |> split_list; val thms = @{map 5} (fn m => fn discD => fn sel_thms => fn triv => fn goal => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => mk_collapse_tac ctxt m discD sel_thms ORELSE HEADGOAL (assume_tac ctxt)) |> Thm.close_derivation \<^here> |> not triv ? perhaps (try (fn thm => refl RS thm))) ms discD_thms sel_thmss trivs goals; in (map_filter (fn (true, _) => NONE | (false, thm) => SOME thm) (trivs ~~ thms), thms) end; val swapped_all_collapse_thms = map2 (fn m => fn thm => if m = 0 then thm else thm RS sym) ms all_collapse_thms; val exhaust_sel_thm = let fun mk_prem usel_ctr = mk_imp_p [mk_Trueprop_eq (u, usel_ctr)]; val goal = fold_rev Logic.all [p, u] (mk_imp_p (map mk_prem usel_ctrs)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_exhaust_sel_tac ctxt n exhaust_disc_thm swapped_all_collapse_thms) |> Thm.close_derivation \<^here> end; val expand_thm = let fun mk_prems k udisc usels vdisc vsels = (if k = n then [] else [mk_Trueprop_eq (udisc, vdisc)]) @ (if null usels then [] else [Logic.list_implies (if n = 1 then [] else map HOLogic.mk_Trueprop [udisc, vdisc], HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_eq) usels vsels)))]); val goal = Library.foldr Logic.list_implies (@{map 5} mk_prems ks udiscs uselss vdiscs vselss, uv_eq); val uncollapse_thms = map2 (fn thm => fn [] => thm | _ => thm RS sym) all_collapse_thms uselss; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_expand_tac ctxt n ms (inst_thm u exhaust_disc_thm) (inst_thm v exhaust_disc_thm) uncollapse_thms distinct_disc_thmsss distinct_disc_thmsss') |> Thm.close_derivation \<^here> end; val (split_sel_thm, split_sel_asm_thm) = let val zss = map (K []) xss; val goal = mk_split_goal usel_ctrs zss usel_fs; val asm_goal = mk_split_asm_goal usel_ctrs zss usel_fs; val thm = prove_split sel_thmss goal; val asm_thm = prove_split_asm asm_goal thm; in (thm, asm_thm) end; val case_eq_if_thm = let val goal = mk_Trueprop_eq (ufcase, mk_IfN B udiscs usel_fs); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_case_eq_if_tac ctxt n uexhaust_thm case_thms disc_thmss' sel_thmss) |> Thm.close_derivation \<^here> end; val disc_eq_case_thms = let fun const_of_bool b = if b then \<^const>\True\ else \<^const>\False\; fun mk_case_args n = map_index (fn (k, argTs) => fold_rev Term.absdummy argTs (const_of_bool (n = k))) ctr_Tss; val goals = map_index (fn (n, udisc) => mk_Trueprop_eq (udisc, list_comb (casexBool, mk_case_args n) $ u)) udiscs; val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_disc_eq_case_tac ctxt (Thm.cterm_of ctxt u) exhaust_thm (flat nontriv_disc_thmss) distinct_thms case_thms) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end; in (sel_defs, all_sel_thms, sel_thmss, nontriv_disc_defs, disc_thmss, nontriv_disc_thmss, discI_thms, nontriv_discI_thms, distinct_disc_thms, distinct_disc_thmsss, [exhaust_disc_thm], [exhaust_sel_thm], all_collapse_thms, safe_collapse_thms, [expand_thm], [split_sel_thm], [split_sel_asm_thm], [case_eq_if_thm], disc_eq_case_thms) end; val case_distrib_thm = let val args = @{map 2} (fn f => fn argTs => let val (args, _) = mk_Frees "x" argTs lthy in fold_rev Term.lambda args (h $ list_comb (f, args)) end) fs ctr_Tss; val goal = mk_Trueprop_eq (h $ ufcase, list_comb (casexC, args) $ u); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_case_distrib_tac ctxt (Thm.cterm_of ctxt u) exhaust_thm case_thms) |> Thm.close_derivation \<^here> end; val exhaust_case_names_attr = Attrib.internal (K (Rule_Cases.case_names exhaust_cases)); val cases_type_attr = Attrib.internal (K (Induct.cases_type fcT_name)); val nontriv_disc_eq_thmss = map (map (fn th => th RS @{thm eq_False[THEN iffD2]} handle THM _ => th RS @{thm eq_True[THEN iffD2]})) nontriv_disc_thmss; val anonymous_notes = [(map (fn th => th RS notE) distinct_thms, safe_elim_attrs), (flat nontriv_disc_eq_thmss, nitpicksimp_attrs)] |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])])); val notes = [(caseN, case_thms, nitpicksimp_attrs @ simp_attrs), (case_congN, [case_cong_thm], []), (case_cong_weak_thmsN, [case_cong_weak_thm], cong_attrs), (case_distribN, [case_distrib_thm], []), (case_eq_ifN, case_eq_if_thms, []), (collapseN, safe_collapse_thms, if ms = [0] then [] else simp_attrs), (discN, flat nontriv_disc_thmss, simp_attrs), (disc_eq_caseN, disc_eq_case_thms, []), (discIN, nontriv_discI_thms, []), (distinctN, distinct_thms, simp_attrs @ inductsimp_attrs), (distinct_discN, distinct_disc_thms, dest_attrs), (exhaustN, [exhaust_thm], [exhaust_case_names_attr, cases_type_attr]), (exhaust_discN, exhaust_disc_thms, [exhaust_case_names_attr]), (exhaust_selN, exhaust_sel_thms, [exhaust_case_names_attr]), (expandN, expand_thms, []), (injectN, inject_thms, iff_attrs @ inductsimp_attrs), (nchotomyN, [nchotomy_thm], []), (selN, all_sel_thms, nitpicksimp_attrs @ simp_attrs), (splitN, [split_thm], []), (split_asmN, [split_asm_thm], []), (split_selN, split_sel_thms, []), (split_sel_asmN, split_sel_asm_thms, []), (split_selsN, split_sel_thms @ split_sel_asm_thms, []), (splitsN, [split_thm, split_asm_thm], [])] |> filter_out (null o #2) |> map (fn (thmN, thms, attrs) => ((qualify true (Binding.name thmN), attrs), [(thms, [])])); val (noted, lthy') = lthy |> Spec_Rules.add Binding.empty Spec_Rules.equational [casex] case_thms |> fold (uncurry (Spec_Rules.add Binding.empty Spec_Rules.equational)) (AList.group (eq_list (op aconv)) (map (`(single o lhs_head_of)) all_sel_thms)) |> fold (uncurry (Spec_Rules.add Binding.empty Spec_Rules.equational)) (filter_out (null o snd) (map single discs ~~ nontriv_disc_eq_thmss)) |> Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => Case_Translation.register (Morphism.term phi casex) (map (Morphism.term phi) ctrs)) |> plugins code_plugin ? (Code.declare_default_eqns (map (rpair true) (flat nontriv_disc_eq_thmss @ case_thms @ all_sel_thms)) #> Local_Theory.declaration {syntax = false, pervasive = false} (fn phi => Context.mapping (add_ctr_code fcT_name (map (Morphism.typ phi) As) (map (dest_Const o Morphism.term phi) ctrs) (Morphism.fact phi inject_thms) (Morphism.fact phi distinct_thms) (Morphism.fact phi case_thms)) I)) |> Local_Theory.notes (anonymous_notes @ notes) (* for "datatype_realizer.ML": *) |>> name_noted_thms fcT_name exhaustN; val ctr_sugar = {kind = kind, T = fcT, ctrs = ctrs, casex = casex, discs = discs, selss = selss, exhaust = exhaust_thm, nchotomy = nchotomy_thm, injects = inject_thms, distincts = distinct_thms, case_thms = case_thms, case_cong = case_cong_thm, case_cong_weak = case_cong_weak_thm, case_distribs = [case_distrib_thm], split = split_thm, split_asm = split_asm_thm, disc_defs = nontriv_disc_defs, disc_thmss = disc_thmss, discIs = discI_thms, disc_eq_cases = disc_eq_case_thms, sel_defs = sel_defs, sel_thmss = sel_thmss, distinct_discsss = distinct_disc_thmsss, exhaust_discs = exhaust_disc_thms, exhaust_sels = exhaust_sel_thms, collapses = all_collapse_thms, expands = expand_thms, split_sels = split_sel_thms, split_sel_asms = split_sel_asm_thms, case_eq_ifs = case_eq_if_thms} |> morph_ctr_sugar (substitute_noted_thm noted); in (ctr_sugar, lthy' |> register_ctr_sugar plugins ctr_sugar) end; in (goalss, after_qed, lthy) end; fun free_constructors kind tacss = (fn (goalss, after_qed, lthy) => map2 (map2 (Thm.close_derivation \<^here> oo Goal.prove_sorry lthy [] [])) goalss tacss |> (fn thms => after_qed thms lthy)) oo prepare_free_constructors kind (K I) (K I); fun free_constructors_cmd kind = (fn (goalss, after_qed, lthy) => Proof.theorem NONE (snd oo after_qed) (map (map (rpair [])) goalss) lthy) oo prepare_free_constructors kind Plugin_Name.make_filter Syntax.read_term; val parse_bound_term = Parse.binding --| \<^keyword>\:\ -- Parse.term; type ctr_options = Plugin_Name.filter * bool; type ctr_options_cmd = (Proof.context -> Plugin_Name.filter) * bool; val default_ctr_options : ctr_options = (Plugin_Name.default_filter, false); val default_ctr_options_cmd : ctr_options_cmd = (K Plugin_Name.default_filter, false); val parse_ctr_options = Scan.optional (\<^keyword>\(\ |-- Parse.list1 (Plugin_Name.parse_filter >> (apfst o K) || Parse.reserved "discs_sels" >> (apsnd o K o K true)) --| \<^keyword>\)\ >> (fn fs => fold I fs default_ctr_options_cmd)) default_ctr_options_cmd; fun parse_ctr_spec parse_ctr parse_arg = parse_opt_binding_colon -- parse_ctr -- Scan.repeat parse_arg; val parse_ctr_specs = Parse.enum1 "|" (parse_ctr_spec Parse.term Parse.binding); val parse_sel_default_eqs = Scan.optional (\<^keyword>\where\ |-- Parse.enum1 "|" Parse.prop) []; val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\free_constructors\ "register an existing freely generated type's constructors" (parse_ctr_options -- Parse.binding --| \<^keyword>\for\ -- parse_ctr_specs -- parse_sel_default_eqs >> free_constructors_cmd Unknown); + + +(** external views **) + +(* document antiquotations *) + +local + +fun antiquote_setup binding co = + Thy_Output.antiquotation_pretty_source_embedded binding + ((Scan.ahead (Scan.lift Parse.not_eof) >> Token.pos_of) -- + Args.type_name {proper = true, strict = true}) + (fn ctxt => fn (pos, type_name) => + let + fun err () = + error ("Bad " ^ Binding.name_of binding ^ ": " ^ quote type_name ^ Position.here pos); + in + (case ctr_sugar_of ctxt type_name of + NONE => err () + | SOME {kind, T = T0, ctrs = ctrs0, ...} => + let + val _ = if co = (kind = Codatatype) then () else err (); + + val T = Logic.unvarifyT_global T0; + val ctrs = map Logic.unvarify_global ctrs0; + + val pretty_typ_bracket = Syntax.pretty_typ (Config.put pretty_priority 1001 ctxt); + fun pretty_ctr ctr = + Pretty.block (Pretty.breaks (Syntax.pretty_term ctxt ctr :: + map pretty_typ_bracket (binder_types (fastype_of ctr)))); + in + Pretty.block (Pretty.keyword1 (Binding.name_of binding) :: Pretty.brk 1 :: + Syntax.pretty_typ ctxt T :: Pretty.str " =" :: Pretty.brk 1 :: + flat (separate [Pretty.brk 1, Pretty.str "| "] (map (single o pretty_ctr) ctrs))) + end) + end); + +in + +val _ = + Theory.setup + (antiquote_setup \<^binding>\datatype\ false #> + antiquote_setup \<^binding>\codatatype\ true); + end; + + +(* theory export *) + +val _ = Export_Theory.setup_presentation (fn context => fn thy => + let + val parents = map (Data.get o Context.Theory) (Theory.parents_of thy); + val datatypes = + (Data.get (Context.Theory thy), []) |-> Symtab.fold + (fn (name, (pos, {kind, T, ctrs, ...})) => + if kind = Record orelse exists (fn tab => Symtab.defined tab name) parents then I + else + let + val pos_properties = Thy_Info.adjust_pos_properties context pos; + val typ = Logic.unvarifyT_global T; + val constrs = map Logic.unvarify_global ctrs; + val typargs = rev (fold Term.add_tfrees (Logic.mk_type typ :: constrs) []); + val constructors = map (fn t => (t, Term.type_of t)) constrs; + in + cons (pos_properties, (name, (kind = Codatatype, (typargs, (typ, constructors))))) + end); + in + if null datatypes then () + else + Export_Theory.export_body thy "datatypes" + let open XML.Encode Term_XML.Encode in + list (pair properties (pair string (pair bool (pair (list (pair string sort)) + (pair typ (list (pair (term (Sign.consts_of thy)) typ))))))) datatypes + end + end); + +end; diff --git a/src/Pure/Concurrent/par_exn.ML b/src/Pure/Concurrent/par_exn.ML --- a/src/Pure/Concurrent/par_exn.ML +++ b/src/Pure/Concurrent/par_exn.ML @@ -1,79 +1,79 @@ (* Title: Pure/Concurrent/par_exn.ML Author: Makarius Parallel exceptions as flattened results from acyclic graph of evaluations. Interrupt counts as neutral element. *) signature PAR_EXN = sig - val identify: Properties.entry list -> exn -> exn + val identify: Properties.T -> exn -> exn val the_serial: exn -> int val make: exn list -> exn val dest: exn -> exn list option val is_interrupted: 'a Exn.result list -> bool val release_all: 'a Exn.result list -> 'a list val release_first: 'a Exn.result list -> 'a list end; structure Par_Exn: PAR_EXN = struct (* identification via serial numbers -- NOT portable! *) fun identify default_props exn = let val props = Exn_Properties.get exn; val update_serial = if Properties.defined props Markup.serialN then [] else [(Markup.serialN, serial_string ())]; val update_props = filter_out (Properties.defined props o #1) default_props; in Exn_Properties.update (update_serial @ update_props) exn end; fun the_serial exn = Value.parse_int (the (Properties.get (Exn_Properties.get exn) Markup.serialN)); val exn_ord = rev_order o int_ord o apply2 the_serial; (* parallel exceptions *) exception Par_Exn of exn list; (*non-empty list with unique identified elements sorted by exn_ord; no occurrences of Par_Exn or Exn.Interrupt*) fun par_exns (Par_Exn exns) = exns | par_exns exn = if Exn.is_interrupt exn then [] else [identify [] exn]; fun make exns = let val exnss = map par_exns exns; val exns' = Ord_List.unions exn_ord exnss handle Option.Option => flat exnss; in if null exns' then Exn.Interrupt else Par_Exn exns' end; fun dest (Par_Exn exns) = SOME exns | dest exn = if Exn.is_interrupt exn then SOME [] else NONE; (* parallel results *) fun is_interrupted results = exists (fn Exn.Exn _ => true | _ => false) results andalso Exn.is_interrupt (make (map_filter Exn.get_exn results)); fun release_all results = if forall (fn Exn.Res _ => true | _ => false) results then map Exn.release results else raise make (map_filter Exn.get_exn results); fun plain_exn (Exn.Res _) = NONE | plain_exn (Exn.Exn (Par_Exn _)) = NONE | plain_exn (Exn.Exn exn) = if Exn.is_interrupt exn then NONE else SOME exn; fun release_first results = (case get_first plain_exn results of NONE => release_all results | SOME exn => Exn.reraise exn); end; diff --git a/src/Pure/ML/exn_properties.ML b/src/Pure/ML/exn_properties.ML --- a/src/Pure/ML/exn_properties.ML +++ b/src/Pure/ML/exn_properties.ML @@ -1,71 +1,71 @@ (* Title: Pure/ML/exn_properties.ML Author: Makarius Exception properties. *) signature EXN_PROPERTIES = sig val position_of_polyml_location: ML_Compiler0.polyml_location -> Position.T val position: exn -> Position.T val get: exn -> Properties.T - val update: Properties.entry list -> exn -> exn + val update: Properties.T -> exn -> exn end; structure Exn_Properties: EXN_PROPERTIES = struct (* source locations *) fun props_of_polyml_location (loc: ML_Compiler0.polyml_location) = (case YXML.parse_body (#file loc) of [] => [] | [XML.Text file] => if file = "Standard Basis" then [] else [(Markup.fileN, ML_System.standard_path file)] | body => XML.Decode.properties body); fun position_of_polyml_location loc = Position.make {line = FixedInt.toInt (#startLine loc), offset = FixedInt.toInt (#startPosition loc), end_offset = FixedInt.toInt (#endPosition loc), props = props_of_polyml_location loc}; fun position exn = (case Exn.polyml_location exn of NONE => Position.none | SOME loc => position_of_polyml_location loc); (* exception properties *) fun get exn = (case Exn.polyml_location exn of NONE => [] | SOME loc => props_of_polyml_location loc); fun update entries exn = if Exn.is_interrupt exn then exn else let val loc = the_default {file = "", startLine = 0, endLine = 0, startPosition = 0, endPosition = 0} (Exn.polyml_location exn); val props = props_of_polyml_location loc; val props' = fold Properties.put entries props; in if props = props' then exn else let val loc' = {file = YXML.string_of_body (XML.Encode.properties props'), startLine = #startLine loc, endLine = #endLine loc, startPosition = #startPosition loc, endPosition = #endPosition loc}; in Thread_Attributes.uninterruptible (fn _ => fn () => PolyML.raiseWithLocation (exn, loc') handle exn' => exn') () end end; end; diff --git a/src/Pure/Thy/export_theory.ML b/src/Pure/Thy/export_theory.ML --- a/src/Pure/Thy/export_theory.ML +++ b/src/Pure/Thy/export_theory.ML @@ -1,430 +1,429 @@ (* Title: Pure/Thy/export_theory.ML Author: Makarius Export foundational theory content and locale/class structure. *) signature EXPORT_THEORY = sig val setup_presentation: (Thy_Info.presentation_context -> theory -> unit) -> unit val export_body: theory -> string -> XML.body -> unit end; structure Export_Theory: EXPORT_THEORY = struct (* approximative syntax *) val get_syntax = Syntax.get_approx o Proof_Context.syn_of; fun get_syntax_type ctxt = get_syntax ctxt o Lexicon.mark_type; fun get_syntax_const ctxt = get_syntax ctxt o Lexicon.mark_const; fun get_syntax_fixed ctxt = get_syntax ctxt o Lexicon.mark_fixed; fun get_syntax_param ctxt loc x = let val thy = Proof_Context.theory_of ctxt in if Class.is_class thy loc then (case AList.lookup (op =) (Class.these_params thy [loc]) x of NONE => NONE | SOME (_, (c, _)) => get_syntax_const ctxt c) else get_syntax_fixed ctxt x end; val encode_syntax = XML.Encode.variant [fn NONE => ([], []), fn SOME (Syntax.Prefix delim) => ([delim], []), fn SOME (Syntax.Infix {assoc, delim, pri}) => let val ass = (case assoc of Printer.No_Assoc => 0 | Printer.Left_Assoc => 1 | Printer.Right_Assoc => 2); open XML.Encode Term_XML.Encode; in ([], triple int string int (ass, delim, pri)) end]; (* free variables: not declared in the context *) val is_free = not oo Name.is_declared; fun add_frees used = fold_aterms (fn Free (x, T) => is_free used x ? insert (op =) (x, T) | _ => I); fun add_tfrees used = (fold_types o fold_atyps) (fn TFree (a, S) => is_free used a ? insert (op =) (a, S) | _ => I); (* locales *) fun locale_content thy loc = let val ctxt = Locale.init loc thy; val args = Locale.params_of thy loc |> map (fn ((x, T), _) => ((x, T), get_syntax_param ctxt loc x)); val axioms = let val (asm, defs) = Locale.specification_of thy loc; val cprops = map (Thm.cterm_of ctxt) (the_list asm @ defs); val (intro1, intro2) = Locale.intros_of thy loc; val intros_tac = Method.try_intros_tac ctxt (the_list intro1 @ the_list intro2) []; val res = Goal.init (Conjunction.mk_conjunction_balanced cprops) |> (ALLGOALS Goal.conjunction_tac THEN intros_tac) |> try Seq.hd; in (case res of SOME goal => Thm.prems_of goal | NONE => raise Fail ("Cannot unfold locale " ^ quote loc)) end; val typargs = rev (fold Term.add_tfrees (map (Free o #1) args @ axioms) []); in {typargs = typargs, args = args, axioms = axioms} end; fun get_locales thy = Locale.get_locales thy |> map_filter (fn loc => if Experiment.is_experiment thy loc then NONE else SOME (loc, ())); fun get_dependencies prev_thys thy = Locale.dest_dependencies prev_thys thy |> map_filter (fn dep => if Experiment.is_experiment thy (#source dep) orelse Experiment.is_experiment thy (#target dep) then NONE else let val (type_params, params) = Locale.parameters_of thy (#source dep); val typargs = fold (Term.add_tfreesT o #2 o #1) params type_params; val substT = typargs |> map_filter (fn v => let val T = TFree v; val T' = Morphism.typ (#morphism dep) T; in if T = T' then NONE else SOME (v, T') end); val subst = params |> map_filter (fn (v, _) => let val t = Free v; val t' = Morphism.term (#morphism dep) t; in if t aconv t' then NONE else SOME (v, t') end); in SOME (dep, (substT, subst)) end); (* general setup *) fun setup_presentation f = Theory.setup (Thy_Info.add_presentation (fn context => fn thy => if Options.bool (#options context) "export_theory" then f context thy else ())); fun export_body thy name body = if XML.is_empty_body body then () else Export.export thy (Path.binding0 (Path.make ["theory", name])) body; (* presentation *) -val _ = setup_presentation (fn {adjust_pos, ...} => fn thy => +val _ = setup_presentation (fn context => fn thy => let val parents = Theory.parents_of thy; val rep_tsig = Type.rep_tsig (Sign.tsig_of thy); val thy_ctxt = Proof_Context.init_global thy; + val pos_properties = Thy_Info.adjust_pos_properties context; + (* spec rules *) - fun position_properties pos = - Position.offset_properties_of (adjust_pos pos) @ Position.id_properties_of pos; - fun spec_rule_content {pos, name, rough_classification, terms, rules} = let val spec = terms @ map Thm.plain_prop_of rules |> Term_Subst.zero_var_indexes_list |> map Logic.unvarify_global; in - {props = position_properties pos, + {props = pos_properties pos, name = name, rough_classification = rough_classification, typargs = rev (fold Term.add_tfrees spec []), args = rev (fold Term.add_frees spec []), terms = map (fn t => (t, Term.type_of t)) (take (length terms) spec), rules = drop (length terms) spec} end; (* entities *) fun make_entity_markup name xname pos serial = - let val props = position_properties pos @ Markup.serial_properties serial; + let val props = pos_properties pos @ Markup.serial_properties serial; in (Markup.entityN, (Markup.nameN, name) :: (Markup.xnameN, xname) :: props) end; fun entity_markup space name = let val xname = Name_Space.extern_shortest thy_ctxt space name; val {serial, pos, ...} = Name_Space.the_entry space name; in make_entity_markup name xname pos serial end; fun export_entities export_name export get_space decls = let val parent_spaces = map get_space parents; val space = get_space thy; in (decls, []) |-> fold (fn (name, decl) => if exists (fn space => Name_Space.declared space name) parent_spaces then I else (case export name decl of NONE => I | SOME body => cons (#serial (Name_Space.the_entry space name), XML.Elem (entity_markup space name, body)))) |> sort (int_ord o apply2 #1) |> map #2 |> export_body thy export_name end; (* types *) val encode_type = let open XML.Encode Term_XML.Encode in triple encode_syntax (list string) (option typ) end; fun export_type c (Type.LogicalType n) = SOME (encode_type (get_syntax_type thy_ctxt c, Name.invent Name.context Name.aT n, NONE)) | export_type c (Type.Abbreviation (args, U, false)) = SOME (encode_type (get_syntax_type thy_ctxt c, args, SOME U)) | export_type _ _ = NONE; val _ = export_entities "types" export_type Sign.type_space (Name_Space.dest_table (#types rep_tsig)); (* consts *) val consts = Sign.consts_of thy; val encode_term = Term_XML.Encode.term consts; val encode_const = let open XML.Encode Term_XML.Encode in pair encode_syntax (pair (list string) (pair typ (pair (option encode_term) bool))) end; fun export_const c (T, abbrev) = let val syntax = get_syntax_const thy_ctxt c; val U = Logic.unvarifyT_global T; val U0 = Type.strip_sorts U; val abbrev' = abbrev |> Option.map (Proofterm.standard_vars_term Name.context #> map_types Type.strip_sorts); val args = map (#1 o dest_TFree) (Consts.typargs consts (c, U0)); val propositional = Object_Logic.is_propositional thy_ctxt (Term.body_type U0); in encode_const (syntax, (args, (U0, (abbrev', propositional)))) end; val _ = export_entities "consts" (SOME oo export_const) Sign.const_space (#constants (Consts.dest consts)); (* axioms *) fun standard_prop used extra_shyps raw_prop raw_proof = let val (prop, proof) = Proofterm.standard_vars used (raw_prop, raw_proof); val args = rev (add_frees used prop []); val typargs = rev (add_tfrees used prop []); val used_typargs = fold (Name.declare o #1) typargs used; val sorts = Name.invent used_typargs Name.aT (length extra_shyps) ~~ extra_shyps; in ((sorts @ typargs, args, prop), proof) end; fun standard_prop_of thm = standard_prop Name.context (Thm.extra_shyps thm) (Thm.full_prop_of thm); val encode_prop = let open XML.Encode Term_XML.Encode in triple (list (pair string sort)) (list (pair string typ)) encode_term end; fun encode_axiom used prop = encode_prop (#1 (standard_prop used [] prop NONE)); val _ = export_entities "axioms" (K (SOME o encode_axiom Name.context)) Theory.axiom_space (Theory.all_axioms_of thy); (* theorems and proof terms *) val clean_thm = Thm.check_hyps (Context.Theory thy) #> Thm.strip_shyps; val prep_thm = clean_thm #> Thm.unconstrainT; val lookup_thm_id = Global_Theory.lookup_thm_id thy; fun expand_name thm_id (header: Proofterm.thm_header) = if #serial header = #serial thm_id then "" else (case lookup_thm_id (Proofterm.thm_header_id header) of NONE => "" | SOME thm_name => Thm_Name.print thm_name); fun entity_markup_thm (serial, (name, i)) = let val space = Facts.space_of (Global_Theory.facts_of thy); val xname = Name_Space.extern_shortest thy_ctxt space name; val {pos, ...} = Name_Space.the_entry space name; in make_entity_markup (Thm_Name.print (name, i)) (Thm_Name.print (xname, i)) pos serial end; fun encode_thm thm_id raw_thm = let val deps = map (Thm_Name.print o #2) (Thm_Deps.thm_deps thy [raw_thm]); val thm = prep_thm raw_thm; val proof0 = if Proofterm.export_standard_enabled () then Proof_Syntax.standard_proof_of {full = true, expand_name = SOME o expand_name thm_id} thm else if Proofterm.export_enabled () then Thm.reconstruct_proof_of thm else MinProof; val (prop, SOME proof) = standard_prop_of thm (SOME proof0); val _ = Thm.expose_proofs thy [thm]; in (prop, deps, proof) |> let open XML.Encode Term_XML.Encode; val encode_proof = Proofterm.encode_standard_proof consts; in triple encode_prop (list string) encode_proof end end; fun export_thm (thm_id, thm_name) = let val markup = entity_markup_thm (#serial thm_id, thm_name); val thm = Global_Theory.get_thm_name thy (thm_name, Position.none); in XML.Elem (markup, encode_thm thm_id thm) end; val _ = export_body thy "thms" (map export_thm (Global_Theory.dest_thm_names thy)); (* type classes *) val encode_class = let open XML.Encode Term_XML.Encode in pair (list (pair string typ)) (list (encode_axiom Name.context)) end; fun export_class name = (case try (Axclass.get_info thy) name of NONE => ([], []) | SOME {params, axioms, ...} => (params, map (Thm.plain_prop_of o clean_thm) axioms)) |> encode_class |> SOME; val _ = export_entities "classes" (fn name => fn () => export_class name) Sign.class_space (map (rpair ()) (Graph.keys (Sorts.classes_of (#2 (#classes rep_tsig))))); (* sort algebra *) local val prop = encode_axiom Name.context o Logic.varify_global; val encode_classrel = let open XML.Encode in list (pair prop (pair string string)) end; val encode_arities = let open XML.Encode Term_XML.Encode in list (pair prop (triple string (list sort) string)) end; in val export_classrel = maps (fn (c, cs) => map (pair c) cs) #> map (`Logic.mk_classrel) #> encode_classrel; val export_arities = map (`Logic.mk_arity) #> encode_arities; val {classrel, arities} = Sorts.dest_algebra (map (#2 o #classes o Type.rep_tsig o Sign.tsig_of) parents) (#2 (#classes rep_tsig)); end; val _ = if null classrel then () else export_body thy "classrel" (export_classrel classrel); val _ = if null arities then () else export_body thy "arities" (export_arities arities); (* locales *) fun encode_locale used = let open XML.Encode Term_XML.Encode in triple (list (pair string sort)) (list (pair (pair string typ) encode_syntax)) (list (encode_axiom used)) end; fun export_locale loc = let val {typargs, args, axioms} = locale_content thy loc; val used = fold Name.declare (map #1 typargs @ map (#1 o #1) args) Name.context; in encode_locale used (typargs, args, axioms) end handle ERROR msg => cat_error msg ("The error(s) above occurred in locale " ^ quote (Locale.markup_name thy_ctxt loc)); val _ = export_entities "locales" (fn loc => fn () => SOME (export_locale loc)) Locale.locale_space (get_locales thy); (* locale dependencies *) fun encode_locale_dependency (dep: Locale.locale_dependency, subst) = (#source dep, (#target dep, (#prefix dep, subst))) |> let open XML.Encode Term_XML.Encode; val encode_subst = pair (list (pair (pair string sort) typ)) (list (pair (pair string typ) (term consts))); in pair string (pair string (pair (list (pair string bool)) encode_subst)) end; val _ = get_dependencies parents thy |> map_index (fn (i, dep) => let val xname = string_of_int (i + 1); val name = Long_Name.implode [Context.theory_name thy, xname]; val markup = make_entity_markup name xname (#pos (#1 dep)) (#serial (#1 dep)); val body = encode_locale_dependency dep; in XML.Elem (markup, body) end) |> export_body thy "locale_dependencies"; (* constdefs *) val constdefs = Defs.dest_constdefs (map Theory.defs_of (Theory.parents_of thy)) (Theory.defs_of thy) |> sort_by #1; val encode_constdefs = let open XML.Encode in list (pair string string) end; val _ = if null constdefs then () else export_body thy "constdefs" (encode_constdefs constdefs); (* spec rules *) val encode_specs = let open XML.Encode Term_XML.Encode in list (fn {props, name, rough_classification, typargs, args, terms, rules} => pair properties (pair string (pair Spec_Rules.encode_rough_classification (pair (list (pair string sort)) (pair (list (pair string typ)) (pair (list (pair encode_term typ)) (list encode_term)))))) (props, (name, (rough_classification, (typargs, (args, (terms, rules))))))) end; val _ = (case Spec_Rules.dest_theory thy of [] => () | spec_rules => export_body thy "spec_rules" (encode_specs (map spec_rule_content spec_rules))); (* parents *) val _ = Export.export thy \<^path_binding>\theory/parents\ (XML.Encode.string (cat_lines (map Context.theory_long_name parents))); in () end); end; diff --git a/src/Pure/Thy/export_theory.scala b/src/Pure/Thy/export_theory.scala --- a/src/Pure/Thy/export_theory.scala +++ b/src/Pure/Thy/export_theory.scala @@ -1,745 +1,757 @@ /* Title: Pure/Thy/export_theory.scala Author: Makarius Export foundational theory content. */ package isabelle import scala.collection.immutable.SortedMap object Export_Theory { /** session content **/ sealed case class Session(name: String, theory_graph: Graph[String, Option[Theory]]) { override def toString: String = name def theory(theory_name: String): Option[Theory] = if (theory_graph.defined(theory_name)) theory_graph.get_node(theory_name) else None def theories: List[Theory] = theory_graph.topological_order.flatMap(theory(_)) } def read_session( store: Sessions.Store, sessions_structure: Sessions.Structure, session_name: String, progress: Progress = No_Progress, - types: Boolean = true, - consts: Boolean = true, - axioms: Boolean = true, - thms: Boolean = true, - classes: Boolean = true, - locales: Boolean = true, - locale_dependencies: Boolean = true, - classrel: Boolean = true, - arities: Boolean = true, - constdefs: Boolean = true, - typedefs: Boolean = true, - spec_rules: Boolean = true, cache: Term.Cache = Term.make_cache()): Session = { val thys = sessions_structure.build_requirements(List(session_name)).flatMap(session => using(store.open_database(session))(db => { db.transaction { for (theory <- Export.read_theory_names(db, session)) yield { progress.echo("Reading theory " + theory) read_theory(Export.Provider.database(db, session, theory), - session, theory, types = types, consts = consts, - axioms = axioms, thms = thms, classes = classes, locales = locales, - locale_dependencies = locale_dependencies, classrel = classrel, arities = arities, - constdefs = constdefs, typedefs = typedefs, - spec_rules = spec_rules, cache = Some(cache)) + session, theory, cache = Some(cache)) } } })) val graph0 = (Graph.string[Option[Theory]] /: thys) { case (g, thy) => g.default_node(thy.name, Some(thy)) } val graph1 = (graph0 /: thys) { case (g0, thy) => (g0 /: thy.parents) { case (g1, parent) => g1.default_node(parent, None).add_edge_acyclic(parent, thy.name) } } Session(session_name, graph1) } /** theory content **/ val export_prefix: String = "theory/" val export_prefix_proofs: String = "proofs/" sealed case class Theory(name: String, parents: List[String], types: List[Type], consts: List[Const], axioms: List[Axiom], thms: List[Thm], classes: List[Class], locales: List[Locale], locale_dependencies: List[Locale_Dependency], classrel: List[Classrel], arities: List[Arity], constdefs: List[Constdef], typedefs: List[Typedef], + datatypes: List[Datatype], spec_rules: List[Spec_Rule]) { override def toString: String = name def entity_iterator: Iterator[Entity] = types.iterator.map(_.entity) ++ consts.iterator.map(_.entity) ++ axioms.iterator.map(_.entity) ++ thms.iterator.map(_.entity) ++ classes.iterator.map(_.entity) ++ locales.iterator.map(_.entity) ++ locale_dependencies.iterator.map(_.entity) def cache(cache: Term.Cache): Theory = Theory(cache.string(name), parents.map(cache.string(_)), types.map(_.cache(cache)), consts.map(_.cache(cache)), axioms.map(_.cache(cache)), thms.map(_.cache(cache)), classes.map(_.cache(cache)), locales.map(_.cache(cache)), locale_dependencies.map(_.cache(cache)), classrel.map(_.cache(cache)), arities.map(_.cache(cache)), constdefs.map(_.cache(cache)), typedefs.map(_.cache(cache)), + datatypes.map(_.cache(cache)), spec_rules.map(_.cache(cache))) } def read_theory(provider: Export.Provider, session_name: String, theory_name: String, - types: Boolean = true, - consts: Boolean = true, - axioms: Boolean = true, - thms: Boolean = true, - classes: Boolean = true, - locales: Boolean = true, - locale_dependencies: Boolean = true, - classrel: Boolean = true, - arities: Boolean = true, - constdefs: Boolean = true, - typedefs: Boolean = true, - spec_rules: Boolean = true, cache: Option[Term.Cache] = None): Theory = { val parents = if (theory_name == Thy_Header.PURE) Nil else { provider(export_prefix + "parents") match { case Some(entry) => split_lines(entry.uncompressed().text) case None => error("Missing theory export in session " + quote(session_name) + ": " + quote(theory_name)) } } val theory = Theory(theory_name, parents, - if (types) read_types(provider) else Nil, - if (consts) read_consts(provider) else Nil, - if (axioms) read_axioms(provider) else Nil, - if (thms) read_thms(provider) else Nil, - if (classes) read_classes(provider) else Nil, - if (locales) read_locales(provider) else Nil, - if (locale_dependencies) read_locale_dependencies(provider) else Nil, - if (classrel) read_classrel(provider) else Nil, - if (arities) read_arities(provider) else Nil, - if (constdefs) read_constdefs(provider) else Nil, - if (typedefs) read_typedefs(provider) else Nil, - if (spec_rules) read_spec_rules(provider) else Nil) + read_types(provider), + read_consts(provider), + read_axioms(provider), + read_thms(provider), + read_classes(provider), + read_locales(provider), + read_locale_dependencies(provider), + read_classrel(provider), + read_arities(provider), + read_constdefs(provider), + read_typedefs(provider), + read_datatypes(provider), + read_spec_rules(provider)) if (cache.isDefined) theory.cache(cache.get) else theory } def read_pure[A](store: Sessions.Store, read: (Export.Provider, String, String) => A): A = { val session_name = Thy_Header.PURE val theory_name = Thy_Header.PURE using(store.open_database(session_name))(db => { db.transaction { read(Export.Provider.database(db, session_name, theory_name), session_name, theory_name) } }) } def read_pure_theory(store: Sessions.Store, cache: Option[Term.Cache] = None): Theory = read_pure(store, read_theory(_, _, _, cache = cache)) def read_pure_proof( store: Sessions.Store, id: Thm_Id, cache: Option[Term.Cache] = None): Option[Proof] = read_pure(store, (provider, _, _) => read_proof(provider, id, cache = cache)) /* entities */ object Kind extends Enumeration { val TYPE = Value("type") val CONST = Value("const") val AXIOM = Value("axiom") val THM = Value("thm") val PROOF = Value("proof") val CLASS = Value("class") val LOCALE = Value("locale") val LOCALE_DEPENDENCY = Value("locale_dependency") val DOCUMENT_HEADING = Value("document_heading") val DOCUMENT_TEXT = Value("document_text") val PROOF_TEXT = Value("proof_text") } sealed case class Entity( kind: Kind.Value, name: String, xname: String, pos: Position.T, id: Option[Long], serial: Long) { override def toString: String = kind.toString + " " + quote(name) def cache(cache: Term.Cache): Entity = Entity(kind, cache.string(name), cache.string(xname), cache.position(pos), id, serial) } def decode_entity(kind: Kind.Value, tree: XML.Tree): (Entity, XML.Body) = { def err(): Nothing = throw new XML.XML_Body(List(tree)) tree match { case XML.Elem(Markup(Markup.ENTITY, props), body) => val name = Markup.Name.unapply(props) getOrElse err() val xname = Markup.XName.unapply(props) getOrElse err() val pos = props.filter({ case (a, _) => Markup.POSITION_PROPERTIES(a) && a != Markup.ID }) val id = Position.Id.unapply(props) val serial = Markup.Serial.unapply(props) getOrElse err() (Entity(kind, name, xname, pos, id, serial), body) case _ => err() } } /* approximative syntax */ object Assoc extends Enumeration { val NO_ASSOC, LEFT_ASSOC, RIGHT_ASSOC = Value } sealed abstract class Syntax case object No_Syntax extends Syntax case class Prefix(delim: String) extends Syntax case class Infix(assoc: Assoc.Value, delim: String, pri: Int) extends Syntax def decode_syntax: XML.Decode.T[Syntax] = XML.Decode.variant(List( { case (Nil, Nil) => No_Syntax }, { case (List(delim), Nil) => Prefix(delim) }, { case (Nil, body) => import XML.Decode._ val (ass, delim, pri) = triple(int, string, int)(body) Infix(Assoc(ass), delim, pri) })) /* types */ sealed case class Type( entity: Entity, syntax: Syntax, args: List[String], abbrev: Option[Term.Typ]) { def cache(cache: Term.Cache): Type = Type(entity.cache(cache), syntax, args.map(cache.string(_)), abbrev.map(cache.typ(_))) } def read_types(provider: Export.Provider): List[Type] = provider.uncompressed_yxml(export_prefix + "types").map((tree: XML.Tree) => { val (entity, body) = decode_entity(Kind.TYPE, tree) val (syntax, args, abbrev) = { import XML.Decode._ triple(decode_syntax, list(string), option(Term_XML.Decode.typ))(body) } Type(entity, syntax, args, abbrev) }) /* consts */ sealed case class Const( entity: Entity, syntax: Syntax, typargs: List[String], typ: Term.Typ, abbrev: Option[Term.Term], propositional: Boolean) { def cache(cache: Term.Cache): Const = Const(entity.cache(cache), syntax, typargs.map(cache.string(_)), cache.typ(typ), abbrev.map(cache.term(_)), propositional) } def read_consts(provider: Export.Provider): List[Const] = provider.uncompressed_yxml(export_prefix + "consts").map((tree: XML.Tree) => { val (entity, body) = decode_entity(Kind.CONST, tree) val (syntax, (typargs, (typ, (abbrev, propositional)))) = { import XML.Decode._ pair(decode_syntax, pair(list(string), pair(Term_XML.Decode.typ, pair(option(Term_XML.Decode.term), bool))))(body) } Const(entity, syntax, typargs, typ, abbrev, propositional) }) /* axioms */ sealed case class Prop( typargs: List[(String, Term.Sort)], args: List[(String, Term.Typ)], term: Term.Term) { def cache(cache: Term.Cache): Prop = Prop( typargs.map({ case (name, sort) => (cache.string(name), cache.sort(sort)) }), args.map({ case (name, typ) => (cache.string(name), cache.typ(typ)) }), cache.term(term)) } def decode_prop(body: XML.Body): Prop = { val (typargs, args, t) = { import XML.Decode._ import Term_XML.Decode._ triple(list(pair(string, sort)), list(pair(string, typ)), term)(body) } Prop(typargs, args, t) } sealed case class Axiom(entity: Entity, prop: Prop) { def cache(cache: Term.Cache): Axiom = Axiom(entity.cache(cache), prop.cache(cache)) } def read_axioms(provider: Export.Provider): List[Axiom] = provider.uncompressed_yxml(export_prefix + "axioms").map((tree: XML.Tree) => { val (entity, body) = decode_entity(Kind.AXIOM, tree) val prop = decode_prop(body) Axiom(entity, prop) }) /* theorems */ sealed case class Thm_Id(serial: Long, theory_name: String) { def pure: Boolean = theory_name == Thy_Header.PURE } sealed case class Thm( entity: Entity, prop: Prop, deps: List[String], proof: Term.Proof) { def cache(cache: Term.Cache): Thm = Thm( entity.cache(cache), prop.cache(cache), deps.map(cache.string _), cache.proof(proof)) } def read_thms(provider: Export.Provider): List[Thm] = provider.uncompressed_yxml(export_prefix + "thms").map((tree: XML.Tree) => { val (entity, body) = decode_entity(Kind.THM, tree) val (prop, deps, prf) = { import XML.Decode._ import Term_XML.Decode._ triple(decode_prop, list(string), proof)(body) } Thm(entity, prop, deps, prf) }) sealed case class Proof( typargs: List[(String, Term.Sort)], args: List[(String, Term.Typ)], term: Term.Term, proof: Term.Proof) { def prop: Prop = Prop(typargs, args, term) def cache(cache: Term.Cache): Proof = Proof( typargs.map({ case (name, sort) => (cache.string(name), cache.sort(sort)) }), args.map({ case (name, typ) => (cache.string(name), cache.typ(typ)) }), cache.term(term), cache.proof(proof)) } def read_proof( provider: Export.Provider, id: Thm_Id, cache: Option[Term.Cache] = None): Option[Proof] = { for { entry <- provider.focus(id.theory_name)(export_prefix_proofs + id.serial) } yield { val body = entry.uncompressed_yxml() val (typargs, (args, (prop_body, proof_body))) = { import XML.Decode._ import Term_XML.Decode._ pair(list(pair(string, sort)), pair(list(pair(string, typ)), pair(x => x, x => x)))(body) } val env = args.toMap val prop = Term_XML.Decode.term_env(env)(prop_body) val proof = Term_XML.Decode.proof_env(env)(proof_body) val result = Proof(typargs, args, prop, proof) cache.map(result.cache(_)) getOrElse result } } def read_proof_boxes( store: Sessions.Store, provider: Export.Provider, proof: Term.Proof, suppress: Thm_Id => Boolean = _ => false, cache: Option[Term.Cache] = None): List[(Thm_Id, Proof)] = { var seen = Set.empty[Long] var result = SortedMap.empty[Long, (Thm_Id, Proof)] def boxes(context: Option[(Long, Term.Proof)], prf: Term.Proof) { prf match { case Term.Abst(_, _, p) => boxes(context, p) case Term.AbsP(_, _, p) => boxes(context, p) case Term.Appt(p, _) => boxes(context, p) case Term.AppP(p, q) => boxes(context, p); boxes(context, q) case thm: Term.PThm if !seen(thm.serial) => seen += thm.serial val id = Thm_Id(thm.serial, thm.theory_name) if (!suppress(id)) { val read = if (id.pure) Export_Theory.read_pure_proof(store, id, cache = cache) else Export_Theory.read_proof(provider, id, cache = cache) read match { case Some(p) => result += (thm.serial -> (id -> p)) boxes(Some((thm.serial, p.proof)), p.proof) case None => error("Missing proof " + thm.serial + " (theory " + quote (thm.theory_name) + ")" + (context match { case None => "" case Some((i, p)) => " in proof " + i + ":\n" + p })) } } case _ => } } boxes(None, proof) result.iterator.map(_._2).toList } /* type classes */ sealed case class Class( entity: Entity, params: List[(String, Term.Typ)], axioms: List[Prop]) { def cache(cache: Term.Cache): Class = Class(entity.cache(cache), params.map({ case (name, typ) => (cache.string(name), cache.typ(typ)) }), axioms.map(_.cache(cache))) } def read_classes(provider: Export.Provider): List[Class] = provider.uncompressed_yxml(export_prefix + "classes").map((tree: XML.Tree) => { val (entity, body) = decode_entity(Kind.CLASS, tree) val (params, axioms) = { import XML.Decode._ import Term_XML.Decode._ pair(list(pair(string, typ)), list(decode_prop))(body) } Class(entity, params, axioms) }) /* locales */ sealed case class Locale( entity: Entity, typargs: List[(String, Term.Sort)], args: List[((String, Term.Typ), Syntax)], axioms: List[Prop]) { def cache(cache: Term.Cache): Locale = Locale(entity.cache(cache), typargs.map({ case (name, sort) => (cache.string(name), cache.sort(sort)) }), args.map({ case ((name, typ), syntax) => ((cache.string(name), cache.typ(typ)), syntax) }), axioms.map(_.cache(cache))) } def read_locales(provider: Export.Provider): List[Locale] = provider.uncompressed_yxml(export_prefix + "locales").map((tree: XML.Tree) => { val (entity, body) = decode_entity(Kind.LOCALE, tree) val (typargs, args, axioms) = { import XML.Decode._ import Term_XML.Decode._ triple(list(pair(string, sort)), list(pair(pair(string, typ), decode_syntax)), list(decode_prop))(body) } Locale(entity, typargs, args, axioms) }) /* locale dependencies */ sealed case class Locale_Dependency( entity: Entity, source: String, target: String, prefix: List[(String, Boolean)], subst_types: List[((String, Term.Sort), Term.Typ)], subst_terms: List[((String, Term.Typ), Term.Term)]) { def cache(cache: Term.Cache): Locale_Dependency = Locale_Dependency(entity.cache(cache), cache.string(source), cache.string(target), prefix.map({ case (name, mandatory) => (cache.string(name), mandatory) }), subst_types.map({ case ((a, s), ty) => ((cache.string(a), cache.sort(s)), cache.typ(ty)) }), subst_terms.map({ case ((x, ty), t) => ((cache.string(x), cache.typ(ty)), cache.term(t)) })) def is_inclusion: Boolean = subst_types.isEmpty && subst_terms.isEmpty } def read_locale_dependencies(provider: Export.Provider): List[Locale_Dependency] = provider.uncompressed_yxml(export_prefix + "locale_dependencies").map((tree: XML.Tree) => { val (entity, body) = decode_entity(Kind.LOCALE_DEPENDENCY, tree) val (source, (target, (prefix, (subst_types, subst_terms)))) = { import XML.Decode._ import Term_XML.Decode._ pair(string, pair(string, pair(list(pair(string, bool)), pair(list(pair(pair(string, sort), typ)), list(pair(pair(string, typ), term))))))(body) } Locale_Dependency(entity, source, target, prefix, subst_types, subst_terms) }) /* sort algebra */ sealed case class Classrel(class1: String, class2: String, prop: Prop) { def cache(cache: Term.Cache): Classrel = Classrel(cache.string(class1), cache.string(class2), prop.cache(cache)) } def read_classrel(provider: Export.Provider): List[Classrel] = { val body = provider.uncompressed_yxml(export_prefix + "classrel") val classrel = { import XML.Decode._ list(pair(decode_prop, pair(string, string)))(body) } for ((prop, (c1, c2)) <- classrel) yield Classrel(c1, c2, prop) } sealed case class Arity(type_name: String, domain: List[Term.Sort], codomain: String, prop: Prop) { def cache(cache: Term.Cache): Arity = Arity(cache.string(type_name), domain.map(cache.sort(_)), cache.string(codomain), prop.cache(cache)) } def read_arities(provider: Export.Provider): List[Arity] = { val body = provider.uncompressed_yxml(export_prefix + "arities") val arities = { import XML.Decode._ import Term_XML.Decode._ list(pair(decode_prop, triple(string, list(sort), string)))(body) } for ((prop, (a, b, c)) <- arities) yield Arity(a, b, c, prop) } /* Pure constdefs */ sealed case class Constdef(name: String, axiom_name: String) { def cache(cache: Term.Cache): Constdef = Constdef(cache.string(name), cache.string(axiom_name)) } def read_constdefs(provider: Export.Provider): List[Constdef] = { val body = provider.uncompressed_yxml(export_prefix + "constdefs") val constdefs = { import XML.Decode._ list(pair(string, string))(body) } for ((name, axiom_name) <- constdefs) yield Constdef(name, axiom_name) } /* HOL typedefs */ sealed case class Typedef(name: String, rep_type: Term.Typ, abs_type: Term.Typ, rep_name: String, abs_name: String, axiom_name: String) { def cache(cache: Term.Cache): Typedef = Typedef(cache.string(name), cache.typ(rep_type), cache.typ(abs_type), cache.string(rep_name), cache.string(abs_name), cache.string(axiom_name)) } def read_typedefs(provider: Export.Provider): List[Typedef] = { val body = provider.uncompressed_yxml(export_prefix + "typedefs") val typedefs = { import XML.Decode._ import Term_XML.Decode._ list(pair(string, pair(typ, pair(typ, pair(string, pair(string, string))))))(body) } for { (name, (rep_type, (abs_type, (rep_name, (abs_name, axiom_name))))) <- typedefs } yield Typedef(name, rep_type, abs_type, rep_name, abs_name, axiom_name) } + /* HOL datatypes */ + + sealed case class Datatype( + pos: Position.T, + name: String, + co: Boolean, + typargs: List[(String, Term.Sort)], + typ: Term.Typ, + constructors: List[(Term.Term, Term.Typ)]) + { + def id: Option[Long] = Position.Id.unapply(pos) + + def cache(cache: Term.Cache): Datatype = + Datatype( + cache.position(pos), + cache.string(name), + co, + typargs.map({ case (name, sort) => (cache.string(name), cache.sort(sort)) }), + cache.typ(typ), + constructors.map({ case (term, typ) => (cache.term(term), cache.typ(typ)) })) + } + + def read_datatypes(provider: Export.Provider): List[Datatype] = + { + val body = provider.uncompressed_yxml(export_prefix + "datatypes") + val datatypes = + { + import XML.Decode._ + import Term_XML.Decode._ + list(pair(properties, pair(string, pair(bool, pair(list(pair(string, sort)), + pair(typ, list(pair(term, typ))))))))(body) + } + for ((pos, (name, (co, (typargs, (typ, constructors))))) <- datatypes) + yield Datatype(pos, name, co, typargs, typ, constructors) + } + + /* Pure spec rules */ sealed abstract class Recursion { def cache(cache: Term.Cache): Recursion = this match { case Primrec(types) => Primrec(types.map(cache.string)) case Primcorec(types) => Primcorec(types.map(cache.string)) case _ => this } } case class Primrec(types: List[String]) extends Recursion case object Recdef extends Recursion case class Primcorec(types: List[String]) extends Recursion case object Corec extends Recursion case object Unknown_Recursion extends Recursion val decode_recursion: XML.Decode.T[Recursion] = { import XML.Decode._ variant(List( { case (Nil, a) => Primrec(list(string)(a)) }, { case (Nil, Nil) => Recdef }, { case (Nil, a) => Primcorec(list(string)(a)) }, { case (Nil, Nil) => Corec }, { case (Nil, Nil) => Unknown_Recursion })) } sealed abstract class Rough_Classification { def is_equational: Boolean = this.isInstanceOf[Equational] def is_inductive: Boolean = this == Inductive def is_co_inductive: Boolean = this == Co_Inductive def is_relational: Boolean = is_inductive || is_co_inductive def is_unknown: Boolean = this == Unknown def cache(cache: Term.Cache): Rough_Classification = this match { case Equational(recursion) => Equational(recursion.cache(cache)) case _ => this } } case class Equational(recursion: Recursion) extends Rough_Classification case object Inductive extends Rough_Classification case object Co_Inductive extends Rough_Classification case object Unknown extends Rough_Classification val decode_rough_classification: XML.Decode.T[Rough_Classification] = { import XML.Decode._ variant(List( { case (Nil, a) => Equational(decode_recursion(a)) }, { case (Nil, Nil) => Inductive }, { case (Nil, Nil) => Co_Inductive }, { case (Nil, Nil) => Unknown })) } sealed case class Spec_Rule( pos: Position.T, name: String, rough_classification: Rough_Classification, typargs: List[(String, Term.Sort)], args: List[(String, Term.Typ)], terms: List[(Term.Term, Term.Typ)], rules: List[Term.Term]) { def id: Option[Long] = Position.Id.unapply(pos) def cache(cache: Term.Cache): Spec_Rule = Spec_Rule( cache.position(pos), cache.string(name), rough_classification.cache(cache), typargs.map({ case (name, sort) => (cache.string(name), cache.sort(sort)) }), args.map({ case (name, typ) => (cache.string(name), cache.typ(typ)) }), terms.map({ case (term, typ) => (cache.term(term), cache.typ(typ)) }), rules.map(cache.term(_))) } def read_spec_rules(provider: Export.Provider): List[Spec_Rule] = { val body = provider.uncompressed_yxml(export_prefix + "spec_rules") val spec_rules = { import XML.Decode._ import Term_XML.Decode._ list( pair(properties, pair(string, pair(decode_rough_classification, pair(list(pair(string, sort)), pair(list(pair(string, typ)), pair(list(pair(term, typ)), list(term))))))))(body) } for ((pos, (name, (rough_classification, (typargs, (args, (terms, rules)))))) <- spec_rules) yield Spec_Rule(pos, name, rough_classification, typargs, args, terms, rules) } } diff --git a/src/Pure/Thy/thy_info.ML b/src/Pure/Thy/thy_info.ML --- a/src/Pure/Thy/thy_info.ML +++ b/src/Pure/Thy/thy_info.ML @@ -1,483 +1,487 @@ (* Title: Pure/Thy/thy_info.ML Author: Markus Wenzel, TU Muenchen Global theory info database, with auto-loading according to theory and file dependencies. *) signature THY_INFO = sig type presentation_context = {options: Options.T, file_pos: Position.T, adjust_pos: Position.T -> Position.T, segments: Thy_Output.segment list} + val adjust_pos_properties: presentation_context -> Position.T -> Properties.T val apply_presentation: presentation_context -> theory -> unit val add_presentation: (presentation_context -> theory -> unit) -> theory -> theory val get_names: unit -> string list val lookup_theory: string -> theory option val get_theory: string -> theory val master_directory: string -> Path.T val remove_thy: string -> unit type context = {options: Options.T, symbols: HTML.symbols, bibtex_entries: string list, last_timing: Toplevel.transition -> Time.time} val use_theories: context -> string -> Path.T -> (string * Position.T) list -> unit val use_thy: string -> unit val script_thy: Position.T -> string -> theory -> theory val register_thy: theory -> unit val finish: unit -> unit end; structure Thy_Info: THY_INFO = struct (** presentation of consolidated theory **) type presentation_context = {options: Options.T, file_pos: Position.T, adjust_pos: Position.T -> Position.T, segments: Thy_Output.segment list}; +fun adjust_pos_properties (context: presentation_context) pos = + Position.offset_properties_of (#adjust_pos context pos) @ Position.id_properties_of pos; + structure Presentation = Theory_Data ( type T = ((presentation_context -> theory -> unit) * stamp) list; val empty = []; val extend = I; fun merge data : T = Library.merge (eq_snd op =) data; ); fun apply_presentation (context: presentation_context) thy = ignore (Presentation.get thy |> Par_List.map (fn (f, _) => f context thy)); fun add_presentation f = Presentation.map (cons (f, stamp ())); val _ = Theory.setup (add_presentation (fn {options, file_pos, segments, ...} => fn thy => if exists (Toplevel.is_skipped_proof o #state) segments then () else let val body = Thy_Output.present_thy options thy segments; val option = Present.document_option options; in if #disabled option then () else let val latex = Latex.isabelle_body (Context.theory_name thy) body; val output = [Latex.output_text latex, Latex.output_positions file_pos latex]; val _ = if Options.bool options "export_document" then Export.export thy (Path.explode_binding0 "document.tex") (XML.blob output) else (); val _ = if #enabled option then Present.theory_output thy output else (); in () end end)); (** thy database **) (* messages *) val show_path = space_implode " via " o map quote; fun cycle_msg names = "Cyclic dependency of " ^ show_path names; (* derived graph operations *) fun add_deps name parents G = String_Graph.add_deps_acyclic (name, parents) G handle String_Graph.CYCLES namess => error (cat_lines (map cycle_msg namess)); fun new_entry name parents entry = String_Graph.new_node (name, entry) #> add_deps name parents; (* global thys *) type deps = {master: (Path.T * SHA1.digest), (*master dependencies for thy file*) imports: (string * Position.T) list}; (*source specification of imports (partially qualified)*) fun make_deps master imports : deps = {master = master, imports = imports}; fun master_dir_deps (d: deps option) = the_default Path.current (Option.map (Path.dir o #1 o #master) d); local val global_thys = Synchronized.var "Thy_Info.thys" (String_Graph.empty: (deps option * theory option) String_Graph.T); in fun get_thys () = Synchronized.value global_thys; fun change_thys f = Synchronized.change global_thys f; end; fun get_names () = String_Graph.topological_order (get_thys ()); (* access thy *) fun lookup thys name = try (String_Graph.get_node thys) name; fun lookup_thy name = lookup (get_thys ()) name; fun get thys name = (case lookup thys name of SOME thy => thy | NONE => error ("Theory loader: nothing known about theory " ^ quote name)); fun get_thy name = get (get_thys ()) name; (* access deps *) val lookup_deps = Option.map #1 o lookup_thy; val master_directory = master_dir_deps o #1 o get_thy; (* access theory *) fun lookup_theory name = (case lookup_thy name of SOME (_, SOME theory) => SOME theory | _ => NONE); fun get_theory name = (case lookup_theory name of SOME theory => theory | _ => error ("Theory loader: undefined entry for theory " ^ quote name)); val get_imports = Resources.imports_of o get_theory; (** thy operations **) (* remove *) fun remove name thys = (case lookup thys name of NONE => thys | SOME (NONE, _) => error ("Cannot update finished theory " ^ quote name) | SOME _ => let val succs = String_Graph.all_succs thys [name]; val _ = writeln ("Theory loader: removing " ^ commas_quote succs); in fold String_Graph.del_node succs thys end); val remove_thy = change_thys o remove; (* update *) fun update deps theory thys = let val name = Context.theory_long_name theory; val parents = map Context.theory_long_name (Theory.parents_of theory); val thys' = remove name thys; val _ = map (get thys') parents; in new_entry name parents (SOME deps, SOME theory) thys' end; fun update_thy deps theory = change_thys (update deps theory); (* context *) type context = {options: Options.T, symbols: HTML.symbols, bibtex_entries: string list, last_timing: Toplevel.transition -> Time.time}; fun default_context (): context = {options = Options.default (), symbols = HTML.no_symbols, bibtex_entries = [], last_timing = K Time.zeroTime}; (* scheduling loader tasks *) datatype result = Result of {theory: theory, exec_id: Document_ID.exec, present: unit -> unit, commit: unit -> unit, weight: int}; fun theory_result theory = Result {theory = theory, exec_id = Document_ID.none, present = I, commit = I, weight = 0}; fun result_theory (Result {theory, ...}) = theory; fun result_present (Result {present, ...}) = present; fun result_commit (Result {commit, ...}) = commit; fun result_ord (Result {weight = i, ...}, Result {weight = j, ...}) = int_ord (j, i); fun join_theory (Result {theory, exec_id, ...}) = let val _ = Execution.join [exec_id]; val res = Exn.capture Thm.consolidate_theory theory; val exns = maps Task_Queue.group_status (Execution.peek exec_id); in res :: map Exn.Exn exns end; datatype task = Task of string list * (theory list -> result) | Finished of theory; fun task_finished (Task _) = false | task_finished (Finished _) = true; fun task_parents deps (parents: string list) = map (the o AList.lookup (op =) deps) parents; val schedule_seq = String_Graph.schedule (fn deps => fn (_, task) => (case task of Task (parents, body) => let val result = body (task_parents deps parents); val _ = Par_Exn.release_all (join_theory result); val _ = result_present result (); val _ = result_commit result (); in result_theory result end | Finished thy => thy)) #> ignore; val schedule_futures = Thread_Attributes.uninterruptible (fn _ => fn tasks => let val futures = tasks |> String_Graph.schedule (fn deps => fn (name, task) => (case task of Task (parents, body) => (singleton o Future.forks) {name = "theory:" ^ name, group = NONE, deps = map (Future.task_of o #2) deps, pri = 0, interrupts = true} (fn () => (case filter (not o can Future.join o #2) deps of [] => body (map (result_theory o Future.join) (task_parents deps parents)) | bad => error ("Failed to load theory " ^ quote name ^ " (unresolved " ^ commas_quote (map #1 bad) ^ ")"))) | Finished theory => Future.value (theory_result theory))); val results1 = futures |> maps (fn future => (case Future.join_result future of Exn.Res result => join_theory result | Exn.Exn exn => [Exn.Exn exn])); val results2 = futures |> map_filter (Exn.get_res o Future.join_result) |> sort result_ord |> Par_List.map (fn result => Exn.capture (result_present result) ()); (* FIXME more precise commit order (!?) *) val results3 = futures |> map (fn future => Exn.capture (fn () => result_commit (Future.join future) ()) ()); (* FIXME avoid global Execution.reset (!??) *) val results4 = map Exn.Exn (maps Task_Queue.group_status (Execution.reset ())); val _ = Par_Exn.release_all (results1 @ results2 @ results3 @ results4); in () end); (* eval theory *) fun excursion keywords master_dir last_timing init elements = let fun prepare_span st span = Command_Span.content span |> Command.read keywords (Command.read_thy st) master_dir init ([], ~1) |> (fn tr => Toplevel.timing (last_timing tr) tr); fun element_result span_elem (st, _) = let val elem = Thy_Element.map_element (prepare_span st) span_elem; val (results, st') = Toplevel.element_result keywords elem st; val pos' = Toplevel.pos_of (Thy_Element.last_element elem); in (results, (st', pos')) end; val (results, (end_state, end_pos)) = fold_map element_result elements (Toplevel.init_toplevel (), Position.none); val thy = Toplevel.end_theory end_pos end_state; in (results, thy) end; fun eval_thy (context: context) update_time master_dir header text_pos text parents = let val {options, symbols, bibtex_entries, last_timing} = context; val (name, _) = #name header; val keywords = fold (curry Keyword.merge_keywords o Thy_Header.get_keywords) parents (Keyword.add_keywords (#keywords header) Keyword.empty_keywords); val spans = Outer_Syntax.parse_spans (Token.explode keywords text_pos text); val elements = Thy_Element.parse_elements keywords spans; fun init () = Resources.begin_theory master_dir header parents |> Present.begin_theory bibtex_entries update_time (fn () => implode (map (HTML.present_span symbols keywords) spans)); val (results, thy) = cond_timeit true ("theory " ^ quote name) (fn () => excursion keywords master_dir last_timing init elements); fun present () = let val segments = (spans ~~ maps Toplevel.join_results results) |> map (fn (span, (tr, st')) => {span = span, command = tr, state = st'}); val context: presentation_context = {options = options, file_pos = text_pos, adjust_pos = I, segments = segments}; in apply_presentation context thy end; in (thy, present, size text) end; (* require_thy -- checking database entries wrt. the file-system *) local fun required_by _ [] = "" | required_by s initiators = s ^ "(required by " ^ show_path (rev initiators) ^ ")"; fun load_thy context initiators update_time deps text (name, pos) keywords parents = let val _ = remove_thy name; val _ = writeln ("Loading theory " ^ quote name ^ required_by " " initiators); val _ = Output.try_protocol_message (Markup.loading_theory name) []; val {master = (thy_path, _), imports} = deps; val dir = Path.dir thy_path; val header = Thy_Header.make (name, pos) imports keywords; val _ = Position.reports (map #2 imports ~~ map Theory.get_markup parents); val exec_id = Document_ID.make (); val _ = Execution.running Document_ID.none exec_id [] orelse raise Fail ("Failed to register execution: " ^ Document_ID.print exec_id); val timing_start = Timing.start (); val text_pos = Position.put_id (Document_ID.print exec_id) (Path.position thy_path); val (theory, present, weight) = eval_thy context update_time dir header text_pos text (if name = Context.PureN then [Context.the_global_context ()] else parents); val timing_result = Timing.result timing_start; val timing_props = [Markup.theory_timing, (Markup.nameN, name)]; val _ = Output.try_protocol_message (timing_props @ Markup.timing_properties timing_result) [] fun commit () = update_thy deps theory; in Result {theory = theory, exec_id = exec_id, present = present, commit = commit, weight = weight} end; fun check_thy_deps dir name = (case lookup_deps name of SOME NONE => (true, NONE, Position.none, get_imports name, []) | NONE => let val {master, text, theory_pos, imports, keywords} = Resources.check_thy dir name in (false, SOME (make_deps master imports, text), theory_pos, imports, keywords) end | SOME (SOME {master, ...}) => let val {master = master', text = text', theory_pos = theory_pos', imports = imports', keywords = keywords'} = Resources.check_thy dir name; val deps' = SOME (make_deps master' imports', text'); val current = #2 master = #2 master' andalso (case lookup_theory name of NONE => false | SOME theory => Resources.loaded_files_current theory); in (current, deps', theory_pos', imports', keywords') end); in fun require_thys context initiators qualifier dir strs tasks = fold_map (require_thy context initiators qualifier dir) strs tasks |>> forall I and require_thy context initiators qualifier dir (s, require_pos) tasks = let val {master_dir, theory_name, ...} = Resources.import_name qualifier dir s; in (case try (String_Graph.get_node tasks) theory_name of SOME task => (task_finished task, tasks) | NONE => let val _ = member (op =) initiators theory_name andalso error (cycle_msg initiators); val (current, deps, theory_pos, imports, keywords) = check_thy_deps master_dir theory_name handle ERROR msg => cat_error msg ("The error(s) above occurred for theory " ^ quote theory_name ^ Position.here require_pos ^ required_by "\n" initiators); val qualifier' = Resources.theory_qualifier theory_name; val dir' = Path.append dir (master_dir_deps (Option.map #1 deps)); val parents = map (#theory_name o Resources.import_name qualifier' dir' o #1) imports; val (parents_current, tasks') = require_thys context (theory_name :: initiators) qualifier' dir' imports tasks; val all_current = current andalso parents_current; val task = if all_current then Finished (get_theory theory_name) else (case deps of NONE => raise Fail "Malformed deps" | SOME (dep, text) => let val update_time = serial (); val load = load_thy context initiators update_time dep text (theory_name, theory_pos) keywords; in Task (parents, load) end); val tasks'' = new_entry theory_name parents task tasks'; in (all_current, tasks'') end) end; end; (* use theories *) fun use_theories context qualifier master_dir imports = let val (_, tasks) = require_thys context [] qualifier master_dir imports String_Graph.empty in if Multithreading.max_threads () > 1 then schedule_futures tasks else schedule_seq tasks end; fun use_thy name = use_theories (default_context ()) Resources.default_qualifier Path.current [(name, Position.none)]; (* toplevel scripting -- without maintaining database *) fun script_thy pos txt thy = let val trs = Outer_Syntax.parse_text thy (K thy) pos txt; val end_pos = if null trs then pos else Toplevel.pos_of (List.last trs); val end_state = fold (Toplevel.command_exception true) trs (Toplevel.init_toplevel ()); in Toplevel.end_theory end_pos end_state end; (* register theory *) fun register_thy theory = let val name = Context.theory_long_name theory; val {master, ...} = Resources.check_thy (Resources.master_directory theory) name; val imports = Resources.imports_of theory; in change_thys (fn thys => let val thys' = remove name thys; val _ = writeln ("Registering theory " ^ quote name); in update (make_deps master imports) theory thys' end) end; (* finish all theories *) fun finish () = change_thys (String_Graph.map (fn _ => fn (_, entry) => (NONE, entry))); end; fun use_thy name = Runtime.toplevel_program (fn () => Thy_Info.use_thy name);