renamed "Codatatype" directory "BNF" (and corresponding session) -- this opens the door to no-nonsense session names like "HOL-BNF-LFP"
--- a/Admin/lib/Tools/update_keywords Fri Sep 21 16:34:40 2012 +0200
+++ b/Admin/lib/Tools/update_keywords Fri Sep 21 16:45:06 2012 +0200
@@ -9,7 +9,7 @@
cd "$ISABELLE_HOME/etc"
"$ISABELLE_TOOL" keywords \
- "$LOG/HOLCF.gz" "$LOG/HOL-Boogie.gz" "$LOG/HOL-Codatatype.gz" "$LOG/HOL-Library.gz" "$LOG/HOL-Nominal.gz" \
+ "$LOG/HOLCF.gz" "$LOG/HOL-BNF.gz" "$LOG/HOL-Boogie.gz" "$LOG/HOL-Library.gz" "$LOG/HOL-Nominal.gz" \
"$LOG/HOL-Statespace.gz" "$LOG/HOL-SPARK.gz" "$LOG/HOL-TPTP.gz" "$LOG/HOL-Import.gz"
"$ISABELLE_TOOL" keywords -k ZF "$LOG/ZF.gz"
--- a/CONTRIBUTORS Fri Sep 21 16:34:40 2012 +0200
+++ b/CONTRIBUTORS Fri Sep 21 16:45:06 2012 +0200
@@ -14,7 +14,7 @@
Sublist_Order) w.r.t. prefixes, suffixes, and embedding on lists.
* August 2012: Dmitriy Traytel, Andrei Popescu, Jasmin Blanchette, TUM
- New (co)datatype package.
+ New BNF-based (co)datatype package.
* August 2012: Andrei Popescu and Dmitriy Traytel, TUM
Theories of ordinals and cardinals.
--- a/NEWS Fri Sep 21 16:34:40 2012 +0200
+++ b/NEWS Fri Sep 21 16:45:06 2012 +0200
@@ -100,8 +100,9 @@
INCOMPATIBILITY.
-* HOL/Codatatype: New (co)datatype package with support for mixed,
-nested recursion and interesting non-free datatypes.
+* HOL/BNF: New (co)datatype package based on bounded natural
+functors with support for mixed, nested recursion and interesting
+non-free datatypes.
* HOL/Cardinals: Theories of ordinals and cardinals
(supersedes the AFP entry "Ordinals_and_Cardinals").
--- a/etc/isar-keywords.el Fri Sep 21 16:34:40 2012 +0200
+++ b/etc/isar-keywords.el Fri Sep 21 16:45:06 2012 +0200
@@ -1,6 +1,6 @@
;;
;; Keyword classification tables for Isabelle/Isar.
-;; Generated from HOLCF + HOL-Boogie + HOL-Codatatype + HOL-Library + HOL-Nominal + HOL-Statespace + HOL-SPARK + HOL-TPTP + HOL-Import.
+;; Generated from HOLCF + HOL-BNF + HOL-Boogie + HOL-Library + HOL-Nominal + HOL-Statespace + HOL-SPARK + HOL-TPTP + HOL-Import.
;; *** DO NOT EDIT *** DO NOT EDIT *** DO NOT EDIT ***
;;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/BNF.thy Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,16 @@
+(* Title: HOL/BNF/BNF.thy
+ Author: Dmitriy Traytel, TU Muenchen
+ Author: Andrei Popescu, TU Muenchen
+ Author: Jasmin Blanchette, TU Muenchen
+ Copyright 2012
+
+Bounded natural functors for (co)datatypes.
+*)
+
+header {* Bounded Natural Functors for (Co)datatypes *}
+
+theory BNF
+imports More_BNFs
+begin
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/BNF_Comp.thy Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,91 @@
+(* Title: HOL/BNF/BNF_Comp.thy
+ Author: Dmitriy Traytel, TU Muenchen
+ Copyright 2012
+
+Composition of bounded natural functors.
+*)
+
+header {* Composition of Bounded Natural Functors *}
+
+theory BNF_Comp
+imports Basic_BNFs
+begin
+
+lemma empty_natural: "(\<lambda>_. {}) o f = image g o (\<lambda>_. {})"
+by (rule ext) simp
+
+lemma Union_natural: "Union o image (image f) = image f o Union"
+by (rule ext) (auto simp only: o_apply)
+
+lemma in_Union_o_assoc: "x \<in> (Union o gset o gmap) A \<Longrightarrow> x \<in> (Union o (gset o gmap)) A"
+by (unfold o_assoc)
+
+lemma comp_single_set_bd:
+ assumes fbd_Card_order: "Card_order fbd" and
+ fset_bd: "\<And>x. |fset x| \<le>o fbd" and
+ gset_bd: "\<And>x. |gset x| \<le>o gbd"
+ shows "|\<Union>fset ` gset x| \<le>o gbd *c fbd"
+apply (subst sym[OF SUP_def])
+apply (rule ordLeq_transitive)
+apply (rule card_of_UNION_Sigma)
+apply (subst SIGMA_CSUM)
+apply (rule ordLeq_transitive)
+apply (rule card_of_Csum_Times')
+apply (rule fbd_Card_order)
+apply (rule ballI)
+apply (rule fset_bd)
+apply (rule ordLeq_transitive)
+apply (rule cprod_mono1)
+apply (rule gset_bd)
+apply (rule ordIso_imp_ordLeq)
+apply (rule ordIso_refl)
+apply (rule Card_order_cprod)
+done
+
+lemma Union_image_insert: "\<Union>f ` insert a B = f a \<union> \<Union>f ` B"
+by simp
+
+lemma Union_image_empty: "A \<union> \<Union>f ` {} = A"
+by simp
+
+lemma image_o_collect: "collect ((\<lambda>f. image g o f) ` F) = image g o collect F"
+by (rule ext) (auto simp add: collect_def)
+
+lemma conj_subset_def: "A \<subseteq> {x. P x \<and> Q x} = (A \<subseteq> {x. P x} \<and> A \<subseteq> {x. Q x})"
+by blast
+
+lemma UN_image_subset: "\<Union>f ` g x \<subseteq> X = (g x \<subseteq> {x. f x \<subseteq> X})"
+by blast
+
+lemma comp_set_bd_Union_o_collect: "|\<Union>\<Union>(\<lambda>f. f x) ` X| \<le>o hbd \<Longrightarrow> |(Union \<circ> collect X) x| \<le>o hbd"
+by (unfold o_apply collect_def SUP_def)
+
+lemma wpull_cong:
+"\<lbrakk>A' = A; B1' = B1; B2' = B2; wpull A B1 B2 f1 f2 p1 p2\<rbrakk> \<Longrightarrow> wpull A' B1' B2' f1 f2 p1 p2"
+by simp
+
+lemma Id_def': "Id = {(a,b). a = b}"
+by auto
+
+lemma Gr_fst_snd: "(Gr R fst)^-1 O Gr R snd = R"
+unfolding Gr_def by auto
+
+lemma subst_rel_def: "A = B \<Longrightarrow> (Gr A f)^-1 O Gr A g = (Gr B f)^-1 O Gr B g"
+by simp
+
+lemma abs_pred_def: "\<lbrakk>\<And>x y. (x, y) \<in> rel = pred x y\<rbrakk> \<Longrightarrow> rel = Collect (split pred)"
+by auto
+
+lemma Collect_split_cong: "Collect (split pred) = Collect (split pred') \<Longrightarrow> pred = pred'"
+by blast
+
+lemma pred_def_abs: "rel = Collect (split pred) \<Longrightarrow> pred = (\<lambda>x y. (x, y) \<in> rel)"
+by auto
+
+lemma mem_Id_eq_eq: "(\<lambda>x y. (x, y) \<in> Id) = (op =)"
+by simp
+
+ML_file "Tools/bnf_comp_tactics.ML"
+ML_file "Tools/bnf_comp.ML"
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/BNF_Def.thy Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,151 @@
+(* Title: HOL/BNF/BNF_Def.thy
+ Author: Dmitriy Traytel, TU Muenchen
+ Copyright 2012
+
+Definition of bounded natural functors.
+*)
+
+header {* Definition of Bounded Natural Functors *}
+
+theory BNF_Def
+imports BNF_Util
+keywords
+ "print_bnfs" :: diag and
+ "bnf_def" :: thy_goal
+begin
+
+lemma collect_o: "collect F o g = collect ((\<lambda>f. f o g) ` F)"
+by (rule ext) (auto simp only: o_apply collect_def)
+
+lemma converse_mono:
+"R1 ^-1 \<subseteq> R2 ^-1 \<longleftrightarrow> R1 \<subseteq> R2"
+unfolding converse_def by auto
+
+lemma converse_shift:
+"R1 \<subseteq> R2 ^-1 \<Longrightarrow> R1 ^-1 \<subseteq> R2"
+unfolding converse_def by auto
+
+definition convol ("<_ , _>") where
+"<f , g> \<equiv> %a. (f a, g a)"
+
+lemma fst_convol:
+"fst o <f , g> = f"
+apply(rule ext)
+unfolding convol_def by simp
+
+lemma snd_convol:
+"snd o <f , g> = g"
+apply(rule ext)
+unfolding convol_def by simp
+
+lemma convol_memI:
+"\<lbrakk>f x = f' x; g x = g' x; P x\<rbrakk> \<Longrightarrow> <f , g> x \<in> {(f' a, g' a) |a. P a}"
+unfolding convol_def by auto
+
+definition csquare where
+"csquare A f1 f2 p1 p2 \<longleftrightarrow> (\<forall> a \<in> A. f1 (p1 a) = f2 (p2 a))"
+
+(* The pullback of sets *)
+definition thePull where
+"thePull B1 B2 f1 f2 = {(b1,b2). b1 \<in> B1 \<and> b2 \<in> B2 \<and> f1 b1 = f2 b2}"
+
+lemma wpull_thePull:
+"wpull (thePull B1 B2 f1 f2) B1 B2 f1 f2 fst snd"
+unfolding wpull_def thePull_def by auto
+
+lemma wppull_thePull:
+assumes "wppull A B1 B2 f1 f2 e1 e2 p1 p2"
+shows
+"\<exists> j. \<forall> a' \<in> thePull B1 B2 f1 f2.
+ j a' \<in> A \<and>
+ e1 (p1 (j a')) = e1 (fst a') \<and> e2 (p2 (j a')) = e2 (snd a')"
+(is "\<exists> j. \<forall> a' \<in> ?A'. ?phi a' (j a')")
+proof(rule bchoice[of ?A' ?phi], default)
+ fix a' assume a': "a' \<in> ?A'"
+ hence "fst a' \<in> B1" unfolding thePull_def by auto
+ moreover
+ from a' have "snd a' \<in> B2" unfolding thePull_def by auto
+ moreover have "f1 (fst a') = f2 (snd a')"
+ using a' unfolding csquare_def thePull_def by auto
+ ultimately show "\<exists> ja'. ?phi a' ja'"
+ using assms unfolding wppull_def by blast
+qed
+
+lemma wpull_wppull:
+assumes wp: "wpull A' B1 B2 f1 f2 p1' p2'" and
+1: "\<forall> a' \<in> A'. j a' \<in> A \<and> e1 (p1 (j a')) = e1 (p1' a') \<and> e2 (p2 (j a')) = e2 (p2' a')"
+shows "wppull A B1 B2 f1 f2 e1 e2 p1 p2"
+unfolding wppull_def proof safe
+ fix b1 b2
+ assume b1: "b1 \<in> B1" and b2: "b2 \<in> B2" and f: "f1 b1 = f2 b2"
+ then obtain a' where a': "a' \<in> A'" and b1: "b1 = p1' a'" and b2: "b2 = p2' a'"
+ using wp unfolding wpull_def by blast
+ show "\<exists>a\<in>A. e1 (p1 a) = e1 b1 \<and> e2 (p2 a) = e2 b2"
+ apply (rule bexI[of _ "j a'"]) unfolding b1 b2 using a' 1 by auto
+qed
+
+lemma wppull_id: "\<lbrakk>wpull UNIV UNIV UNIV f1 f2 p1 p2; e1 = id; e2 = id\<rbrakk> \<Longrightarrow>
+ wppull UNIV UNIV UNIV f1 f2 e1 e2 p1 p2"
+by (erule wpull_wppull) auto
+
+lemma Id_alt: "Id = Gr UNIV id"
+unfolding Gr_def by auto
+
+lemma Gr_UNIV_id: "f = id \<Longrightarrow> (Gr UNIV f)^-1 O Gr UNIV f = Gr UNIV f"
+unfolding Gr_def by auto
+
+lemma Gr_mono: "A \<subseteq> B \<Longrightarrow> Gr A f \<subseteq> Gr B f"
+unfolding Gr_def by auto
+
+lemma wpull_Gr:
+"wpull (Gr A f) A (f ` A) f id fst snd"
+unfolding wpull_def Gr_def by auto
+
+definition "pick_middle P Q a c = (SOME b. (a,b) \<in> P \<and> (b,c) \<in> Q)"
+
+lemma pick_middle:
+"(a,c) \<in> P O Q \<Longrightarrow> (a, pick_middle P Q a c) \<in> P \<and> (pick_middle P Q a c, c) \<in> Q"
+unfolding pick_middle_def apply(rule someI_ex)
+using assms unfolding relcomp_def by auto
+
+definition fstO where "fstO P Q ac = (fst ac, pick_middle P Q (fst ac) (snd ac))"
+definition sndO where "sndO P Q ac = (pick_middle P Q (fst ac) (snd ac), snd ac)"
+
+lemma fstO_in: "ac \<in> P O Q \<Longrightarrow> fstO P Q ac \<in> P"
+unfolding fstO_def
+by (subst (asm) surjective_pairing) (rule pick_middle[THEN conjunct1])
+
+lemma fst_fstO: "fst bc = (fst \<circ> fstO P Q) bc"
+unfolding comp_def fstO_def by simp
+
+lemma snd_sndO: "snd bc = (snd \<circ> sndO P Q) bc"
+unfolding comp_def sndO_def by simp
+
+lemma sndO_in: "ac \<in> P O Q \<Longrightarrow> sndO P Q ac \<in> Q"
+unfolding sndO_def
+by (subst (asm) surjective_pairing) (rule pick_middle[THEN conjunct2])
+
+lemma csquare_fstO_sndO:
+"csquare (P O Q) snd fst (fstO P Q) (sndO P Q)"
+unfolding csquare_def fstO_def sndO_def using pick_middle by simp
+
+lemma wppull_fstO_sndO:
+shows "wppull (P O Q) P Q snd fst fst snd (fstO P Q) (sndO P Q)"
+using pick_middle unfolding wppull_def fstO_def sndO_def relcomp_def by auto
+
+lemma snd_fst_flip: "snd xy = (fst o (%(x, y). (y, x))) xy"
+by (simp split: prod.split)
+
+lemma fst_snd_flip: "fst xy = (snd o (%(x, y). (y, x))) xy"
+by (simp split: prod.split)
+
+lemma flip_rel: "A \<subseteq> (R ^-1) \<Longrightarrow> (%(x, y). (y, x)) ` A \<subseteq> R"
+by auto
+
+lemma pointfreeE: "f o g = f' o g' \<Longrightarrow> f (g x) = f' (g' x)"
+unfolding o_def fun_eq_iff by simp
+
+ML_file "Tools/bnf_def_tactics.ML"
+ML_file"Tools/bnf_def.ML"
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/BNF_FP.thy Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,113 @@
+(* Title: HOL/BNF/BNF_FP.thy
+ Author: Dmitriy Traytel, TU Muenchen
+ Author: Jasmin Blanchette, TU Muenchen
+ Copyright 2012
+
+Composition of bounded natural functors.
+*)
+
+header {* Composition of Bounded Natural Functors *}
+
+theory BNF_FP
+imports BNF_Comp BNF_Wrap
+keywords
+ "defaults"
+begin
+
+lemma case_unit: "(case u of () => f) = f"
+by (cases u) (hypsubst, rule unit.cases)
+
+lemma unit_all_impI: "(P () \<Longrightarrow> Q ()) \<Longrightarrow> \<forall>x. P x \<longrightarrow> Q x"
+by simp
+
+lemma prod_all_impI: "(\<And>x y. P (x, y) \<Longrightarrow> Q (x, y)) \<Longrightarrow> \<forall>x. P x \<longrightarrow> Q x"
+by clarify
+
+lemma prod_all_impI_step: "(\<And>x. \<forall>y. P (x, y) \<longrightarrow> Q (x, y)) \<Longrightarrow> \<forall>x. P x \<longrightarrow> Q x"
+by auto
+
+lemma all_unit_eq: "(\<And>x. PROP P x) \<equiv> PROP P ()"
+by simp
+
+lemma all_prod_eq: "(\<And>x. PROP P x) \<equiv> (\<And>a b. PROP P (a, b))"
+by clarsimp
+
+lemma rev_bspec: "a \<in> A \<Longrightarrow> \<forall>z \<in> A. P z \<Longrightarrow> P a"
+by simp
+
+lemma Un_cong: "\<lbrakk>A = B; C = D\<rbrakk> \<Longrightarrow> A \<union> C = B \<union> D"
+by simp
+
+lemma pointfree_idE: "f o g = id \<Longrightarrow> f (g x) = x"
+unfolding o_def fun_eq_iff by simp
+
+lemma o_bij:
+ assumes gf: "g o f = id" and fg: "f o g = id"
+ shows "bij f"
+unfolding bij_def inj_on_def surj_def proof safe
+ fix a1 a2 assume "f a1 = f a2"
+ hence "g ( f a1) = g (f a2)" by simp
+ thus "a1 = a2" using gf unfolding fun_eq_iff by simp
+next
+ fix b
+ have "b = f (g b)"
+ using fg unfolding fun_eq_iff by simp
+ thus "EX a. b = f a" by blast
+qed
+
+lemma ssubst_mem: "\<lbrakk>t = s; s \<in> X\<rbrakk> \<Longrightarrow> t \<in> X" by simp
+
+lemma sum_case_step:
+ "sum_case (sum_case f' g') g (Inl p) = sum_case f' g' p"
+ "sum_case f (sum_case f' g') (Inr p) = sum_case f' g' p"
+by auto
+
+lemma one_pointE: "\<lbrakk>\<And>x. s = x \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
+by simp
+
+lemma obj_one_pointE: "\<forall>x. s = x \<longrightarrow> P \<Longrightarrow> P"
+by blast
+
+lemma obj_sumE_f':
+"\<lbrakk>\<forall>x. s = f (Inl x) \<longrightarrow> P; \<forall>x. s = f (Inr x) \<longrightarrow> P\<rbrakk> \<Longrightarrow> s = f x \<longrightarrow> P"
+by (cases x) blast+
+
+lemma obj_sumE_f:
+"\<lbrakk>\<forall>x. s = f (Inl x) \<longrightarrow> P; \<forall>x. s = f (Inr x) \<longrightarrow> P\<rbrakk> \<Longrightarrow> \<forall>x. s = f x \<longrightarrow> P"
+by (rule allI) (rule obj_sumE_f')
+
+lemma obj_sumE: "\<lbrakk>\<forall>x. s = Inl x \<longrightarrow> P; \<forall>x. s = Inr x \<longrightarrow> P\<rbrakk> \<Longrightarrow> P"
+by (cases s) auto
+
+lemma obj_sum_step':
+"\<lbrakk>\<forall>x. s = f (Inr (Inl x)) \<longrightarrow> P; \<forall>x. s = f (Inr (Inr x)) \<longrightarrow> P\<rbrakk> \<Longrightarrow> s = f (Inr x) \<longrightarrow> P"
+by (cases x) blast+
+
+lemma obj_sum_step:
+"\<lbrakk>\<forall>x. s = f (Inr (Inl x)) \<longrightarrow> P; \<forall>x. s = f (Inr (Inr x)) \<longrightarrow> P\<rbrakk> \<Longrightarrow> \<forall>x. s = f (Inr x) \<longrightarrow> P"
+by (rule allI) (rule obj_sum_step')
+
+lemma sum_case_if:
+"sum_case f g (if p then Inl x else Inr y) = (if p then f x else g y)"
+by simp
+
+lemma mem_UN_compreh_eq: "(z : \<Union>{y. \<exists>x\<in>A. y = F x}) = (\<exists>x\<in>A. z : F x)"
+by blast
+
+lemma prod_set_simps:
+"fsts (x, y) = {x}"
+"snds (x, y) = {y}"
+unfolding fsts_def snds_def by simp+
+
+lemma sum_set_simps:
+"setl (Inl x) = {x}"
+"setl (Inr x) = {}"
+"setr (Inl x) = {}"
+"setr (Inr x) = {x}"
+unfolding sum_set_defs by simp+
+
+ML_file "Tools/bnf_fp.ML"
+ML_file "Tools/bnf_fp_sugar_tactics.ML"
+ML_file "Tools/bnf_fp_sugar.ML"
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/BNF_GFP.thy Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,331 @@
+(* Title: HOL/BNF/BNF_GFP.thy
+ Author: Dmitriy Traytel, TU Muenchen
+ Copyright 2012
+
+Greatest fixed point operation on bounded natural functors.
+*)
+
+header {* Greatest Fixed Point Operation on Bounded Natural Functors *}
+
+theory BNF_GFP
+imports BNF_FP Equiv_Relations_More "~~/src/HOL/Library/Prefix_Order"
+keywords
+ "codata_raw" :: thy_decl and
+ "codata" :: thy_decl
+begin
+
+lemma sum_case_comp_Inl:
+"sum_case f g \<circ> Inl = f"
+unfolding comp_def by simp
+
+lemma sum_case_expand_Inr: "f o Inl = g \<Longrightarrow> f x = sum_case g (f o Inr) x"
+by (auto split: sum.splits)
+
+lemma converse_Times: "(A \<times> B) ^-1 = B \<times> A"
+by auto
+
+lemma equiv_triv1:
+assumes "equiv A R" and "(a, b) \<in> R" and "(a, c) \<in> R"
+shows "(b, c) \<in> R"
+using assms unfolding equiv_def sym_def trans_def by blast
+
+lemma equiv_triv2:
+assumes "equiv A R" and "(a, b) \<in> R" and "(b, c) \<in> R"
+shows "(a, c) \<in> R"
+using assms unfolding equiv_def trans_def by blast
+
+lemma equiv_proj:
+ assumes e: "equiv A R" and "z \<in> R"
+ shows "(proj R o fst) z = (proj R o snd) z"
+proof -
+ from assms(2) have z: "(fst z, snd z) \<in> R" by auto
+ have P: "\<And>x. (fst z, x) \<in> R \<Longrightarrow> (snd z, x) \<in> R" by (erule equiv_triv1[OF e z])
+ have "\<And>x. (snd z, x) \<in> R \<Longrightarrow> (fst z, x) \<in> R" by (erule equiv_triv2[OF e z])
+ with P show ?thesis unfolding proj_def[abs_def] by auto
+qed
+
+(* Operators: *)
+definition diag where "diag A \<equiv> {(a,a) | a. a \<in> A}"
+definition image2 where "image2 A f g = {(f a, g a) | a. a \<in> A}"
+
+lemma diagI: "x \<in> A \<Longrightarrow> (x, x) \<in> diag A"
+unfolding diag_def by simp
+
+lemma diagE: "(a, b) \<in> diag A \<Longrightarrow> a = b"
+unfolding diag_def by simp
+
+lemma diagE': "x \<in> diag A \<Longrightarrow> fst x = snd x"
+unfolding diag_def by auto
+
+lemma diag_fst: "x \<in> diag A \<Longrightarrow> fst x \<in> A"
+unfolding diag_def by auto
+
+lemma diag_UNIV: "diag UNIV = Id"
+unfolding diag_def by auto
+
+lemma diag_converse: "diag A = (diag A) ^-1"
+unfolding diag_def by auto
+
+lemma diag_Comp: "diag A = diag A O diag A"
+unfolding diag_def by auto
+
+lemma diag_Gr: "diag A = Gr A id"
+unfolding diag_def Gr_def by simp
+
+lemma diag_UNIV_I: "x = y \<Longrightarrow> (x, y) \<in> diag UNIV"
+unfolding diag_def by auto
+
+lemma image2_eqI: "\<lbrakk>b = f x; c = g x; x \<in> A\<rbrakk> \<Longrightarrow> (b, c) \<in> image2 A f g"
+unfolding image2_def by auto
+
+lemma Id_subset: "Id \<subseteq> {(a, b). P a b \<or> a = b}"
+by auto
+
+lemma IdD: "(a, b) \<in> Id \<Longrightarrow> a = b"
+by auto
+
+lemma image2_Gr: "image2 A f g = (Gr A f)^-1 O (Gr A g)"
+unfolding image2_def Gr_def by auto
+
+lemma GrI: "\<lbrakk>x \<in> A; f x = fx\<rbrakk> \<Longrightarrow> (x, fx) \<in> Gr A f"
+unfolding Gr_def by simp
+
+lemma GrE: "(x, fx) \<in> Gr A f \<Longrightarrow> (x \<in> A \<Longrightarrow> f x = fx \<Longrightarrow> P) \<Longrightarrow> P"
+unfolding Gr_def by simp
+
+lemma GrD1: "(x, fx) \<in> Gr A f \<Longrightarrow> x \<in> A"
+unfolding Gr_def by simp
+
+lemma GrD2: "(x, fx) \<in> Gr A f \<Longrightarrow> f x = fx"
+unfolding Gr_def by simp
+
+lemma Gr_incl: "Gr A f \<subseteq> A <*> B \<longleftrightarrow> f ` A \<subseteq> B"
+unfolding Gr_def by auto
+
+definition relImage where
+"relImage R f \<equiv> {(f a1, f a2) | a1 a2. (a1,a2) \<in> R}"
+
+definition relInvImage where
+"relInvImage A R f \<equiv> {(a1, a2) | a1 a2. a1 \<in> A \<and> a2 \<in> A \<and> (f a1, f a2) \<in> R}"
+
+lemma relImage_Gr:
+"\<lbrakk>R \<subseteq> A \<times> A\<rbrakk> \<Longrightarrow> relImage R f = (Gr A f)^-1 O R O Gr A f"
+unfolding relImage_def Gr_def relcomp_def by auto
+
+lemma relInvImage_Gr: "\<lbrakk>R \<subseteq> B \<times> B\<rbrakk> \<Longrightarrow> relInvImage A R f = Gr A f O R O (Gr A f)^-1"
+unfolding Gr_def relcomp_def image_def relInvImage_def by auto
+
+lemma relImage_mono:
+"R1 \<subseteq> R2 \<Longrightarrow> relImage R1 f \<subseteq> relImage R2 f"
+unfolding relImage_def by auto
+
+lemma relInvImage_mono:
+"R1 \<subseteq> R2 \<Longrightarrow> relInvImage A R1 f \<subseteq> relInvImage A R2 f"
+unfolding relInvImage_def by auto
+
+lemma relInvImage_diag:
+"(\<And>a1 a2. f a1 = f a2 \<longleftrightarrow> a1 = a2) \<Longrightarrow> relInvImage A (diag B) f \<subseteq> Id"
+unfolding relInvImage_def diag_def by auto
+
+lemma relInvImage_UNIV_relImage:
+"R \<subseteq> relInvImage UNIV (relImage R f) f"
+unfolding relInvImage_def relImage_def by auto
+
+lemma equiv_Image: "equiv A R \<Longrightarrow> (\<And>a b. (a, b) \<in> R \<Longrightarrow> a \<in> A \<and> b \<in> A \<and> R `` {a} = R `` {b})"
+unfolding equiv_def refl_on_def Image_def by (auto intro: transD symD)
+
+lemma relImage_proj:
+assumes "equiv A R"
+shows "relImage R (proj R) \<subseteq> diag (A//R)"
+unfolding relImage_def diag_def apply safe
+using proj_iff[OF assms]
+by (metis assms equiv_Image proj_def proj_preserves)
+
+lemma relImage_relInvImage:
+assumes "R \<subseteq> f ` A <*> f ` A"
+shows "relImage (relInvImage A R f) f = R"
+using assms unfolding relImage_def relInvImage_def by fastforce
+
+lemma subst_Pair: "P x y \<Longrightarrow> a = (x, y) \<Longrightarrow> P (fst a) (snd a)"
+by simp
+
+lemma fst_diag_id: "(fst \<circ> (%x. (x, x))) z = id z"
+by simp
+
+lemma snd_diag_id: "(snd \<circ> (%x. (x, x))) z = id z"
+by simp
+
+lemma Collect_restrict': "{(x, y) | x y. phi x y \<and> P x y} \<subseteq> {(x, y) | x y. phi x y}"
+by auto
+
+lemma image_convolD: "\<lbrakk>(a, b) \<in> <f, g> ` X\<rbrakk> \<Longrightarrow> \<exists>x. x \<in> X \<and> a = f x \<and> b = g x"
+unfolding convol_def by auto
+
+(*Extended Sublist*)
+
+definition prefCl where
+ "prefCl Kl = (\<forall> kl1 kl2. kl1 \<le> kl2 \<and> kl2 \<in> Kl \<longrightarrow> kl1 \<in> Kl)"
+definition PrefCl where
+ "PrefCl A n = (\<forall>kl kl'. kl \<in> A n \<and> kl' \<le> kl \<longrightarrow> (\<exists>m\<le>n. kl' \<in> A m))"
+
+lemma prefCl_UN:
+ "\<lbrakk>\<And>n. PrefCl A n\<rbrakk> \<Longrightarrow> prefCl (\<Union>n. A n)"
+unfolding prefCl_def PrefCl_def by fastforce
+
+definition Succ where "Succ Kl kl = {k . kl @ [k] \<in> Kl}"
+definition Shift where "Shift Kl k = {kl. k # kl \<in> Kl}"
+definition shift where "shift lab k = (\<lambda>kl. lab (k # kl))"
+
+lemma empty_Shift: "\<lbrakk>[] \<in> Kl; k \<in> Succ Kl []\<rbrakk> \<Longrightarrow> [] \<in> Shift Kl k"
+unfolding Shift_def Succ_def by simp
+
+lemma Shift_clists: "Kl \<subseteq> Field (clists r) \<Longrightarrow> Shift Kl k \<subseteq> Field (clists r)"
+unfolding Shift_def clists_def Field_card_of by auto
+
+lemma Shift_prefCl: "prefCl Kl \<Longrightarrow> prefCl (Shift Kl k)"
+unfolding prefCl_def Shift_def
+proof safe
+ fix kl1 kl2
+ assume "\<forall>kl1 kl2. kl1 \<le> kl2 \<and> kl2 \<in> Kl \<longrightarrow> kl1 \<in> Kl"
+ "kl1 \<le> kl2" "k # kl2 \<in> Kl"
+ thus "k # kl1 \<in> Kl" using Cons_prefix_Cons[of k kl1 k kl2] by blast
+qed
+
+lemma not_in_Shift: "kl \<notin> Shift Kl x \<Longrightarrow> x # kl \<notin> Kl"
+unfolding Shift_def by simp
+
+lemma prefCl_Succ: "\<lbrakk>prefCl Kl; k # kl \<in> Kl\<rbrakk> \<Longrightarrow> k \<in> Succ Kl []"
+unfolding Succ_def proof
+ assume "prefCl Kl" "k # kl \<in> Kl"
+ moreover have "k # [] \<le> k # kl" by auto
+ ultimately have "k # [] \<in> Kl" unfolding prefCl_def by blast
+ thus "[] @ [k] \<in> Kl" by simp
+qed
+
+lemma SuccD: "k \<in> Succ Kl kl \<Longrightarrow> kl @ [k] \<in> Kl"
+unfolding Succ_def by simp
+
+lemmas SuccE = SuccD[elim_format]
+
+lemma SuccI: "kl @ [k] \<in> Kl \<Longrightarrow> k \<in> Succ Kl kl"
+unfolding Succ_def by simp
+
+lemma ShiftD: "kl \<in> Shift Kl k \<Longrightarrow> k # kl \<in> Kl"
+unfolding Shift_def by simp
+
+lemma Succ_Shift: "Succ (Shift Kl k) kl = Succ Kl (k # kl)"
+unfolding Succ_def Shift_def by auto
+
+lemma ShiftI: "k # kl \<in> Kl \<Longrightarrow> kl \<in> Shift Kl k"
+unfolding Shift_def by simp
+
+lemma Func_cexp: "|Func A B| =o |B| ^c |A|"
+unfolding cexp_def Field_card_of by (simp only: card_of_refl)
+
+lemma clists_bound: "A \<in> Field (cpow (clists r)) - {{}} \<Longrightarrow> |A| \<le>o clists r"
+unfolding cpow_def clists_def Field_card_of by (auto simp: card_of_mono1)
+
+lemma cpow_clists_czero: "\<lbrakk>A \<in> Field (cpow (clists r)) - {{}}; |A| =o czero\<rbrakk> \<Longrightarrow> False"
+unfolding cpow_def clists_def
+by (auto simp add: card_of_ordIso_czero_iff_empty[symmetric])
+ (erule notE, erule ordIso_transitive, rule czero_ordIso)
+
+lemma incl_UNION_I:
+assumes "i \<in> I" and "A \<subseteq> F i"
+shows "A \<subseteq> UNION I F"
+using assms by auto
+
+lemma Nil_clists: "{[]} \<subseteq> Field (clists r)"
+unfolding clists_def Field_card_of by auto
+
+lemma Cons_clists:
+ "\<lbrakk>x \<in> Field r; xs \<in> Field (clists r)\<rbrakk> \<Longrightarrow> x # xs \<in> Field (clists r)"
+unfolding clists_def Field_card_of by auto
+
+lemma length_Cons: "length (x # xs) = Suc (length xs)"
+by simp
+
+lemma length_append_singleton: "length (xs @ [x]) = Suc (length xs)"
+by simp
+
+(*injection into the field of a cardinal*)
+definition "toCard_pred A r f \<equiv> inj_on f A \<and> f ` A \<subseteq> Field r \<and> Card_order r"
+definition "toCard A r \<equiv> SOME f. toCard_pred A r f"
+
+lemma ex_toCard_pred:
+"\<lbrakk>|A| \<le>o r; Card_order r\<rbrakk> \<Longrightarrow> \<exists> f. toCard_pred A r f"
+unfolding toCard_pred_def
+using card_of_ordLeq[of A "Field r"]
+ ordLeq_ordIso_trans[OF _ card_of_unique[of "Field r" r], of "|A|"]
+by blast
+
+lemma toCard_pred_toCard:
+ "\<lbrakk>|A| \<le>o r; Card_order r\<rbrakk> \<Longrightarrow> toCard_pred A r (toCard A r)"
+unfolding toCard_def using someI_ex[OF ex_toCard_pred] .
+
+lemma toCard_inj: "\<lbrakk>|A| \<le>o r; Card_order r; x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow>
+ toCard A r x = toCard A r y \<longleftrightarrow> x = y"
+using toCard_pred_toCard unfolding inj_on_def toCard_pred_def by blast
+
+lemma toCard: "\<lbrakk>|A| \<le>o r; Card_order r; b \<in> A\<rbrakk> \<Longrightarrow> toCard A r b \<in> Field r"
+using toCard_pred_toCard unfolding toCard_pred_def by blast
+
+definition "fromCard A r k \<equiv> SOME b. b \<in> A \<and> toCard A r b = k"
+
+lemma fromCard_toCard:
+"\<lbrakk>|A| \<le>o r; Card_order r; b \<in> A\<rbrakk> \<Longrightarrow> fromCard A r (toCard A r b) = b"
+unfolding fromCard_def by (rule some_equality) (auto simp add: toCard_inj)
+
+(* pick according to the weak pullback *)
+definition pickWP_pred where
+"pickWP_pred A p1 p2 b1 b2 a \<equiv> a \<in> A \<and> p1 a = b1 \<and> p2 a = b2"
+
+definition pickWP where
+"pickWP A p1 p2 b1 b2 \<equiv> SOME a. pickWP_pred A p1 p2 b1 b2 a"
+
+lemma pickWP_pred:
+assumes "wpull A B1 B2 f1 f2 p1 p2" and
+"b1 \<in> B1" and "b2 \<in> B2" and "f1 b1 = f2 b2"
+shows "\<exists> a. pickWP_pred A p1 p2 b1 b2 a"
+using assms unfolding wpull_def pickWP_pred_def by blast
+
+lemma pickWP_pred_pickWP:
+assumes "wpull A B1 B2 f1 f2 p1 p2" and
+"b1 \<in> B1" and "b2 \<in> B2" and "f1 b1 = f2 b2"
+shows "pickWP_pred A p1 p2 b1 b2 (pickWP A p1 p2 b1 b2)"
+unfolding pickWP_def using assms by(rule someI_ex[OF pickWP_pred])
+
+lemma pickWP:
+assumes "wpull A B1 B2 f1 f2 p1 p2" and
+"b1 \<in> B1" and "b2 \<in> B2" and "f1 b1 = f2 b2"
+shows "pickWP A p1 p2 b1 b2 \<in> A"
+ "p1 (pickWP A p1 p2 b1 b2) = b1"
+ "p2 (pickWP A p1 p2 b1 b2) = b2"
+using assms pickWP_pred_pickWP unfolding pickWP_pred_def by fastforce+
+
+lemma Inl_Field_csum: "a \<in> Field r \<Longrightarrow> Inl a \<in> Field (r +c s)"
+unfolding Field_card_of csum_def by auto
+
+lemma Inr_Field_csum: "a \<in> Field s \<Longrightarrow> Inr a \<in> Field (r +c s)"
+unfolding Field_card_of csum_def by auto
+
+lemma nat_rec_0: "f = nat_rec f1 (%n rec. f2 n rec) \<Longrightarrow> f 0 = f1"
+by auto
+
+lemma nat_rec_Suc: "f = nat_rec f1 (%n rec. f2 n rec) \<Longrightarrow> f (Suc n) = f2 n (f n)"
+by auto
+
+lemma list_rec_Nil: "f = list_rec f1 (%x xs rec. f2 x xs rec) \<Longrightarrow> f [] = f1"
+by auto
+
+lemma list_rec_Cons: "f = list_rec f1 (%x xs rec. f2 x xs rec) \<Longrightarrow> f (x # xs) = f2 x xs (f xs)"
+by auto
+
+lemma not_arg_cong_Inr: "x \<noteq> y \<Longrightarrow> Inr x \<noteq> Inr y"
+by simp
+
+ML_file "Tools/bnf_gfp_util.ML"
+ML_file "Tools/bnf_gfp_tactics.ML"
+ML_file "Tools/bnf_gfp.ML"
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/BNF_LFP.thy Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,228 @@
+(* Title: HOL/BNF/BNF_LFP.thy
+ Author: Dmitriy Traytel, TU Muenchen
+ Copyright 2012
+
+Least fixed point operation on bounded natural functors.
+*)
+
+header {* Least Fixed Point Operation on Bounded Natural Functors *}
+
+theory BNF_LFP
+imports BNF_FP
+keywords
+ "data_raw" :: thy_decl and
+ "data" :: thy_decl
+begin
+
+lemma subset_emptyI: "(\<And>x. x \<in> A \<Longrightarrow> False) \<Longrightarrow> A \<subseteq> {}"
+by blast
+
+lemma image_Collect_subsetI:
+ "(\<And>x. P x \<Longrightarrow> f x \<in> B) \<Longrightarrow> f ` {x. P x} \<subseteq> B"
+by blast
+
+lemma Collect_restrict: "{x. x \<in> X \<and> P x} \<subseteq> X"
+by auto
+
+lemma prop_restrict: "\<lbrakk>x \<in> Z; Z \<subseteq> {x. x \<in> X \<and> P x}\<rbrakk> \<Longrightarrow> P x"
+by auto
+
+lemma underS_I: "\<lbrakk>i \<noteq> j; (i, j) \<in> R\<rbrakk> \<Longrightarrow> i \<in> rel.underS R j"
+unfolding rel.underS_def by simp
+
+lemma underS_E: "i \<in> rel.underS R j \<Longrightarrow> i \<noteq> j \<and> (i, j) \<in> R"
+unfolding rel.underS_def by simp
+
+lemma underS_Field: "i \<in> rel.underS R j \<Longrightarrow> i \<in> Field R"
+unfolding rel.underS_def Field_def by auto
+
+lemma FieldI2: "(i, j) \<in> R \<Longrightarrow> j \<in> Field R"
+unfolding Field_def by auto
+
+lemma fst_convol': "fst (<f, g> x) = f x"
+using fst_convol unfolding convol_def by simp
+
+lemma snd_convol': "snd (<f, g> x) = g x"
+using snd_convol unfolding convol_def by simp
+
+lemma convol_expand_snd: "fst o f = g \<Longrightarrow> <g, snd o f> = f"
+unfolding convol_def by auto
+
+definition inver where
+ "inver g f A = (ALL a : A. g (f a) = a)"
+
+lemma bij_betw_iff_ex:
+ "bij_betw f A B = (EX g. g ` B = A \<and> inver g f A \<and> inver f g B)" (is "?L = ?R")
+proof (rule iffI)
+ assume ?L
+ hence f: "f ` A = B" and inj_f: "inj_on f A" unfolding bij_betw_def by auto
+ let ?phi = "% b a. a : A \<and> f a = b"
+ have "ALL b : B. EX a. ?phi b a" using f by blast
+ then obtain g where g: "ALL b : B. g b : A \<and> f (g b) = b"
+ using bchoice[of B ?phi] by blast
+ hence gg: "ALL b : f ` A. g b : A \<and> f (g b) = b" using f by blast
+ have gf: "inver g f A" unfolding inver_def
+ by (metis (no_types) gg imageI[of _ A f] the_inv_into_f_f[OF inj_f])
+ moreover have "g ` B \<le> A \<and> inver f g B" using g unfolding inver_def by blast
+ moreover have "A \<le> g ` B"
+ proof safe
+ fix a assume a: "a : A"
+ hence "f a : B" using f by auto
+ moreover have "a = g (f a)" using a gf unfolding inver_def by auto
+ ultimately show "a : g ` B" by blast
+ qed
+ ultimately show ?R by blast
+next
+ assume ?R
+ then obtain g where g: "g ` B = A \<and> inver g f A \<and> inver f g B" by blast
+ show ?L unfolding bij_betw_def
+ proof safe
+ show "inj_on f A" unfolding inj_on_def
+ proof safe
+ fix a1 a2 assume a: "a1 : A" "a2 : A" and "f a1 = f a2"
+ hence "g (f a1) = g (f a2)" by simp
+ thus "a1 = a2" using a g unfolding inver_def by simp
+ qed
+ next
+ fix a assume "a : A"
+ then obtain b where b: "b : B" and a: "a = g b" using g by blast
+ hence "b = f (g b)" using g unfolding inver_def by auto
+ thus "f a : B" unfolding a using b by simp
+ next
+ fix b assume "b : B"
+ hence "g b : A \<and> b = f (g b)" using g unfolding inver_def by auto
+ thus "b : f ` A" by auto
+ qed
+qed
+
+lemma bij_betw_ex_weakE:
+ "\<lbrakk>bij_betw f A B\<rbrakk> \<Longrightarrow> \<exists>g. g ` B \<subseteq> A \<and> inver g f A \<and> inver f g B"
+by (auto simp only: bij_betw_iff_ex)
+
+lemma inver_surj: "\<lbrakk>g ` B \<subseteq> A; f ` A \<subseteq> B; inver g f A\<rbrakk> \<Longrightarrow> g ` B = A"
+unfolding inver_def by auto (rule rev_image_eqI, auto)
+
+lemma inver_mono: "\<lbrakk>A \<subseteq> B; inver f g B\<rbrakk> \<Longrightarrow> inver f g A"
+unfolding inver_def by auto
+
+lemma inver_pointfree: "inver f g A = (\<forall>a \<in> A. (f o g) a = a)"
+unfolding inver_def by simp
+
+lemma bij_betwE: "bij_betw f A B \<Longrightarrow> \<forall>a\<in>A. f a \<in> B"
+unfolding bij_betw_def by auto
+
+lemma bij_betw_imageE: "bij_betw f A B \<Longrightarrow> f ` A = B"
+unfolding bij_betw_def by auto
+
+lemma inverE: "\<lbrakk>inver f f' A; x \<in> A\<rbrakk> \<Longrightarrow> f (f' x) = x"
+unfolding inver_def by auto
+
+lemma bij_betw_inver1: "bij_betw f A B \<Longrightarrow> inver (inv_into A f) f A"
+unfolding bij_betw_def inver_def by auto
+
+lemma bij_betw_inver2: "bij_betw f A B \<Longrightarrow> inver f (inv_into A f) B"
+unfolding bij_betw_def inver_def by auto
+
+lemma bij_betwI: "\<lbrakk>bij_betw g B A; inver g f A; inver f g B\<rbrakk> \<Longrightarrow> bij_betw f A B"
+by (drule bij_betw_imageE, unfold bij_betw_iff_ex) blast
+
+lemma bij_betwI':
+ "\<lbrakk>\<And>x y. \<lbrakk>x \<in> X; y \<in> X\<rbrakk> \<Longrightarrow> (f x = f y) = (x = y);
+ \<And>x. x \<in> X \<Longrightarrow> f x \<in> Y;
+ \<And>y. y \<in> Y \<Longrightarrow> \<exists>x \<in> X. y = f x\<rbrakk> \<Longrightarrow> bij_betw f X Y"
+unfolding bij_betw_def inj_on_def
+apply (rule conjI)
+ apply blast
+by (erule thin_rl) blast
+
+lemma surj_fun_eq:
+ assumes surj_on: "f ` X = UNIV" and eq_on: "\<forall>x \<in> X. (g1 o f) x = (g2 o f) x"
+ shows "g1 = g2"
+proof (rule ext)
+ fix y
+ from surj_on obtain x where "x \<in> X" and "y = f x" by blast
+ thus "g1 y = g2 y" using eq_on by simp
+qed
+
+lemma Card_order_wo_rel: "Card_order r \<Longrightarrow> wo_rel r"
+unfolding wo_rel_def card_order_on_def by blast
+
+lemma Cinfinite_limit: "\<lbrakk>x \<in> Field r; Cinfinite r\<rbrakk> \<Longrightarrow>
+ \<exists>y \<in> Field r. x \<noteq> y \<and> (x, y) \<in> r"
+unfolding cinfinite_def by (auto simp add: infinite_Card_order_limit)
+
+lemma Card_order_trans:
+ "\<lbrakk>Card_order r; x \<noteq> y; (x, y) \<in> r; y \<noteq> z; (y, z) \<in> r\<rbrakk> \<Longrightarrow> x \<noteq> z \<and> (x, z) \<in> r"
+unfolding card_order_on_def well_order_on_def linear_order_on_def
+ partial_order_on_def preorder_on_def trans_def antisym_def by blast
+
+lemma Cinfinite_limit2:
+ assumes x1: "x1 \<in> Field r" and x2: "x2 \<in> Field r" and r: "Cinfinite r"
+ shows "\<exists>y \<in> Field r. (x1 \<noteq> y \<and> (x1, y) \<in> r) \<and> (x2 \<noteq> y \<and> (x2, y) \<in> r)"
+proof -
+ from r have trans: "trans r" and total: "Total r" and antisym: "antisym r"
+ unfolding card_order_on_def well_order_on_def linear_order_on_def
+ partial_order_on_def preorder_on_def by auto
+ obtain y1 where y1: "y1 \<in> Field r" "x1 \<noteq> y1" "(x1, y1) \<in> r"
+ using Cinfinite_limit[OF x1 r] by blast
+ obtain y2 where y2: "y2 \<in> Field r" "x2 \<noteq> y2" "(x2, y2) \<in> r"
+ using Cinfinite_limit[OF x2 r] by blast
+ show ?thesis
+ proof (cases "y1 = y2")
+ case True with y1 y2 show ?thesis by blast
+ next
+ case False
+ with y1(1) y2(1) total have "(y1, y2) \<in> r \<or> (y2, y1) \<in> r"
+ unfolding total_on_def by auto
+ thus ?thesis
+ proof
+ assume *: "(y1, y2) \<in> r"
+ with trans y1(3) have "(x1, y2) \<in> r" unfolding trans_def by blast
+ with False y1 y2 * antisym show ?thesis by (cases "x1 = y2") (auto simp: antisym_def)
+ next
+ assume *: "(y2, y1) \<in> r"
+ with trans y2(3) have "(x2, y1) \<in> r" unfolding trans_def by blast
+ with False y1 y2 * antisym show ?thesis by (cases "x2 = y1") (auto simp: antisym_def)
+ qed
+ qed
+qed
+
+lemma Cinfinite_limit_finite: "\<lbrakk>finite X; X \<subseteq> Field r; Cinfinite r\<rbrakk>
+ \<Longrightarrow> \<exists>y \<in> Field r. \<forall>x \<in> X. (x \<noteq> y \<and> (x, y) \<in> r)"
+proof (induct X rule: finite_induct)
+ case empty thus ?case unfolding cinfinite_def using ex_in_conv[of "Field r"] finite.emptyI by auto
+next
+ case (insert x X)
+ then obtain y where y: "y \<in> Field r" "\<forall>x \<in> X. (x \<noteq> y \<and> (x, y) \<in> r)" by blast
+ then obtain z where z: "z \<in> Field r" "x \<noteq> z \<and> (x, z) \<in> r" "y \<noteq> z \<and> (y, z) \<in> r"
+ using Cinfinite_limit2[OF _ y(1) insert(5), of x] insert(4) by blast
+ show ?case
+ apply (intro bexI ballI)
+ apply (erule insertE)
+ apply hypsubst
+ apply (rule z(2))
+ using Card_order_trans[OF insert(5)[THEN conjunct2]] y(2) z(3)
+ apply blast
+ apply (rule z(1))
+ done
+qed
+
+lemma insert_subsetI: "\<lbrakk>x \<in> A; X \<subseteq> A\<rbrakk> \<Longrightarrow> insert x X \<subseteq> A"
+by auto
+
+(*helps resolution*)
+lemma well_order_induct_imp:
+ "wo_rel r \<Longrightarrow> (\<And>x. \<forall>y. y \<noteq> x \<and> (y, x) \<in> r \<longrightarrow> y \<in> Field r \<longrightarrow> P y \<Longrightarrow> x \<in> Field r \<longrightarrow> P x) \<Longrightarrow>
+ x \<in> Field r \<longrightarrow> P x"
+by (erule wo_rel.well_order_induct)
+
+lemma meta_spec2:
+ assumes "(\<And>x y. PROP P x y)"
+ shows "PROP P x y"
+by (rule `(\<And>x y. PROP P x y)`)
+
+ML_file "Tools/bnf_lfp_util.ML"
+ML_file "Tools/bnf_lfp_tactics.ML"
+ML_file "Tools/bnf_lfp.ML"
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/BNF_Util.thy Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,66 @@
+(* Title: HOL/BNF/BNF_Util.thy
+ Author: Dmitriy Traytel, TU Muenchen
+ Author: Jasmin Blanchette, TU Muenchen
+ Copyright 2012
+
+Library for bounded natural functors.
+*)
+
+header {* Library for Bounded Natural Functors *}
+
+theory BNF_Util
+imports "../Cardinals/Cardinal_Arithmetic"
+begin
+
+lemma subset_Collect_iff: "B \<subseteq> A \<Longrightarrow> (B \<subseteq> {x \<in> A. P x}) = (\<forall>x \<in> B. P x)"
+by blast
+
+lemma subset_CollectI: "B \<subseteq> A \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> Q x \<Longrightarrow> P x) \<Longrightarrow> ({x \<in> B. Q x} \<subseteq> {x \<in> A. P x})"
+by blast
+
+definition collect where
+"collect F x = (\<Union>f \<in> F. f x)"
+
+(* Weak pullbacks: *)
+definition wpull where
+"wpull A B1 B2 f1 f2 p1 p2 \<longleftrightarrow>
+ (\<forall> b1 b2. b1 \<in> B1 \<and> b2 \<in> B2 \<and> f1 b1 = f2 b2 \<longrightarrow> (\<exists> a \<in> A. p1 a = b1 \<and> p2 a = b2))"
+
+(* Weak pseudo-pullbacks *)
+definition wppull where
+"wppull A B1 B2 f1 f2 e1 e2 p1 p2 \<longleftrightarrow>
+ (\<forall> b1 b2. b1 \<in> B1 \<and> b2 \<in> B2 \<and> f1 b1 = f2 b2 \<longrightarrow>
+ (\<exists> a \<in> A. e1 (p1 a) = e1 b1 \<and> e2 (p2 a) = e2 b2))"
+
+lemma fst_snd: "\<lbrakk>snd x = (y, z)\<rbrakk> \<Longrightarrow> fst (snd x) = y"
+by simp
+
+lemma snd_snd: "\<lbrakk>snd x = (y, z)\<rbrakk> \<Longrightarrow> snd (snd x) = z"
+by simp
+
+lemma fstI: "x = (y, z) \<Longrightarrow> fst x = y"
+by simp
+
+lemma sndI: "x = (y, z) \<Longrightarrow> snd x = z"
+by simp
+
+lemma bijI: "\<lbrakk>\<And>x y. (f x = f y) = (x = y); \<And>y. \<exists>x. y = f x\<rbrakk> \<Longrightarrow> bij f"
+unfolding bij_def inj_on_def by auto blast
+
+lemma pair_mem_Collect_split:
+"(\<lambda>x y. (x, y) \<in> {(x, y). P x y}) = P"
+by simp
+
+lemma Collect_pair_mem_eq: "{(x, y). (x, y) \<in> R} = R"
+by simp
+
+lemma Collect_fst_snd_mem_eq: "{p. (fst p, snd p) \<in> A} = A"
+by simp
+
+(* Operator: *)
+definition "Gr A f = {(a, f a) | a. a \<in> A}"
+
+ML_file "Tools/bnf_util.ML"
+ML_file "Tools/bnf_tactics.ML"
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/BNF_Wrap.thy Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,28 @@
+(* Title: HOL/BNF/BNF_Wrap.thy
+ Author: Jasmin Blanchette, TU Muenchen
+ Copyright 2012
+
+Wrapping datatypes.
+*)
+
+header {* Wrapping Datatypes *}
+
+theory BNF_Wrap
+imports BNF_Util
+keywords
+ "wrap_data" :: thy_goal and
+ "no_dests"
+begin
+
+lemma iffI_np: "\<lbrakk>x \<Longrightarrow> \<not> y; \<not> x \<Longrightarrow> y\<rbrakk> \<Longrightarrow> \<not> x \<longleftrightarrow> y"
+by (erule iffI) (erule contrapos_pn)
+
+lemma iff_contradict:
+"\<not> P \<Longrightarrow> P \<longleftrightarrow> Q \<Longrightarrow> Q \<Longrightarrow> R"
+"\<not> Q \<Longrightarrow> P \<longleftrightarrow> Q \<Longrightarrow> P \<Longrightarrow> R"
+by blast+
+
+ML_file "Tools/bnf_wrap_tactics.ML"
+ML_file "Tools/bnf_wrap.ML"
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Basic_BNFs.thy Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,417 @@
+(* Title: HOL/BNF/Basic_BNFs.thy
+ Author: Dmitriy Traytel, TU Muenchen
+ Author: Andrei Popescu, TU Muenchen
+ Author: Jasmin Blanchette, TU Muenchen
+ Copyright 2012
+
+Registration of basic types as bounded natural functors.
+*)
+
+header {* Registration of Basic Types as Bounded Natural Functors *}
+
+theory Basic_BNFs
+imports BNF_Def
+begin
+
+lemma wpull_id: "wpull UNIV B1 B2 id id id id"
+unfolding wpull_def by simp
+
+lemmas natLeq_card_order = natLeq_Card_order[unfolded Field_natLeq]
+
+lemma ctwo_card_order: "card_order ctwo"
+using Card_order_ctwo by (unfold ctwo_def Field_card_of)
+
+lemma natLeq_cinfinite: "cinfinite natLeq"
+unfolding cinfinite_def Field_natLeq by (rule nat_infinite)
+
+bnf_def ID: "id :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" ["\<lambda>x. {x}"] "\<lambda>_:: 'a. natLeq" ["id :: 'a \<Rightarrow> 'a"]
+ "\<lambda>x :: 'a \<Rightarrow> 'b \<Rightarrow> bool. x"
+apply auto
+apply (rule natLeq_card_order)
+apply (rule natLeq_cinfinite)
+apply (rule ordLess_imp_ordLeq[OF finite_ordLess_infinite[OF _ natLeq_Well_order]])
+apply (auto simp add: Field_card_of Field_natLeq card_of_well_order_on)[3]
+apply (rule ordLeq_transitive)
+apply (rule ordLeq_cexp1[of natLeq])
+apply (rule Cinfinite_Cnotzero)
+apply (rule conjI)
+apply (rule natLeq_cinfinite)
+apply (rule natLeq_Card_order)
+apply (rule card_of_Card_order)
+apply (rule cexp_mono1)
+apply (rule ordLeq_csum1)
+apply (rule card_of_Card_order)
+apply (rule disjI2)
+apply (rule cone_ordLeq_cexp)
+apply (rule ordLeq_transitive)
+apply (rule cone_ordLeq_ctwo)
+apply (rule ordLeq_csum2)
+apply (rule Card_order_ctwo)
+apply (rule natLeq_Card_order)
+apply (auto simp: Gr_def fun_eq_iff)
+done
+
+bnf_def DEADID: "id :: 'a \<Rightarrow> 'a" [] "\<lambda>_:: 'a. natLeq +c |UNIV :: 'a set|" ["SOME x :: 'a. True"]
+ "op =::'a \<Rightarrow> 'a \<Rightarrow> bool"
+apply (auto simp add: wpull_id)
+apply (rule card_order_csum)
+apply (rule natLeq_card_order)
+apply (rule card_of_card_order_on)
+apply (rule cinfinite_csum)
+apply (rule disjI1)
+apply (rule natLeq_cinfinite)
+apply (rule ordLess_imp_ordLeq)
+apply (rule ordLess_ordLeq_trans)
+apply (rule ordLess_ctwo_cexp)
+apply (rule card_of_Card_order)
+apply (rule cexp_mono2'')
+apply (rule ordLeq_csum2)
+apply (rule card_of_Card_order)
+apply (rule ctwo_Cnotzero)
+apply (rule card_of_Card_order)
+apply (auto simp: Id_def Gr_def fun_eq_iff)
+done
+
+definition setl :: "'a + 'b \<Rightarrow> 'a set" where
+"setl x = (case x of Inl z => {z} | _ => {})"
+
+definition setr :: "'a + 'b \<Rightarrow> 'b set" where
+"setr x = (case x of Inr z => {z} | _ => {})"
+
+lemmas sum_set_defs = setl_def[abs_def] setr_def[abs_def]
+
+definition sum_rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('c \<Rightarrow> 'd \<Rightarrow> bool) \<Rightarrow> 'a + 'c \<Rightarrow> 'b + 'd \<Rightarrow> bool" where
+"sum_rel \<phi> \<psi> x y =
+ (case x of Inl a1 \<Rightarrow> (case y of Inl a2 \<Rightarrow> \<phi> a1 a2 | Inr _ \<Rightarrow> False)
+ | Inr b1 \<Rightarrow> (case y of Inl _ \<Rightarrow> False | Inr b2 \<Rightarrow> \<psi> b1 b2))"
+
+bnf_def sum_map [setl, setr] "\<lambda>_::'a + 'b. natLeq" [Inl, Inr] sum_rel
+proof -
+ show "sum_map id id = id" by (rule sum_map.id)
+next
+ fix f1 f2 g1 g2
+ show "sum_map (g1 o f1) (g2 o f2) = sum_map g1 g2 o sum_map f1 f2"
+ by (rule sum_map.comp[symmetric])
+next
+ fix x f1 f2 g1 g2
+ assume a1: "\<And>z. z \<in> setl x \<Longrightarrow> f1 z = g1 z" and
+ a2: "\<And>z. z \<in> setr x \<Longrightarrow> f2 z = g2 z"
+ thus "sum_map f1 f2 x = sum_map g1 g2 x"
+ proof (cases x)
+ case Inl thus ?thesis using a1 by (clarsimp simp: setl_def)
+ next
+ case Inr thus ?thesis using a2 by (clarsimp simp: setr_def)
+ qed
+next
+ fix f1 f2
+ show "setl o sum_map f1 f2 = image f1 o setl"
+ by (rule ext, unfold o_apply) (simp add: setl_def split: sum.split)
+next
+ fix f1 f2
+ show "setr o sum_map f1 f2 = image f2 o setr"
+ by (rule ext, unfold o_apply) (simp add: setr_def split: sum.split)
+next
+ show "card_order natLeq" by (rule natLeq_card_order)
+next
+ show "cinfinite natLeq" by (rule natLeq_cinfinite)
+next
+ fix x
+ show "|setl x| \<le>o natLeq"
+ apply (rule ordLess_imp_ordLeq)
+ apply (rule finite_iff_ordLess_natLeq[THEN iffD1])
+ by (simp add: setl_def split: sum.split)
+next
+ fix x
+ show "|setr x| \<le>o natLeq"
+ apply (rule ordLess_imp_ordLeq)
+ apply (rule finite_iff_ordLess_natLeq[THEN iffD1])
+ by (simp add: setr_def split: sum.split)
+next
+ fix A1 :: "'a set" and A2 :: "'b set"
+ have in_alt: "{x. (case x of Inl z => {z} | _ => {}) \<subseteq> A1 \<and>
+ (case x of Inr z => {z} | _ => {}) \<subseteq> A2} = A1 <+> A2" (is "?L = ?R")
+ proof safe
+ fix x :: "'a + 'b"
+ assume "(case x of Inl z \<Rightarrow> {z} | _ \<Rightarrow> {}) \<subseteq> A1" "(case x of Inr z \<Rightarrow> {z} | _ \<Rightarrow> {}) \<subseteq> A2"
+ hence "x \<in> Inl ` A1 \<or> x \<in> Inr ` A2" by (cases x) simp+
+ thus "x \<in> A1 <+> A2" by blast
+ qed (auto split: sum.split)
+ show "|{x. setl x \<subseteq> A1 \<and> setr x \<subseteq> A2}| \<le>o
+ (( |A1| +c |A2| ) +c ctwo) ^c natLeq"
+ apply (rule ordIso_ordLeq_trans)
+ apply (rule card_of_ordIso_subst)
+ apply (unfold sum_set_defs)
+ apply (rule in_alt)
+ apply (rule ordIso_ordLeq_trans)
+ apply (rule Plus_csum)
+ apply (rule ordLeq_transitive)
+ apply (rule ordLeq_csum1)
+ apply (rule Card_order_csum)
+ apply (rule ordLeq_cexp1)
+ apply (rule conjI)
+ using Field_natLeq UNIV_not_empty czeroE apply fast
+ apply (rule natLeq_Card_order)
+ by (rule Card_order_csum)
+next
+ fix A1 A2 B11 B12 B21 B22 f11 f12 f21 f22 p11 p12 p21 p22
+ assume "wpull A1 B11 B21 f11 f21 p11 p21" "wpull A2 B12 B22 f12 f22 p12 p22"
+ hence
+ pull1: "\<And>b1 b2. \<lbrakk>b1 \<in> B11; b2 \<in> B21; f11 b1 = f21 b2\<rbrakk> \<Longrightarrow> \<exists>a \<in> A1. p11 a = b1 \<and> p21 a = b2"
+ and pull2: "\<And>b1 b2. \<lbrakk>b1 \<in> B12; b2 \<in> B22; f12 b1 = f22 b2\<rbrakk> \<Longrightarrow> \<exists>a \<in> A2. p12 a = b1 \<and> p22 a = b2"
+ unfolding wpull_def by blast+
+ show "wpull {x. setl x \<subseteq> A1 \<and> setr x \<subseteq> A2}
+ {x. setl x \<subseteq> B11 \<and> setr x \<subseteq> B12} {x. setl x \<subseteq> B21 \<and> setr x \<subseteq> B22}
+ (sum_map f11 f12) (sum_map f21 f22) (sum_map p11 p12) (sum_map p21 p22)"
+ (is "wpull ?in ?in1 ?in2 ?mapf1 ?mapf2 ?mapp1 ?mapp2")
+ proof (unfold wpull_def)
+ { fix B1 B2
+ assume *: "B1 \<in> ?in1" "B2 \<in> ?in2" "?mapf1 B1 = ?mapf2 B2"
+ have "\<exists>A \<in> ?in. ?mapp1 A = B1 \<and> ?mapp2 A = B2"
+ proof (cases B1)
+ case (Inl b1)
+ { fix b2 assume "B2 = Inr b2"
+ with Inl *(3) have False by simp
+ } then obtain b2 where Inl': "B2 = Inl b2" by (cases B2) (simp, blast)
+ with Inl * have "b1 \<in> B11" "b2 \<in> B21" "f11 b1 = f21 b2"
+ by (simp add: setl_def)+
+ with pull1 obtain a where "a \<in> A1" "p11 a = b1" "p21 a = b2" by blast+
+ with Inl Inl' have "Inl a \<in> ?in" "?mapp1 (Inl a) = B1 \<and> ?mapp2 (Inl a) = B2"
+ by (simp add: sum_set_defs)+
+ thus ?thesis by blast
+ next
+ case (Inr b1)
+ { fix b2 assume "B2 = Inl b2"
+ with Inr *(3) have False by simp
+ } then obtain b2 where Inr': "B2 = Inr b2" by (cases B2) (simp, blast)
+ with Inr * have "b1 \<in> B12" "b2 \<in> B22" "f12 b1 = f22 b2"
+ by (simp add: sum_set_defs)+
+ with pull2 obtain a where "a \<in> A2" "p12 a = b1" "p22 a = b2" by blast+
+ with Inr Inr' have "Inr a \<in> ?in" "?mapp1 (Inr a) = B1 \<and> ?mapp2 (Inr a) = B2"
+ by (simp add: sum_set_defs)+
+ thus ?thesis by blast
+ qed
+ }
+ thus "\<forall>B1 B2. B1 \<in> ?in1 \<and> B2 \<in> ?in2 \<and> ?mapf1 B1 = ?mapf2 B2 \<longrightarrow>
+ (\<exists>A \<in> ?in. ?mapp1 A = B1 \<and> ?mapp2 A = B2)" by fastforce
+ qed
+next
+ fix R S
+ show "{p. sum_rel (\<lambda>x y. (x, y) \<in> R) (\<lambda>x y. (x, y) \<in> S) (fst p) (snd p)} =
+ (Gr {x. setl x \<subseteq> R \<and> setr x \<subseteq> S} (sum_map fst fst))\<inverse> O
+ Gr {x. setl x \<subseteq> R \<and> setr x \<subseteq> S} (sum_map snd snd)"
+ unfolding setl_def setr_def sum_rel_def Gr_def relcomp_unfold converse_unfold
+ by (fastforce split: sum.splits)
+qed (auto simp: sum_set_defs)
+
+lemma singleton_ordLeq_ctwo_natLeq: "|{x}| \<le>o ctwo *c natLeq"
+ apply (rule ordLeq_transitive)
+ apply (rule ordLeq_cprod2)
+ apply (rule ctwo_Cnotzero)
+ apply (auto simp: Field_card_of intro: card_of_card_order_on)
+ apply (rule cprod_mono2)
+ apply (rule ordLess_imp_ordLeq)
+ apply (unfold finite_iff_ordLess_natLeq[symmetric])
+ by simp
+
+definition fsts :: "'a \<times> 'b \<Rightarrow> 'a set" where
+"fsts x = {fst x}"
+
+definition snds :: "'a \<times> 'b \<Rightarrow> 'b set" where
+"snds x = {snd x}"
+
+lemmas prod_set_defs = fsts_def[abs_def] snds_def[abs_def]
+
+definition prod_rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('c \<Rightarrow> 'd \<Rightarrow> bool) \<Rightarrow> 'a \<times> 'c \<Rightarrow> 'b \<times> 'd \<Rightarrow> bool" where
+"prod_rel \<phi> \<psi> p1 p2 = (case p1 of (a1, b1) \<Rightarrow> case p2 of (a2, b2) \<Rightarrow> \<phi> a1 a2 \<and> \<psi> b1 b2)"
+
+bnf_def map_pair [fsts, snds] "\<lambda>_::'a \<times> 'b. ctwo *c natLeq" [Pair] prod_rel
+proof (unfold prod_set_defs)
+ show "map_pair id id = id" by (rule map_pair.id)
+next
+ fix f1 f2 g1 g2
+ show "map_pair (g1 o f1) (g2 o f2) = map_pair g1 g2 o map_pair f1 f2"
+ by (rule map_pair.comp[symmetric])
+next
+ fix x f1 f2 g1 g2
+ assume "\<And>z. z \<in> {fst x} \<Longrightarrow> f1 z = g1 z" "\<And>z. z \<in> {snd x} \<Longrightarrow> f2 z = g2 z"
+ thus "map_pair f1 f2 x = map_pair g1 g2 x" by (cases x) simp
+next
+ fix f1 f2
+ show "(\<lambda>x. {fst x}) o map_pair f1 f2 = image f1 o (\<lambda>x. {fst x})"
+ by (rule ext, unfold o_apply) simp
+next
+ fix f1 f2
+ show "(\<lambda>x. {snd x}) o map_pair f1 f2 = image f2 o (\<lambda>x. {snd x})"
+ by (rule ext, unfold o_apply) simp
+next
+ show "card_order (ctwo *c natLeq)"
+ apply (rule card_order_cprod)
+ apply (rule ctwo_card_order)
+ by (rule natLeq_card_order)
+next
+ show "cinfinite (ctwo *c natLeq)"
+ apply (rule cinfinite_cprod2)
+ apply (rule ctwo_Cnotzero)
+ apply (rule conjI[OF _ natLeq_Card_order])
+ by (rule natLeq_cinfinite)
+next
+ fix x
+ show "|{fst x}| \<le>o ctwo *c natLeq"
+ by (rule singleton_ordLeq_ctwo_natLeq)
+next
+ fix x
+ show "|{snd x}| \<le>o ctwo *c natLeq"
+ by (rule singleton_ordLeq_ctwo_natLeq)
+next
+ fix A1 :: "'a set" and A2 :: "'b set"
+ have in_alt: "{x. {fst x} \<subseteq> A1 \<and> {snd x} \<subseteq> A2} = A1 \<times> A2" by auto
+ show "|{x. {fst x} \<subseteq> A1 \<and> {snd x} \<subseteq> A2}| \<le>o
+ ( ( |A1| +c |A2| ) +c ctwo) ^c (ctwo *c natLeq)"
+ apply (rule ordIso_ordLeq_trans)
+ apply (rule card_of_ordIso_subst)
+ apply (rule in_alt)
+ apply (rule ordIso_ordLeq_trans)
+ apply (rule Times_cprod)
+ apply (rule ordLeq_transitive)
+ apply (rule cprod_csum_cexp)
+ apply (rule cexp_mono)
+ apply (rule ordLeq_csum1)
+ apply (rule Card_order_csum)
+ apply (rule ordLeq_cprod1)
+ apply (rule Card_order_ctwo)
+ apply (rule Cinfinite_Cnotzero)
+ apply (rule conjI[OF _ natLeq_Card_order])
+ apply (rule natLeq_cinfinite)
+ apply (rule disjI2)
+ apply (rule cone_ordLeq_cexp)
+ apply (rule ordLeq_transitive)
+ apply (rule cone_ordLeq_ctwo)
+ apply (rule ordLeq_csum2)
+ apply (rule Card_order_ctwo)
+ apply (rule notE)
+ apply (rule ctwo_not_czero)
+ apply assumption
+ by (rule Card_order_ctwo)
+next
+ fix A1 A2 B11 B12 B21 B22 f11 f12 f21 f22 p11 p12 p21 p22
+ assume "wpull A1 B11 B21 f11 f21 p11 p21" "wpull A2 B12 B22 f12 f22 p12 p22"
+ thus "wpull {x. {fst x} \<subseteq> A1 \<and> {snd x} \<subseteq> A2}
+ {x. {fst x} \<subseteq> B11 \<and> {snd x} \<subseteq> B12} {x. {fst x} \<subseteq> B21 \<and> {snd x} \<subseteq> B22}
+ (map_pair f11 f12) (map_pair f21 f22) (map_pair p11 p12) (map_pair p21 p22)"
+ unfolding wpull_def by simp fast
+next
+ fix R S
+ show "{p. prod_rel (\<lambda>x y. (x, y) \<in> R) (\<lambda>x y. (x, y) \<in> S) (fst p) (snd p)} =
+ (Gr {x. {fst x} \<subseteq> R \<and> {snd x} \<subseteq> S} (map_pair fst fst))\<inverse> O
+ Gr {x. {fst x} \<subseteq> R \<and> {snd x} \<subseteq> S} (map_pair snd snd)"
+ unfolding prod_set_defs prod_rel_def Gr_def relcomp_unfold converse_unfold
+ by auto
+qed simp+
+
+(* Categorical version of pullback: *)
+lemma wpull_cat:
+assumes p: "wpull A B1 B2 f1 f2 p1 p2"
+and c: "f1 o q1 = f2 o q2"
+and r: "range q1 \<subseteq> B1" "range q2 \<subseteq> B2"
+obtains h where "range h \<subseteq> A \<and> q1 = p1 o h \<and> q2 = p2 o h"
+proof-
+ have *: "\<forall>d. \<exists>a \<in> A. p1 a = q1 d & p2 a = q2 d"
+ proof safe
+ fix d
+ have "f1 (q1 d) = f2 (q2 d)" using c unfolding comp_def[abs_def] by (rule fun_cong)
+ moreover
+ have "q1 d : B1" "q2 d : B2" using r unfolding image_def by auto
+ ultimately show "\<exists>a \<in> A. p1 a = q1 d \<and> p2 a = q2 d"
+ using p unfolding wpull_def by auto
+ qed
+ then obtain h where "!! d. h d \<in> A & p1 (h d) = q1 d & p2 (h d) = q2 d" by metis
+ thus ?thesis using that by fastforce
+qed
+
+lemma card_of_bounded_range:
+ "|{f :: 'd \<Rightarrow> 'a. range f \<subseteq> B}| \<le>o |Func (UNIV :: 'd set) B|" (is "|?LHS| \<le>o |?RHS|")
+proof -
+ let ?f = "\<lambda>f. %x. if f x \<in> B then Some (f x) else None"
+ have "inj_on ?f ?LHS" unfolding inj_on_def
+ proof (unfold fun_eq_iff, safe)
+ fix g :: "'d \<Rightarrow> 'a" and f :: "'d \<Rightarrow> 'a" and x
+ assume "range f \<subseteq> B" "range g \<subseteq> B" and eq: "\<forall>x. ?f f x = ?f g x"
+ hence "f x \<in> B" "g x \<in> B" by auto
+ with eq have "Some (f x) = Some (g x)" by metis
+ thus "f x = g x" by simp
+ qed
+ moreover have "?f ` ?LHS \<subseteq> ?RHS" unfolding Func_def by fastforce
+ ultimately show ?thesis using card_of_ordLeq by fast
+qed
+
+definition fun_rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('c \<Rightarrow> 'a) \<Rightarrow> ('c \<Rightarrow> 'b) \<Rightarrow> bool" where
+"fun_rel \<phi> f g = (\<forall>x. \<phi> (f x) (g x))"
+
+bnf_def "op \<circ>" [range] "\<lambda>_:: 'a \<Rightarrow> 'b. natLeq +c |UNIV :: 'a set|" ["%c x::'b::type. c::'a::type"]
+ fun_rel
+proof
+ fix f show "id \<circ> f = id f" by simp
+next
+ fix f g show "op \<circ> (g \<circ> f) = op \<circ> g \<circ> op \<circ> f"
+ unfolding comp_def[abs_def] ..
+next
+ fix x f g
+ assume "\<And>z. z \<in> range x \<Longrightarrow> f z = g z"
+ thus "f \<circ> x = g \<circ> x" by auto
+next
+ fix f show "range \<circ> op \<circ> f = op ` f \<circ> range"
+ unfolding image_def comp_def[abs_def] by auto
+next
+ show "card_order (natLeq +c |UNIV| )" (is "_ (_ +c ?U)")
+ apply (rule card_order_csum)
+ apply (rule natLeq_card_order)
+ by (rule card_of_card_order_on)
+(* *)
+ show "cinfinite (natLeq +c ?U)"
+ apply (rule cinfinite_csum)
+ apply (rule disjI1)
+ by (rule natLeq_cinfinite)
+next
+ fix f :: "'d => 'a"
+ have "|range f| \<le>o | (UNIV::'d set) |" (is "_ \<le>o ?U") by (rule card_of_image)
+ also have "?U \<le>o natLeq +c ?U" by (rule ordLeq_csum2) (rule card_of_Card_order)
+ finally show "|range f| \<le>o natLeq +c ?U" .
+next
+ fix B :: "'a set"
+ have "|{f::'d => 'a. range f \<subseteq> B}| \<le>o |Func (UNIV :: 'd set) B|" by (rule card_of_bounded_range)
+ also have "|Func (UNIV :: 'd set) B| =o |B| ^c |UNIV :: 'd set|"
+ unfolding cexp_def Field_card_of by (rule card_of_refl)
+ also have "|B| ^c |UNIV :: 'd set| \<le>o
+ ( |B| +c ctwo) ^c (natLeq +c |UNIV :: 'd set| )"
+ apply (rule cexp_mono)
+ apply (rule ordLeq_csum1) apply (rule card_of_Card_order)
+ apply (rule ordLeq_csum2) apply (rule card_of_Card_order)
+ apply (rule disjI2) apply (rule cone_ordLeq_cexp)
+ apply (rule ordLeq_transitive) apply (rule cone_ordLeq_ctwo) apply (rule ordLeq_csum2)
+ apply (rule Card_order_ctwo)
+ apply (rule notE) apply (rule conjunct1) apply (rule Cnotzero_UNIV) apply blast
+ apply (rule card_of_Card_order)
+ done
+ finally
+ show "|{f::'d => 'a. range f \<subseteq> B}| \<le>o
+ ( |B| +c ctwo) ^c (natLeq +c |UNIV :: 'd set| )" .
+next
+ fix A B1 B2 f1 f2 p1 p2 assume p: "wpull A B1 B2 f1 f2 p1 p2"
+ show "wpull {h. range h \<subseteq> A} {g1. range g1 \<subseteq> B1} {g2. range g2 \<subseteq> B2}
+ (op \<circ> f1) (op \<circ> f2) (op \<circ> p1) (op \<circ> p2)"
+ unfolding wpull_def
+ proof safe
+ fix g1 g2 assume r: "range g1 \<subseteq> B1" "range g2 \<subseteq> B2"
+ and c: "f1 \<circ> g1 = f2 \<circ> g2"
+ show "\<exists>h \<in> {h. range h \<subseteq> A}. p1 \<circ> h = g1 \<and> p2 \<circ> h = g2"
+ using wpull_cat[OF p c r] by simp metis
+ qed
+next
+ fix R
+ show "{p. fun_rel (\<lambda>x y. (x, y) \<in> R) (fst p) (snd p)} =
+ (Gr {x. range x \<subseteq> R} (op \<circ> fst))\<inverse> O Gr {x. range x \<subseteq> R} (op \<circ> snd)"
+ unfolding fun_rel_def Gr_def relcomp_unfold converse_unfold
+ by (auto intro!: exI dest!: in_mono)
+qed auto
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Countable_Set.thy Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,366 @@
+(* Title: HOL/BNF/Countable_Set.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+(At most) countable sets.
+*)
+
+header {* (At Most) Countable Sets *}
+
+theory Countable_Set
+imports "../Cardinals/Cardinals" "~~/src/HOL/Library/Countable"
+begin
+
+
+subsection{* Basics *}
+
+definition "countable A \<equiv> |A| \<le>o natLeq"
+
+lemma countable_card_of_nat:
+"countable A \<longleftrightarrow> |A| \<le>o |UNIV::nat set|"
+unfolding countable_def using card_of_nat
+using ordLeq_ordIso_trans ordIso_symmetric by blast
+
+lemma countable_ex_to_nat:
+fixes A :: "'a set"
+shows "countable A \<longleftrightarrow> (\<exists> f::'a\<Rightarrow>nat. inj_on f A)"
+unfolding countable_card_of_nat card_of_ordLeq[symmetric] by auto
+
+lemma countable_or_card_of:
+assumes "countable A"
+shows "(finite A \<and> |A| <o |UNIV::nat set| ) \<or>
+ (infinite A \<and> |A| =o |UNIV::nat set| )"
+apply (cases "finite A")
+ apply(metis finite_iff_cardOf_nat)
+ by (metis assms countable_card_of_nat infinite_iff_card_of_nat ordIso_iff_ordLeq)
+
+lemma countable_or:
+assumes "countable A"
+shows "(\<exists> f::'a\<Rightarrow>nat. finite A \<and> inj_on f A) \<or>
+ (\<exists> f::'a\<Rightarrow>nat. infinite A \<and> bij_betw f A UNIV)"
+using countable_or_card_of[OF assms]
+by (metis assms card_of_ordIso countable_ex_to_nat)
+
+lemma countable_cases_card_of[elim, consumes 1, case_names Fin Inf]:
+assumes "countable A"
+and "\<lbrakk>finite A; |A| <o |UNIV::nat set|\<rbrakk> \<Longrightarrow> phi"
+and "\<lbrakk>infinite A; |A| =o |UNIV::nat set|\<rbrakk> \<Longrightarrow> phi"
+shows phi
+using assms countable_or_card_of by blast
+
+lemma countable_cases[elim, consumes 1, case_names Fin Inf]:
+assumes "countable A"
+and "\<And> f::'a\<Rightarrow>nat. \<lbrakk>finite A; inj_on f A\<rbrakk> \<Longrightarrow> phi"
+and "\<And> f::'a\<Rightarrow>nat. \<lbrakk>infinite A; bij_betw f A UNIV\<rbrakk> \<Longrightarrow> phi"
+shows phi
+using assms countable_or by metis
+
+definition toNat_pred :: "'a set \<Rightarrow> ('a \<Rightarrow> nat) \<Rightarrow> bool"
+where
+"toNat_pred (A::'a set) f \<equiv>
+ (finite A \<and> inj_on f A) \<or> (infinite A \<and> bij_betw f A UNIV)"
+definition toNat where "toNat A \<equiv> SOME f. toNat_pred A f"
+
+lemma toNat_pred:
+assumes "countable A"
+shows "\<exists> f. toNat_pred A f"
+using assms countable_ex_to_nat toNat_pred_def by (cases rule: countable_cases) auto
+
+lemma toNat_pred_toNat:
+assumes "countable A"
+shows "toNat_pred A (toNat A)"
+unfolding toNat_def apply(rule someI_ex[of "toNat_pred A"])
+using toNat_pred[OF assms] .
+
+lemma bij_betw_toNat:
+assumes c: "countable A" and i: "infinite A"
+shows "bij_betw (toNat A) A (UNIV::nat set)"
+using toNat_pred_toNat[OF c] unfolding toNat_pred_def using i by auto
+
+lemma inj_on_toNat:
+assumes c: "countable A"
+shows "inj_on (toNat A) A"
+using c apply(cases rule: countable_cases)
+using bij_betw_toNat[OF c] toNat_pred_toNat[OF c]
+unfolding toNat_pred_def unfolding bij_betw_def by auto
+
+lemma toNat_inj[simp]:
+assumes c: "countable A" and a: "a \<in> A" and b: "b \<in> A"
+shows "toNat A a = toNat A b \<longleftrightarrow> a = b"
+using inj_on_toNat[OF c] using a b unfolding inj_on_def by auto
+
+lemma image_toNat:
+assumes c: "countable A" and i: "infinite A"
+shows "toNat A ` A = UNIV"
+using bij_betw_toNat[OF assms] unfolding bij_betw_def by simp
+
+lemma toNat_surj:
+assumes "countable A" and i: "infinite A"
+shows "\<exists> a. a \<in> A \<and> toNat A a = n"
+using image_toNat[OF assms]
+by (metis (no_types) image_iff iso_tuple_UNIV_I)
+
+definition
+"fromNat A n \<equiv>
+ if n \<in> toNat A ` A then inv_into A (toNat A) n
+ else (SOME a. a \<in> A)"
+
+lemma fromNat:
+assumes "A \<noteq> {}"
+shows "fromNat A n \<in> A"
+unfolding fromNat_def by (metis assms equals0I inv_into_into someI_ex)
+
+lemma toNat_fromNat[simp]:
+assumes "n \<in> toNat A ` A"
+shows "toNat A (fromNat A n) = n"
+by (metis assms f_inv_into_f fromNat_def)
+
+lemma infinite_toNat_fromNat[simp]:
+assumes c: "countable A" and i: "infinite A"
+shows "toNat A (fromNat A n) = n"
+apply(rule toNat_fromNat) using toNat_surj[OF assms]
+by (metis image_iff)
+
+lemma fromNat_toNat[simp]:
+assumes c: "countable A" and a: "a \<in> A"
+shows "fromNat A (toNat A a) = a"
+by (metis a c equals0D fromNat imageI toNat_fromNat toNat_inj)
+
+lemma fromNat_inj:
+assumes c: "countable A" and i: "infinite A"
+shows "fromNat A m = fromNat A n \<longleftrightarrow> m = n" (is "?L = ?R \<longleftrightarrow> ?K")
+proof-
+ have "?L = ?R \<longleftrightarrow> toNat A ?L = toNat A ?R"
+ unfolding toNat_inj[OF c fromNat[OF infinite_imp_nonempty[OF i]]
+ fromNat[OF infinite_imp_nonempty[OF i]]] ..
+ also have "... \<longleftrightarrow> ?K" using c i by simp
+ finally show ?thesis .
+qed
+
+lemma fromNat_surj:
+assumes c: "countable A" and a: "a \<in> A"
+shows "\<exists> n. fromNat A n = a"
+apply(rule exI[of _ "toNat A a"]) using assms by simp
+
+lemma fromNat_image_incl:
+assumes "A \<noteq> {}"
+shows "fromNat A ` UNIV \<subseteq> A"
+using fromNat[OF assms] by auto
+
+lemma incl_fromNat_image:
+assumes "countable A"
+shows "A \<subseteq> fromNat A ` UNIV"
+unfolding image_def using fromNat_surj[OF assms] by auto
+
+lemma fromNat_image[simp]:
+assumes "A \<noteq> {}" and "countable A"
+shows "fromNat A ` UNIV = A"
+by (metis assms equalityI fromNat_image_incl incl_fromNat_image)
+
+lemma fromNat_inject[simp]:
+assumes A: "A \<noteq> {}" "countable A" and B: "B \<noteq> {}" "countable B"
+shows "fromNat A = fromNat B \<longleftrightarrow> A = B"
+by (metis assms fromNat_image)
+
+lemma inj_on_fromNat:
+"inj_on fromNat ({A. A \<noteq> {} \<and> countable A})"
+unfolding inj_on_def by auto
+
+
+subsection {* Preservation under the set theoretic operations *}
+
+lemma contable_empty[simp,intro]:
+"countable {}"
+by (metis countable_ex_to_nat inj_on_empty)
+
+lemma incl_countable:
+assumes "A \<subseteq> B" and "countable B"
+shows "countable A"
+by (metis assms countable_ex_to_nat subset_inj_on)
+
+lemma countable_diff:
+assumes "countable A"
+shows "countable (A - B)"
+by (metis Diff_subset assms incl_countable)
+
+lemma finite_countable[simp]:
+assumes "finite A"
+shows "countable A"
+by (metis assms countable_ex_to_nat finite_imp_inj_to_nat_seg)
+
+lemma countable_singl[simp]:
+"countable {a}"
+by simp
+
+lemma countable_insert[simp]:
+"countable (insert a A) \<longleftrightarrow> countable A"
+proof
+ assume c: "countable A"
+ thus "countable (insert a A)"
+ apply (cases rule: countable_cases_card_of)
+ apply (metis finite_countable finite_insert)
+ unfolding countable_card_of_nat
+ by (metis infinite_card_of_insert ordIso_imp_ordLeq ordIso_transitive)
+qed(insert incl_countable, metis incl_countable subset_insertI)
+
+lemma contable_IntL[simp]:
+assumes "countable A"
+shows "countable (A \<inter> B)"
+by (metis Int_lower1 assms incl_countable)
+
+lemma contable_IntR[simp]:
+assumes "countable B"
+shows "countable (A \<inter> B)"
+by (metis assms contable_IntL inf.commute)
+
+lemma countable_UN[simp]:
+assumes cI: "countable I" and cA: "\<And> i. i \<in> I \<Longrightarrow> countable (A i)"
+shows "countable (\<Union> i \<in> I. A i)"
+using assms unfolding countable_card_of_nat
+apply(intro card_of_UNION_ordLeq_infinite) by auto
+
+lemma contable_Un[simp]:
+"countable (A \<union> B) \<longleftrightarrow> countable A \<and> countable B"
+proof safe
+ assume cA: "countable A" and cB: "countable B"
+ let ?I = "{0,Suc 0}" let ?As = "\<lambda> i. case i of 0 \<Rightarrow> A|Suc 0 \<Rightarrow> B"
+ have AB: "A \<union> B = (\<Union> i \<in> ?I. ?As i)" by simp
+ show "countable (A \<union> B)" unfolding AB apply(rule countable_UN)
+ using cA cB by auto
+qed (metis Un_upper1 incl_countable, metis Un_upper2 incl_countable)
+
+lemma countable_INT[simp]:
+assumes "i \<in> I" and "countable (A i)"
+shows "countable (\<Inter> i \<in> I. A i)"
+by (metis INF_insert assms contable_IntL insert_absorb)
+
+lemma countable_class[simp]:
+fixes A :: "('a::countable) set"
+shows "countable A"
+proof-
+ have "inj_on to_nat A" by (metis inj_on_to_nat)
+ thus ?thesis by (metis countable_ex_to_nat)
+qed
+
+lemma countable_image[simp]:
+assumes "countable A"
+shows "countable (f ` A)"
+using assms unfolding countable_card_of_nat
+by (metis card_of_image ordLeq_transitive)
+
+lemma countable_ordLeq:
+assumes "|A| \<le>o |B|" and "countable B"
+shows "countable A"
+using assms unfolding countable_card_of_nat by(rule ordLeq_transitive)
+
+lemma countable_ordLess:
+assumes AB: "|A| <o |B|" and B: "countable B"
+shows "countable A"
+using countable_ordLeq[OF ordLess_imp_ordLeq[OF AB] B] .
+
+lemma countable_vimage:
+assumes "B \<subseteq> range f" and "countable (f -` B)"
+shows "countable B"
+by (metis Int_absorb2 assms countable_image image_vimage_eq)
+
+lemma surj_countable_vimage:
+assumes s: "surj f" and c: "countable (f -` B)"
+shows "countable B"
+apply(rule countable_vimage[OF _ c]) using s by auto
+
+lemma countable_Collect[simp]:
+assumes "countable A"
+shows "countable {a \<in> A. \<phi> a}"
+by (metis Collect_conj_eq Int_absorb Int_commute Int_def assms contable_IntR)
+
+lemma countable_Plus[simp]:
+assumes A: "countable A" and B: "countable B"
+shows "countable (A <+> B)"
+proof-
+ let ?U = "UNIV::nat set"
+ have "|A| \<le>o |?U|" and "|B| \<le>o |?U|" using A B
+ using card_of_nat[THEN ordIso_symmetric] ordLeq_ordIso_trans
+ unfolding countable_def by blast+
+ hence "|A <+> B| \<le>o |?U|" by (intro card_of_Plus_ordLeq_infinite) auto
+ thus ?thesis using card_of_nat unfolding countable_def by(rule ordLeq_ordIso_trans)
+qed
+
+lemma countable_Times[simp]:
+assumes A: "countable A" and B: "countable B"
+shows "countable (A \<times> B)"
+proof-
+ let ?U = "UNIV::nat set"
+ have "|A| \<le>o |?U|" and "|B| \<le>o |?U|" using A B
+ using card_of_nat[THEN ordIso_symmetric] ordLeq_ordIso_trans
+ unfolding countable_def by blast+
+ hence "|A \<times> B| \<le>o |?U|" by (intro card_of_Times_ordLeq_infinite) auto
+ thus ?thesis using card_of_nat unfolding countable_def by(rule ordLeq_ordIso_trans)
+qed
+
+lemma ordLeq_countable:
+assumes "|A| \<le>o |B|" and "countable B"
+shows "countable A"
+using assms unfolding countable_def by(rule ordLeq_transitive)
+
+lemma ordLess_countable:
+assumes A: "|A| <o |B|" and B: "countable B"
+shows "countable A"
+by (rule ordLeq_countable[OF ordLess_imp_ordLeq[OF A] B])
+
+lemma countable_def2: "countable A \<longleftrightarrow> |A| \<le>o |UNIV :: nat set|"
+unfolding countable_def using card_of_nat[THEN ordIso_symmetric]
+by (metis (lifting) Field_card_of Field_natLeq card_of_mono2 card_of_nat
+ countable_def ordIso_imp_ordLeq ordLeq_countable)
+
+
+subsection{* The type of countable sets *}
+
+typedef (open) 'a cset = "{A :: 'a set. countable A}"
+apply(rule exI[of _ "{}"]) by simp
+
+abbreviation rcset where "rcset \<equiv> Rep_cset"
+abbreviation acset where "acset \<equiv> Abs_cset"
+
+lemmas acset_rcset = Rep_cset_inverse
+declare acset_rcset[simp]
+
+lemma acset_surj:
+"\<exists> A. countable A \<and> acset A = C"
+apply(cases rule: Abs_cset_cases[of C]) by auto
+
+lemma rcset_acset[simp]:
+assumes "countable A"
+shows "rcset (acset A) = A"
+using Abs_cset_inverse assms by auto
+
+lemma acset_inj[simp]:
+assumes "countable A" and "countable B"
+shows "acset A = acset B \<longleftrightarrow> A = B"
+using assms Abs_cset_inject by auto
+
+lemma rcset[simp]:
+"countable (rcset C)"
+using Rep_cset by simp
+
+lemma rcset_inj[simp]:
+"rcset C = rcset D \<longleftrightarrow> C = D"
+by (metis acset_rcset)
+
+lemma rcset_surj:
+assumes "countable A"
+shows "\<exists> C. rcset C = A"
+apply(cases rule: Rep_cset_cases[of A])
+using assms by auto
+
+definition "cIn a C \<equiv> (a \<in> rcset C)"
+definition "cEmp \<equiv> acset {}"
+definition "cIns a C \<equiv> acset (insert a (rcset C))"
+abbreviation cSingl where "cSingl a \<equiv> cIns a cEmp"
+definition "cUn C D \<equiv> acset (rcset C \<union> rcset D)"
+definition "cInt C D \<equiv> acset (rcset C \<inter> rcset D)"
+definition "cDif C D \<equiv> acset (rcset C - rcset D)"
+definition "cIm f C \<equiv> acset (f ` rcset C)"
+definition "cVim f D \<equiv> acset (f -` rcset D)"
+(* TODO eventually: nice setup for these operations, copied from the set setup *)
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Equiv_Relations_More.thy Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,161 @@
+(* Title: HOL/BNF/Equiv_Relations_More.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Some preliminaries on equivalence relations and quotients.
+*)
+
+header {* Some Preliminaries on Equivalence Relations and Quotients *}
+
+theory Equiv_Relations_More
+imports Equiv_Relations Hilbert_Choice
+begin
+
+
+(* Recall the following constants and lemmas:
+
+term Eps
+term "A//r"
+lemmas equiv_def
+lemmas refl_on_def
+ -- note that "reflexivity on" also assumes inclusion of the relation's field into r
+
+*)
+
+definition proj where "proj r x = r `` {x}"
+
+definition univ where "univ f X == f (Eps (%x. x \<in> X))"
+
+lemma proj_preserves:
+"x \<in> A \<Longrightarrow> proj r x \<in> A//r"
+unfolding proj_def by (rule quotientI)
+
+lemma proj_in_iff:
+assumes "equiv A r"
+shows "(proj r x \<in> A//r) = (x \<in> A)"
+apply(rule iffI, auto simp add: proj_preserves)
+unfolding proj_def quotient_def proof clarsimp
+ fix y assume y: "y \<in> A" and "r `` {x} = r `` {y}"
+ moreover have "y \<in> r `` {y}" using assms y unfolding equiv_def refl_on_def by blast
+ ultimately have "(x,y) \<in> r" by blast
+ thus "x \<in> A" using assms unfolding equiv_def refl_on_def by blast
+qed
+
+lemma proj_iff:
+"\<lbrakk>equiv A r; {x,y} \<subseteq> A\<rbrakk> \<Longrightarrow> (proj r x = proj r y) = ((x,y) \<in> r)"
+by (simp add: proj_def eq_equiv_class_iff)
+
+(*
+lemma in_proj: "\<lbrakk>equiv A r; x \<in> A\<rbrakk> \<Longrightarrow> x \<in> proj r x"
+unfolding proj_def equiv_def refl_on_def by blast
+*)
+
+lemma proj_image: "(proj r) ` A = A//r"
+unfolding proj_def[abs_def] quotient_def by blast
+
+lemma in_quotient_imp_non_empty:
+"\<lbrakk>equiv A r; X \<in> A//r\<rbrakk> \<Longrightarrow> X \<noteq> {}"
+unfolding quotient_def using equiv_class_self by fast
+
+lemma in_quotient_imp_in_rel:
+"\<lbrakk>equiv A r; X \<in> A//r; {x,y} \<subseteq> X\<rbrakk> \<Longrightarrow> (x,y) \<in> r"
+using quotient_eq_iff by fastforce
+
+lemma in_quotient_imp_closed:
+"\<lbrakk>equiv A r; X \<in> A//r; x \<in> X; (x,y) \<in> r\<rbrakk> \<Longrightarrow> y \<in> X"
+unfolding quotient_def equiv_def trans_def by blast
+
+lemma in_quotient_imp_subset:
+"\<lbrakk>equiv A r; X \<in> A//r\<rbrakk> \<Longrightarrow> X \<subseteq> A"
+using assms in_quotient_imp_in_rel equiv_type by fastforce
+
+lemma equiv_Eps_in:
+"\<lbrakk>equiv A r; X \<in> A//r\<rbrakk> \<Longrightarrow> Eps (%x. x \<in> X) \<in> X"
+apply (rule someI2_ex)
+using in_quotient_imp_non_empty by blast
+
+lemma equiv_Eps_preserves:
+assumes ECH: "equiv A r" and X: "X \<in> A//r"
+shows "Eps (%x. x \<in> X) \<in> A"
+apply (rule in_mono[rule_format])
+ using assms apply (rule in_quotient_imp_subset)
+by (rule equiv_Eps_in) (rule assms)+
+
+lemma proj_Eps:
+assumes "equiv A r" and "X \<in> A//r"
+shows "proj r (Eps (%x. x \<in> X)) = X"
+unfolding proj_def proof auto
+ fix x assume x: "x \<in> X"
+ thus "(Eps (%x. x \<in> X), x) \<in> r" using assms equiv_Eps_in in_quotient_imp_in_rel by fast
+next
+ fix x assume "(Eps (%x. x \<in> X),x) \<in> r"
+ thus "x \<in> X" using in_quotient_imp_closed[OF assms equiv_Eps_in[OF assms]] by fast
+qed
+
+(*
+lemma Eps_proj:
+assumes "equiv A r" and "x \<in> A"
+shows "(Eps (%y. y \<in> proj r x), x) \<in> r"
+proof-
+ have 1: "proj r x \<in> A//r" using assms proj_preserves by fastforce
+ hence "Eps(%y. y \<in> proj r x) \<in> proj r x" using assms equiv_Eps_in by auto
+ moreover have "x \<in> proj r x" using assms in_proj by fastforce
+ ultimately show ?thesis using assms 1 in_quotient_imp_in_rel by fastforce
+qed
+
+lemma equiv_Eps_iff:
+assumes "equiv A r" and "{X,Y} \<subseteq> A//r"
+shows "((Eps (%x. x \<in> X),Eps (%y. y \<in> Y)) \<in> r) = (X = Y)"
+proof-
+ have "Eps (%x. x \<in> X) \<in> X \<and> Eps (%y. y \<in> Y) \<in> Y" using assms equiv_Eps_in by auto
+ thus ?thesis using assms quotient_eq_iff by fastforce
+qed
+
+lemma equiv_Eps_inj_on:
+assumes "equiv A r"
+shows "inj_on (%X. Eps (%x. x \<in> X)) (A//r)"
+unfolding inj_on_def proof clarify
+ fix X Y assume X: "X \<in> A//r" and Y: "Y \<in> A//r" and Eps: "Eps (%x. x \<in> X) = Eps (%y. y \<in> Y)"
+ hence "Eps (%x. x \<in> X) \<in> A" using assms equiv_Eps_preserves by auto
+ hence "(Eps (%x. x \<in> X), Eps (%y. y \<in> Y)) \<in> r"
+ using assms Eps unfolding quotient_def equiv_def refl_on_def by auto
+ thus "X= Y" using X Y assms equiv_Eps_iff by auto
+qed
+*)
+
+lemma univ_commute:
+assumes ECH: "equiv A r" and RES: "f respects r" and x: "x \<in> A"
+shows "(univ f) (proj r x) = f x"
+unfolding univ_def proof -
+ have prj: "proj r x \<in> A//r" using x proj_preserves by fast
+ hence "Eps (%y. y \<in> proj r x) \<in> A" using ECH equiv_Eps_preserves by fast
+ moreover have "proj r (Eps (%y. y \<in> proj r x)) = proj r x" using ECH prj proj_Eps by fast
+ ultimately have "(x, Eps (%y. y \<in> proj r x)) \<in> r" using x ECH proj_iff by fast
+ thus "f (Eps (%y. y \<in> proj r x)) = f x" using RES unfolding congruent_def by fastforce
+qed
+
+(*
+lemma univ_unique:
+assumes ECH: "equiv A r" and
+ RES: "f respects r" and COM: "\<forall> x \<in> A. G (proj r x) = f x"
+shows "\<forall> X \<in> A//r. G X = univ f X"
+proof
+ fix X assume "X \<in> A//r"
+ then obtain x where x: "x \<in> A" and X: "X = proj r x" using ECH proj_image[of r A] by blast
+ have "G X = f x" unfolding X using x COM by simp
+ thus "G X = univ f X" unfolding X using ECH RES x univ_commute by fastforce
+qed
+*)
+
+lemma univ_preserves:
+assumes ECH: "equiv A r" and RES: "f respects r" and
+ PRES: "\<forall> x \<in> A. f x \<in> B"
+shows "\<forall> X \<in> A//r. univ f X \<in> B"
+proof
+ fix X assume "X \<in> A//r"
+ then obtain x where x: "x \<in> A" and X: "X = proj r x" using ECH proj_image[of r A] by blast
+ hence "univ f X = f x" using assms univ_commute by fastforce
+ thus "univ f X \<in> B" using x PRES by simp
+qed
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Examples/HFset.thy Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,60 @@
+(* Title: HOL/BNF/Examples/HFset.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Hereditary sets.
+*)
+
+header {* Hereditary Sets *}
+
+theory HFset
+imports "../BNF"
+begin
+
+
+section {* Datatype definition *}
+
+data_raw hfset: 'hfset = "'hfset fset"
+
+
+section {* Customization of terms *}
+
+subsection{* Constructors *}
+
+definition "Fold hs \<equiv> hfset_ctor hs"
+
+lemma hfset_simps[simp]:
+"\<And>hs1 hs2. Fold hs1 = Fold hs2 \<longrightarrow> hs1 = hs2"
+unfolding Fold_def hfset.ctor_inject by auto
+
+theorem hfset_cases[elim, case_names Fold]:
+assumes Fold: "\<And> hs. h = Fold hs \<Longrightarrow> phi"
+shows phi
+using Fold unfolding Fold_def
+by (cases rule: hfset.ctor_exhaust[of h]) simp
+
+lemma hfset_induct[case_names Fold, induct type: hfset]:
+assumes Fold: "\<And> hs. (\<And> h. h |\<in>| hs \<Longrightarrow> phi h) \<Longrightarrow> phi (Fold hs)"
+shows "phi t"
+apply (induct rule: hfset.ctor_induct)
+using Fold unfolding Fold_def fset_fset_member mem_Collect_eq ..
+
+(* alternative induction principle, using fset: *)
+lemma hfset_induct_fset[case_names Fold, induct type: hfset]:
+assumes Fold: "\<And> hs. (\<And> h. h \<in> fset hs \<Longrightarrow> phi h) \<Longrightarrow> phi (Fold hs)"
+shows "phi t"
+apply (induct rule: hfset_induct)
+using Fold by (metis notin_fset)
+
+subsection{* Recursion and iteration (fold) *}
+
+lemma hfset_ctor_rec:
+"hfset_ctor_rec R (Fold hs) = R (map_fset <id, hfset_ctor_rec R> hs)"
+using hfset.ctor_recs unfolding Fold_def .
+
+(* The iterator has a simpler form: *)
+lemma hfset_ctor_fold:
+"hfset_ctor_fold R (Fold hs) = R (map_fset (hfset_ctor_fold R) hs)"
+using hfset.ctor_folds unfolding Fold_def .
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Examples/Infinite_Derivation_Trees/Gram_Lang.thy Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,1366 @@
+(* Title: HOL/BNF/Examples/Infinite_Derivation_Trees/Gram_Lang.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Language of a grammar.
+*)
+
+header {* Language of a Grammar *}
+
+theory Gram_Lang
+imports Tree
+begin
+
+
+consts P :: "(N \<times> (T + N) set) set"
+axiomatization where
+ finite_N: "finite (UNIV::N set)"
+and finite_in_P: "\<And> n tns. (n,tns) \<in> P \<longrightarrow> finite tns"
+and used: "\<And> n. \<exists> tns. (n,tns) \<in> P"
+
+
+subsection{* Tree basics: frontier, interior, etc. *}
+
+lemma Tree_cong:
+assumes "root tr = root tr'" and "cont tr = cont tr'"
+shows "tr = tr'"
+by (metis Node_root_cont assms)
+
+inductive finiteT where
+Node: "\<lbrakk>finite as; (finiteT^#) as\<rbrakk> \<Longrightarrow> finiteT (Node a as)"
+monos lift_mono
+
+lemma finiteT_induct[consumes 1, case_names Node, induct pred: finiteT]:
+assumes 1: "finiteT tr"
+and IH: "\<And>as n. \<lbrakk>finite as; (\<phi>^#) as\<rbrakk> \<Longrightarrow> \<phi> (Node n as)"
+shows "\<phi> tr"
+using 1 apply(induct rule: finiteT.induct)
+apply(rule IH) apply assumption apply(elim mono_lift) by simp
+
+
+(* Frontier *)
+
+inductive inFr :: "N set \<Rightarrow> Tree \<Rightarrow> T \<Rightarrow> bool" where
+Base: "\<lbrakk>root tr \<in> ns; Inl t \<in> cont tr\<rbrakk> \<Longrightarrow> inFr ns tr t"
+|
+Ind: "\<lbrakk>root tr \<in> ns; Inr tr1 \<in> cont tr; inFr ns tr1 t\<rbrakk> \<Longrightarrow> inFr ns tr t"
+
+definition "Fr ns tr \<equiv> {t. inFr ns tr t}"
+
+lemma inFr_root_in: "inFr ns tr t \<Longrightarrow> root tr \<in> ns"
+by (metis inFr.simps)
+
+lemma inFr_mono:
+assumes "inFr ns tr t" and "ns \<subseteq> ns'"
+shows "inFr ns' tr t"
+using assms apply(induct arbitrary: ns' rule: inFr.induct)
+using Base Ind by (metis inFr.simps set_mp)+
+
+lemma inFr_Ind_minus:
+assumes "inFr ns1 tr1 t" and "Inr tr1 \<in> cont tr"
+shows "inFr (insert (root tr) ns1) tr t"
+using assms apply(induct rule: inFr.induct)
+ apply (metis inFr.simps insert_iff)
+ by (metis inFr.simps inFr_mono insertI1 subset_insertI)
+
+(* alternative definition *)
+inductive inFr2 :: "N set \<Rightarrow> Tree \<Rightarrow> T \<Rightarrow> bool" where
+Base: "\<lbrakk>root tr \<in> ns; Inl t \<in> cont tr\<rbrakk> \<Longrightarrow> inFr2 ns tr t"
+|
+Ind: "\<lbrakk>Inr tr1 \<in> cont tr; inFr2 ns1 tr1 t\<rbrakk>
+ \<Longrightarrow> inFr2 (insert (root tr) ns1) tr t"
+
+lemma inFr2_root_in: "inFr2 ns tr t \<Longrightarrow> root tr \<in> ns"
+apply(induct rule: inFr2.induct) by auto
+
+lemma inFr2_mono:
+assumes "inFr2 ns tr t" and "ns \<subseteq> ns'"
+shows "inFr2 ns' tr t"
+using assms apply(induct arbitrary: ns' rule: inFr2.induct)
+using Base Ind
+apply (metis subsetD) by (metis inFr2.simps insert_absorb insert_subset)
+
+lemma inFr2_Ind:
+assumes "inFr2 ns tr1 t" and "root tr \<in> ns" and "Inr tr1 \<in> cont tr"
+shows "inFr2 ns tr t"
+using assms apply(induct rule: inFr2.induct)
+ apply (metis inFr2.simps insert_absorb)
+ by (metis inFr2.simps insert_absorb)
+
+lemma inFr_inFr2:
+"inFr = inFr2"
+apply (rule ext)+ apply(safe)
+ apply(erule inFr.induct)
+ apply (metis (lifting) inFr2.Base)
+ apply (metis (lifting) inFr2_Ind)
+ apply(erule inFr2.induct)
+ apply (metis (lifting) inFr.Base)
+ apply (metis (lifting) inFr_Ind_minus)
+done
+
+lemma not_root_inFr:
+assumes "root tr \<notin> ns"
+shows "\<not> inFr ns tr t"
+by (metis assms inFr_root_in)
+
+theorem not_root_Fr:
+assumes "root tr \<notin> ns"
+shows "Fr ns tr = {}"
+using not_root_inFr[OF assms] unfolding Fr_def by auto
+
+
+(* Interior *)
+
+inductive inItr :: "N set \<Rightarrow> Tree \<Rightarrow> N \<Rightarrow> bool" where
+Base: "root tr \<in> ns \<Longrightarrow> inItr ns tr (root tr)"
+|
+Ind: "\<lbrakk>root tr \<in> ns; Inr tr1 \<in> cont tr; inItr ns tr1 n\<rbrakk> \<Longrightarrow> inItr ns tr n"
+
+definition "Itr ns tr \<equiv> {n. inItr ns tr n}"
+
+lemma inItr_root_in: "inItr ns tr n \<Longrightarrow> root tr \<in> ns"
+by (metis inItr.simps)
+
+lemma inItr_mono:
+assumes "inItr ns tr n" and "ns \<subseteq> ns'"
+shows "inItr ns' tr n"
+using assms apply(induct arbitrary: ns' rule: inItr.induct)
+using Base Ind by (metis inItr.simps set_mp)+
+
+
+(* The subtree relation *)
+
+inductive subtr where
+Refl: "root tr \<in> ns \<Longrightarrow> subtr ns tr tr"
+|
+Step: "\<lbrakk>root tr3 \<in> ns; subtr ns tr1 tr2; Inr tr2 \<in> cont tr3\<rbrakk> \<Longrightarrow> subtr ns tr1 tr3"
+
+lemma subtr_rootL_in:
+assumes "subtr ns tr1 tr2"
+shows "root tr1 \<in> ns"
+using assms apply(induct rule: subtr.induct) by auto
+
+lemma subtr_rootR_in:
+assumes "subtr ns tr1 tr2"
+shows "root tr2 \<in> ns"
+using assms apply(induct rule: subtr.induct) by auto
+
+lemmas subtr_roots_in = subtr_rootL_in subtr_rootR_in
+
+lemma subtr_mono:
+assumes "subtr ns tr1 tr2" and "ns \<subseteq> ns'"
+shows "subtr ns' tr1 tr2"
+using assms apply(induct arbitrary: ns' rule: subtr.induct)
+using Refl Step by (metis subtr.simps set_mp)+
+
+lemma subtr_trans_Un:
+assumes "subtr ns12 tr1 tr2" and "subtr ns23 tr2 tr3"
+shows "subtr (ns12 \<union> ns23) tr1 tr3"
+proof-
+ have "subtr ns23 tr2 tr3 \<Longrightarrow>
+ (\<forall> ns12 tr1. subtr ns12 tr1 tr2 \<longrightarrow> subtr (ns12 \<union> ns23) tr1 tr3)"
+ apply(induct rule: subtr.induct, safe)
+ apply (metis subtr_mono sup_commute sup_ge2)
+ by (metis (lifting) Step UnI2)
+ thus ?thesis using assms by auto
+qed
+
+lemma subtr_trans:
+assumes "subtr ns tr1 tr2" and "subtr ns tr2 tr3"
+shows "subtr ns tr1 tr3"
+using subtr_trans_Un[OF assms] by simp
+
+lemma subtr_StepL:
+assumes r: "root tr1 \<in> ns" and tr12: "Inr tr1 \<in> cont tr2" and s: "subtr ns tr2 tr3"
+shows "subtr ns tr1 tr3"
+apply(rule subtr_trans[OF _ s]) apply(rule Step[of tr2 ns tr1 tr1])
+by (metis assms subtr_rootL_in Refl)+
+
+(* alternative definition: *)
+inductive subtr2 where
+Refl: "root tr \<in> ns \<Longrightarrow> subtr2 ns tr tr"
+|
+Step: "\<lbrakk>root tr1 \<in> ns; Inr tr1 \<in> cont tr2; subtr2 ns tr2 tr3\<rbrakk> \<Longrightarrow> subtr2 ns tr1 tr3"
+
+lemma subtr2_rootL_in:
+assumes "subtr2 ns tr1 tr2"
+shows "root tr1 \<in> ns"
+using assms apply(induct rule: subtr2.induct) by auto
+
+lemma subtr2_rootR_in:
+assumes "subtr2 ns tr1 tr2"
+shows "root tr2 \<in> ns"
+using assms apply(induct rule: subtr2.induct) by auto
+
+lemmas subtr2_roots_in = subtr2_rootL_in subtr2_rootR_in
+
+lemma subtr2_mono:
+assumes "subtr2 ns tr1 tr2" and "ns \<subseteq> ns'"
+shows "subtr2 ns' tr1 tr2"
+using assms apply(induct arbitrary: ns' rule: subtr2.induct)
+using Refl Step by (metis subtr2.simps set_mp)+
+
+lemma subtr2_trans_Un:
+assumes "subtr2 ns12 tr1 tr2" and "subtr2 ns23 tr2 tr3"
+shows "subtr2 (ns12 \<union> ns23) tr1 tr3"
+proof-
+ have "subtr2 ns12 tr1 tr2 \<Longrightarrow>
+ (\<forall> ns23 tr3. subtr2 ns23 tr2 tr3 \<longrightarrow> subtr2 (ns12 \<union> ns23) tr1 tr3)"
+ apply(induct rule: subtr2.induct, safe)
+ apply (metis subtr2_mono sup_commute sup_ge2)
+ by (metis Un_iff subtr2.simps)
+ thus ?thesis using assms by auto
+qed
+
+lemma subtr2_trans:
+assumes "subtr2 ns tr1 tr2" and "subtr2 ns tr2 tr3"
+shows "subtr2 ns tr1 tr3"
+using subtr2_trans_Un[OF assms] by simp
+
+lemma subtr2_StepR:
+assumes r: "root tr3 \<in> ns" and tr23: "Inr tr2 \<in> cont tr3" and s: "subtr2 ns tr1 tr2"
+shows "subtr2 ns tr1 tr3"
+apply(rule subtr2_trans[OF s]) apply(rule Step[of _ _ tr3])
+by (metis assms subtr2_rootR_in Refl)+
+
+lemma subtr_subtr2:
+"subtr = subtr2"
+apply (rule ext)+ apply(safe)
+ apply(erule subtr.induct)
+ apply (metis (lifting) subtr2.Refl)
+ apply (metis (lifting) subtr2_StepR)
+ apply(erule subtr2.induct)
+ apply (metis (lifting) subtr.Refl)
+ apply (metis (lifting) subtr_StepL)
+done
+
+lemma subtr_inductL[consumes 1, case_names Refl Step]:
+assumes s: "subtr ns tr1 tr2" and Refl: "\<And>ns tr. \<phi> ns tr tr"
+and Step:
+"\<And>ns tr1 tr2 tr3.
+ \<lbrakk>root tr1 \<in> ns; Inr tr1 \<in> cont tr2; subtr ns tr2 tr3; \<phi> ns tr2 tr3\<rbrakk> \<Longrightarrow> \<phi> ns tr1 tr3"
+shows "\<phi> ns tr1 tr2"
+using s unfolding subtr_subtr2 apply(rule subtr2.induct)
+using Refl Step unfolding subtr_subtr2 by auto
+
+lemma subtr_UNIV_inductL[consumes 1, case_names Refl Step]:
+assumes s: "subtr UNIV tr1 tr2" and Refl: "\<And>tr. \<phi> tr tr"
+and Step:
+"\<And>tr1 tr2 tr3.
+ \<lbrakk>Inr tr1 \<in> cont tr2; subtr UNIV tr2 tr3; \<phi> tr2 tr3\<rbrakk> \<Longrightarrow> \<phi> tr1 tr3"
+shows "\<phi> tr1 tr2"
+using s apply(induct rule: subtr_inductL)
+apply(rule Refl) using Step subtr_mono by (metis subset_UNIV)
+
+(* Subtree versus frontier: *)
+lemma subtr_inFr:
+assumes "inFr ns tr t" and "subtr ns tr tr1"
+shows "inFr ns tr1 t"
+proof-
+ have "subtr ns tr tr1 \<Longrightarrow> (\<forall> t. inFr ns tr t \<longrightarrow> inFr ns tr1 t)"
+ apply(induct rule: subtr.induct, safe) by (metis inFr.Ind)
+ thus ?thesis using assms by auto
+qed
+
+corollary Fr_subtr:
+"Fr ns tr = \<Union> {Fr ns tr' | tr'. subtr ns tr' tr}"
+unfolding Fr_def proof safe
+ fix t assume t: "inFr ns tr t" hence "root tr \<in> ns" by (rule inFr_root_in)
+ thus "t \<in> \<Union>{{t. inFr ns tr' t} |tr'. subtr ns tr' tr}"
+ apply(intro UnionI[of "{t. inFr ns tr t}" _ t]) using t subtr.Refl by auto
+qed(metis subtr_inFr)
+
+lemma inFr_subtr:
+assumes "inFr ns tr t"
+shows "\<exists> tr'. subtr ns tr' tr \<and> Inl t \<in> cont tr'"
+using assms apply(induct rule: inFr.induct) apply safe
+ apply (metis subtr.Refl)
+ by (metis (lifting) subtr.Step)
+
+corollary Fr_subtr_cont:
+"Fr ns tr = \<Union> {Inl -` cont tr' | tr'. subtr ns tr' tr}"
+unfolding Fr_def
+apply safe
+apply (frule inFr_subtr)
+apply auto
+by (metis inFr.Base subtr_inFr subtr_rootL_in)
+
+(* Subtree versus interior: *)
+lemma subtr_inItr:
+assumes "inItr ns tr n" and "subtr ns tr tr1"
+shows "inItr ns tr1 n"
+proof-
+ have "subtr ns tr tr1 \<Longrightarrow> (\<forall> t. inItr ns tr n \<longrightarrow> inItr ns tr1 n)"
+ apply(induct rule: subtr.induct, safe) by (metis inItr.Ind)
+ thus ?thesis using assms by auto
+qed
+
+corollary Itr_subtr:
+"Itr ns tr = \<Union> {Itr ns tr' | tr'. subtr ns tr' tr}"
+unfolding Itr_def apply safe
+apply (metis (lifting, mono_tags) UnionI inItr_root_in mem_Collect_eq subtr.Refl)
+by (metis subtr_inItr)
+
+lemma inItr_subtr:
+assumes "inItr ns tr n"
+shows "\<exists> tr'. subtr ns tr' tr \<and> root tr' = n"
+using assms apply(induct rule: inItr.induct) apply safe
+ apply (metis subtr.Refl)
+ by (metis (lifting) subtr.Step)
+
+corollary Itr_subtr_cont:
+"Itr ns tr = {root tr' | tr'. subtr ns tr' tr}"
+unfolding Itr_def apply safe
+ apply (metis (lifting, mono_tags) UnionI inItr_subtr mem_Collect_eq vimageI2)
+ by (metis inItr.Base subtr_inItr subtr_rootL_in)
+
+
+subsection{* The immediate subtree function *}
+
+(* production of: *)
+abbreviation "prodOf tr \<equiv> (id \<oplus> root) ` (cont tr)"
+(* subtree of: *)
+definition "subtrOf tr n \<equiv> SOME tr'. Inr tr' \<in> cont tr \<and> root tr' = n"
+
+lemma subtrOf:
+assumes n: "Inr n \<in> prodOf tr"
+shows "Inr (subtrOf tr n) \<in> cont tr \<and> root (subtrOf tr n) = n"
+proof-
+ obtain tr' where "Inr tr' \<in> cont tr \<and> root tr' = n"
+ using n unfolding image_def by (metis (lifting) Inr_oplus_elim assms)
+ thus ?thesis unfolding subtrOf_def by(rule someI)
+qed
+
+lemmas Inr_subtrOf = subtrOf[THEN conjunct1]
+lemmas root_subtrOf[simp] = subtrOf[THEN conjunct2]
+
+lemma Inl_prodOf: "Inl -` (prodOf tr) = Inl -` (cont tr)"
+proof safe
+ fix t ttr assume "Inl t = (id \<oplus> root) ttr" and "ttr \<in> cont tr"
+ thus "t \<in> Inl -` cont tr" by(cases ttr, auto)
+next
+ fix t assume "Inl t \<in> cont tr" thus "t \<in> Inl -` prodOf tr"
+ by (metis (lifting) id_def image_iff sum_map.simps(1) vimageI2)
+qed
+
+lemma root_prodOf:
+assumes "Inr tr' \<in> cont tr"
+shows "Inr (root tr') \<in> prodOf tr"
+by (metis (lifting) assms image_iff sum_map.simps(2))
+
+
+subsection{* Derivation trees *}
+
+coinductive dtree where
+Tree: "\<lbrakk>(root tr, (id \<oplus> root) ` (cont tr)) \<in> P; inj_on root (Inr -` cont tr);
+ lift dtree (cont tr)\<rbrakk> \<Longrightarrow> dtree tr"
+monos lift_mono
+
+(* destruction rules: *)
+lemma dtree_P:
+assumes "dtree tr"
+shows "(root tr, (id \<oplus> root) ` (cont tr)) \<in> P"
+using assms unfolding dtree.simps by auto
+
+lemma dtree_inj_on:
+assumes "dtree tr"
+shows "inj_on root (Inr -` cont tr)"
+using assms unfolding dtree.simps by auto
+
+lemma dtree_inj[simp]:
+assumes "dtree tr" and "Inr tr1 \<in> cont tr" and "Inr tr2 \<in> cont tr"
+shows "root tr1 = root tr2 \<longleftrightarrow> tr1 = tr2"
+using assms dtree_inj_on unfolding inj_on_def by auto
+
+lemma dtree_lift:
+assumes "dtree tr"
+shows "lift dtree (cont tr)"
+using assms unfolding dtree.simps by auto
+
+
+(* coinduction:*)
+lemma dtree_coind[elim, consumes 1, case_names Hyp]:
+assumes phi: "\<phi> tr"
+and Hyp:
+"\<And> tr. \<phi> tr \<Longrightarrow>
+ (root tr, image (id \<oplus> root) (cont tr)) \<in> P \<and>
+ inj_on root (Inr -` cont tr) \<and>
+ lift (\<lambda> tr. \<phi> tr \<or> dtree tr) (cont tr)"
+shows "dtree tr"
+apply(rule dtree.coinduct[of \<phi> tr, OF phi])
+using Hyp by blast
+
+lemma dtree_raw_coind[elim, consumes 1, case_names Hyp]:
+assumes phi: "\<phi> tr"
+and Hyp:
+"\<And> tr. \<phi> tr \<Longrightarrow>
+ (root tr, image (id \<oplus> root) (cont tr)) \<in> P \<and>
+ inj_on root (Inr -` cont tr) \<and>
+ lift \<phi> (cont tr)"
+shows "dtree tr"
+using phi apply(induct rule: dtree_coind)
+using Hyp mono_lift
+by (metis (mono_tags) mono_lift)
+
+lemma dtree_subtr_inj_on:
+assumes d: "dtree tr1" and s: "subtr ns tr tr1"
+shows "inj_on root (Inr -` cont tr)"
+using s d apply(induct rule: subtr.induct)
+apply (metis (lifting) dtree_inj_on) by (metis dtree_lift lift_def)
+
+lemma dtree_subtr_P:
+assumes d: "dtree tr1" and s: "subtr ns tr tr1"
+shows "(root tr, (id \<oplus> root) ` cont tr) \<in> P"
+using s d apply(induct rule: subtr.induct)
+apply (metis (lifting) dtree_P) by (metis dtree_lift lift_def)
+
+lemma subtrOf_root[simp]:
+assumes tr: "dtree tr" and cont: "Inr tr' \<in> cont tr"
+shows "subtrOf tr (root tr') = tr'"
+proof-
+ have 0: "Inr (subtrOf tr (root tr')) \<in> cont tr" using Inr_subtrOf
+ by (metis (lifting) cont root_prodOf)
+ have "root (subtrOf tr (root tr')) = root tr'"
+ using root_subtrOf by (metis (lifting) cont root_prodOf)
+ thus ?thesis unfolding dtree_inj[OF tr 0 cont] .
+qed
+
+lemma surj_subtrOf:
+assumes "dtree tr" and 0: "Inr tr' \<in> cont tr"
+shows "\<exists> n. Inr n \<in> prodOf tr \<and> subtrOf tr n = tr'"
+apply(rule exI[of _ "root tr'"])
+using root_prodOf[OF 0] subtrOf_root[OF assms] by simp
+
+lemma dtree_subtr:
+assumes "dtree tr1" and "subtr ns tr tr1"
+shows "dtree tr"
+proof-
+ have "(\<exists> ns tr1. dtree tr1 \<and> subtr ns tr tr1) \<Longrightarrow> dtree tr"
+ proof (induct rule: dtree_raw_coind)
+ case (Hyp tr)
+ then obtain ns tr1 where tr1: "dtree tr1" and tr_tr1: "subtr ns tr tr1" by auto
+ show ?case unfolding lift_def proof safe
+ show "(root tr, (id \<oplus> root) ` cont tr) \<in> P" using dtree_subtr_P[OF tr1 tr_tr1] .
+ next
+ show "inj_on root (Inr -` cont tr)" using dtree_subtr_inj_on[OF tr1 tr_tr1] .
+ next
+ fix tr' assume tr': "Inr tr' \<in> cont tr"
+ have tr_tr1: "subtr (ns \<union> {root tr'}) tr tr1" using subtr_mono[OF tr_tr1] by auto
+ have "subtr (ns \<union> {root tr'}) tr' tr1" using subtr_StepL[OF _ tr' tr_tr1] by auto
+ thus "\<exists>ns' tr1. dtree tr1 \<and> subtr ns' tr' tr1" using tr1 by blast
+ qed
+ qed
+ thus ?thesis using assms by auto
+qed
+
+
+subsection{* Default trees *}
+
+(* Pick a left-hand side of a production for each nonterminal *)
+definition S where "S n \<equiv> SOME tns. (n,tns) \<in> P"
+
+lemma S_P: "(n, S n) \<in> P"
+using used unfolding S_def by(rule someI_ex)
+
+lemma finite_S: "finite (S n)"
+using S_P finite_in_P by auto
+
+
+(* The default tree of a nonterminal *)
+definition deftr :: "N \<Rightarrow> Tree" where
+"deftr \<equiv> unfold id S"
+
+lemma deftr_simps[simp]:
+"root (deftr n) = n"
+"cont (deftr n) = image (id \<oplus> deftr) (S n)"
+using unfold(1)[of id S n] unfold(2)[of S n id, OF finite_S]
+unfolding deftr_def by simp_all
+
+lemmas root_deftr = deftr_simps(1)
+lemmas cont_deftr = deftr_simps(2)
+
+lemma root_o_deftr[simp]: "root o deftr = id"
+by (rule ext, auto)
+
+lemma dtree_deftr: "dtree (deftr n)"
+proof-
+ {fix tr assume "\<exists> n. tr = deftr n" hence "dtree tr"
+ apply(induct rule: dtree_raw_coind) apply safe
+ unfolding deftr_simps image_compose[symmetric] sum_map.comp id_o
+ root_o_deftr sum_map.id image_id id_apply apply(rule S_P)
+ unfolding inj_on_def lift_def by auto
+ }
+ thus ?thesis by auto
+qed
+
+
+subsection{* Hereditary substitution *}
+
+(* Auxiliary concept: The root-ommiting frontier: *)
+definition "inFrr ns tr t \<equiv> \<exists> tr'. Inr tr' \<in> cont tr \<and> inFr ns tr' t"
+definition "Frr ns tr \<equiv> {t. \<exists> tr'. Inr tr' \<in> cont tr \<and> t \<in> Fr ns tr'}"
+
+context
+fixes tr0 :: Tree
+begin
+
+definition "hsubst_r tr \<equiv> root tr"
+definition "hsubst_c tr \<equiv> if root tr = root tr0 then cont tr0 else cont tr"
+
+(* Hereditary substitution: *)
+definition hsubst :: "Tree \<Rightarrow> Tree" where
+"hsubst \<equiv> unfold hsubst_r hsubst_c"
+
+lemma finite_hsubst_c: "finite (hsubst_c n)"
+unfolding hsubst_c_def by (metis finite_cont)
+
+lemma root_hsubst[simp]: "root (hsubst tr) = root tr"
+using unfold(1)[of hsubst_r hsubst_c tr] unfolding hsubst_def hsubst_r_def by simp
+
+lemma root_o_subst[simp]: "root o hsubst = root"
+unfolding comp_def root_hsubst ..
+
+lemma cont_hsubst_eq[simp]:
+assumes "root tr = root tr0"
+shows "cont (hsubst tr) = (id \<oplus> hsubst) ` (cont tr0)"
+apply(subst id_o[symmetric, of id]) unfolding id_o
+using unfold(2)[of hsubst_c tr hsubst_r, OF finite_hsubst_c]
+unfolding hsubst_def hsubst_c_def using assms by simp
+
+lemma hsubst_eq:
+assumes "root tr = root tr0"
+shows "hsubst tr = hsubst tr0"
+apply(rule Tree_cong) using assms cont_hsubst_eq by auto
+
+lemma cont_hsubst_neq[simp]:
+assumes "root tr \<noteq> root tr0"
+shows "cont (hsubst tr) = (id \<oplus> hsubst) ` (cont tr)"
+apply(subst id_o[symmetric, of id]) unfolding id_o
+using unfold(2)[of hsubst_c tr hsubst_r, OF finite_hsubst_c]
+unfolding hsubst_def hsubst_c_def using assms by simp
+
+lemma Inl_cont_hsubst_eq[simp]:
+assumes "root tr = root tr0"
+shows "Inl -` cont (hsubst tr) = Inl -` (cont tr0)"
+unfolding cont_hsubst_eq[OF assms] by simp
+
+lemma Inr_cont_hsubst_eq[simp]:
+assumes "root tr = root tr0"
+shows "Inr -` cont (hsubst tr) = hsubst ` Inr -` cont tr0"
+unfolding cont_hsubst_eq[OF assms] by simp
+
+lemma Inl_cont_hsubst_neq[simp]:
+assumes "root tr \<noteq> root tr0"
+shows "Inl -` cont (hsubst tr) = Inl -` (cont tr)"
+unfolding cont_hsubst_neq[OF assms] by simp
+
+lemma Inr_cont_hsubst_neq[simp]:
+assumes "root tr \<noteq> root tr0"
+shows "Inr -` cont (hsubst tr) = hsubst ` Inr -` cont tr"
+unfolding cont_hsubst_neq[OF assms] by simp
+
+lemma dtree_hsubst:
+assumes tr0: "dtree tr0" and tr: "dtree tr"
+shows "dtree (hsubst tr)"
+proof-
+ {fix tr1 have "(\<exists> tr. dtree tr \<and> tr1 = hsubst tr) \<Longrightarrow> dtree tr1"
+ proof (induct rule: dtree_raw_coind)
+ case (Hyp tr1) then obtain tr
+ where dtr: "dtree tr" and tr1: "tr1 = hsubst tr" by auto
+ show ?case unfolding lift_def tr1 proof safe
+ show "(root (hsubst tr), prodOf (hsubst tr)) \<in> P"
+ unfolding tr1 apply(cases "root tr = root tr0")
+ using dtree_P[OF dtr] dtree_P[OF tr0]
+ by (auto simp add: image_compose[symmetric] sum_map.comp)
+ show "inj_on root (Inr -` cont (hsubst tr))"
+ apply(cases "root tr = root tr0") using dtree_inj_on[OF dtr] dtree_inj_on[OF tr0]
+ unfolding inj_on_def by (auto, blast)
+ fix tr' assume "Inr tr' \<in> cont (hsubst tr)"
+ thus "\<exists>tra. dtree tra \<and> tr' = hsubst tra"
+ apply(cases "root tr = root tr0", simp_all)
+ apply (metis dtree_lift lift_def tr0)
+ by (metis dtr dtree_lift lift_def)
+ qed
+ qed
+ }
+ thus ?thesis using assms by blast
+qed
+
+lemma Frr: "Frr ns tr = {t. inFrr ns tr t}"
+unfolding inFrr_def Frr_def Fr_def by auto
+
+lemma inFr_hsubst_imp:
+assumes "inFr ns (hsubst tr) t"
+shows "t \<in> Inl -` (cont tr0) \<or> inFrr (ns - {root tr0}) tr0 t \<or>
+ inFr (ns - {root tr0}) tr t"
+proof-
+ {fix tr1
+ have "inFr ns tr1 t \<Longrightarrow>
+ (\<And> tr. tr1 = hsubst tr \<Longrightarrow> (t \<in> Inl -` (cont tr0) \<or> inFrr (ns - {root tr0}) tr0 t \<or>
+ inFr (ns - {root tr0}) tr t))"
+ proof(induct rule: inFr.induct)
+ case (Base tr1 ns t tr)
+ hence rtr: "root tr1 \<in> ns" and t_tr1: "Inl t \<in> cont tr1" and tr1: "tr1 = hsubst tr"
+ by auto
+ show ?case
+ proof(cases "root tr1 = root tr0")
+ case True
+ hence "t \<in> Inl -` (cont tr0)" using t_tr1 unfolding tr1 by auto
+ thus ?thesis by simp
+ next
+ case False
+ hence "inFr (ns - {root tr0}) tr t" using t_tr1 unfolding tr1 apply simp
+ by (metis Base.prems Diff_iff root_hsubst inFr.Base rtr singletonE)
+ thus ?thesis by simp
+ qed
+ next
+ case (Ind tr1 ns tr1' t) note IH = Ind(4)
+ have rtr1: "root tr1 \<in> ns" and tr1'_tr1: "Inr tr1' \<in> cont tr1"
+ and t_tr1': "inFr ns tr1' t" and tr1: "tr1 = hsubst tr" using Ind by auto
+ have rtr1: "root tr1 = root tr" unfolding tr1 by simp
+ show ?case
+ proof(cases "root tr1 = root tr0")
+ case True
+ then obtain tr' where tr'_tr0: "Inr tr' \<in> cont tr0" and tr1': "tr1' = hsubst tr'"
+ using tr1'_tr1 unfolding tr1 by auto
+ show ?thesis using IH[OF tr1'] proof (elim disjE)
+ assume "inFr (ns - {root tr0}) tr' t"
+ thus ?thesis using tr'_tr0 unfolding inFrr_def by auto
+ qed auto
+ next
+ case False
+ then obtain tr' where tr'_tr: "Inr tr' \<in> cont tr" and tr1': "tr1' = hsubst tr'"
+ using tr1'_tr1 unfolding tr1 by auto
+ show ?thesis using IH[OF tr1'] proof (elim disjE)
+ assume "inFr (ns - {root tr0}) tr' t"
+ thus ?thesis using tr'_tr unfolding inFrr_def
+ by (metis Diff_iff False Ind(1) empty_iff inFr2_Ind inFr_inFr2 insert_iff rtr1)
+ qed auto
+ qed
+ qed
+ }
+ thus ?thesis using assms by auto
+qed
+
+lemma inFr_hsubst_notin:
+assumes "inFr ns tr t" and "root tr0 \<notin> ns"
+shows "inFr ns (hsubst tr) t"
+using assms apply(induct rule: inFr.induct)
+apply (metis Inl_cont_hsubst_neq inFr2.Base inFr_inFr2 root_hsubst vimageD vimageI2)
+by (metis (lifting) Inr_cont_hsubst_neq inFr.Ind rev_image_eqI root_hsubst vimageD vimageI2)
+
+lemma inFr_hsubst_minus:
+assumes "inFr (ns - {root tr0}) tr t"
+shows "inFr ns (hsubst tr) t"
+proof-
+ have 1: "inFr (ns - {root tr0}) (hsubst tr) t"
+ using inFr_hsubst_notin[OF assms] by simp
+ show ?thesis using inFr_mono[OF 1] by auto
+qed
+
+lemma inFr_self_hsubst:
+assumes "root tr0 \<in> ns"
+shows
+"inFr ns (hsubst tr0) t \<longleftrightarrow>
+ t \<in> Inl -` (cont tr0) \<or> inFrr (ns - {root tr0}) tr0 t"
+(is "?A \<longleftrightarrow> ?B \<or> ?C")
+apply(intro iffI)
+apply (metis inFr_hsubst_imp Diff_iff inFr_root_in insertI1) proof(elim disjE)
+ assume ?B thus ?A apply(intro inFr.Base) using assms by auto
+next
+ assume ?C then obtain tr where
+ tr_tr0: "Inr tr \<in> cont tr0" and t_tr: "inFr (ns - {root tr0}) tr t"
+ unfolding inFrr_def by auto
+ def tr1 \<equiv> "hsubst tr"
+ have 1: "inFr ns tr1 t" using t_tr unfolding tr1_def using inFr_hsubst_minus by auto
+ have "Inr tr1 \<in> cont (hsubst tr0)" unfolding tr1_def using tr_tr0 by auto
+ thus ?A using 1 inFr.Ind assms by (metis root_hsubst)
+qed
+
+theorem Fr_self_hsubst:
+assumes "root tr0 \<in> ns"
+shows "Fr ns (hsubst tr0) = Inl -` (cont tr0) \<union> Frr (ns - {root tr0}) tr0"
+using inFr_self_hsubst[OF assms] unfolding Frr Fr_def by auto
+
+end (* context *)
+
+
+subsection{* Regular trees *}
+
+hide_const regular
+
+definition "reg f tr \<equiv> \<forall> tr'. subtr UNIV tr' tr \<longrightarrow> tr' = f (root tr')"
+definition "regular tr \<equiv> \<exists> f. reg f tr"
+
+lemma reg_def2: "reg f tr \<longleftrightarrow> (\<forall> ns tr'. subtr ns tr' tr \<longrightarrow> tr' = f (root tr'))"
+unfolding reg_def using subtr_mono by (metis subset_UNIV)
+
+lemma regular_def2: "regular tr \<longleftrightarrow> (\<exists> f. reg f tr \<and> (\<forall> n. root (f n) = n))"
+unfolding regular_def proof safe
+ fix f assume f: "reg f tr"
+ def g \<equiv> "\<lambda> n. if inItr UNIV tr n then f n else deftr n"
+ show "\<exists>g. reg g tr \<and> (\<forall>n. root (g n) = n)"
+ apply(rule exI[of _ g])
+ using f deftr_simps(1) unfolding g_def reg_def apply safe
+ apply (metis (lifting) inItr.Base subtr_inItr subtr_rootL_in)
+ by (metis (full_types) inItr_subtr subtr_subtr2)
+qed auto
+
+lemma reg_root:
+assumes "reg f tr"
+shows "f (root tr) = tr"
+using assms unfolding reg_def
+by (metis (lifting) iso_tuple_UNIV_I subtr.Refl)
+
+
+lemma reg_Inr_cont:
+assumes "reg f tr" and "Inr tr' \<in> cont tr"
+shows "reg f tr'"
+by (metis (lifting) assms iso_tuple_UNIV_I reg_def subtr.Step)
+
+lemma reg_subtr:
+assumes "reg f tr" and "subtr ns tr' tr"
+shows "reg f tr'"
+using assms unfolding reg_def using subtr_trans[of UNIV tr] UNIV_I
+by (metis UNIV_eq_I UnCI Un_upper1 iso_tuple_UNIV_I subtr_mono subtr_trans)
+
+lemma regular_subtr:
+assumes r: "regular tr" and s: "subtr ns tr' tr"
+shows "regular tr'"
+using r reg_subtr[OF _ s] unfolding regular_def by auto
+
+lemma subtr_deftr:
+assumes "subtr ns tr' (deftr n)"
+shows "tr' = deftr (root tr')"
+proof-
+ {fix tr have "subtr ns tr' tr \<Longrightarrow> (\<forall> n. tr = deftr n \<longrightarrow> tr' = deftr (root tr'))"
+ apply (induct rule: subtr.induct)
+ proof(metis (lifting) deftr_simps(1), safe)
+ fix tr3 ns tr1 tr2 n
+ assume 1: "root (deftr n) \<in> ns" and 2: "subtr ns tr1 tr2"
+ and IH: "\<forall>n. tr2 = deftr n \<longrightarrow> tr1 = deftr (root tr1)"
+ and 3: "Inr tr2 \<in> cont (deftr n)"
+ have "tr2 \<in> deftr ` UNIV"
+ using 3 unfolding deftr_simps image_def
+ by (metis (lifting, full_types) 3 CollectI Inr_oplus_iff cont_deftr
+ iso_tuple_UNIV_I)
+ then obtain n where "tr2 = deftr n" by auto
+ thus "tr1 = deftr (root tr1)" using IH by auto
+ qed
+ }
+ thus ?thesis using assms by auto
+qed
+
+lemma reg_deftr: "reg deftr (deftr n)"
+unfolding reg_def using subtr_deftr by auto
+
+lemma dtree_subtrOf_Union:
+assumes "dtree tr"
+shows "\<Union>{K tr' |tr'. Inr tr' \<in> cont tr} =
+ \<Union>{K (subtrOf tr n) |n. Inr n \<in> prodOf tr}"
+unfolding Union_eq Bex_def mem_Collect_eq proof safe
+ fix x xa tr'
+ assume x: "x \<in> K tr'" and tr'_tr: "Inr tr' \<in> cont tr"
+ show "\<exists>X. (\<exists>n. X = K (subtrOf tr n) \<and> Inr n \<in> prodOf tr) \<and> x \<in> X"
+ apply(rule exI[of _ "K (subtrOf tr (root tr'))"]) apply(intro conjI)
+ apply(rule exI[of _ "root tr'"]) apply (metis (lifting) root_prodOf tr'_tr)
+ by (metis (lifting) assms subtrOf_root tr'_tr x)
+next
+ fix x X n ttr
+ assume x: "x \<in> K (subtrOf tr n)" and n: "Inr n = (id \<oplus> root) ttr" and ttr: "ttr \<in> cont tr"
+ show "\<exists>X. (\<exists>tr'. X = K tr' \<and> Inr tr' \<in> cont tr) \<and> x \<in> X"
+ apply(rule exI[of _ "K (subtrOf tr n)"]) apply(intro conjI)
+ apply(rule exI[of _ "subtrOf tr n"]) apply (metis imageI n subtrOf ttr)
+ using x .
+qed
+
+
+
+
+subsection {* Paths in a regular tree *}
+
+inductive path :: "(N \<Rightarrow> Tree) \<Rightarrow> N list \<Rightarrow> bool" for f where
+Base: "path f [n]"
+|
+Ind: "\<lbrakk>path f (n1 # nl); Inr (f n1) \<in> cont (f n)\<rbrakk>
+ \<Longrightarrow> path f (n # n1 # nl)"
+
+lemma path_NE:
+assumes "path f nl"
+shows "nl \<noteq> Nil"
+using assms apply(induct rule: path.induct) by auto
+
+lemma path_post:
+assumes f: "path f (n # nl)" and nl: "nl \<noteq> []"
+shows "path f nl"
+proof-
+ obtain n1 nl1 where nl: "nl = n1 # nl1" using nl by (cases nl, auto)
+ show ?thesis using assms unfolding nl using path.simps by (metis (lifting) list.inject)
+qed
+
+lemma path_post_concat:
+assumes "path f (nl1 @ nl2)" and "nl2 \<noteq> Nil"
+shows "path f nl2"
+using assms apply (induct nl1)
+apply (metis append_Nil) by (metis Nil_is_append_conv append_Cons path_post)
+
+lemma path_concat:
+assumes "path f nl1" and "path f ((last nl1) # nl2)"
+shows "path f (nl1 @ nl2)"
+using assms apply(induct rule: path.induct) apply simp
+by (metis append_Cons last.simps list.simps(3) path.Ind)
+
+lemma path_distinct:
+assumes "path f nl"
+shows "\<exists> nl'. path f nl' \<and> hd nl' = hd nl \<and> last nl' = last nl \<and>
+ set nl' \<subseteq> set nl \<and> distinct nl'"
+using assms proof(induct rule: length_induct)
+ case (1 nl) hence p_nl: "path f nl" by simp
+ then obtain n nl1 where nl: "nl = n # nl1" by (metis list.exhaust path_NE)
+ show ?case
+ proof(cases nl1)
+ case Nil
+ show ?thesis apply(rule exI[of _ nl]) using path.Base unfolding nl Nil by simp
+ next
+ case (Cons n1 nl2)
+ hence p1: "path f nl1" by (metis list.simps nl p_nl path_post)
+ show ?thesis
+ proof(cases "n \<in> set nl1")
+ case False
+ obtain nl1' where p1': "path f nl1'" and hd_nl1': "hd nl1' = hd nl1" and
+ l_nl1': "last nl1' = last nl1" and d_nl1': "distinct nl1'"
+ and s_nl1': "set nl1' \<subseteq> set nl1"
+ using 1(1)[THEN allE[of _ nl1]] p1 unfolding nl by auto
+ obtain nl2' where nl1': "nl1' = n1 # nl2'" using path_NE[OF p1'] hd_nl1'
+ unfolding Cons by(cases nl1', auto)
+ show ?thesis apply(intro exI[of _ "n # nl1'"]) unfolding nl proof safe
+ show "path f (n # nl1')" unfolding nl1'
+ apply(rule path.Ind, metis nl1' p1')
+ by (metis (lifting) Cons list.inject nl p1 p_nl path.simps path_NE)
+ qed(insert l_nl1' Cons nl1' s_nl1' d_nl1' False, auto)
+ next
+ case True
+ then obtain nl11 nl12 where nl1: "nl1 = nl11 @ n # nl12"
+ by (metis split_list)
+ have p12: "path f (n # nl12)"
+ apply(rule path_post_concat[of _ "n # nl11"]) using p_nl[unfolded nl nl1] by auto
+ obtain nl12' where p1': "path f nl12'" and hd_nl12': "hd nl12' = n" and
+ l_nl12': "last nl12' = last (n # nl12)" and d_nl12': "distinct nl12'"
+ and s_nl12': "set nl12' \<subseteq> {n} \<union> set nl12"
+ using 1(1)[THEN allE[of _ "n # nl12"]] p12 unfolding nl nl1 by auto
+ thus ?thesis apply(intro exI[of _ nl12']) unfolding nl nl1 by auto
+ qed
+ qed
+qed
+
+lemma path_subtr:
+assumes f: "\<And> n. root (f n) = n"
+and p: "path f nl"
+shows "subtr (set nl) (f (last nl)) (f (hd nl))"
+using p proof (induct rule: path.induct)
+ case (Ind n1 nl n) let ?ns1 = "insert n1 (set nl)"
+ have "path f (n1 # nl)"
+ and "subtr ?ns1 (f (last (n1 # nl))) (f n1)"
+ and fn1: "Inr (f n1) \<in> cont (f n)" using Ind by simp_all
+ hence fn1_flast: "subtr (insert n ?ns1) (f (last (n1 # nl))) (f n1)"
+ by (metis subset_insertI subtr_mono)
+ have 1: "last (n # n1 # nl) = last (n1 # nl)" by auto
+ have "subtr (insert n ?ns1) (f (last (n1 # nl))) (f n)"
+ using f subtr.Step[OF _ fn1_flast fn1] by auto
+ thus ?case unfolding 1 by simp
+qed (metis f hd.simps last_ConsL last_in_set not_Cons_self2 subtr.Refl)
+
+lemma reg_subtr_path_aux:
+assumes f: "reg f tr" and n: "subtr ns tr1 tr"
+shows "\<exists> nl. path f nl \<and> f (hd nl) = tr \<and> f (last nl) = tr1 \<and> set nl \<subseteq> ns"
+using n f proof(induct rule: subtr.induct)
+ case (Refl tr ns)
+ thus ?case
+ apply(intro exI[of _ "[root tr]"]) apply simp by (metis (lifting) path.Base reg_root)
+next
+ case (Step tr ns tr2 tr1)
+ hence rtr: "root tr \<in> ns" and tr1_tr: "Inr tr1 \<in> cont tr"
+ and tr2_tr1: "subtr ns tr2 tr1" and tr: "reg f tr" by auto
+ have tr1: "reg f tr1" using reg_subtr[OF tr] rtr tr1_tr
+ by (metis (lifting) Step.prems iso_tuple_UNIV_I reg_def subtr.Step)
+ obtain nl where nl: "path f nl" and f_nl: "f (hd nl) = tr1"
+ and last_nl: "f (last nl) = tr2" and set: "set nl \<subseteq> ns" using Step(3)[OF tr1] by auto
+ have 0: "path f (root tr # nl)" apply (subst path.simps)
+ using f_nl nl reg_root tr tr1_tr by (metis hd.simps neq_Nil_conv)
+ show ?case apply(rule exI[of _ "(root tr) # nl"])
+ using 0 reg_root tr last_nl nl path_NE rtr set by auto
+qed
+
+lemma reg_subtr_path:
+assumes f: "reg f tr" and n: "subtr ns tr1 tr"
+shows "\<exists> nl. distinct nl \<and> path f nl \<and> f (hd nl) = tr \<and> f (last nl) = tr1 \<and> set nl \<subseteq> ns"
+using reg_subtr_path_aux[OF assms] path_distinct[of f]
+by (metis (lifting) order_trans)
+
+lemma subtr_iff_path:
+assumes r: "reg f tr" and f: "\<And> n. root (f n) = n"
+shows "subtr ns tr1 tr \<longleftrightarrow>
+ (\<exists> nl. distinct nl \<and> path f nl \<and> f (hd nl) = tr \<and> f (last nl) = tr1 \<and> set nl \<subseteq> ns)"
+proof safe
+ fix nl assume p: "path f nl" and nl: "set nl \<subseteq> ns"
+ have "subtr (set nl) (f (last nl)) (f (hd nl))"
+ apply(rule path_subtr) using p f by simp_all
+ thus "subtr ns (f (last nl)) (f (hd nl))"
+ using subtr_mono nl by auto
+qed(insert reg_subtr_path[OF r], auto)
+
+lemma inFr_iff_path:
+assumes r: "reg f tr" and f: "\<And> n. root (f n) = n"
+shows
+"inFr ns tr t \<longleftrightarrow>
+ (\<exists> nl tr1. distinct nl \<and> path f nl \<and> f (hd nl) = tr \<and> f (last nl) = tr1 \<and>
+ set nl \<subseteq> ns \<and> Inl t \<in> cont tr1)"
+apply safe
+apply (metis (no_types) inFr_subtr r reg_subtr_path)
+by (metis f inFr.Base path_subtr subtr_inFr subtr_mono subtr_rootL_in)
+
+
+
+subsection{* The regular cut of a tree *}
+
+context fixes tr0 :: Tree
+begin
+
+(* Picking a subtree of a certain root: *)
+definition "pick n \<equiv> SOME tr. subtr UNIV tr tr0 \<and> root tr = n"
+
+lemma pick:
+assumes "inItr UNIV tr0 n"
+shows "subtr UNIV (pick n) tr0 \<and> root (pick n) = n"
+proof-
+ have "\<exists> tr. subtr UNIV tr tr0 \<and> root tr = n"
+ using assms by (metis (lifting) inItr_subtr)
+ thus ?thesis unfolding pick_def by(rule someI_ex)
+qed
+
+lemmas subtr_pick = pick[THEN conjunct1]
+lemmas root_pick = pick[THEN conjunct2]
+
+lemma dtree_pick:
+assumes tr0: "dtree tr0" and n: "inItr UNIV tr0 n"
+shows "dtree (pick n)"
+using dtree_subtr[OF tr0 subtr_pick[OF n]] .
+
+definition "regOf_r n \<equiv> root (pick n)"
+definition "regOf_c n \<equiv> (id \<oplus> root) ` cont (pick n)"
+
+(* The regular tree of a function: *)
+definition regOf :: "N \<Rightarrow> Tree" where
+"regOf \<equiv> unfold regOf_r regOf_c"
+
+lemma finite_regOf_c: "finite (regOf_c n)"
+unfolding regOf_c_def by (metis finite_cont finite_imageI)
+
+lemma root_regOf_pick: "root (regOf n) = root (pick n)"
+using unfold(1)[of regOf_r regOf_c n] unfolding regOf_def regOf_r_def by simp
+
+lemma root_regOf[simp]:
+assumes "inItr UNIV tr0 n"
+shows "root (regOf n) = n"
+unfolding root_regOf_pick root_pick[OF assms] ..
+
+lemma cont_regOf[simp]:
+"cont (regOf n) = (id \<oplus> (regOf o root)) ` cont (pick n)"
+apply(subst id_o[symmetric, of id]) unfolding sum_map.comp[symmetric]
+unfolding image_compose unfolding regOf_c_def[symmetric]
+using unfold(2)[of regOf_c n regOf_r, OF finite_regOf_c]
+unfolding regOf_def ..
+
+lemma Inl_cont_regOf[simp]:
+"Inl -` (cont (regOf n)) = Inl -` (cont (pick n))"
+unfolding cont_regOf by simp
+
+lemma Inr_cont_regOf:
+"Inr -` (cont (regOf n)) = (regOf \<circ> root) ` (Inr -` cont (pick n))"
+unfolding cont_regOf by simp
+
+lemma subtr_regOf:
+assumes n: "inItr UNIV tr0 n" and "subtr UNIV tr1 (regOf n)"
+shows "\<exists> n1. inItr UNIV tr0 n1 \<and> tr1 = regOf n1"
+proof-
+ {fix tr ns assume "subtr UNIV tr1 tr"
+ hence "tr = regOf n \<longrightarrow> (\<exists> n1. inItr UNIV tr0 n1 \<and> tr1 = regOf n1)"
+ proof (induct rule: subtr_UNIV_inductL)
+ case (Step tr2 tr1 tr)
+ show ?case proof
+ assume "tr = regOf n"
+ then obtain n1 where tr2: "Inr tr2 \<in> cont tr1"
+ and tr1_tr: "subtr UNIV tr1 tr" and n1: "inItr UNIV tr0 n1" and tr1: "tr1 = regOf n1"
+ using Step by auto
+ obtain tr2' where tr2: "tr2 = regOf (root tr2')"
+ and tr2': "Inr tr2' \<in> cont (pick n1)"
+ using tr2 Inr_cont_regOf[of n1]
+ unfolding tr1 image_def o_def using vimage_eq by auto
+ have "inItr UNIV tr0 (root tr2')"
+ using inItr.Base inItr.Ind n1 pick subtr_inItr tr2' by (metis iso_tuple_UNIV_I)
+ thus "\<exists>n2. inItr UNIV tr0 n2 \<and> tr2 = regOf n2" using tr2 by blast
+ qed
+ qed(insert n, auto)
+ }
+ thus ?thesis using assms by auto
+qed
+
+lemma root_regOf_root:
+assumes n: "inItr UNIV tr0 n" and t_tr: "t_tr \<in> cont (pick n)"
+shows "(id \<oplus> (root \<circ> regOf \<circ> root)) t_tr = (id \<oplus> root) t_tr"
+using assms apply(cases t_tr)
+ apply (metis (lifting) sum_map.simps(1))
+ using pick regOf_def regOf_r_def unfold(1)
+ inItr.Base o_apply subtr_StepL subtr_inItr sum_map.simps(2)
+ by (metis UNIV_I)
+
+lemma regOf_P:
+assumes tr0: "dtree tr0" and n: "inItr UNIV tr0 n"
+shows "(n, (id \<oplus> root) ` cont (regOf n)) \<in> P" (is "?L \<in> P")
+proof-
+ have "?L = (n, (id \<oplus> root) ` cont (pick n))"
+ unfolding cont_regOf image_compose[symmetric] sum_map.comp id_o o_assoc
+ unfolding Pair_eq apply(rule conjI[OF refl]) apply(rule image_cong[OF refl])
+ by(rule root_regOf_root[OF n])
+ moreover have "... \<in> P" by (metis (lifting) dtree_pick root_pick dtree_P n tr0)
+ ultimately show ?thesis by simp
+qed
+
+lemma dtree_regOf:
+assumes tr0: "dtree tr0" and "inItr UNIV tr0 n"
+shows "dtree (regOf n)"
+proof-
+ {fix tr have "\<exists> n. inItr UNIV tr0 n \<and> tr = regOf n \<Longrightarrow> dtree tr"
+ proof (induct rule: dtree_raw_coind)
+ case (Hyp tr)
+ then obtain n where n: "inItr UNIV tr0 n" and tr: "tr = regOf n" by auto
+ show ?case unfolding lift_def apply safe
+ apply (metis (lifting) regOf_P root_regOf n tr tr0)
+ unfolding tr Inr_cont_regOf unfolding inj_on_def apply clarsimp using root_regOf
+ apply (metis UNIV_I inItr.Base n pick subtr2.simps subtr_inItr subtr_subtr2)
+ by (metis n subtr.Refl subtr_StepL subtr_regOf tr UNIV_I)
+ qed
+ }
+ thus ?thesis using assms by blast
+qed
+
+(* The regular cut of a tree: *)
+definition "rcut \<equiv> regOf (root tr0)"
+
+theorem reg_rcut: "reg regOf rcut"
+unfolding reg_def rcut_def
+by (metis inItr.Base root_regOf subtr_regOf UNIV_I)
+
+lemma rcut_reg:
+assumes "reg regOf tr0"
+shows "rcut = tr0"
+using assms unfolding rcut_def reg_def by (metis subtr.Refl UNIV_I)
+
+theorem rcut_eq: "rcut = tr0 \<longleftrightarrow> reg regOf tr0"
+using reg_rcut rcut_reg by metis
+
+theorem regular_rcut: "regular rcut"
+using reg_rcut unfolding regular_def by blast
+
+theorem Fr_rcut: "Fr UNIV rcut \<subseteq> Fr UNIV tr0"
+proof safe
+ fix t assume "t \<in> Fr UNIV rcut"
+ then obtain tr where t: "Inl t \<in> cont tr" and tr: "subtr UNIV tr (regOf (root tr0))"
+ using Fr_subtr[of UNIV "regOf (root tr0)"] unfolding rcut_def
+ by (metis (full_types) Fr_def inFr_subtr mem_Collect_eq)
+ obtain n where n: "inItr UNIV tr0 n" and tr: "tr = regOf n" using tr
+ by (metis (lifting) inItr.Base subtr_regOf UNIV_I)
+ have "Inl t \<in> cont (pick n)" using t using Inl_cont_regOf[of n] unfolding tr
+ by (metis (lifting) vimageD vimageI2)
+ moreover have "subtr UNIV (pick n) tr0" using subtr_pick[OF n] ..
+ ultimately show "t \<in> Fr UNIV tr0" unfolding Fr_subtr_cont by auto
+qed
+
+theorem dtree_rcut:
+assumes "dtree tr0"
+shows "dtree rcut"
+unfolding rcut_def using dtree_regOf[OF assms inItr.Base] by simp
+
+theorem root_rcut[simp]: "root rcut = root tr0"
+unfolding rcut_def
+by (metis (lifting) root_regOf inItr.Base reg_def reg_root subtr_rootR_in)
+
+end (* context *)
+
+
+subsection{* Recursive description of the regular tree frontiers *}
+
+lemma regular_inFr:
+assumes r: "regular tr" and In: "root tr \<in> ns"
+and t: "inFr ns tr t"
+shows "t \<in> Inl -` (cont tr) \<or>
+ (\<exists> tr'. Inr tr' \<in> cont tr \<and> inFr (ns - {root tr}) tr' t)"
+(is "?L \<or> ?R")
+proof-
+ obtain f where r: "reg f tr" and f: "\<And>n. root (f n) = n"
+ using r unfolding regular_def2 by auto
+ obtain nl tr1 where d_nl: "distinct nl" and p: "path f nl" and hd_nl: "f (hd nl) = tr"
+ and l_nl: "f (last nl) = tr1" and s_nl: "set nl \<subseteq> ns" and t_tr1: "Inl t \<in> cont tr1"
+ using t unfolding inFr_iff_path[OF r f] by auto
+ obtain n nl1 where nl: "nl = n # nl1" by (metis (lifting) p path.simps)
+ hence f_n: "f n = tr" using hd_nl by simp
+ have n_nl1: "n \<notin> set nl1" using d_nl unfolding nl by auto
+ show ?thesis
+ proof(cases nl1)
+ case Nil hence "tr = tr1" using f_n l_nl unfolding nl by simp
+ hence ?L using t_tr1 by simp thus ?thesis by simp
+ next
+ case (Cons n1 nl2) note nl1 = Cons
+ have 1: "last nl1 = last nl" "hd nl1 = n1" unfolding nl nl1 by simp_all
+ have p1: "path f nl1" and n1_tr: "Inr (f n1) \<in> cont tr"
+ using path.simps[of f nl] p f_n unfolding nl nl1 by auto
+ have r1: "reg f (f n1)" using reg_Inr_cont[OF r n1_tr] .
+ have 0: "inFr (set nl1) (f n1) t" unfolding inFr_iff_path[OF r1 f]
+ apply(intro exI[of _ nl1], intro exI[of _ tr1])
+ using d_nl unfolding 1 l_nl unfolding nl using p1 t_tr1 by auto
+ have root_tr: "root tr = n" by (metis f f_n)
+ have "inFr (ns - {root tr}) (f n1) t" apply(rule inFr_mono[OF 0])
+ using s_nl unfolding root_tr unfolding nl using n_nl1 by auto
+ thus ?thesis using n1_tr by auto
+ qed
+qed
+
+theorem regular_Fr:
+assumes r: "regular tr" and In: "root tr \<in> ns"
+shows "Fr ns tr =
+ Inl -` (cont tr) \<union>
+ \<Union> {Fr (ns - {root tr}) tr' | tr'. Inr tr' \<in> cont tr}"
+unfolding Fr_def
+using In inFr.Base regular_inFr[OF assms] apply safe
+apply (simp, metis (full_types) UnionI mem_Collect_eq)
+apply simp
+by (simp, metis (lifting) inFr_Ind_minus insert_Diff)
+
+
+subsection{* The generated languages *}
+
+(* The (possibly inifinite tree) generated language *)
+definition "L ns n \<equiv> {Fr ns tr | tr. dtree tr \<and> root tr = n}"
+
+(* The regular-tree generated language *)
+definition "Lr ns n \<equiv> {Fr ns tr | tr. dtree tr \<and> root tr = n \<and> regular tr}"
+
+theorem L_rec_notin:
+assumes "n \<notin> ns"
+shows "L ns n = {{}}"
+using assms unfolding L_def apply safe
+ using not_root_Fr apply force
+ apply(rule exI[of _ "deftr n"])
+ by (metis (no_types) dtree_deftr not_root_Fr root_deftr)
+
+theorem Lr_rec_notin:
+assumes "n \<notin> ns"
+shows "Lr ns n = {{}}"
+using assms unfolding Lr_def apply safe
+ using not_root_Fr apply force
+ apply(rule exI[of _ "deftr n"])
+ by (metis (no_types) regular_def dtree_deftr not_root_Fr reg_deftr root_deftr)
+
+lemma dtree_subtrOf:
+assumes "dtree tr" and "Inr n \<in> prodOf tr"
+shows "dtree (subtrOf tr n)"
+by (metis assms dtree_lift lift_def subtrOf)
+
+theorem Lr_rec_in:
+assumes n: "n \<in> ns"
+shows "Lr ns n \<subseteq>
+{Inl -` tns \<union> (\<Union> {K n' | n'. Inr n' \<in> tns}) | tns K.
+ (n,tns) \<in> P \<and>
+ (\<forall> n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> Lr (ns - {n}) n')}"
+(is "Lr ns n \<subseteq> {?F tns K | tns K. (n,tns) \<in> P \<and> ?\<phi> tns K}")
+proof safe
+ fix ts assume "ts \<in> Lr ns n"
+ then obtain tr where dtr: "dtree tr" and r: "root tr = n" and tr: "regular tr"
+ and ts: "ts = Fr ns tr" unfolding Lr_def by auto
+ def tns \<equiv> "(id \<oplus> root) ` (cont tr)"
+ def K \<equiv> "\<lambda> n'. Fr (ns - {n}) (subtrOf tr n')"
+ show "\<exists>tns K. ts = ?F tns K \<and> (n, tns) \<in> P \<and> ?\<phi> tns K"
+ apply(rule exI[of _ tns], rule exI[of _ K]) proof(intro conjI allI impI)
+ show "ts = Inl -` tns \<union> \<Union>{K n' |n'. Inr n' \<in> tns}"
+ unfolding ts regular_Fr[OF tr n[unfolded r[symmetric]]]
+ unfolding tns_def K_def r[symmetric]
+ unfolding Inl_prodOf dtree_subtrOf_Union[OF dtr] ..
+ show "(n, tns) \<in> P" unfolding tns_def r[symmetric] using dtree_P[OF dtr] .
+ fix n' assume "Inr n' \<in> tns" thus "K n' \<in> Lr (ns - {n}) n'"
+ unfolding K_def Lr_def mem_Collect_eq apply(intro exI[of _ "subtrOf tr n'"])
+ using dtr tr apply(intro conjI refl) unfolding tns_def
+ apply(erule dtree_subtrOf[OF dtr])
+ apply (metis subtrOf)
+ by (metis Inr_subtrOf UNIV_I regular_subtr subtr.simps)
+ qed
+qed
+
+lemma hsubst_aux:
+fixes n ftr tns
+assumes n: "n \<in> ns" and tns: "finite tns" and
+1: "\<And> n'. Inr n' \<in> tns \<Longrightarrow> dtree (ftr n')"
+defines "tr \<equiv> Node n ((id \<oplus> ftr) ` tns)" defines "tr' \<equiv> hsubst tr tr"
+shows "Fr ns tr' = Inl -` tns \<union> \<Union>{Fr (ns - {n}) (ftr n') |n'. Inr n' \<in> tns}"
+(is "_ = ?B") proof-
+ have rtr: "root tr = n" and ctr: "cont tr = (id \<oplus> ftr) ` tns"
+ unfolding tr_def using tns by auto
+ have Frr: "Frr (ns - {n}) tr = \<Union>{Fr (ns - {n}) (ftr n') |n'. Inr n' \<in> tns}"
+ unfolding Frr_def ctr by auto
+ have "Fr ns tr' = Inl -` (cont tr) \<union> Frr (ns - {n}) tr"
+ using Fr_self_hsubst[OF n[unfolded rtr[symmetric]]] unfolding tr'_def rtr ..
+ also have "... = ?B" unfolding ctr Frr by simp
+ finally show ?thesis .
+qed
+
+theorem L_rec_in:
+assumes n: "n \<in> ns"
+shows "
+{Inl -` tns \<union> (\<Union> {K n' | n'. Inr n' \<in> tns}) | tns K.
+ (n,tns) \<in> P \<and>
+ (\<forall> n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> L (ns - {n}) n')}
+ \<subseteq> L ns n"
+proof safe
+ fix tns K
+ assume P: "(n, tns) \<in> P" and 0: "\<forall>n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> L (ns - {n}) n'"
+ {fix n' assume "Inr n' \<in> tns"
+ hence "K n' \<in> L (ns - {n}) n'" using 0 by auto
+ hence "\<exists> tr'. K n' = Fr (ns - {n}) tr' \<and> dtree tr' \<and> root tr' = n'"
+ unfolding L_def mem_Collect_eq by auto
+ }
+ then obtain ftr where 0: "\<And> n'. Inr n' \<in> tns \<Longrightarrow>
+ K n' = Fr (ns - {n}) (ftr n') \<and> dtree (ftr n') \<and> root (ftr n') = n'"
+ by metis
+ def tr \<equiv> "Node n ((id \<oplus> ftr) ` tns)" def tr' \<equiv> "hsubst tr tr"
+ have rtr: "root tr = n" and ctr: "cont tr = (id \<oplus> ftr) ` tns"
+ unfolding tr_def by (simp, metis P cont_Node finite_imageI finite_in_P)
+ have prtr: "prodOf tr = tns" apply(rule Inl_Inr_image_cong)
+ unfolding ctr apply simp apply simp apply safe
+ using 0 unfolding image_def apply force apply simp by (metis 0 vimageI2)
+ have 1: "{K n' |n'. Inr n' \<in> tns} = {Fr (ns - {n}) (ftr n') |n'. Inr n' \<in> tns}"
+ using 0 by auto
+ have dtr: "dtree tr" apply(rule dtree.Tree)
+ apply (metis (lifting) P prtr rtr)
+ unfolding inj_on_def ctr lift_def using 0 by auto
+ hence dtr': "dtree tr'" unfolding tr'_def by (metis dtree_hsubst)
+ have tns: "finite tns" using finite_in_P P by simp
+ have "Inl -` tns \<union> \<Union>{Fr (ns - {n}) (ftr n') |n'. Inr n' \<in> tns} \<in> L ns n"
+ unfolding L_def mem_Collect_eq apply(intro exI[of _ tr'] conjI)
+ using dtr' 0 hsubst_aux[OF assms tns, of ftr] unfolding tr_def tr'_def by auto
+ thus "Inl -` tns \<union> \<Union>{K n' |n'. Inr n' \<in> tns} \<in> L ns n" unfolding 1 .
+qed
+
+lemma card_N: "(n::N) \<in> ns \<Longrightarrow> card (ns - {n}) < card ns"
+by (metis finite_N Diff_UNIV Diff_infinite_finite card_Diff1_less finite.emptyI)
+
+function LL where
+"LL ns n =
+ (if n \<notin> ns then {{}} else
+ {Inl -` tns \<union> (\<Union> {K n' | n'. Inr n' \<in> tns}) | tns K.
+ (n,tns) \<in> P \<and>
+ (\<forall> n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> LL (ns - {n}) n')})"
+by(pat_completeness, auto)
+termination apply(relation "inv_image (measure card) fst")
+using card_N by auto
+
+declare LL.simps[code] (* TODO: Does code generation for LL work? *)
+declare LL.simps[simp del]
+
+theorem Lr_LL: "Lr ns n \<subseteq> LL ns n"
+proof (induct ns arbitrary: n rule: measure_induct[of card])
+ case (1 ns n) show ?case proof(cases "n \<in> ns")
+ case False thus ?thesis unfolding Lr_rec_notin[OF False] by (simp add: LL.simps)
+ next
+ case True show ?thesis apply(rule subset_trans)
+ using Lr_rec_in[OF True] apply assumption
+ unfolding LL.simps[of ns n] using True 1 card_N proof clarsimp
+ fix tns K
+ assume "n \<in> ns" hence c: "card (ns - {n}) < card ns" using card_N by blast
+ assume "(n, tns) \<in> P"
+ and "\<forall>n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> Lr (ns - {n}) n'"
+ thus "\<exists>tnsa Ka.
+ Inl -` tns \<union> \<Union>{K n' |n'. Inr n' \<in> tns} =
+ Inl -` tnsa \<union> \<Union>{Ka n' |n'. Inr n' \<in> tnsa} \<and>
+ (n, tnsa) \<in> P \<and> (\<forall>n'. Inr n' \<in> tnsa \<longrightarrow> Ka n' \<in> LL (ns - {n}) n')"
+ apply(intro exI[of _ tns] exI[of _ K]) using c 1 by auto
+ qed
+ qed
+qed
+
+theorem LL_L: "LL ns n \<subseteq> L ns n"
+proof (induct ns arbitrary: n rule: measure_induct[of card])
+ case (1 ns n) show ?case proof(cases "n \<in> ns")
+ case False thus ?thesis unfolding L_rec_notin[OF False] by (simp add: LL.simps)
+ next
+ case True show ?thesis apply(rule subset_trans)
+ prefer 2 using L_rec_in[OF True] apply assumption
+ unfolding LL.simps[of ns n] using True 1 card_N proof clarsimp
+ fix tns K
+ assume "n \<in> ns" hence c: "card (ns - {n}) < card ns" using card_N by blast
+ assume "(n, tns) \<in> P"
+ and "\<forall>n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> LL (ns - {n}) n'"
+ thus "\<exists>tnsa Ka.
+ Inl -` tns \<union> \<Union>{K n' |n'. Inr n' \<in> tns} =
+ Inl -` tnsa \<union> \<Union>{Ka n' |n'. Inr n' \<in> tnsa} \<and>
+ (n, tnsa) \<in> P \<and> (\<forall>n'. Inr n' \<in> tnsa \<longrightarrow> Ka n' \<in> L (ns - {n}) n')"
+ apply(intro exI[of _ tns] exI[of _ K]) using c 1 by auto
+ qed
+ qed
+qed
+
+(* The subsumpsion relation between languages *)
+definition "subs L1 L2 \<equiv> \<forall> ts2 \<in> L2. \<exists> ts1 \<in> L1. ts1 \<subseteq> ts2"
+
+lemma incl_subs[simp]: "L2 \<subseteq> L1 \<Longrightarrow> subs L1 L2"
+unfolding subs_def by auto
+
+lemma subs_refl[simp]: "subs L1 L1" unfolding subs_def by auto
+
+lemma subs_trans: "\<lbrakk>subs L1 L2; subs L2 L3\<rbrakk> \<Longrightarrow> subs L1 L3"
+unfolding subs_def by (metis subset_trans)
+
+(* Language equivalence *)
+definition "leqv L1 L2 \<equiv> subs L1 L2 \<and> subs L2 L1"
+
+lemma subs_leqv[simp]: "leqv L1 L2 \<Longrightarrow> subs L1 L2"
+unfolding leqv_def by auto
+
+lemma subs_leqv_sym[simp]: "leqv L1 L2 \<Longrightarrow> subs L2 L1"
+unfolding leqv_def by auto
+
+lemma leqv_refl[simp]: "leqv L1 L1" unfolding leqv_def by auto
+
+lemma leqv_trans:
+assumes 12: "leqv L1 L2" and 23: "leqv L2 L3"
+shows "leqv L1 L3"
+using assms unfolding leqv_def by (metis (lifting) subs_trans)
+
+lemma leqv_sym: "leqv L1 L2 \<Longrightarrow> leqv L2 L1"
+unfolding leqv_def by auto
+
+lemma leqv_Sym: "leqv L1 L2 \<longleftrightarrow> leqv L2 L1"
+unfolding leqv_def by auto
+
+lemma Lr_incl_L: "Lr ns ts \<subseteq> L ns ts"
+unfolding Lr_def L_def by auto
+
+lemma Lr_subs_L: "subs (Lr UNIV ts) (L UNIV ts)"
+unfolding subs_def proof safe
+ fix ts2 assume "ts2 \<in> L UNIV ts"
+ then obtain tr where ts2: "ts2 = Fr UNIV tr" and dtr: "dtree tr" and rtr: "root tr = ts"
+ unfolding L_def by auto
+ thus "\<exists>ts1\<in>Lr UNIV ts. ts1 \<subseteq> ts2"
+ apply(intro bexI[of _ "Fr UNIV (rcut tr)"])
+ unfolding Lr_def L_def using Fr_rcut dtree_rcut root_rcut regular_rcut by auto
+qed
+
+theorem Lr_leqv_L: "leqv (Lr UNIV ts) (L UNIV ts)"
+using Lr_subs_L unfolding leqv_def by (metis (lifting) Lr_incl_L incl_subs)
+
+theorem LL_leqv_L: "leqv (LL UNIV ts) (L UNIV ts)"
+by (metis (lifting) LL_L Lr_LL Lr_subs_L incl_subs leqv_def subs_trans)
+
+theorem LL_leqv_Lr: "leqv (LL UNIV ts) (Lr UNIV ts)"
+using Lr_leqv_L LL_leqv_L by (metis leqv_Sym leqv_trans)
+
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Examples/Infinite_Derivation_Trees/Parallel.thy Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,152 @@
+(* Title: HOL/BNF/Examples/Infinite_Derivation_Trees/Parallel.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Parallel composition.
+*)
+
+header {* Parallel Composition *}
+
+theory Parallel
+imports Tree
+begin
+
+
+consts Nplus :: "N \<Rightarrow> N \<Rightarrow> N" (infixl "+" 60)
+
+axiomatization where
+ Nplus_comm: "(a::N) + b = b + (a::N)"
+and Nplus_assoc: "((a::N) + b) + c = a + (b + c)"
+
+
+
+section{* Parallel composition *}
+
+fun par_r where "par_r (tr1,tr2) = root tr1 + root tr2"
+fun par_c where
+"par_c (tr1,tr2) =
+ Inl ` (Inl -` (cont tr1 \<union> cont tr2)) \<union>
+ Inr ` (Inr -` cont tr1 \<times> Inr -` cont tr2)"
+
+declare par_r.simps[simp del] declare par_c.simps[simp del]
+
+definition par :: "Tree \<times> Tree \<Rightarrow> Tree" where
+"par \<equiv> unfold par_r par_c"
+
+abbreviation par_abbr (infixr "\<parallel>" 80) where "tr1 \<parallel> tr2 \<equiv> par (tr1, tr2)"
+
+lemma finite_par_c: "finite (par_c (tr1, tr2))"
+unfolding par_c.simps apply(rule finite_UnI)
+ apply (metis finite_Un finite_cont finite_imageI finite_vimageI inj_Inl)
+ apply(intro finite_imageI finite_cartesian_product finite_vimageI)
+ using finite_cont by auto
+
+lemma root_par: "root (tr1 \<parallel> tr2) = root tr1 + root tr2"
+using unfold(1)[of par_r par_c "(tr1,tr2)"] unfolding par_def par_r.simps by simp
+
+lemma cont_par:
+"cont (tr1 \<parallel> tr2) = (id \<oplus> par) ` par_c (tr1,tr2)"
+using unfold(2)[of par_c "(tr1,tr2)" par_r, OF finite_par_c]
+unfolding par_def ..
+
+lemma Inl_cont_par[simp]:
+"Inl -` (cont (tr1 \<parallel> tr2)) = Inl -` (cont tr1 \<union> cont tr2)"
+unfolding cont_par par_c.simps by auto
+
+lemma Inr_cont_par[simp]:
+"Inr -` (cont (tr1 \<parallel> tr2)) = par ` (Inr -` cont tr1 \<times> Inr -` cont tr2)"
+unfolding cont_par par_c.simps by auto
+
+lemma Inl_in_cont_par:
+"Inl t \<in> cont (tr1 \<parallel> tr2) \<longleftrightarrow> (Inl t \<in> cont tr1 \<or> Inl t \<in> cont tr2)"
+using Inl_cont_par[of tr1 tr2] unfolding vimage_def by auto
+
+lemma Inr_in_cont_par:
+"Inr t \<in> cont (tr1 \<parallel> tr2) \<longleftrightarrow> (t \<in> par ` (Inr -` cont tr1 \<times> Inr -` cont tr2))"
+using Inr_cont_par[of tr1 tr2] unfolding vimage_def by auto
+
+
+section{* =-coinductive proofs *}
+
+(* Detailed proofs of commutativity and associativity: *)
+theorem par_com: "tr1 \<parallel> tr2 = tr2 \<parallel> tr1"
+proof-
+ let ?\<phi> = "\<lambda> trA trB. \<exists> tr1 tr2. trA = tr1 \<parallel> tr2 \<and> trB = tr2 \<parallel> tr1"
+ {fix trA trB
+ assume "?\<phi> trA trB" hence "trA = trB"
+ proof (induct rule: Tree_coind, safe)
+ fix tr1 tr2
+ show "root (tr1 \<parallel> tr2) = root (tr2 \<parallel> tr1)"
+ unfolding root_par by (rule Nplus_comm)
+ next
+ fix tr1 tr2 :: Tree
+ let ?trA = "tr1 \<parallel> tr2" let ?trB = "tr2 \<parallel> tr1"
+ show "(?\<phi> ^#2) (cont ?trA) (cont ?trB)"
+ unfolding lift2_def proof(intro conjI allI impI)
+ fix n show "Inl n \<in> cont (tr1 \<parallel> tr2) \<longleftrightarrow> Inl n \<in> cont (tr2 \<parallel> tr1)"
+ unfolding Inl_in_cont_par by auto
+ next
+ fix trA' assume "Inr trA' \<in> cont ?trA"
+ then obtain tr1' tr2' where "trA' = tr1' \<parallel> tr2'"
+ and "Inr tr1' \<in> cont tr1" and "Inr tr2' \<in> cont tr2"
+ unfolding Inr_in_cont_par by auto
+ thus "\<exists> trB'. Inr trB' \<in> cont ?trB \<and> ?\<phi> trA' trB'"
+ apply(intro exI[of _ "tr2' \<parallel> tr1'"]) unfolding Inr_in_cont_par by auto
+ next
+ fix trB' assume "Inr trB' \<in> cont ?trB"
+ then obtain tr1' tr2' where "trB' = tr2' \<parallel> tr1'"
+ and "Inr tr1' \<in> cont tr1" and "Inr tr2' \<in> cont tr2"
+ unfolding Inr_in_cont_par by auto
+ thus "\<exists> trA'. Inr trA' \<in> cont ?trA \<and> ?\<phi> trA' trB'"
+ apply(intro exI[of _ "tr1' \<parallel> tr2'"]) unfolding Inr_in_cont_par by auto
+ qed
+ qed
+ }
+ thus ?thesis by blast
+qed
+
+theorem par_assoc: "(tr1 \<parallel> tr2) \<parallel> tr3 = tr1 \<parallel> (tr2 \<parallel> tr3)"
+proof-
+ let ?\<phi> =
+ "\<lambda> trA trB. \<exists> tr1 tr2 tr3. trA = (tr1 \<parallel> tr2) \<parallel> tr3 \<and>
+ trB = tr1 \<parallel> (tr2 \<parallel> tr3)"
+ {fix trA trB
+ assume "?\<phi> trA trB" hence "trA = trB"
+ proof (induct rule: Tree_coind, safe)
+ fix tr1 tr2 tr3
+ show "root ((tr1 \<parallel> tr2) \<parallel> tr3) = root (tr1 \<parallel> (tr2 \<parallel> tr3))"
+ unfolding root_par by (rule Nplus_assoc)
+ next
+ fix tr1 tr2 tr3
+ let ?trA = "(tr1 \<parallel> tr2) \<parallel> tr3" let ?trB = "tr1 \<parallel> (tr2 \<parallel> tr3)"
+ show "(?\<phi> ^#2) (cont ?trA) (cont ?trB)"
+ unfolding lift2_def proof(intro conjI allI impI)
+ fix n show "Inl n \<in> (cont ?trA) \<longleftrightarrow> Inl n \<in> (cont ?trB)"
+ unfolding Inl_in_cont_par by simp
+ next
+ fix trA' assume "Inr trA' \<in> cont ?trA"
+ then obtain tr1' tr2' tr3' where "trA' = (tr1' \<parallel> tr2') \<parallel> tr3'"
+ and "Inr tr1' \<in> cont tr1" and "Inr tr2' \<in> cont tr2"
+ and "Inr tr3' \<in> cont tr3" unfolding Inr_in_cont_par by auto
+ thus "\<exists> trB'. Inr trB' \<in> cont ?trB \<and> ?\<phi> trA' trB'"
+ apply(intro exI[of _ "tr1' \<parallel> (tr2' \<parallel> tr3')"])
+ unfolding Inr_in_cont_par by auto
+ next
+ fix trB' assume "Inr trB' \<in> cont ?trB"
+ then obtain tr1' tr2' tr3' where "trB' = tr1' \<parallel> (tr2' \<parallel> tr3')"
+ and "Inr tr1' \<in> cont tr1" and "Inr tr2' \<in> cont tr2"
+ and "Inr tr3' \<in> cont tr3" unfolding Inr_in_cont_par by auto
+ thus "\<exists> trA'. Inr trA' \<in> cont ?trA \<and> ?\<phi> trA' trB'"
+ apply(intro exI[of _ "(tr1' \<parallel> tr2') \<parallel> tr3'"])
+ unfolding Inr_in_cont_par by auto
+ qed
+ qed
+ }
+ thus ?thesis by blast
+qed
+
+
+
+
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Examples/Infinite_Derivation_Trees/Prelim.thy Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,67 @@
+(* Title: HOL/BNF/Examples/Infinite_Derivation_Trees/Prelim.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Preliminaries.
+*)
+
+header {* Preliminaries *}
+
+theory Prelim
+imports "../../BNF"
+begin
+
+declare fset_to_fset[simp]
+
+lemma fst_snd_convol_o[simp]: "<fst o s, snd o s> = s"
+apply(rule ext) by (simp add: convol_def)
+
+abbreviation sm_abbrev (infix "\<oplus>" 60)
+where "f \<oplus> g \<equiv> Sum_Type.sum_map f g"
+
+lemma sum_map_InlD: "(f \<oplus> g) z = Inl x \<Longrightarrow> \<exists>y. z = Inl y \<and> f y = x"
+by (cases z) auto
+
+lemma sum_map_InrD: "(f \<oplus> g) z = Inr x \<Longrightarrow> \<exists>y. z = Inr y \<and> g y = x"
+by (cases z) auto
+
+abbreviation sum_case_abbrev ("[[_,_]]" 800)
+where "[[f,g]] \<equiv> Sum_Type.sum_case f g"
+
+lemma inj_Inl[simp]: "inj Inl" unfolding inj_on_def by auto
+lemma inj_Inr[simp]: "inj Inr" unfolding inj_on_def by auto
+
+lemma Inl_oplus_elim:
+assumes "Inl tr \<in> (id \<oplus> f) ` tns"
+shows "Inl tr \<in> tns"
+using assms apply clarify by (case_tac x, auto)
+
+lemma Inl_oplus_iff[simp]: "Inl tr \<in> (id \<oplus> f) ` tns \<longleftrightarrow> Inl tr \<in> tns"
+using Inl_oplus_elim
+by (metis id_def image_iff sum_map.simps(1))
+
+lemma Inl_m_oplus[simp]: "Inl -` (id \<oplus> f) ` tns = Inl -` tns"
+using Inl_oplus_iff unfolding vimage_def by auto
+
+lemma Inr_oplus_elim:
+assumes "Inr tr \<in> (id \<oplus> f) ` tns"
+shows "\<exists> n. Inr n \<in> tns \<and> f n = tr"
+using assms apply clarify by (case_tac x, auto)
+
+lemma Inr_oplus_iff[simp]:
+"Inr tr \<in> (id \<oplus> f) ` tns \<longleftrightarrow> (\<exists> n. Inr n \<in> tns \<and> f n = tr)"
+apply (rule iffI)
+ apply (metis Inr_oplus_elim)
+by (metis image_iff sum_map.simps(2))
+
+lemma Inr_m_oplus[simp]: "Inr -` (id \<oplus> f) ` tns = f ` (Inr -` tns)"
+using Inr_oplus_iff unfolding vimage_def by auto
+
+lemma Inl_Inr_image_cong:
+assumes "Inl -` A = Inl -` B" and "Inr -` A = Inr -` B"
+shows "A = B"
+apply safe using assms apply(case_tac x, auto) by(case_tac x, auto)
+
+
+
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Examples/Infinite_Derivation_Trees/Tree.thy Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,326 @@
+(* Title: HOL/BNF/Examples/Infinite_Derivation_Trees/Tree.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Trees with nonterminal internal nodes and terminal leaves.
+*)
+
+header {* Trees with Nonterminal Internal Nodes and Terminal Leaves *}
+
+theory Tree
+imports Prelim
+begin
+
+hide_fact (open) Quotient_Product.prod_rel_def
+
+typedecl N typedecl T
+
+codata_raw Tree: 'Tree = "N \<times> (T + 'Tree) fset"
+
+
+section {* Sugar notations for Tree *}
+
+subsection{* Setup for map, set, rel *}
+
+(* These should be eventually inferred from compositionality *)
+
+lemma pre_Tree_map:
+"pre_Tree_map f (n, as) = (n, map_fset (id \<oplus> f) as)"
+unfolding pre_Tree_map_def id_apply
+sum_map_def by simp
+
+lemma pre_Tree_map':
+"pre_Tree_map f n_as = (fst n_as, map_fset (id \<oplus> f) (snd n_as))"
+using pre_Tree_map by(cases n_as, simp)
+
+
+definition
+"llift2 \<phi> as1 as2 \<longleftrightarrow>
+ (\<forall> n. Inl n \<in> fset as1 \<longleftrightarrow> Inl n \<in> fset as2) \<and>
+ (\<forall> tr1. Inr tr1 \<in> fset as1 \<longrightarrow> (\<exists> tr2. Inr tr2 \<in> fset as2 \<and> \<phi> tr1 tr2)) \<and>
+ (\<forall> tr2. Inr tr2 \<in> fset as2 \<longrightarrow> (\<exists> tr1. Inr tr1 \<in> fset as1 \<and> \<phi> tr1 tr2))"
+
+lemma pre_Tree_rel: "pre_Tree_rel \<phi> (n1,as1) (n2,as2) \<longleftrightarrow> n1 = n2 \<and> llift2 \<phi> as1 as2"
+unfolding llift2_def pre_Tree_rel_def sum_rel_def[abs_def] prod_rel_def fset_rel_def split_conv
+apply (auto split: sum.splits)
+apply (metis sumE)
+apply (metis sumE)
+apply (metis sumE)
+apply (metis sumE)
+apply (metis sumE sum.simps(1,2,4))
+apply (metis sumE sum.simps(1,2,4))
+done
+
+
+subsection{* Constructors *}
+
+definition NNode :: "N \<Rightarrow> (T + Tree)fset \<Rightarrow> Tree"
+where "NNode n as \<equiv> Tree_ctor (n,as)"
+
+lemmas ctor_defs = NNode_def
+
+
+subsection {* Pre-selectors *}
+
+(* These are mere auxiliaries *)
+
+definition "asNNode tr \<equiv> SOME n_as. NNode (fst n_as) (snd n_as) = tr"
+lemmas pre_sel_defs = asNNode_def
+
+
+subsection {* Selectors *}
+
+(* One for each pair (constructor, constructor argument) *)
+
+(* For NNode: *)
+definition root :: "Tree \<Rightarrow> N" where "root tr = fst (asNNode tr)"
+definition ccont :: "Tree \<Rightarrow> (T + Tree)fset" where "ccont tr = snd (asNNode tr)"
+
+lemmas sel_defs = root_def ccont_def
+
+
+subsection {* Basic properties *}
+
+(* Constructors versus selectors *)
+lemma NNode_surj: "\<exists> n as. NNode n as = tr"
+unfolding NNode_def
+by (metis Tree.ctor_dtor pair_collapse)
+
+lemma NNode_asNNode:
+"NNode (fst (asNNode tr)) (snd (asNNode tr)) = tr"
+proof-
+ obtain n as where "NNode n as = tr" using NNode_surj[of tr] by blast
+ hence "NNode (fst (n,as)) (snd (n,as)) = tr" by simp
+ thus ?thesis unfolding asNNode_def by(rule someI)
+qed
+
+theorem NNode_root_ccont[simp]:
+"NNode (root tr) (ccont tr) = tr"
+using NNode_asNNode unfolding root_def ccont_def .
+
+(* Constructors *)
+theorem TTree_simps[simp]:
+"NNode n as = NNode n' as' \<longleftrightarrow> n = n' \<and> as = as'"
+unfolding ctor_defs Tree.ctor_inject by auto
+
+theorem TTree_cases[elim, case_names NNode Choice]:
+assumes NNode: "\<And> n as. tr = NNode n as \<Longrightarrow> phi"
+shows phi
+proof(cases rule: Tree.ctor_exhaust[of tr])
+ fix x assume "tr = Tree_ctor x"
+ thus ?thesis
+ apply(cases x)
+ using NNode unfolding ctor_defs apply blast
+ done
+qed
+
+(* Constructors versus selectors *)
+theorem TTree_sel_ctor[simp]:
+"root (NNode n as) = n"
+"ccont (NNode n as) = as"
+unfolding root_def ccont_def
+by (metis (no_types) NNode_asNNode TTree_simps)+
+
+
+subsection{* Coinduction *}
+
+theorem TTree_coind_Node[elim, consumes 1, case_names NNode, induct pred: "HOL.eq"]:
+assumes phi: "\<phi> tr1 tr2" and
+NNode: "\<And> n1 n2 as1 as2.
+ \<lbrakk>\<phi> (NNode n1 as1) (NNode n2 as2)\<rbrakk> \<Longrightarrow>
+ n1 = n2 \<and> llift2 \<phi> as1 as2"
+shows "tr1 = tr2"
+apply(rule mp[OF Tree.rel_coinduct[of \<phi> tr1 tr2] phi]) proof clarify
+ fix tr1 tr2 assume \<phi>: "\<phi> tr1 tr2"
+ show "pre_Tree_rel \<phi> (Tree_dtor tr1) (Tree_dtor tr2)"
+ apply(cases rule: Tree.ctor_exhaust[of tr1], cases rule: Tree.ctor_exhaust[of tr2])
+ apply (simp add: Tree.dtor_ctor)
+ apply(case_tac x, case_tac xa, simp)
+ unfolding pre_Tree_rel apply(rule NNode) using \<phi> unfolding NNode_def by simp
+qed
+
+theorem TTree_coind[elim, consumes 1, case_names LLift]:
+assumes phi: "\<phi> tr1 tr2" and
+LLift: "\<And> tr1 tr2. \<phi> tr1 tr2 \<Longrightarrow>
+ root tr1 = root tr2 \<and> llift2 \<phi> (ccont tr1) (ccont tr2)"
+shows "tr1 = tr2"
+using phi apply(induct rule: TTree_coind_Node)
+using LLift by (metis TTree_sel_ctor)
+
+
+
+subsection {* Coiteration *}
+
+(* Preliminaries: *)
+declare Tree.dtor_ctor[simp]
+declare Tree.ctor_dtor[simp]
+
+lemma Tree_dtor_NNode[simp]:
+"Tree_dtor (NNode n as) = (n,as)"
+unfolding NNode_def Tree.dtor_ctor ..
+
+lemma Tree_dtor_root_ccont:
+"Tree_dtor tr = (root tr, ccont tr)"
+unfolding root_def ccont_def
+by (metis (lifting) NNode_asNNode Tree_dtor_NNode)
+
+(* Coiteration *)
+definition TTree_unfold ::
+"('b \<Rightarrow> N) \<Rightarrow> ('b \<Rightarrow> (T + 'b) fset) \<Rightarrow> 'b \<Rightarrow> Tree"
+where "TTree_unfold rt ct \<equiv> Tree_dtor_unfold <rt,ct>"
+
+lemma Tree_unfold_unfold:
+"Tree_dtor_unfold s = TTree_unfold (fst o s) (snd o s)"
+apply(rule ext)
+unfolding TTree_unfold_def by simp
+
+theorem TTree_unfold:
+"root (TTree_unfold rt ct b) = rt b"
+"ccont (TTree_unfold rt ct b) = map_fset (id \<oplus> TTree_unfold rt ct) (ct b)"
+using Tree.dtor_unfolds[of "<rt,ct>" b] unfolding Tree_unfold_unfold fst_convol snd_convol
+unfolding pre_Tree_map' fst_convol' snd_convol'
+unfolding Tree_dtor_root_ccont by simp_all
+
+(* Corecursion, stronger than coiteration (unfold) *)
+definition TTree_corec ::
+"('b \<Rightarrow> N) \<Rightarrow> ('b \<Rightarrow> (T + (Tree + 'b)) fset) \<Rightarrow> 'b \<Rightarrow> Tree"
+where "TTree_corec rt ct \<equiv> Tree_dtor_corec <rt,ct>"
+
+lemma Tree_dtor_corec_corec:
+"Tree_dtor_corec s = TTree_corec (fst o s) (snd o s)"
+apply(rule ext)
+unfolding TTree_corec_def by simp
+
+theorem TTree_corec:
+"root (TTree_corec rt ct b) = rt b"
+"ccont (TTree_corec rt ct b) = map_fset (id \<oplus> ([[id, TTree_corec rt ct]]) ) (ct b)"
+using Tree.dtor_corecs[of "<rt,ct>" b] unfolding Tree_dtor_corec_corec fst_convol snd_convol
+unfolding pre_Tree_map' fst_convol' snd_convol'
+unfolding Tree_dtor_root_ccont by simp_all
+
+
+subsection{* The characteristic theorems transported from fset to set *}
+
+definition "Node n as \<equiv> NNode n (the_inv fset as)"
+definition "cont \<equiv> fset o ccont"
+definition "unfold rt ct \<equiv> TTree_unfold rt (the_inv fset o ct)"
+definition "corec rt ct \<equiv> TTree_corec rt (the_inv fset o ct)"
+
+definition lift ("_ ^#" 200) where
+"lift \<phi> as \<longleftrightarrow> (\<forall> tr. Inr tr \<in> as \<longrightarrow> \<phi> tr)"
+
+definition lift2 ("_ ^#2" 200) where
+"lift2 \<phi> as1 as2 \<longleftrightarrow>
+ (\<forall> n. Inl n \<in> as1 \<longleftrightarrow> Inl n \<in> as2) \<and>
+ (\<forall> tr1. Inr tr1 \<in> as1 \<longrightarrow> (\<exists> tr2. Inr tr2 \<in> as2 \<and> \<phi> tr1 tr2)) \<and>
+ (\<forall> tr2. Inr tr2 \<in> as2 \<longrightarrow> (\<exists> tr1. Inr tr1 \<in> as1 \<and> \<phi> tr1 tr2))"
+
+definition liftS ("_ ^#s" 200) where
+"liftS trs = {as. Inr -` as \<subseteq> trs}"
+
+lemma lift2_llift2:
+"\<lbrakk>finite as1; finite as2\<rbrakk> \<Longrightarrow>
+ lift2 \<phi> as1 as2 \<longleftrightarrow> llift2 \<phi> (the_inv fset as1) (the_inv fset as2)"
+unfolding lift2_def llift2_def by auto
+
+lemma llift2_lift2:
+"llift2 \<phi> as1 as2 \<longleftrightarrow> lift2 \<phi> (fset as1) (fset as2)"
+using lift2_llift2 by (metis finite_fset fset_cong fset_to_fset)
+
+lemma mono_lift:
+assumes "(\<phi>^#) as"
+and "\<And> tr. \<phi> tr \<Longrightarrow> \<phi>' tr"
+shows "(\<phi>'^#) as"
+using assms unfolding lift_def[abs_def] by blast
+
+lemma mono_liftS:
+assumes "trs1 \<subseteq> trs2 "
+shows "(trs1 ^#s) \<subseteq> (trs2 ^#s)"
+using assms unfolding liftS_def[abs_def] by blast
+
+lemma lift_mono:
+assumes "\<phi> \<le> \<phi>'"
+shows "(\<phi>^#) \<le> (\<phi>'^#)"
+using assms unfolding lift_def[abs_def] by blast
+
+lemma mono_lift2:
+assumes "(\<phi>^#2) as1 as2"
+and "\<And> tr1 tr2. \<phi> tr1 tr2 \<Longrightarrow> \<phi>' tr1 tr2"
+shows "(\<phi>'^#2) as1 as2"
+using assms unfolding lift2_def[abs_def] by blast
+
+lemma lift2_mono:
+assumes "\<phi> \<le> \<phi>'"
+shows "(\<phi>^#2) \<le> (\<phi>'^#2)"
+using assms unfolding lift2_def[abs_def] by blast
+
+lemma finite_cont[simp]: "finite (cont tr)"
+unfolding cont_def by auto
+
+theorem Node_root_cont[simp]:
+"Node (root tr) (cont tr) = tr"
+using NNode_root_ccont unfolding Node_def cont_def
+by (metis cont_def finite_cont fset_cong fset_to_fset o_def)
+
+theorem Tree_simps[simp]:
+assumes "finite as" and "finite as'"
+shows "Node n as = Node n' as' \<longleftrightarrow> n = n' \<and> as = as'"
+using assms TTree_simps unfolding Node_def
+by (metis fset_to_fset)
+
+theorem Tree_cases[elim, case_names Node Choice]:
+assumes Node: "\<And> n as. \<lbrakk>finite as; tr = Node n as\<rbrakk> \<Longrightarrow> phi"
+shows phi
+apply(cases rule: TTree_cases[of tr])
+using Node unfolding Node_def
+by (metis Node Node_root_cont finite_cont)
+
+theorem Tree_sel_ctor[simp]:
+"root (Node n as) = n"
+"finite as \<Longrightarrow> cont (Node n as) = as"
+unfolding Node_def cont_def by auto
+
+theorems root_Node = Tree_sel_ctor(1)
+theorems cont_Node = Tree_sel_ctor(2)
+
+theorem Tree_coind_Node[elim, consumes 1, case_names Node]:
+assumes phi: "\<phi> tr1 tr2" and
+Node:
+"\<And> n1 n2 as1 as2.
+ \<lbrakk>finite as1; finite as2; \<phi> (Node n1 as1) (Node n2 as2)\<rbrakk>
+ \<Longrightarrow> n1 = n2 \<and> (\<phi>^#2) as1 as2"
+shows "tr1 = tr2"
+using phi apply(induct rule: TTree_coind_Node)
+unfolding llift2_lift2 apply(rule Node)
+unfolding Node_def
+apply (metis finite_fset)
+apply (metis finite_fset)
+by (metis finite_fset fset_cong fset_to_fset)
+
+theorem Tree_coind[elim, consumes 1, case_names Lift, induct pred: "HOL.eq"]:
+assumes phi: "\<phi> tr1 tr2" and
+Lift: "\<And> tr1 tr2. \<phi> tr1 tr2 \<Longrightarrow>
+ root tr1 = root tr2 \<and> (\<phi>^#2) (cont tr1) (cont tr2)"
+shows "tr1 = tr2"
+using phi apply(induct rule: TTree_coind)
+unfolding llift2_lift2 apply(rule Lift[unfolded cont_def comp_def]) .
+
+theorem unfold:
+"root (unfold rt ct b) = rt b"
+"finite (ct b) \<Longrightarrow> cont (unfold rt ct b) = image (id \<oplus> unfold rt ct) (ct b)"
+using TTree_unfold[of rt "the_inv fset \<circ> ct" b] unfolding unfold_def
+apply - apply metis
+unfolding cont_def comp_def
+by (metis (no_types) fset_to_fset map_fset_image)
+
+
+theorem corec:
+"root (corec rt ct b) = rt b"
+"finite (ct b) \<Longrightarrow> cont (corec rt ct b) = image (id \<oplus> ([[id, corec rt ct]])) (ct b)"
+using TTree_corec[of rt "the_inv fset \<circ> ct" b] unfolding corec_def
+apply - apply metis
+unfolding cont_def comp_def
+by (metis (no_types) fset_to_fset map_fset_image)
+
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Examples/Lambda_Term.thy Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,259 @@
+(* Title: HOL/BNF/Examples/Lambda_Term.thy
+ Author: Dmitriy Traytel, TU Muenchen
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Lambda-terms.
+*)
+
+header {* Lambda-Terms *}
+
+theory Lambda_Term
+imports "../BNF"
+begin
+
+
+section {* Datatype definition *}
+
+data_raw trm: 'trm = "'a + 'trm \<times> 'trm + 'a \<times> 'trm + ('a \<times> 'trm) fset \<times> 'trm"
+
+
+section {* Customization of terms *}
+
+subsection{* Set and map *}
+
+lemma pre_trm_set2_Lt: "pre_trm_set2 (Inr (Inr (Inr (xts, t)))) = snd ` (fset xts) \<union> {t}"
+unfolding pre_trm_set2_def sum_set_defs prod_set_defs collect_def[abs_def]
+by auto
+
+lemma pre_trm_set2_Var: "\<And>x. pre_trm_set2 (Inl x) = {}"
+and pre_trm_set2_App:
+"\<And>t1 t2. pre_trm_set2 (Inr (Inl t1t2)) = {fst t1t2, snd t1t2}"
+and pre_trm_set2_Lam:
+"\<And>x t. pre_trm_set2 (Inr (Inr (Inl (x, t)))) = {t}"
+unfolding pre_trm_set2_def sum_set_defs prod_set_defs collect_def[abs_def]
+by auto
+
+lemma pre_trm_map:
+"\<And> a1. pre_trm_map f1 f2 (Inl a1) = Inl (f1 a1)"
+"\<And> a2 b2. pre_trm_map f1 f2 (Inr (Inl (a2,b2))) = Inr (Inl (f2 a2, f2 b2))"
+"\<And> a1 a2. pre_trm_map f1 f2 (Inr (Inr (Inl (a1,a2)))) = Inr (Inr (Inl (f1 a1, f2 a2)))"
+"\<And> a1a2s a2.
+ pre_trm_map f1 f2 (Inr (Inr (Inr (a1a2s, a2)))) =
+ Inr (Inr (Inr (map_fset (\<lambda> (a1', a2'). (f1 a1', f2 a2')) a1a2s, f2 a2)))"
+unfolding pre_trm_map_def collect_def[abs_def] map_pair_def by auto
+
+
+subsection{* Constructors *}
+
+definition "Var x \<equiv> trm_ctor (Inl x)"
+definition "App t1 t2 \<equiv> trm_ctor (Inr (Inl (t1,t2)))"
+definition "Lam x t \<equiv> trm_ctor (Inr (Inr (Inl (x,t))))"
+definition "Lt xts t \<equiv> trm_ctor (Inr (Inr (Inr (xts,t))))"
+
+lemmas ctor_defs = Var_def App_def Lam_def Lt_def
+
+theorem trm_simps[simp]:
+"\<And>x y. Var x = Var y \<longleftrightarrow> x = y"
+"\<And>t1 t2 t1' t2'. App t1 t2 = App t1' t2' \<longleftrightarrow> t1 = t1' \<and> t2 = t2'"
+"\<And>x x' t t'. Lam x t = Lam x' t' \<longleftrightarrow> x = x' \<and> t = t'"
+"\<And> xts xts' t t'. Lt xts t = Lt xts' t' \<longleftrightarrow> xts = xts' \<and> t = t'"
+(* *)
+"\<And> x t1 t2. Var x \<noteq> App t1 t2" "\<And>x y t. Var x \<noteq> Lam y t" "\<And> x xts t. Var x \<noteq> Lt xts t"
+"\<And> t1 t2 x t. App t1 t2 \<noteq> Lam x t" "\<And> t1 t2 xts t. App t1 t2 \<noteq> Lt xts t"
+"\<And>x t xts t1. Lam x t \<noteq> Lt xts t1"
+unfolding ctor_defs trm.ctor_inject by auto
+
+theorem trm_cases[elim, case_names Var App Lam Lt]:
+assumes Var: "\<And> x. t = Var x \<Longrightarrow> phi"
+and App: "\<And> t1 t2. t = App t1 t2 \<Longrightarrow> phi"
+and Lam: "\<And> x t1. t = Lam x t1 \<Longrightarrow> phi"
+and Lt: "\<And> xts t1. t = Lt xts t1 \<Longrightarrow> phi"
+shows phi
+proof(cases rule: trm.ctor_exhaust[of t])
+ fix x assume "t = trm_ctor x"
+ thus ?thesis
+ apply(cases x) using Var unfolding ctor_defs apply blast
+ apply(case_tac b) using App unfolding ctor_defs apply(case_tac a, blast)
+ apply(case_tac ba) using Lam unfolding ctor_defs apply(case_tac a, blast)
+ apply(case_tac bb) using Lt unfolding ctor_defs by blast
+qed
+
+lemma trm_induct[case_names Var App Lam Lt, induct type: trm]:
+assumes Var: "\<And> (x::'a). phi (Var x)"
+and App: "\<And> t1 t2. \<lbrakk>phi t1; phi t2\<rbrakk> \<Longrightarrow> phi (App t1 t2)"
+and Lam: "\<And> x t. phi t \<Longrightarrow> phi (Lam x t)"
+and Lt: "\<And> xts t. \<lbrakk>\<And> x1 t1. (x1,t1) |\<in>| xts \<Longrightarrow> phi t1; phi t\<rbrakk> \<Longrightarrow> phi (Lt xts t)"
+shows "phi t"
+proof(induct rule: trm.ctor_induct)
+ fix u :: "'a + 'a trm \<times> 'a trm + 'a \<times> 'a trm + ('a \<times> 'a trm) fset \<times> 'a trm"
+ assume IH: "\<And>t. t \<in> pre_trm_set2 u \<Longrightarrow> phi t"
+ show "phi (trm_ctor u)"
+ proof(cases u)
+ case (Inl x)
+ show ?thesis using Var unfolding Var_def Inl .
+ next
+ case (Inr uu) note Inr1 = Inr
+ show ?thesis
+ proof(cases uu)
+ case (Inl t1t2)
+ obtain t1 t2 where t1t2: "t1t2 = (t1,t2)" by (cases t1t2, blast)
+ show ?thesis unfolding Inr1 Inl t1t2 App_def[symmetric] apply(rule App)
+ using IH unfolding Inr1 Inl pre_trm_set2_App t1t2 fst_conv snd_conv by blast+
+ next
+ case (Inr uuu) note Inr2 = Inr
+ show ?thesis
+ proof(cases uuu)
+ case (Inl xt)
+ obtain x t where xt: "xt = (x,t)" by (cases xt, blast)
+ show ?thesis unfolding Inr1 Inr2 Inl xt Lam_def[symmetric] apply(rule Lam)
+ using IH unfolding Inr1 Inr2 Inl pre_trm_set2_Lam xt by blast
+ next
+ case (Inr xts_t)
+ obtain xts t where xts_t: "xts_t = (xts,t)" by (cases xts_t, blast)
+ show ?thesis unfolding Inr1 Inr2 Inr xts_t Lt_def[symmetric] apply(rule Lt) using IH
+ unfolding Inr1 Inr2 Inr pre_trm_set2_Lt xts_t fset_fset_member image_def by auto
+ qed
+ qed
+ qed
+qed
+
+
+subsection{* Recursion and iteration (fold) *}
+
+definition
+"sumJoin4 f1 f2 f3 f4 \<equiv>
+\<lambda> k. (case k of
+ Inl x1 \<Rightarrow> f1 x1
+|Inr k1 \<Rightarrow> (case k1 of
+ Inl ((s2,a2),(t2,b2)) \<Rightarrow> f2 s2 a2 t2 b2
+|Inr k2 \<Rightarrow> (case k2 of Inl (x3,(t3,b3)) \<Rightarrow> f3 x3 t3 b3
+|Inr (xts,(t4,b4)) \<Rightarrow> f4 xts t4 b4)))"
+
+lemma sumJoin4_simps[simp]:
+"\<And>x. sumJoin4 var app lam lt (Inl x) = var x"
+"\<And> t1 a1 t2 a2. sumJoin4 var app lam lt (Inr (Inl ((t1,a1),(t2,a2)))) = app t1 a1 t2 a2"
+"\<And> x t a. sumJoin4 var app lam lt (Inr (Inr (Inl (x,(t,a))))) = lam x t a"
+"\<And> xtas t a. sumJoin4 var app lam lt (Inr (Inr (Inr (xtas,(t,a))))) = lt xtas t a"
+unfolding sumJoin4_def by auto
+
+definition "trmrec var app lam lt \<equiv> trm_ctor_rec (sumJoin4 var app lam lt)"
+
+lemma trmrec_Var[simp]:
+"trmrec var app lam lt (Var x) = var x"
+unfolding trmrec_def Var_def trm.ctor_recs pre_trm_map(1) by simp
+
+lemma trmrec_App[simp]:
+"trmrec var app lam lt (App t1 t2) =
+ app t1 (trmrec var app lam lt t1) t2 (trmrec var app lam lt t2)"
+unfolding trmrec_def App_def trm.ctor_recs pre_trm_map(2) convol_def by simp
+
+lemma trmrec_Lam[simp]:
+"trmrec var app lam lt (Lam x t) = lam x t (trmrec var app lam lt t)"
+unfolding trmrec_def Lam_def trm.ctor_recs pre_trm_map(3) convol_def by simp
+
+lemma trmrec_Lt[simp]:
+"trmrec var app lam lt (Lt xts t) =
+ lt (map_fset (\<lambda> (x,t). (x,t,trmrec var app lam lt t)) xts) t (trmrec var app lam lt t)"
+unfolding trmrec_def Lt_def trm.ctor_recs pre_trm_map(4) convol_def by simp
+
+definition
+"sumJoinI4 f1 f2 f3 f4 \<equiv>
+\<lambda> k. (case k of
+ Inl x1 \<Rightarrow> f1 x1
+|Inr k1 \<Rightarrow> (case k1 of
+ Inl (a2,b2) \<Rightarrow> f2 a2 b2
+|Inr k2 \<Rightarrow> (case k2 of Inl (x3,b3) \<Rightarrow> f3 x3 b3
+|Inr (xts,b4) \<Rightarrow> f4 xts b4)))"
+
+lemma sumJoinI4_simps[simp]:
+"\<And>x. sumJoinI4 var app lam lt (Inl x) = var x"
+"\<And> a1 a2. sumJoinI4 var app lam lt (Inr (Inl (a1,a2))) = app a1 a2"
+"\<And> x a. sumJoinI4 var app lam lt (Inr (Inr (Inl (x,a)))) = lam x a"
+"\<And> xtas a. sumJoinI4 var app lam lt (Inr (Inr (Inr (xtas,a)))) = lt xtas a"
+unfolding sumJoinI4_def by auto
+
+(* The iterator has a simpler, hence more manageable type. *)
+definition "trmfold var app lam lt \<equiv> trm_ctor_fold (sumJoinI4 var app lam lt)"
+
+lemma trmfold_Var[simp]:
+"trmfold var app lam lt (Var x) = var x"
+unfolding trmfold_def Var_def trm.ctor_folds pre_trm_map(1) by simp
+
+lemma trmfold_App[simp]:
+"trmfold var app lam lt (App t1 t2) =
+ app (trmfold var app lam lt t1) (trmfold var app lam lt t2)"
+unfolding trmfold_def App_def trm.ctor_folds pre_trm_map(2) by simp
+
+lemma trmfold_Lam[simp]:
+"trmfold var app lam lt (Lam x t) = lam x (trmfold var app lam lt t)"
+unfolding trmfold_def Lam_def trm.ctor_folds pre_trm_map(3) by simp
+
+lemma trmfold_Lt[simp]:
+"trmfold var app lam lt (Lt xts t) =
+ lt (map_fset (\<lambda> (x,t). (x,trmfold var app lam lt t)) xts) (trmfold var app lam lt t)"
+unfolding trmfold_def Lt_def trm.ctor_folds pre_trm_map(4) by simp
+
+
+subsection{* Example: The set of all variables varsOf and free variables fvarsOf of a term: *}
+
+definition "varsOf = trmfold
+(\<lambda> x. {x})
+(\<lambda> X1 X2. X1 \<union> X2)
+(\<lambda> x X. X \<union> {x})
+(\<lambda> xXs Y. Y \<union> (\<Union> { {x} \<union> X | x X. (x,X) |\<in>| xXs}))"
+
+lemma varsOf_simps[simp]:
+"varsOf (Var x) = {x}"
+"varsOf (App t1 t2) = varsOf t1 \<union> varsOf t2"
+"varsOf (Lam x t) = varsOf t \<union> {x}"
+"varsOf (Lt xts t) =
+ varsOf t \<union> (\<Union> { {x} \<union> X | x X. (x,X) |\<in>| map_fset (\<lambda> (x,t1). (x,varsOf t1)) xts})"
+unfolding varsOf_def by simp_all
+
+definition "fvarsOf = trmfold
+(\<lambda> x. {x})
+(\<lambda> X1 X2. X1 \<union> X2)
+(\<lambda> x X. X - {x})
+(\<lambda> xtXs Y. Y - {x | x X. (x,X) |\<in>| xtXs} \<union> (\<Union> {X | x X. (x,X) |\<in>| xtXs}))"
+
+lemma fvarsOf_simps[simp]:
+"fvarsOf (Var x) = {x}"
+"fvarsOf (App t1 t2) = fvarsOf t1 \<union> fvarsOf t2"
+"fvarsOf (Lam x t) = fvarsOf t - {x}"
+"fvarsOf (Lt xts t) =
+ fvarsOf t
+ - {x | x X. (x,X) |\<in>| map_fset (\<lambda> (x,t1). (x,fvarsOf t1)) xts}
+ \<union> (\<Union> {X | x X. (x,X) |\<in>| map_fset (\<lambda> (x,t1). (x,fvarsOf t1)) xts})"
+unfolding fvarsOf_def by simp_all
+
+lemma diff_Un_incl_triv: "\<lbrakk>A \<subseteq> D; C \<subseteq> E\<rbrakk> \<Longrightarrow> A - B \<union> C \<subseteq> D \<union> E" by blast
+
+lemma in_map_fset_iff:
+"(x, X) |\<in>| map_fset (\<lambda>(x, t1). (x, f t1)) xts \<longleftrightarrow>
+ (\<exists> t1. (x,t1) |\<in>| xts \<and> X = f t1)"
+unfolding map_fset_def2_raw in_fset fset_afset unfolding fset_def2_raw by auto
+
+lemma fvarsOf_varsOf: "fvarsOf t \<subseteq> varsOf t"
+proof induct
+ case (Lt xts t)
+ thus ?case unfolding fvarsOf_simps varsOf_simps
+ proof (elim diff_Un_incl_triv)
+ show
+ "\<Union>{X | x X. (x, X) |\<in>| map_fset (\<lambda>(x, t1). (x, fvarsOf t1)) xts}
+ \<subseteq> \<Union>{{x} \<union> X |x X. (x, X) |\<in>| map_fset (\<lambda>(x, t1). (x, varsOf t1)) xts}"
+ (is "_ \<subseteq> \<Union> ?L")
+ proof(rule Sup_mono, safe)
+ fix a x X
+ assume "(x, X) |\<in>| map_fset (\<lambda>(x, t1). (x, fvarsOf t1)) xts"
+ then obtain t1 where x_t1: "(x,t1) |\<in>| xts" and X: "X = fvarsOf t1"
+ unfolding in_map_fset_iff by auto
+ let ?Y = "varsOf t1"
+ have x_Y: "(x,?Y) |\<in>| map_fset (\<lambda>(x, t1). (x, varsOf t1)) xts"
+ using x_t1 unfolding in_map_fset_iff by auto
+ show "\<exists> Y \<in> ?L. X \<subseteq> Y" unfolding X using Lt(1) x_Y x_t1 by auto
+ qed
+ qed
+qed auto
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Examples/ListF.thy Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,171 @@
+(* Title: HOL/BNF/Examples/ListF.thy
+ Author: Dmitriy Traytel, TU Muenchen
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Finite lists.
+*)
+
+header {* Finite Lists *}
+
+theory ListF
+imports "../BNF"
+begin
+
+data_raw listF: 'list = "unit + 'a \<times> 'list"
+
+definition "NilF = listF_ctor (Inl ())"
+definition "Conss a as \<equiv> listF_ctor (Inr (a, as))"
+
+lemma listF_map_NilF[simp]: "listF_map f NilF = NilF"
+unfolding listF_map_def pre_listF_map_def NilF_def listF.ctor_folds by simp
+
+lemma listF_map_Conss[simp]:
+ "listF_map f (Conss x xs) = Conss (f x) (listF_map f xs)"
+unfolding listF_map_def pre_listF_map_def Conss_def listF.ctor_folds by simp
+
+lemma listF_set_NilF[simp]: "listF_set NilF = {}"
+unfolding listF_set_def NilF_def listF.ctor_folds pre_listF_set1_def pre_listF_set2_def
+ sum_set_defs pre_listF_map_def collect_def[abs_def] by simp
+
+lemma listF_set_Conss[simp]: "listF_set (Conss x xs) = {x} \<union> listF_set xs"
+unfolding listF_set_def Conss_def listF.ctor_folds pre_listF_set1_def pre_listF_set2_def
+ sum_set_defs prod_set_defs pre_listF_map_def collect_def[abs_def] by simp
+
+lemma fold_sum_case_NilF: "listF_ctor_fold (sum_case f g) NilF = f ()"
+unfolding NilF_def listF.ctor_folds pre_listF_map_def by simp
+
+
+lemma fold_sum_case_Conss:
+ "listF_ctor_fold (sum_case f g) (Conss y ys) = g (y, listF_ctor_fold (sum_case f g) ys)"
+unfolding Conss_def listF.ctor_folds pre_listF_map_def by simp
+
+(* familiar induction principle *)
+lemma listF_induct:
+ fixes xs :: "'a listF"
+ assumes IB: "P NilF" and IH: "\<And>x xs. P xs \<Longrightarrow> P (Conss x xs)"
+ shows "P xs"
+proof (rule listF.ctor_induct)
+ fix xs :: "unit + 'a \<times> 'a listF"
+ assume raw_IH: "\<And>a. a \<in> pre_listF_set2 xs \<Longrightarrow> P a"
+ show "P (listF_ctor xs)"
+ proof (cases xs)
+ case (Inl a) with IB show ?thesis unfolding NilF_def by simp
+ next
+ case (Inr b)
+ then obtain y ys where yys: "listF_ctor xs = Conss y ys"
+ unfolding Conss_def listF.ctor_inject by (blast intro: prod.exhaust)
+ hence "ys \<in> pre_listF_set2 xs"
+ unfolding pre_listF_set2_def Conss_def listF.ctor_inject sum_set_defs prod_set_defs
+ collect_def[abs_def] by simp
+ with raw_IH have "P ys" by blast
+ with IH have "P (Conss y ys)" by blast
+ with yys show ?thesis by simp
+ qed
+qed
+
+rep_datatype NilF Conss
+by (blast intro: listF_induct) (auto simp add: NilF_def Conss_def listF.ctor_inject)
+
+definition Singll ("[[_]]") where
+ [simp]: "Singll a \<equiv> Conss a NilF"
+
+definition appendd (infixr "@@" 65) where
+ "appendd \<equiv> listF_ctor_fold (sum_case (\<lambda> _. id) (\<lambda> (a,f) bs. Conss a (f bs)))"
+
+definition "lrev \<equiv> listF_ctor_fold (sum_case (\<lambda> _. NilF) (\<lambda> (b,bs). bs @@ [[b]]))"
+
+lemma lrev_NilF[simp]: "lrev NilF = NilF"
+unfolding lrev_def by (simp add: fold_sum_case_NilF)
+
+lemma lrev_Conss[simp]: "lrev (Conss y ys) = lrev ys @@ [[y]]"
+unfolding lrev_def by (simp add: fold_sum_case_Conss)
+
+lemma NilF_appendd[simp]: "NilF @@ ys = ys"
+unfolding appendd_def by (simp add: fold_sum_case_NilF)
+
+lemma Conss_append[simp]: "Conss x xs @@ ys = Conss x (xs @@ ys)"
+unfolding appendd_def by (simp add: fold_sum_case_Conss)
+
+lemma appendd_NilF[simp]: "xs @@ NilF = xs"
+by (rule listF_induct) auto
+
+lemma appendd_assoc[simp]: "(xs @@ ys) @@ zs = xs @@ ys @@ zs"
+by (rule listF_induct) auto
+
+lemma lrev_appendd[simp]: "lrev (xs @@ ys) = lrev ys @@ lrev xs"
+by (rule listF_induct[of _ xs]) auto
+
+lemma listF_map_appendd[simp]:
+ "listF_map f (xs @@ ys) = listF_map f xs @@ listF_map f ys"
+by (rule listF_induct[of _ xs]) auto
+
+lemma lrev_listF_map[simp]: "lrev (listF_map f xs) = listF_map f (lrev xs)"
+by (rule listF_induct[of _ xs]) auto
+
+lemma lrev_lrev[simp]: "lrev (lrev as) = as"
+by (rule listF_induct) auto
+
+fun lengthh where
+ "lengthh NilF = 0"
+| "lengthh (Conss x xs) = Suc (lengthh xs)"
+
+fun nthh where
+ "nthh (Conss x xs) 0 = x"
+| "nthh (Conss x xs) (Suc n) = nthh xs n"
+| "nthh xs i = undefined"
+
+lemma lengthh_listF_map[simp]: "lengthh (listF_map f xs) = lengthh xs"
+by (rule listF_induct[of _ xs]) auto
+
+lemma nthh_listF_map[simp]:
+ "i < lengthh xs \<Longrightarrow> nthh (listF_map f xs) i = f (nthh xs i)"
+by (induct rule: nthh.induct) auto
+
+lemma nthh_listF_set[simp]: "i < lengthh xs \<Longrightarrow> nthh xs i \<in> listF_set xs"
+by (induct rule: nthh.induct) auto
+
+lemma NilF_iff[iff]: "(lengthh xs = 0) = (xs = NilF)"
+by (induct xs) auto
+
+lemma Conss_iff[iff]:
+ "(lengthh xs = Suc n) = (\<exists>y ys. xs = Conss y ys \<and> lengthh ys = n)"
+by (induct xs) auto
+
+lemma Conss_iff'[iff]:
+ "(Suc n = lengthh xs) = (\<exists>y ys. xs = Conss y ys \<and> lengthh ys = n)"
+by (induct xs) (simp, simp, blast)
+
+lemma listF_induct2: "\<lbrakk>lengthh xs = lengthh ys; P NilF NilF;
+ \<And>x xs y ys. P xs ys \<Longrightarrow> P (Conss x xs) (Conss y ys)\<rbrakk> \<Longrightarrow> P xs ys"
+by (induct xs arbitrary: ys rule: listF_induct) auto
+
+fun zipp where
+ "zipp NilF NilF = NilF"
+| "zipp (Conss x xs) (Conss y ys) = Conss (x, y) (zipp xs ys)"
+| "zipp xs ys = undefined"
+
+lemma listF_map_fst_zip[simp]:
+ "lengthh xs = lengthh ys \<Longrightarrow> listF_map fst (zipp xs ys) = xs"
+by (erule listF_induct2) auto
+
+lemma listF_map_snd_zip[simp]:
+ "lengthh xs = lengthh ys \<Longrightarrow> listF_map snd (zipp xs ys) = ys"
+by (erule listF_induct2) auto
+
+lemma lengthh_zip[simp]:
+ "lengthh xs = lengthh ys \<Longrightarrow> lengthh (zipp xs ys) = lengthh xs"
+by (erule listF_induct2) auto
+
+lemma nthh_zip[simp]:
+ assumes *: "lengthh xs = lengthh ys"
+ shows "i < lengthh xs \<Longrightarrow> nthh (zipp xs ys) i = (nthh xs i, nthh ys i)"
+proof (induct arbitrary: i rule: listF_induct2[OF *])
+ case (2 x xs y ys) thus ?case by (induct i) auto
+qed simp
+
+lemma list_set_nthh[simp]:
+ "(x \<in> listF_set xs) \<Longrightarrow> (\<exists>i < lengthh xs. nthh xs i = x)"
+by (induct xs) (auto, induct rule: nthh.induct, auto)
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Examples/Misc_Codata.thy Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,110 @@
+(* Title: HOL/BNF/Examples/Misc_Data.thy
+ Author: Dmitriy Traytel, TU Muenchen
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Miscellaneous codatatype declarations.
+*)
+
+header {* Miscellaneous Codatatype Declarations *}
+
+theory Misc_Codata
+imports "../BNF"
+begin
+
+codata simple = X1 | X2 | X3 | X4
+
+codata simple' = X1' unit | X2' unit | X3' unit | X4' unit
+
+codata 'a stream = Stream 'a "'a stream"
+
+codata 'a mylist = MyNil | MyCons 'a "'a mylist"
+
+codata ('b, 'c, 'd, 'e) some_passive =
+ SP1 "('b, 'c, 'd, 'e) some_passive" | SP2 'b | SP3 'c | SP4 'd | SP5 'e
+
+codata lambda =
+ Var string |
+ App lambda lambda |
+ Abs string lambda |
+ Let "(string \<times> lambda) fset" lambda
+
+codata 'a par_lambda =
+ PVar 'a |
+ PApp "'a par_lambda" "'a par_lambda" |
+ PAbs 'a "'a par_lambda" |
+ PLet "('a \<times> 'a par_lambda) fset" "'a par_lambda"
+
+(*
+ ('a, 'b1, 'b2) F1 = 'a * 'b1 + 'a * 'b2
+ ('a, 'b1, 'b2) F2 = unit + 'b1 * 'b2
+*)
+
+codata 'a J1 = J11 'a "'a J1" | J12 'a "'a J2"
+ and 'a J2 = J21 | J22 "'a J1" "'a J2"
+
+codata 'a tree = TEmpty | TNode 'a "'a forest"
+ and 'a forest = FNil | FCons "'a tree" "'a forest"
+
+codata 'a tree' = TEmpty' | TNode' "'a branch" "'a branch"
+ and 'a branch = Branch 'a "'a tree'"
+
+codata ('a, 'b) exp = Term "('a, 'b) trm" | Sum "('a, 'b) trm" "('a, 'b) exp"
+ and ('a, 'b) trm = Factor "('a, 'b) factor" | Prod "('a, 'b) factor" "('a, 'b) trm"
+ and ('a, 'b) factor = C 'a | V 'b | Paren "('a, 'b) exp"
+
+codata ('a, 'b, 'c) some_killing =
+ SK "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b, 'c) some_killing + ('a, 'b, 'c) in_here"
+ and ('a, 'b, 'c) in_here =
+ IH1 'b 'a | IH2 'c
+
+codata_raw some_killing': 'a = "'b \<Rightarrow> 'd \<Rightarrow> ('a + 'c)"
+and in_here': 'c = "'d + 'e"
+
+codata_raw some_killing'': 'a = "'b \<Rightarrow> 'c"
+and in_here'': 'c = "'d \<times> 'b + 'e"
+
+codata ('b, 'c) less_killing = LK "'b \<Rightarrow> 'c"
+
+codata 'b cps = CPS1 'b | CPS2 "'b \<Rightarrow> 'b cps"
+
+codata ('b1, 'b2, 'b3, 'b4, 'b5, 'b6, 'b7, 'b8, 'b9) fun_rhs =
+ FR "'b1 \<Rightarrow> 'b2 \<Rightarrow> 'b3 \<Rightarrow> 'b4 \<Rightarrow> 'b5 \<Rightarrow> 'b6 \<Rightarrow> 'b7 \<Rightarrow> 'b8 \<Rightarrow> 'b9 \<Rightarrow>
+ ('b1, 'b2, 'b3, 'b4, 'b5, 'b6, 'b7, 'b8, 'b9) fun_rhs"
+
+codata ('b1, 'b2, 'b3, 'b4, 'b5, 'b6, 'b7, 'b8, 'b9, 'b10, 'b11, 'b12, 'b13, 'b14, 'b15, 'b16, 'b17,
+ 'b18, 'b19, 'b20) fun_rhs' =
+ FR' "'b1 \<Rightarrow> 'b2 \<Rightarrow> 'b3 \<Rightarrow> 'b4 \<Rightarrow> 'b5 \<Rightarrow> 'b6 \<Rightarrow> 'b7 \<Rightarrow> 'b8 \<Rightarrow> 'b9 \<Rightarrow> 'b10 \<Rightarrow> 'b11 \<Rightarrow> 'b12 \<Rightarrow> 'b13 \<Rightarrow> 'b14 \<Rightarrow>
+ 'b15 \<Rightarrow> 'b16 \<Rightarrow> 'b17 \<Rightarrow> 'b18 \<Rightarrow> 'b19 \<Rightarrow> 'b20 \<Rightarrow>
+ ('b1, 'b2, 'b3, 'b4, 'b5, 'b6, 'b7, 'b8, 'b9, 'b10, 'b11, 'b12, 'b13, 'b14, 'b15, 'b16, 'b17,
+ 'b18, 'b19, 'b20) fun_rhs'"
+
+codata ('a, 'b, 'c) wit3_F1 = W1 'a "('a, 'b, 'c) wit3_F1" "('a, 'b, 'c) wit3_F2"
+ and ('a, 'b, 'c) wit3_F2 = W2 'b "('a, 'b, 'c) wit3_F2"
+ and ('a, 'b, 'c) wit3_F3 = W31 'a 'b "('a, 'b, 'c) wit3_F1" | W32 'c 'a 'b "('a, 'b, 'c) wit3_F1"
+
+codata ('c, 'e, 'g) coind_wit1 =
+ CW1 'c "('c, 'e, 'g) coind_wit1" "('c, 'e, 'g) ind_wit" "('c, 'e, 'g) coind_wit2"
+ and ('c, 'e, 'g) coind_wit2 =
+ CW21 "('c, 'e, 'g) coind_wit2" 'e | CW22 'c 'g
+ and ('c, 'e, 'g) ind_wit =
+ IW1 | IW2 'c
+
+codata ('b, 'a) bar = BAR "'a \<Rightarrow> 'b"
+codata ('a, 'b, 'c, 'd) foo = FOO "'d + 'b \<Rightarrow> 'c + 'a"
+
+codata 'a dead_foo = A
+codata ('a, 'b) use_dead_foo = Y "'a" "'b dead_foo"
+
+(* SLOW, MEMORY-HUNGRY
+codata ('a, 'c) D1 = A1 "('a, 'c) D2" | B1 "'a list"
+ and ('a, 'c) D2 = A2 "('a, 'c) D3" | B2 "'c list"
+ and ('a, 'c) D3 = A3 "('a, 'c) D3" | B3 "('a, 'c) D4" | C3 "('a, 'c) D4" "('a, 'c) D5"
+ and ('a, 'c) D4 = A4 "('a, 'c) D5" | B4 "'a list list list"
+ and ('a, 'c) D5 = A5 "('a, 'c) D6"
+ and ('a, 'c) D6 = A6 "('a, 'c) D7"
+ and ('a, 'c) D7 = A7 "('a, 'c) D8"
+ and ('a, 'c) D8 = A8 "('a, 'c) D1 list"
+*)
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Examples/Misc_Data.thy Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,154 @@
+(* Title: HOL/BNF/Examples/Misc_Data.thy
+ Author: Dmitriy Traytel, TU Muenchen
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Miscellaneous datatype declarations.
+*)
+
+header {* Miscellaneous Datatype Declarations *}
+
+theory Misc_Data
+imports "../BNF"
+begin
+
+data simple = X1 | X2 | X3 | X4
+
+data simple' = X1' unit | X2' unit | X3' unit | X4' unit
+
+data 'a mylist = MyNil | MyCons 'a "'a mylist"
+
+data ('b, 'c, 'd, 'e) some_passive =
+ SP1 "('b, 'c, 'd, 'e) some_passive" | SP2 'b | SP3 'c | SP4 'd | SP5 'e
+
+data lambda =
+ Var string |
+ App lambda lambda |
+ Abs string lambda |
+ Let "(string \<times> lambda) fset" lambda
+
+data 'a par_lambda =
+ PVar 'a |
+ PApp "'a par_lambda" "'a par_lambda" |
+ PAbs 'a "'a par_lambda" |
+ PLet "('a \<times> 'a par_lambda) fset" "'a par_lambda"
+
+(*
+ ('a, 'b1, 'b2) F1 = 'a * 'b1 + 'a * 'b2
+ ('a, 'b1, 'b2) F2 = unit + 'b1 * 'b2
+*)
+
+data 'a I1 = I11 'a "'a I1" | I12 'a "'a I2"
+ and 'a I2 = I21 | I22 "'a I1" "'a I2"
+
+data 'a tree = TEmpty | TNode 'a "'a forest"
+ and 'a forest = FNil | FCons "'a tree" "'a forest"
+
+data 'a tree' = TEmpty' | TNode' "'a branch" "'a branch"
+ and 'a branch = Branch 'a "'a tree'"
+
+data ('a, 'b) exp = Term "('a, 'b) trm" | Sum "('a, 'b) trm" "('a, 'b) exp"
+ and ('a, 'b) trm = Factor "('a, 'b) factor" | Prod "('a, 'b) factor" "('a, 'b) trm"
+ and ('a, 'b) factor = C 'a | V 'b | Paren "('a, 'b) exp"
+
+data ('a, 'b, 'c) some_killing =
+ SK "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b, 'c) some_killing + ('a, 'b, 'c) in_here"
+ and ('a, 'b, 'c) in_here =
+ IH1 'b 'a | IH2 'c
+
+data 'b nofail1 = NF11 "'b nofail1" 'b | NF12 'b
+data 'b nofail2 = NF2 "('b nofail2 \<times> 'b \<times> 'b nofail2 \<times> 'b) list"
+data 'b nofail3 = NF3 'b "('b nofail3 \<times> 'b \<times> 'b nofail3 \<times> 'b) fset"
+data 'b nofail4 = NF4 "('b nofail4 \<times> ('b nofail4 \<times> 'b \<times> 'b nofail4 \<times> 'b) fset) list"
+
+(*
+data 'b fail = F "'b fail" 'b "'b fail" "'b list"
+data 'b fail = F "'b fail" 'b "'b fail" 'b
+data 'b fail = F1 "'b fail" 'b | F2 "'b fail"
+data 'b fail = F "'b fail" 'b
+*)
+
+data l1 = L1 "l2 list"
+ and l2 = L21 "l1 fset" | L22 l2
+
+data kk1 = KK1 kk2
+ and kk2 = KK2 kk3
+ and kk3 = KK3 "kk1 list"
+
+data t1 = T11 t3 | T12 t2
+ and t2 = T2 t1
+ and t3 = T3
+
+data t1' = T11' t2' | T12' t3'
+ and t2' = T2' t1'
+ and t3' = T3'
+
+(*
+data fail1 = F1 fail2
+ and fail2 = F2 fail3
+ and fail3 = F3 fail1
+
+data fail1 = F1 "fail2 list" fail2
+ and fail2 = F2 "fail2 fset" fail3
+ and fail3 = F3 fail1
+
+data fail1 = F1 "fail2 list" fail2
+ and fail2 = F2 "fail1 fset" fail1
+*)
+
+(* SLOW
+data ('a, 'c) D1 = A1 "('a, 'c) D2" | B1 "'a list"
+ and ('a, 'c) D2 = A2 "('a, 'c) D3" | B2 "'c list"
+ and ('a, 'c) D3 = A3 "('a, 'c) D3" | B3 "('a, 'c) D4" | C3 "('a, 'c) D4" "('a, 'c) D5"
+ and ('a, 'c) D4 = A4 "('a, 'c) D5" | B4 "'a list list list"
+ and ('a, 'c) D5 = A5 "('a, 'c) D6"
+ and ('a, 'c) D6 = A6 "('a, 'c) D7"
+ and ('a, 'c) D7 = A7 "('a, 'c) D8"
+ and ('a, 'c) D8 = A8 "('a, 'c) D1 list"
+
+(*time comparison*)
+datatype ('a, 'c) D1' = A1' "('a, 'c) D2'" | B1' "'a list"
+ and ('a, 'c) D2' = A2' "('a, 'c) D3'" | B2' "'c list"
+ and ('a, 'c) D3' = A3' "('a, 'c) D3'" | B3' "('a, 'c) D4'" | C3' "('a, 'c) D4'" "('a, 'c) D5'"
+ and ('a, 'c) D4' = A4' "('a, 'c) D5'" | B4' "'a list list list"
+ and ('a, 'c) D5' = A5' "('a, 'c) D6'"
+ and ('a, 'c) D6' = A6' "('a, 'c) D7'"
+ and ('a, 'c) D7' = A7' "('a, 'c) D8'"
+ and ('a, 'c) D8' = A8' "('a, 'c) D1' list"
+*)
+
+(* fail:
+data tt1 = TT11 tt2 tt3 | TT12 tt2 tt4
+ and tt2 = TT2
+ and tt3 = TT3 tt4
+ and tt4 = TT4 tt1
+*)
+
+data k1 = K11 k2 k3 | K12 k2 k4
+ and k2 = K2
+ and k3 = K3 k4
+ and k4 = K4
+
+data tt1 = TT11 tt3 tt2 | TT12 tt2 tt4
+ and tt2 = TT2
+ and tt3 = TT3 tt1
+ and tt4 = TT4
+
+(* SLOW
+data s1 = S11 s2 s3 s4 | S12 s3 | S13 s2 s6 | S14 s4 s2 | S15 s2 s2
+ and s2 = S21 s7 s5 | S22 s5 s4 s6
+ and s3 = S31 s1 s7 s2 | S32 s3 s3 | S33 s4 s5
+ and s4 = S4 s5
+ and s5 = S5
+ and s6 = S61 s6 | S62 s1 s2 | S63 s6
+ and s7 = S71 s8 | S72 s5
+ and s8 = S8 nat
+*)
+
+data ('a, 'b) bar = Bar "'b \<Rightarrow> 'a"
+data ('a, 'b, 'c, 'd) foo = Foo "'d + 'b \<Rightarrow> 'c + 'a"
+
+data 'a dead_foo = A
+data ('a, 'b) use_dead_foo = Y "'a" "'b dead_foo"
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Examples/Process.thy Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,367 @@
+(* Title: HOL/BNF/Examples/Process.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Processes.
+*)
+
+header {* Processes *}
+
+theory Process
+imports "../BNF"
+begin
+
+hide_fact (open) Quotient_Product.prod_rel_def
+
+codata 'a process =
+ isAction: Action (prefOf: 'a) (contOf: "'a process") |
+ isChoice: Choice (ch1Of: "'a process") (ch2Of: "'a process")
+
+(* Read: prefix of, continuation of, choice 1 of, choice 2 of *)
+
+section {* Customization *}
+
+subsection {* Basic properties *}
+
+declare
+ pre_process_rel_def[simp]
+ sum_rel_def[simp]
+ prod_rel_def[simp]
+
+(* Constructors versus discriminators *)
+theorem isAction_isChoice:
+"isAction p \<or> isChoice p"
+by (rule process.disc_exhaust) auto
+
+theorem not_isAction_isChoice: "\<not> (isAction p \<and> isChoice p)"
+by (cases rule: process.exhaust[of p]) auto
+
+
+subsection{* Coinduction *}
+
+theorem process_coind[elim, consumes 1, case_names iss Action Choice, induct pred: "HOL.eq"]:
+assumes phi: "\<phi> p p'" and
+iss: "\<And>p p'. \<phi> p p' \<Longrightarrow> (isAction p \<longleftrightarrow> isAction p') \<and> (isChoice p \<longleftrightarrow> isChoice p')" and
+Act: "\<And> a a' p p'. \<phi> (Action a p) (Action a' p') \<Longrightarrow> a = a' \<and> \<phi> p p'" and
+Ch: "\<And> p q p' q'. \<phi> (Choice p q) (Choice p' q') \<Longrightarrow> \<phi> p p' \<and> \<phi> q q'"
+shows "p = p'"
+proof(intro mp[OF process.rel_coinduct, of \<phi>, OF _ phi], clarify)
+ fix p p' assume \<phi>: "\<phi> p p'"
+ show "pre_process_rel (op =) \<phi> (process_dtor p) (process_dtor p')"
+ proof(cases rule: process.exhaust[of p])
+ case (Action a q) note p = Action
+ hence "isAction p'" using iss[OF \<phi>] by (cases rule: process.exhaust[of p'], auto)
+ then obtain a' q' where p': "p' = Action a' q'" by (cases rule: process.exhaust[of p'], auto)
+ have 0: "a = a' \<and> \<phi> q q'" using Act[OF \<phi>[unfolded p p']] .
+ have dtor: "process_dtor p = Inl (a,q)" "process_dtor p' = Inl (a',q')"
+ unfolding p p' Action_def process.dtor_ctor by simp_all
+ show ?thesis using 0 unfolding dtor by simp
+ next
+ case (Choice p1 p2) note p = Choice
+ hence "isChoice p'" using iss[OF \<phi>] by (cases rule: process.exhaust[of p'], auto)
+ then obtain p1' p2' where p': "p' = Choice p1' p2'"
+ by (cases rule: process.exhaust[of p'], auto)
+ have 0: "\<phi> p1 p1' \<and> \<phi> p2 p2'" using Ch[OF \<phi>[unfolded p p']] .
+ have dtor: "process_dtor p = Inr (p1,p2)" "process_dtor p' = Inr (p1',p2')"
+ unfolding p p' Choice_def process.dtor_ctor by simp_all
+ show ?thesis using 0 unfolding dtor by simp
+ qed
+qed
+
+(* Stronger coinduction, up to equality: *)
+theorem process_strong_coind[elim, consumes 1, case_names iss Action Choice]:
+assumes phi: "\<phi> p p'" and
+iss: "\<And>p p'. \<phi> p p' \<Longrightarrow> (isAction p \<longleftrightarrow> isAction p') \<and> (isChoice p \<longleftrightarrow> isChoice p')" and
+Act: "\<And> a a' p p'. \<phi> (Action a p) (Action a' p') \<Longrightarrow> a = a' \<and> (\<phi> p p' \<or> p = p')" and
+Ch: "\<And> p q p' q'. \<phi> (Choice p q) (Choice p' q') \<Longrightarrow> (\<phi> p p' \<or> p = p') \<and> (\<phi> q q' \<or> q = q')"
+shows "p = p'"
+proof(intro mp[OF process.rel_strong_coinduct, of \<phi>, OF _ phi], clarify)
+ fix p p' assume \<phi>: "\<phi> p p'"
+ show "pre_process_rel (op =) (\<lambda>a b. \<phi> a b \<or> a = b) (process_dtor p) (process_dtor p')"
+ proof(cases rule: process.exhaust[of p])
+ case (Action a q) note p = Action
+ hence "isAction p'" using iss[OF \<phi>] by (cases rule: process.exhaust[of p'], auto)
+ then obtain a' q' where p': "p' = Action a' q'" by (cases rule: process.exhaust[of p'], auto)
+ have 0: "a = a' \<and> (\<phi> q q' \<or> q = q')" using Act[OF \<phi>[unfolded p p']] .
+ have dtor: "process_dtor p = Inl (a,q)" "process_dtor p' = Inl (a',q')"
+ unfolding p p' Action_def process.dtor_ctor by simp_all
+ show ?thesis using 0 unfolding dtor by simp
+ next
+ case (Choice p1 p2) note p = Choice
+ hence "isChoice p'" using iss[OF \<phi>] by (cases rule: process.exhaust[of p'], auto)
+ then obtain p1' p2' where p': "p' = Choice p1' p2'"
+ by (cases rule: process.exhaust[of p'], auto)
+ have 0: "(\<phi> p1 p1' \<or> p1 = p1') \<and> (\<phi> p2 p2' \<or> p2 = p2')" using Ch[OF \<phi>[unfolded p p']] .
+ have dtor: "process_dtor p = Inr (p1,p2)" "process_dtor p' = Inr (p1',p2')"
+ unfolding p p' Choice_def process.dtor_ctor by simp_all
+ show ?thesis using 0 unfolding dtor by simp
+ qed
+qed
+
+
+subsection {* Coiteration (unfold) *}
+
+
+section{* Coinductive definition of the notion of trace *}
+
+(* Say we have a type of streams: *)
+
+typedecl 'a stream
+
+consts Ccons :: "'a \<Rightarrow> 'a stream \<Rightarrow> 'a stream"
+
+(* Use the existing coinductive package (distinct from our
+new codatatype package, but highly compatible with it): *)
+
+coinductive trace where
+"trace p as \<Longrightarrow> trace (Action a p) (Ccons a as)"
+|
+"trace p as \<or> trace q as \<Longrightarrow> trace (Choice p q) as"
+
+
+section{* Examples of corecursive definitions: *}
+
+subsection{* Single-guard fixpoint definition *}
+
+definition
+"BX \<equiv>
+ process_unfold
+ (\<lambda> P. True)
+ (\<lambda> P. ''a'')
+ (\<lambda> P. P)
+ undefined
+ undefined
+ ()"
+
+lemma BX: "BX = Action ''a'' BX"
+unfolding BX_def
+using process.unfolds(1)[of "\<lambda> P. True" "()" "\<lambda> P. ''a''" "\<lambda> P. P"] by simp
+
+
+subsection{* Multi-guard fixpoint definitions, simulated with auxiliary arguments *}
+
+datatype x_y_ax = x | y | ax
+
+definition "isA \<equiv> \<lambda> K. case K of x \<Rightarrow> False |y \<Rightarrow> True |ax \<Rightarrow> True"
+definition "pr \<equiv> \<lambda> K. case K of x \<Rightarrow> undefined |y \<Rightarrow> ''b'' |ax \<Rightarrow> ''a''"
+definition "co \<equiv> \<lambda> K. case K of x \<Rightarrow> undefined |y \<Rightarrow> x |ax \<Rightarrow> x"
+lemmas Action_defs = isA_def pr_def co_def
+
+definition "c1 \<equiv> \<lambda> K. case K of x \<Rightarrow> ax |y \<Rightarrow> undefined |ax \<Rightarrow> undefined"
+definition "c2 \<equiv> \<lambda> K. case K of x \<Rightarrow> y |y \<Rightarrow> undefined |ax \<Rightarrow> undefined"
+lemmas Choice_defs = c1_def c2_def
+
+definition "F \<equiv> process_unfold isA pr co c1 c2"
+definition "X = F x" definition "Y = F y" definition "AX = F ax"
+
+lemma X_Y_AX: "X = Choice AX Y" "Y = Action ''b'' X" "AX = Action ''a'' X"
+unfolding X_def Y_def AX_def F_def
+using process.unfolds(2)[of isA x "pr" co c1 c2]
+ process.unfolds(1)[of isA y "pr" co c1 c2]
+ process.unfolds(1)[of isA ax "pr" co c1 c2]
+unfolding Action_defs Choice_defs by simp_all
+
+(* end product: *)
+lemma X_AX:
+"X = Choice AX (Action ''b'' X)"
+"AX = Action ''a'' X"
+using X_Y_AX by simp_all
+
+
+
+section{* Case study: Multi-guard fixpoint definitions, without auxiliary arguments *}
+
+hide_const x y ax X Y AX
+
+(* Process terms *)
+datatype ('a,'pvar) process_term =
+ VAR 'pvar |
+ PROC "'a process" |
+ ACT 'a "('a,'pvar) process_term" | CH "('a,'pvar) process_term" "('a,'pvar) process_term"
+
+(* below, sys represents a system of equations *)
+fun isACT where
+"isACT sys (VAR X) =
+ (case sys X of ACT a T \<Rightarrow> True |PROC p \<Rightarrow> isAction p |_ \<Rightarrow> False)"
+|
+"isACT sys (PROC p) = isAction p"
+|
+"isACT sys (ACT a T) = True"
+|
+"isACT sys (CH T1 T2) = False"
+
+fun PREF where
+"PREF sys (VAR X) =
+ (case sys X of ACT a T \<Rightarrow> a | PROC p \<Rightarrow> prefOf p)"
+|
+"PREF sys (PROC p) = prefOf p"
+|
+"PREF sys (ACT a T) = a"
+
+fun CONT where
+"CONT sys (VAR X) =
+ (case sys X of ACT a T \<Rightarrow> T | PROC p \<Rightarrow> PROC (contOf p))"
+|
+"CONT sys (PROC p) = PROC (contOf p)"
+|
+"CONT sys (ACT a T) = T"
+
+fun CH1 where
+"CH1 sys (VAR X) =
+ (case sys X of CH T1 T2 \<Rightarrow> T1 |PROC p \<Rightarrow> PROC (ch1Of p))"
+|
+"CH1 sys (PROC p) = PROC (ch1Of p)"
+|
+"CH1 sys (CH T1 T2) = T1"
+
+fun CH2 where
+"CH2 sys (VAR X) =
+ (case sys X of CH T1 T2 \<Rightarrow> T2 |PROC p \<Rightarrow> PROC (ch2Of p))"
+|
+"CH2 sys (PROC p) = PROC (ch2Of p)"
+|
+"CH2 sys (CH T1 T2) = T2"
+
+definition "guarded sys \<equiv> \<forall> X Y. sys X \<noteq> VAR Y"
+
+definition
+"solution sys \<equiv>
+ process_unfold
+ (isACT sys)
+ (PREF sys)
+ (CONT sys)
+ (CH1 sys)
+ (CH2 sys)"
+
+lemma solution_Action:
+assumes "isACT sys T"
+shows "solution sys T = Action (PREF sys T) (solution sys (CONT sys T))"
+unfolding solution_def
+using process.unfolds(1)[of "isACT sys" T "PREF sys" "CONT sys" "CH1 sys" "CH2 sys"]
+ assms by simp
+
+lemma solution_Choice:
+assumes "\<not> isACT sys T"
+shows "solution sys T = Choice (solution sys (CH1 sys T)) (solution sys (CH2 sys T))"
+unfolding solution_def
+using process.unfolds(2)[of "isACT sys" T "PREF sys" "CONT sys" "CH1 sys" "CH2 sys"]
+ assms by simp
+
+lemma isACT_VAR:
+assumes g: "guarded sys"
+shows "isACT sys (VAR X) \<longleftrightarrow> isACT sys (sys X)"
+using g unfolding guarded_def by (cases "sys X") auto
+
+lemma solution_VAR:
+assumes g: "guarded sys"
+shows "solution sys (VAR X) = solution sys (sys X)"
+proof(cases "isACT sys (VAR X)")
+ case True
+ hence T: "isACT sys (sys X)" unfolding isACT_VAR[OF g] .
+ show ?thesis
+ unfolding solution_Action[OF T] using solution_Action[of sys "VAR X"] True g
+ unfolding guarded_def by (cases "sys X", auto)
+next
+ case False note FFalse = False
+ hence TT: "\<not> isACT sys (sys X)" unfolding isACT_VAR[OF g] .
+ show ?thesis
+ unfolding solution_Choice[OF TT] using solution_Choice[of sys "VAR X"] FFalse g
+ unfolding guarded_def by (cases "sys X", auto)
+qed
+
+lemma solution_PROC[simp]:
+"solution sys (PROC p) = p"
+proof-
+ {fix q assume "q = solution sys (PROC p)"
+ hence "p = q"
+ proof(induct rule: process_coind)
+ case (iss p p')
+ from isAction_isChoice[of p] show ?case
+ proof
+ assume p: "isAction p"
+ hence 0: "isACT sys (PROC p)" by simp
+ thus ?thesis using iss not_isAction_isChoice
+ unfolding solution_Action[OF 0] by auto
+ next
+ assume "isChoice p"
+ hence 0: "\<not> isACT sys (PROC p)"
+ using not_isAction_isChoice by auto
+ thus ?thesis using iss isAction_isChoice
+ unfolding solution_Choice[OF 0] by auto
+ qed
+ next
+ case (Action a a' p p')
+ hence 0: "isACT sys (PROC (Action a p))" by simp
+ show ?case using Action unfolding solution_Action[OF 0] by simp
+ next
+ case (Choice p q p' q')
+ hence 0: "\<not> isACT sys (PROC (Choice p q))" using not_isAction_isChoice by auto
+ show ?case using Choice unfolding solution_Choice[OF 0] by simp
+ qed
+ }
+ thus ?thesis by metis
+qed
+
+lemma solution_ACT[simp]:
+"solution sys (ACT a T) = Action a (solution sys T)"
+by (metis CONT.simps(3) PREF.simps(3) isACT.simps(3) solution_Action)
+
+lemma solution_CH[simp]:
+"solution sys (CH T1 T2) = Choice (solution sys T1) (solution sys T2)"
+by (metis CH1.simps(3) CH2.simps(3) isACT.simps(4) solution_Choice)
+
+
+(* Example: *)
+
+fun sys where
+"sys 0 = CH (VAR (Suc 0)) (ACT ''b'' (VAR 0))"
+|
+"sys (Suc 0) = ACT ''a'' (VAR 0)"
+| (* dummy guarded term for variables outside the system: *)
+"sys X = ACT ''a'' (VAR 0)"
+
+lemma guarded_sys:
+"guarded sys"
+unfolding guarded_def proof (intro allI)
+ fix X Y show "sys X \<noteq> VAR Y" by (cases X, simp, case_tac nat, auto)
+qed
+
+(* the actual processes: *)
+definition "x \<equiv> solution sys (VAR 0)"
+definition "ax \<equiv> solution sys (VAR (Suc 0))"
+
+(* end product: *)
+lemma x_ax:
+"x = Choice ax (Action ''b'' x)"
+"ax = Action ''a'' x"
+unfolding x_def ax_def by (subst solution_VAR[OF guarded_sys], simp)+
+
+
+(* Thanks to the inclusion of processes as process terms, one can
+also consider parametrized systems of equations---here, x is a (semantic)
+process parameter: *)
+
+fun sys' where
+"sys' 0 = CH (PROC x) (ACT ''b'' (VAR 0))"
+|
+"sys' (Suc 0) = CH (ACT ''a'' (VAR 0)) (PROC x)"
+| (* dummy guarded term : *)
+"sys' X = ACT ''a'' (VAR 0)"
+
+lemma guarded_sys':
+"guarded sys'"
+unfolding guarded_def proof (intro allI)
+ fix X Y show "sys' X \<noteq> VAR Y" by (cases X, simp, case_tac nat, auto)
+qed
+
+(* the actual processes: *)
+definition "y \<equiv> solution sys' (VAR 0)"
+definition "ay \<equiv> solution sys' (VAR (Suc 0))"
+
+(* end product: *)
+lemma y_ay:
+"y = Choice x (Action ''b'' y)"
+"ay = Choice (Action ''a'' y) x"
+unfolding y_def ay_def by (subst solution_VAR[OF guarded_sys'], simp)+
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Examples/Stream.thy Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,157 @@
+(* Title: HOL/BNF/Examples/Stream.thy
+ Author: Dmitriy Traytel, TU Muenchen
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Infinite streams.
+*)
+
+header {* Infinite Streams *}
+
+theory Stream
+imports TreeFI
+begin
+
+hide_const (open) Quotient_Product.prod_rel
+hide_fact (open) Quotient_Product.prod_rel_def
+
+codata_raw stream: 's = "'a \<times> 's"
+
+(* selectors for streams *)
+definition "hdd as \<equiv> fst (stream_dtor as)"
+definition "tll as \<equiv> snd (stream_dtor as)"
+
+lemma unfold_pair_fun_hdd[simp]: "hdd (stream_dtor_unfold (f \<odot> g) t) = f t"
+unfolding hdd_def pair_fun_def stream.dtor_unfolds by simp
+
+lemma unfold_pair_fun_tll[simp]: "tll (stream_dtor_unfold (f \<odot> g) t) =
+ stream_dtor_unfold (f \<odot> g) (g t)"
+unfolding tll_def pair_fun_def stream.dtor_unfolds by simp
+
+(* infinite trees: *)
+coinductive infiniteTr where
+"\<lbrakk>tr' \<in> listF_set (sub tr); infiniteTr tr'\<rbrakk> \<Longrightarrow> infiniteTr tr"
+
+lemma infiniteTr_strong_coind[consumes 1, case_names sub]:
+assumes *: "phi tr" and
+**: "\<And> tr. phi tr \<Longrightarrow> \<exists> tr' \<in> listF_set (sub tr). phi tr' \<or> infiniteTr tr'"
+shows "infiniteTr tr"
+using assms by (elim infiniteTr.coinduct) blast
+
+lemma infiniteTr_coind[consumes 1, case_names sub, induct pred: infiniteTr]:
+assumes *: "phi tr" and
+**: "\<And> tr. phi tr \<Longrightarrow> \<exists> tr' \<in> listF_set (sub tr). phi tr'"
+shows "infiniteTr tr"
+using assms by (elim infiniteTr.coinduct) blast
+
+lemma infiniteTr_sub[simp]:
+"infiniteTr tr \<Longrightarrow> (\<exists> tr' \<in> listF_set (sub tr). infiniteTr tr')"
+by (erule infiniteTr.cases) blast
+
+definition "konigPath \<equiv> stream_dtor_unfold
+ (lab \<odot> (\<lambda>tr. SOME tr'. tr' \<in> listF_set (sub tr) \<and> infiniteTr tr'))"
+
+lemma hdd_simps1[simp]: "hdd (konigPath t) = lab t"
+unfolding konigPath_def by simp
+
+lemma tll_simps2[simp]: "tll (konigPath t) =
+ konigPath (SOME tr. tr \<in> listF_set (sub t) \<and> infiniteTr tr)"
+unfolding konigPath_def by simp
+
+(* proper paths in trees: *)
+coinductive properPath where
+"\<lbrakk>hdd as = lab tr; tr' \<in> listF_set (sub tr); properPath (tll as) tr'\<rbrakk> \<Longrightarrow>
+ properPath as tr"
+
+lemma properPath_strong_coind[consumes 1, case_names hdd_lab sub]:
+assumes *: "phi as tr" and
+**: "\<And> as tr. phi as tr \<Longrightarrow> hdd as = lab tr" and
+***: "\<And> as tr.
+ phi as tr \<Longrightarrow>
+ \<exists> tr' \<in> listF_set (sub tr). phi (tll as) tr' \<or> properPath (tll as) tr'"
+shows "properPath as tr"
+using assms by (elim properPath.coinduct) blast
+
+lemma properPath_coind[consumes 1, case_names hdd_lab sub, induct pred: properPath]:
+assumes *: "phi as tr" and
+**: "\<And> as tr. phi as tr \<Longrightarrow> hdd as = lab tr" and
+***: "\<And> as tr.
+ phi as tr \<Longrightarrow>
+ \<exists> tr' \<in> listF_set (sub tr). phi (tll as) tr'"
+shows "properPath as tr"
+using properPath_strong_coind[of phi, OF * **] *** by blast
+
+lemma properPath_hdd_lab:
+"properPath as tr \<Longrightarrow> hdd as = lab tr"
+by (erule properPath.cases) blast
+
+lemma properPath_sub:
+"properPath as tr \<Longrightarrow>
+ \<exists> tr' \<in> listF_set (sub tr). phi (tll as) tr' \<or> properPath (tll as) tr'"
+by (erule properPath.cases) blast
+
+(* prove the following by coinduction *)
+theorem Konig:
+ assumes "infiniteTr tr"
+ shows "properPath (konigPath tr) tr"
+proof-
+ {fix as
+ assume "infiniteTr tr \<and> as = konigPath tr" hence "properPath as tr"
+ proof (induct rule: properPath_coind, safe)
+ fix t
+ let ?t = "SOME t'. t' \<in> listF_set (sub t) \<and> infiniteTr t'"
+ assume "infiniteTr t"
+ hence "\<exists>t' \<in> listF_set (sub t). infiniteTr t'" by simp
+ hence "\<exists>t'. t' \<in> listF_set (sub t) \<and> infiniteTr t'" by blast
+ hence "?t \<in> listF_set (sub t) \<and> infiniteTr ?t" by (elim someI_ex)
+ moreover have "tll (konigPath t) = konigPath ?t" by simp
+ ultimately show "\<exists>t' \<in> listF_set (sub t).
+ infiniteTr t' \<and> tll (konigPath t) = konigPath t'" by blast
+ qed simp
+ }
+ thus ?thesis using assms by blast
+qed
+
+(* some more stream theorems *)
+
+lemma stream_map[simp]: "stream_map f = stream_dtor_unfold (f o hdd \<odot> tll)"
+unfolding stream_map_def pair_fun_def hdd_def[abs_def] tll_def[abs_def]
+ map_pair_def o_def prod_case_beta by simp
+
+lemma prod_rel[simp]: "prod_rel \<phi>1 \<phi>2 a b = (\<phi>1 (fst a) (fst b) \<and> \<phi>2 (snd a) (snd b))"
+unfolding prod_rel_def by auto
+
+lemmas stream_coind =
+ mp[OF stream.rel_coinduct, unfolded prod_rel[abs_def], folded hdd_def tll_def]
+
+definition plus :: "nat stream \<Rightarrow> nat stream \<Rightarrow> nat stream" (infixr "\<oplus>" 66) where
+ [simp]: "plus xs ys =
+ stream_dtor_unfold ((%(xs, ys). hdd xs + hdd ys) \<odot> (%(xs, ys). (tll xs, tll ys))) (xs, ys)"
+
+definition scalar :: "nat \<Rightarrow> nat stream \<Rightarrow> nat stream" (infixr "\<cdot>" 68) where
+ [simp]: "scalar n = stream_map (\<lambda>x. n * x)"
+
+definition ones :: "nat stream" where [simp]: "ones = stream_dtor_unfold ((%x. 1) \<odot> id) ()"
+definition twos :: "nat stream" where [simp]: "twos = stream_dtor_unfold ((%x. 2) \<odot> id) ()"
+definition ns :: "nat \<Rightarrow> nat stream" where [simp]: "ns n = scalar n ones"
+
+lemma "ones \<oplus> ones = twos"
+by (intro stream_coind[where P="%x1 x2. \<exists>x. x1 = ones \<oplus> ones \<and> x2 = twos"]) auto
+
+lemma "n \<cdot> twos = ns (2 * n)"
+by (intro stream_coind[where P="%x1 x2. \<exists>n. x1 = n \<cdot> twos \<and> x2 = ns (2 * n)"]) force+
+
+lemma prod_scalar: "(n * m) \<cdot> xs = n \<cdot> m \<cdot> xs"
+by (intro stream_coind[where P="%x1 x2. \<exists>n m xs. x1 = (n * m) \<cdot> xs \<and> x2 = n \<cdot> m \<cdot> xs"]) force+
+
+lemma scalar_plus: "n \<cdot> (xs \<oplus> ys) = n \<cdot> xs \<oplus> n \<cdot> ys"
+by (intro stream_coind[where P="%x1 x2. \<exists>n xs ys. x1 = n \<cdot> (xs \<oplus> ys) \<and> x2 = n \<cdot> xs \<oplus> n \<cdot> ys"])
+ (force simp: add_mult_distrib2)+
+
+lemma plus_comm: "xs \<oplus> ys = ys \<oplus> xs"
+by (intro stream_coind[where P="%x1 x2. \<exists>xs ys. x1 = xs \<oplus> ys \<and> x2 = ys \<oplus> xs"]) force+
+
+lemma plus_assoc: "(xs \<oplus> ys) \<oplus> zs = xs \<oplus> ys \<oplus> zs"
+by (intro stream_coind[where P="%x1 x2. \<exists>xs ys zs. x1 = (xs \<oplus> ys) \<oplus> zs \<and> x2 = xs \<oplus> ys \<oplus> zs"]) force+
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Examples/TreeFI.thy Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,83 @@
+(* Title: HOL/BNF/Examples/TreeFI.thy
+ Author: Dmitriy Traytel, TU Muenchen
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Finitely branching possibly infinite trees.
+*)
+
+header {* Finitely Branching Possibly Infinite Trees *}
+
+theory TreeFI
+imports ListF
+begin
+
+hide_const (open) Sublist.sub
+
+codata_raw treeFI: 'tree = "'a \<times> 'tree listF"
+
+lemma pre_treeFI_listF_set[simp]: "pre_treeFI_set2 (i, xs) = listF_set xs"
+unfolding pre_treeFI_set2_def collect_def[abs_def] prod_set_defs
+by (auto simp add: listF.set_natural')
+
+(* selectors for trees *)
+definition "lab tr \<equiv> fst (treeFI_dtor tr)"
+definition "sub tr \<equiv> snd (treeFI_dtor tr)"
+
+lemma dtor[simp]: "treeFI_dtor tr = (lab tr, sub tr)"
+unfolding lab_def sub_def by simp
+
+definition pair_fun (infixr "\<odot>" 50) where
+ "f \<odot> g \<equiv> \<lambda>x. (f x, g x)"
+
+lemma unfold_pair_fun_lab: "lab (treeFI_dtor_unfold (f \<odot> g) t) = f t"
+unfolding lab_def pair_fun_def treeFI.dtor_unfolds pre_treeFI_map_def by simp
+
+lemma unfold_pair_fun_sub: "sub (treeFI_dtor_unfold (f \<odot> g) t) = listF_map (treeFI_dtor_unfold (f \<odot> g)) (g t)"
+unfolding sub_def pair_fun_def treeFI.dtor_unfolds pre_treeFI_map_def by simp
+
+(* Tree reverse:*)
+definition "trev \<equiv> treeFI_dtor_unfold (lab \<odot> lrev o sub)"
+
+lemma trev_simps1[simp]: "lab (trev t) = lab t"
+unfolding trev_def by (simp add: unfold_pair_fun_lab)
+
+lemma trev_simps2[simp]: "sub (trev t) = listF_map trev (lrev (sub t))"
+unfolding trev_def by (simp add: unfold_pair_fun_sub)
+
+lemma treeFI_coinduct:
+assumes *: "phi x y"
+and step: "\<And>a b. phi a b \<Longrightarrow>
+ lab a = lab b \<and>
+ lengthh (sub a) = lengthh (sub b) \<and>
+ (\<forall>i < lengthh (sub a). phi (nthh (sub a) i) (nthh (sub b) i))"
+shows "x = y"
+proof (rule mp[OF treeFI.dtor_coinduct, of phi, OF _ *])
+ fix a b :: "'a treeFI"
+ let ?zs = "zipp (sub a) (sub b)"
+ let ?z = "(lab a, ?zs)"
+ assume "phi a b"
+ with step have step': "lab a = lab b" "lengthh (sub a) = lengthh (sub b)"
+ "\<forall>i < lengthh (sub a). phi (nthh (sub a) i) (nthh (sub b) i)" by auto
+ hence "pre_treeFI_map id fst ?z = treeFI_dtor a" "pre_treeFI_map id snd ?z = treeFI_dtor b"
+ unfolding pre_treeFI_map_def by auto
+ moreover have "\<forall>(x, y) \<in> pre_treeFI_set2 ?z. phi x y"
+ proof safe
+ fix z1 z2
+ assume "(z1, z2) \<in> pre_treeFI_set2 ?z"
+ hence "(z1, z2) \<in> listF_set ?zs" by auto
+ hence "\<exists>i < lengthh ?zs. nthh ?zs i = (z1, z2)" by auto
+ with step'(2) obtain i where "i < lengthh (sub a)"
+ "nthh (sub a) i = z1" "nthh (sub b) i = z2" by auto
+ with step'(3) show "phi z1 z2" by auto
+ qed
+ ultimately show "\<exists>z.
+ (pre_treeFI_map id fst z = treeFI_dtor a \<and>
+ pre_treeFI_map id snd z = treeFI_dtor b) \<and>
+ (\<forall>x y. (x, y) \<in> pre_treeFI_set2 z \<longrightarrow> phi x y)" by blast
+qed
+
+lemma trev_trev: "trev (trev tr) = tr"
+by (rule treeFI_coinduct[of "%a b. trev (trev b) = a"]) auto
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Examples/TreeFsetI.thy Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,59 @@
+(* Title: HOL/BNF/Examples/TreeFsetI.thy
+ Author: Dmitriy Traytel, TU Muenchen
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Finitely branching possibly infinite trees, with sets of children.
+*)
+
+header {* Finitely Branching Possibly Infinite Trees, with Sets of Children *}
+
+theory TreeFsetI
+imports "../BNF"
+begin
+
+hide_const (open) Sublist.sub
+hide_fact (open) Quotient_Product.prod_rel_def
+
+definition pair_fun (infixr "\<odot>" 50) where
+ "f \<odot> g \<equiv> \<lambda>x. (f x, g x)"
+
+codata_raw treeFsetI: 't = "'a \<times> 't fset"
+
+(* selectors for trees *)
+definition "lab t \<equiv> fst (treeFsetI_dtor t)"
+definition "sub t \<equiv> snd (treeFsetI_dtor t)"
+
+lemma dtor[simp]: "treeFsetI_dtor t = (lab t, sub t)"
+unfolding lab_def sub_def by simp
+
+lemma unfold_pair_fun_lab: "lab (treeFsetI_dtor_unfold (f \<odot> g) t) = f t"
+unfolding lab_def pair_fun_def treeFsetI.dtor_unfolds pre_treeFsetI_map_def by simp
+
+lemma unfold_pair_fun_sub: "sub (treeFsetI_dtor_unfold (f \<odot> g) t) = map_fset (treeFsetI_dtor_unfold (f \<odot> g)) (g t)"
+unfolding sub_def pair_fun_def treeFsetI.dtor_unfolds pre_treeFsetI_map_def by simp
+
+(* tree map (contrived example): *)
+definition "tmap f \<equiv> treeFsetI_dtor_unfold (f o lab \<odot> sub)"
+
+lemma tmap_simps1[simp]: "lab (tmap f t) = f (lab t)"
+unfolding tmap_def by (simp add: unfold_pair_fun_lab)
+
+lemma trev_simps2[simp]: "sub (tmap f t) = map_fset (tmap f) (sub t)"
+unfolding tmap_def by (simp add: unfold_pair_fun_sub)
+
+lemma pre_treeFsetI_rel[simp]: "pre_treeFsetI_rel R1 R2 a b = (R1 (fst a) (fst b) \<and>
+ (\<forall>t \<in> fset (snd a). (\<exists>u \<in> fset (snd b). R2 t u)) \<and>
+ (\<forall>t \<in> fset (snd b). (\<exists>u \<in> fset (snd a). R2 u t)))"
+apply (cases a)
+apply (cases b)
+apply (simp add: pre_treeFsetI_rel_def prod_rel_def fset_rel_def)
+done
+
+lemmas treeFsetI_coind = mp[OF treeFsetI.rel_coinduct]
+
+lemma "tmap (f o g) x = tmap f (tmap g x)"
+by (intro treeFsetI_coind[where P="%x1 x2. \<exists>x. x1 = tmap (f o g) x \<and> x2 = tmap f (tmap g x)"])
+ force+
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/More_BNFs.thy Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,1511 @@
+(* Title: HOL/BNF/More_BNFs.thy
+ Author: Dmitriy Traytel, TU Muenchen
+ Author: Andrei Popescu, TU Muenchen
+ Author: Andreas Lochbihler, Karlsruhe Institute of Technology
+ Author: Jasmin Blanchette, TU Muenchen
+ Copyright 2012
+
+Registration of various types as bounded natural functors.
+*)
+
+header {* Registration of Various Types as Bounded Natural Functors *}
+
+theory More_BNFs
+imports
+ BNF_LFP
+ BNF_GFP
+ "~~/src/HOL/Quotient_Examples/FSet"
+ "~~/src/HOL/Library/Multiset"
+ Countable_Set
+begin
+
+lemma option_rec_conv_option_case: "option_rec = option_case"
+by (simp add: fun_eq_iff split: option.split)
+
+definition option_rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a option \<Rightarrow> 'b option \<Rightarrow> bool" where
+"option_rel R x_opt y_opt =
+ (case (x_opt, y_opt) of
+ (None, None) \<Rightarrow> True
+ | (Some x, Some y) \<Rightarrow> R x y
+ | _ \<Rightarrow> False)"
+
+bnf_def Option.map [Option.set] "\<lambda>_::'a option. natLeq" ["None"] option_rel
+proof -
+ show "Option.map id = id" by (simp add: fun_eq_iff Option.map_def split: option.split)
+next
+ fix f g
+ show "Option.map (g \<circ> f) = Option.map g \<circ> Option.map f"
+ by (auto simp add: fun_eq_iff Option.map_def split: option.split)
+next
+ fix f g x
+ assume "\<And>z. z \<in> Option.set x \<Longrightarrow> f z = g z"
+ thus "Option.map f x = Option.map g x"
+ by (simp cong: Option.map_cong)
+next
+ fix f
+ show "Option.set \<circ> Option.map f = op ` f \<circ> Option.set"
+ by fastforce
+next
+ show "card_order natLeq" by (rule natLeq_card_order)
+next
+ show "cinfinite natLeq" by (rule natLeq_cinfinite)
+next
+ fix x
+ show "|Option.set x| \<le>o natLeq"
+ by (cases x) (simp_all add: ordLess_imp_ordLeq finite_iff_ordLess_natLeq[symmetric])
+next
+ fix A
+ have unfold: "{x. Option.set x \<subseteq> A} = Some ` A \<union> {None}"
+ by (auto simp add: option_rec_conv_option_case Option.set_def split: option.split_asm)
+ show "|{x. Option.set x \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq"
+ apply (rule ordIso_ordLeq_trans)
+ apply (rule card_of_ordIso_subst[OF unfold])
+ apply (rule ordLeq_transitive)
+ apply (rule Un_csum)
+ apply (rule ordLeq_transitive)
+ apply (rule csum_mono)
+ apply (rule card_of_image)
+ apply (rule ordIso_ordLeq_trans)
+ apply (rule single_cone)
+ apply (rule cone_ordLeq_ctwo)
+ apply (rule ordLeq_cexp1)
+ apply (simp_all add: natLeq_cinfinite natLeq_Card_order cinfinite_not_czero Card_order_csum)
+ done
+next
+ fix A B1 B2 f1 f2 p1 p2
+ assume wpull: "wpull A B1 B2 f1 f2 p1 p2"
+ show "wpull {x. Option.set x \<subseteq> A} {x. Option.set x \<subseteq> B1} {x. Option.set x \<subseteq> B2}
+ (Option.map f1) (Option.map f2) (Option.map p1) (Option.map p2)"
+ (is "wpull ?A ?B1 ?B2 ?f1 ?f2 ?p1 ?p2")
+ unfolding wpull_def
+ proof (intro strip, elim conjE)
+ fix b1 b2
+ assume "b1 \<in> ?B1" "b2 \<in> ?B2" "?f1 b1 = ?f2 b2"
+ thus "\<exists>a \<in> ?A. ?p1 a = b1 \<and> ?p2 a = b2" using wpull
+ unfolding wpull_def by (cases b2) (auto 4 5)
+ qed
+next
+ fix z
+ assume "z \<in> Option.set None"
+ thus False by simp
+next
+ fix R
+ show "{p. option_rel (\<lambda>x y. (x, y) \<in> R) (fst p) (snd p)} =
+ (Gr {x. Option.set x \<subseteq> R} (Option.map fst))\<inverse> O Gr {x. Option.set x \<subseteq> R} (Option.map snd)"
+ unfolding option_rel_def Gr_def relcomp_unfold converse_unfold
+ by (auto simp: trans[OF eq_commute option_map_is_None] trans[OF eq_commute option_map_eq_Some]
+ split: option.splits) blast
+qed
+
+lemma card_of_list_in:
+ "|{xs. set xs \<subseteq> A}| \<le>o |Pfunc (UNIV :: nat set) A|" (is "|?LHS| \<le>o |?RHS|")
+proof -
+ let ?f = "%xs. %i. if i < length xs \<and> set xs \<subseteq> A then Some (nth xs i) else None"
+ have "inj_on ?f ?LHS" unfolding inj_on_def fun_eq_iff
+ proof safe
+ fix xs :: "'a list" and ys :: "'a list"
+ assume su: "set xs \<subseteq> A" "set ys \<subseteq> A" and eq: "\<forall>i. ?f xs i = ?f ys i"
+ hence *: "length xs = length ys"
+ by (metis linorder_cases option.simps(2) order_less_irrefl)
+ thus "xs = ys" by (rule nth_equalityI) (metis * eq su option.inject)
+ qed
+ moreover have "?f ` ?LHS \<subseteq> ?RHS" unfolding Pfunc_def by fastforce
+ ultimately show ?thesis using card_of_ordLeq by blast
+qed
+
+lemma list_in_empty: "A = {} \<Longrightarrow> {x. set x \<subseteq> A} = {[]}"
+by simp
+
+lemma card_of_Func: "|Func A B| =o |B| ^c |A|"
+unfolding cexp_def Field_card_of by (rule card_of_refl)
+
+lemma not_emp_czero_notIn_ordIso_Card_order:
+"A \<noteq> {} \<Longrightarrow> ( |A|, czero) \<notin> ordIso \<and> Card_order |A|"
+ apply (rule conjI)
+ apply (metis Field_card_of czeroE)
+ by (rule card_of_Card_order)
+
+lemma list_in_bd: "|{x. set x \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq"
+proof -
+ fix A :: "'a set"
+ show "|{x. set x \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq"
+ proof (cases "A = {}")
+ case False thus ?thesis
+ apply -
+ apply (rule ordLeq_transitive)
+ apply (rule card_of_list_in)
+ apply (rule ordLeq_transitive)
+ apply (erule card_of_Pfunc_Pow_Func)
+ apply (rule ordIso_ordLeq_trans)
+ apply (rule Times_cprod)
+ apply (rule cprod_cinfinite_bound)
+ apply (rule ordIso_ordLeq_trans)
+ apply (rule Pow_cexp_ctwo)
+ apply (rule ordIso_ordLeq_trans)
+ apply (rule cexp_cong2)
+ apply (rule card_of_nat)
+ apply (rule Card_order_ctwo)
+ apply (rule card_of_Card_order)
+ apply (rule natLeq_Card_order)
+ apply (rule disjI1)
+ apply (rule ctwo_Cnotzero)
+ apply (rule cexp_mono1)
+ apply (rule ordLeq_csum2)
+ apply (rule Card_order_ctwo)
+ apply (rule disjI1)
+ apply (rule ctwo_Cnotzero)
+ apply (rule natLeq_Card_order)
+ apply (rule ordIso_ordLeq_trans)
+ apply (rule card_of_Func)
+ apply (rule ordIso_ordLeq_trans)
+ apply (rule cexp_cong2)
+ apply (rule card_of_nat)
+ apply (rule card_of_Card_order)
+ apply (rule card_of_Card_order)
+ apply (rule natLeq_Card_order)
+ apply (rule disjI1)
+ apply (erule not_emp_czero_notIn_ordIso_Card_order)
+ apply (rule cexp_mono1)
+ apply (rule ordLeq_csum1)
+ apply (rule card_of_Card_order)
+ apply (rule disjI1)
+ apply (erule not_emp_czero_notIn_ordIso_Card_order)
+ apply (rule natLeq_Card_order)
+ apply (rule card_of_Card_order)
+ apply (rule card_of_Card_order)
+ apply (rule Cinfinite_cexp)
+ apply (rule ordLeq_csum2)
+ apply (rule Card_order_ctwo)
+ apply (rule conjI)
+ apply (rule natLeq_cinfinite)
+ by (rule natLeq_Card_order)
+ next
+ case True thus ?thesis
+ apply -
+ apply (rule ordIso_ordLeq_trans)
+ apply (rule card_of_ordIso_subst)
+ apply (erule list_in_empty)
+ apply (rule ordIso_ordLeq_trans)
+ apply (rule single_cone)
+ apply (rule cone_ordLeq_cexp)
+ apply (rule ordLeq_transitive)
+ apply (rule cone_ordLeq_ctwo)
+ apply (rule ordLeq_csum2)
+ by (rule Card_order_ctwo)
+ qed
+qed
+
+bnf_def map [set] "\<lambda>_::'a list. natLeq" ["[]"]
+proof -
+ show "map id = id" by (rule List.map.id)
+next
+ fix f g
+ show "map (g o f) = map g o map f" by (rule List.map.comp[symmetric])
+next
+ fix x f g
+ assume "\<And>z. z \<in> set x \<Longrightarrow> f z = g z"
+ thus "map f x = map g x" by simp
+next
+ fix f
+ show "set o map f = image f o set" by (rule ext, unfold o_apply, rule set_map)
+next
+ show "card_order natLeq" by (rule natLeq_card_order)
+next
+ show "cinfinite natLeq" by (rule natLeq_cinfinite)
+next
+ fix x
+ show "|set x| \<le>o natLeq"
+ apply (rule ordLess_imp_ordLeq)
+ apply (rule finite_ordLess_infinite[OF _ natLeq_Well_order])
+ unfolding Field_natLeq Field_card_of by (auto simp: card_of_well_order_on)
+next
+ fix A :: "'a set"
+ show "|{x. set x \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq" by (rule list_in_bd)
+next
+ fix A B1 B2 f1 f2 p1 p2
+ assume "wpull A B1 B2 f1 f2 p1 p2"
+ hence pull: "\<And>b1 b2. b1 \<in> B1 \<and> b2 \<in> B2 \<and> f1 b1 = f2 b2 \<Longrightarrow> \<exists>a \<in> A. p1 a = b1 \<and> p2 a = b2"
+ unfolding wpull_def by auto
+ show "wpull {x. set x \<subseteq> A} {x. set x \<subseteq> B1} {x. set x \<subseteq> B2} (map f1) (map f2) (map p1) (map p2)"
+ (is "wpull ?A ?B1 ?B2 _ _ _ _")
+ proof (unfold wpull_def)
+ { fix as bs assume *: "as \<in> ?B1" "bs \<in> ?B2" "map f1 as = map f2 bs"
+ hence "length as = length bs" by (metis length_map)
+ hence "\<exists>zs \<in> ?A. map p1 zs = as \<and> map p2 zs = bs" using *
+ proof (induct as bs rule: list_induct2)
+ case (Cons a as b bs)
+ hence "a \<in> B1" "b \<in> B2" "f1 a = f2 b" by auto
+ with pull obtain z where "z \<in> A" "p1 z = a" "p2 z = b" by blast
+ moreover
+ from Cons obtain zs where "zs \<in> ?A" "map p1 zs = as" "map p2 zs = bs" by auto
+ ultimately have "z # zs \<in> ?A" "map p1 (z # zs) = a # as \<and> map p2 (z # zs) = b # bs" by auto
+ thus ?case by (rule_tac x = "z # zs" in bexI)
+ qed simp
+ }
+ thus "\<forall>as bs. as \<in> ?B1 \<and> bs \<in> ?B2 \<and> map f1 as = map f2 bs \<longrightarrow>
+ (\<exists>zs \<in> ?A. map p1 zs = as \<and> map p2 zs = bs)" by blast
+ qed
+qed simp+
+
+(* Finite sets *)
+abbreviation afset where "afset \<equiv> abs_fset"
+abbreviation rfset where "rfset \<equiv> rep_fset"
+
+lemma fset_fset_member:
+"fset A = {a. a |\<in>| A}"
+unfolding fset_def fset_member_def by auto
+
+lemma afset_rfset:
+"afset (rfset x) = x"
+by (rule Quotient_fset[unfolded Quotient_def, THEN conjunct1, rule_format])
+
+lemma afset_rfset_id:
+"afset o rfset = id"
+unfolding comp_def afset_rfset id_def ..
+
+lemma rfset:
+"rfset A = rfset B \<longleftrightarrow> A = B"
+by (metis afset_rfset)
+
+lemma afset_set:
+"afset as = afset bs \<longleftrightarrow> set as = set bs"
+using Quotient_fset unfolding Quotient_def list_eq_def by auto
+
+lemma surj_afset:
+"\<exists> as. A = afset as"
+by (metis afset_rfset)
+
+lemma fset_def2:
+"fset = set o rfset"
+unfolding fset_def map_fun_def[abs_def] by simp
+
+lemma fset_def2_raw:
+"fset A = set (rfset A)"
+unfolding fset_def2 by simp
+
+lemma fset_comp_afset:
+"fset o afset = set"
+unfolding fset_def2 comp_def apply(rule ext)
+unfolding afset_set[symmetric] afset_rfset ..
+
+lemma fset_afset:
+"fset (afset as) = set as"
+unfolding fset_comp_afset[symmetric] by simp
+
+lemma set_rfset_afset:
+"set (rfset (afset as)) = set as"
+unfolding afset_set[symmetric] afset_rfset ..
+
+lemma map_fset_comp_afset:
+"(map_fset f) o afset = afset o (map f)"
+unfolding map_fset_def map_fun_def[abs_def] comp_def apply(rule ext)
+unfolding afset_set set_map set_rfset_afset id_apply ..
+
+lemma map_fset_afset:
+"(map_fset f) (afset as) = afset (map f as)"
+using map_fset_comp_afset unfolding comp_def fun_eq_iff by auto
+
+lemma fset_map_fset:
+"fset (map_fset f A) = (image f) (fset A)"
+apply(subst afset_rfset[symmetric, of A])
+unfolding map_fset_afset fset_afset set_map
+unfolding fset_def2_raw ..
+
+lemma map_fset_def2:
+"map_fset f = afset o (map f) o rfset"
+unfolding map_fset_def map_fun_def[abs_def] by simp
+
+lemma map_fset_def2_raw:
+"map_fset f A = afset (map f (rfset A))"
+unfolding map_fset_def2 by simp
+
+lemma finite_ex_fset:
+assumes "finite A"
+shows "\<exists> B. fset B = A"
+by (metis assms finite_list fset_afset)
+
+lemma wpull_image:
+assumes "wpull A B1 B2 f1 f2 p1 p2"
+shows "wpull (Pow A) (Pow B1) (Pow B2) (image f1) (image f2) (image p1) (image p2)"
+unfolding wpull_def Pow_def Bex_def mem_Collect_eq proof clarify
+ fix Y1 Y2 assume Y1: "Y1 \<subseteq> B1" and Y2: "Y2 \<subseteq> B2" and EQ: "f1 ` Y1 = f2 ` Y2"
+ def X \<equiv> "{a \<in> A. p1 a \<in> Y1 \<and> p2 a \<in> Y2}"
+ show "\<exists>X\<subseteq>A. p1 ` X = Y1 \<and> p2 ` X = Y2"
+ proof (rule exI[of _ X], intro conjI)
+ show "p1 ` X = Y1"
+ proof
+ show "Y1 \<subseteq> p1 ` X"
+ proof safe
+ fix y1 assume y1: "y1 \<in> Y1"
+ then obtain y2 where y2: "y2 \<in> Y2" and eq: "f1 y1 = f2 y2" using EQ by auto
+ then obtain x where "x \<in> A" and "p1 x = y1" and "p2 x = y2"
+ using assms y1 Y1 Y2 unfolding wpull_def by blast
+ thus "y1 \<in> p1 ` X" unfolding X_def using y1 y2 by auto
+ qed
+ qed(unfold X_def, auto)
+ show "p2 ` X = Y2"
+ proof
+ show "Y2 \<subseteq> p2 ` X"
+ proof safe
+ fix y2 assume y2: "y2 \<in> Y2"
+ then obtain y1 where y1: "y1 \<in> Y1" and eq: "f1 y1 = f2 y2" using EQ by force
+ then obtain x where "x \<in> A" and "p1 x = y1" and "p2 x = y2"
+ using assms y2 Y1 Y2 unfolding wpull_def by blast
+ thus "y2 \<in> p2 ` X" unfolding X_def using y1 y2 by auto
+ qed
+ qed(unfold X_def, auto)
+ qed(unfold X_def, auto)
+qed
+
+lemma fset_to_fset: "finite A \<Longrightarrow> fset (the_inv fset A) = A"
+by (rule f_the_inv_into_f) (auto simp: inj_on_def fset_cong dest!: finite_ex_fset)
+
+definition fset_rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> 'b fset \<Rightarrow> bool" where
+"fset_rel R a b \<longleftrightarrow>
+ (\<forall>t \<in> fset a. \<exists>u \<in> fset b. R t u) \<and>
+ (\<forall>t \<in> fset b. \<exists>u \<in> fset a. R u t)"
+
+lemma fset_rel_aux:
+"(\<forall>t \<in> fset a. \<exists>u \<in> fset b. R t u) \<and> (\<forall>u \<in> fset b. \<exists>t \<in> fset a. R t u) \<longleftrightarrow>
+ (a, b) \<in> (Gr {a. fset a \<subseteq> {(a, b). R a b}} (map_fset fst))\<inverse> O
+ Gr {a. fset a \<subseteq> {(a, b). R a b}} (map_fset snd)" (is "?L = ?R")
+proof
+ assume ?L
+ def R' \<equiv> "the_inv fset (Collect (split R) \<inter> (fset a \<times> fset b))" (is "the_inv fset ?L'")
+ have "finite ?L'" by (intro finite_Int[OF disjI2] finite_cartesian_product) auto
+ hence *: "fset R' = ?L'" unfolding R'_def by (intro fset_to_fset)
+ show ?R unfolding Gr_def relcomp_unfold converse_unfold
+ proof (intro CollectI prod_caseI exI conjI)
+ from * show "(R', a) = (R', map_fset fst R')" using conjunct1[OF `?L`]
+ by (auto simp add: fset_cong[symmetric] image_def Int_def split: prod.splits)
+ from * show "(R', b) = (R', map_fset snd R')" using conjunct2[OF `?L`]
+ by (auto simp add: fset_cong[symmetric] image_def Int_def split: prod.splits)
+ qed (auto simp add: *)
+next
+ assume ?R thus ?L unfolding Gr_def relcomp_unfold converse_unfold
+ apply (simp add: subset_eq Ball_def)
+ apply (rule conjI)
+ apply (clarsimp, metis snd_conv)
+ by (clarsimp, metis fst_conv)
+qed
+
+bnf_def map_fset [fset] "\<lambda>_::'a fset. natLeq" ["{||}"] fset_rel
+proof -
+ show "map_fset id = id"
+ unfolding map_fset_def2 map_id o_id afset_rfset_id ..
+next
+ fix f g
+ show "map_fset (g o f) = map_fset g o map_fset f"
+ unfolding map_fset_def2 map.comp[symmetric] comp_def apply(rule ext)
+ unfolding afset_set set_map fset_def2_raw[symmetric] image_image[symmetric]
+ unfolding map_fset_afset[symmetric] map_fset_image afset_rfset
+ by (rule refl)
+next
+ fix x f g
+ assume "\<And>z. z \<in> fset x \<Longrightarrow> f z = g z"
+ hence "map f (rfset x) = map g (rfset x)"
+ apply(intro map_cong) unfolding fset_def2_raw by auto
+ thus "map_fset f x = map_fset g x" unfolding map_fset_def2_raw
+ by (rule arg_cong)
+next
+ fix f
+ show "fset o map_fset f = image f o fset"
+ unfolding comp_def fset_map_fset ..
+next
+ show "card_order natLeq" by (rule natLeq_card_order)
+next
+ show "cinfinite natLeq" by (rule natLeq_cinfinite)
+next
+ fix x
+ show "|fset x| \<le>o natLeq"
+ unfolding fset_def2_raw
+ apply (rule ordLess_imp_ordLeq)
+ apply (rule finite_iff_ordLess_natLeq[THEN iffD1])
+ by (rule finite_set)
+next
+ fix A :: "'a set"
+ have "|{x. fset x \<subseteq> A}| \<le>o |afset ` {as. set as \<subseteq> A}|"
+ apply(rule card_of_mono1) unfolding fset_def2_raw apply auto
+ apply (rule image_eqI)
+ by (auto simp: afset_rfset)
+ also have "|afset ` {as. set as \<subseteq> A}| \<le>o |{as. set as \<subseteq> A}|" using card_of_image .
+ also have "|{as. set as \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq" by (rule list_in_bd)
+ finally show "|{x. fset x \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq" .
+next
+ fix A B1 B2 f1 f2 p1 p2
+ assume wp: "wpull A B1 B2 f1 f2 p1 p2"
+ hence "wpull (Pow A) (Pow B1) (Pow B2) (image f1) (image f2) (image p1) (image p2)"
+ by (rule wpull_image)
+ show "wpull {x. fset x \<subseteq> A} {x. fset x \<subseteq> B1} {x. fset x \<subseteq> B2}
+ (map_fset f1) (map_fset f2) (map_fset p1) (map_fset p2)"
+ unfolding wpull_def Pow_def Bex_def mem_Collect_eq proof clarify
+ fix y1 y2
+ assume Y1: "fset y1 \<subseteq> B1" and Y2: "fset y2 \<subseteq> B2"
+ assume "map_fset f1 y1 = map_fset f2 y2"
+ hence EQ: "f1 ` (fset y1) = f2 ` (fset y2)" unfolding map_fset_def2_raw
+ unfolding afset_set set_map fset_def2_raw .
+ with Y1 Y2 obtain X where X: "X \<subseteq> A"
+ and Y1: "p1 ` X = fset y1" and Y2: "p2 ` X = fset y2"
+ using wpull_image[OF wp] unfolding wpull_def Pow_def
+ unfolding Bex_def mem_Collect_eq apply -
+ apply(erule allE[of _ "fset y1"], erule allE[of _ "fset y2"]) by auto
+ have "\<forall> y1' \<in> fset y1. \<exists> x. x \<in> X \<and> y1' = p1 x" using Y1 by auto
+ then obtain q1 where q1: "\<forall> y1' \<in> fset y1. q1 y1' \<in> X \<and> y1' = p1 (q1 y1')" by metis
+ have "\<forall> y2' \<in> fset y2. \<exists> x. x \<in> X \<and> y2' = p2 x" using Y2 by auto
+ then obtain q2 where q2: "\<forall> y2' \<in> fset y2. q2 y2' \<in> X \<and> y2' = p2 (q2 y2')" by metis
+ def X' \<equiv> "q1 ` (fset y1) \<union> q2 ` (fset y2)"
+ have X': "X' \<subseteq> A" and Y1: "p1 ` X' = fset y1" and Y2: "p2 ` X' = fset y2"
+ using X Y1 Y2 q1 q2 unfolding X'_def by auto
+ have fX': "finite X'" unfolding X'_def by simp
+ then obtain x where X'eq: "X' = fset x" by (auto dest: finite_ex_fset)
+ show "\<exists>x. fset x \<subseteq> A \<and> map_fset p1 x = y1 \<and> map_fset p2 x = y2"
+ apply(intro exI[of _ "x"]) using X' Y1 Y2
+ unfolding X'eq map_fset_def2_raw fset_def2_raw set_map[symmetric]
+ afset_set[symmetric] afset_rfset by simp
+ qed
+next
+ fix R
+ show "{p. fset_rel (\<lambda>x y. (x, y) \<in> R) (fst p) (snd p)} =
+ (Gr {x. fset x \<subseteq> R} (map_fset fst))\<inverse> O Gr {x. fset x \<subseteq> R} (map_fset snd)"
+ unfolding fset_rel_def fset_rel_aux by simp
+qed auto
+
+(* Countable sets *)
+
+lemma card_of_countable_sets_range:
+fixes A :: "'a set"
+shows "|{X. X \<subseteq> A \<and> countable X \<and> X \<noteq> {}}| \<le>o |{f::nat \<Rightarrow> 'a. range f \<subseteq> A}|"
+apply(rule card_of_ordLeqI[of fromNat]) using inj_on_fromNat
+unfolding inj_on_def by auto
+
+lemma card_of_countable_sets_Func:
+"|{X. X \<subseteq> A \<and> countable X \<and> X \<noteq> {}}| \<le>o |A| ^c natLeq"
+using card_of_countable_sets_range card_of_Func_UNIV[THEN ordIso_symmetric]
+unfolding cexp_def Field_natLeq Field_card_of
+by (rule ordLeq_ordIso_trans)
+
+lemma ordLeq_countable_subsets:
+"|A| \<le>o |{X. X \<subseteq> A \<and> countable X}|"
+apply (rule card_of_ordLeqI[of "\<lambda> a. {a}"]) unfolding inj_on_def by auto
+
+lemma finite_countable_subset:
+"finite {X. X \<subseteq> A \<and> countable X} \<longleftrightarrow> finite A"
+apply default
+ apply (erule contrapos_pp)
+ apply (rule card_of_ordLeq_infinite)
+ apply (rule ordLeq_countable_subsets)
+ apply assumption
+apply (rule finite_Collect_conjI)
+apply (rule disjI1)
+by (erule finite_Collect_subsets)
+
+lemma card_of_countable_sets:
+"|{X. X \<subseteq> A \<and> countable X}| \<le>o ( |A| +c ctwo) ^c natLeq"
+(is "|?L| \<le>o _")
+proof(cases "finite A")
+ let ?R = "Func (UNIV::nat set) (A <+> (UNIV::bool set))"
+ case True hence "finite ?L" by simp
+ moreover have "infinite ?R"
+ apply(rule infinite_Func[of _ "Inr True" "Inr False"]) by auto
+ ultimately show ?thesis unfolding cexp_def csum_def ctwo_def Field_natLeq Field_card_of
+ apply(intro ordLess_imp_ordLeq) by (rule finite_ordLess_infinite2)
+next
+ case False
+ hence "|{X. X \<subseteq> A \<and> countable X}| =o |{X. X \<subseteq> A \<and> countable X} - {{}}|"
+ by (intro card_of_infinite_diff_finitte finite.emptyI finite.insertI ordIso_symmetric)
+ (unfold finite_countable_subset)
+ also have "|{X. X \<subseteq> A \<and> countable X} - {{}}| \<le>o |A| ^c natLeq"
+ using card_of_countable_sets_Func[of A] unfolding set_diff_eq by auto
+ also have "|A| ^c natLeq \<le>o ( |A| +c ctwo) ^c natLeq"
+ apply(rule cexp_mono1_cone_ordLeq)
+ apply(rule ordLeq_csum1, rule card_of_Card_order)
+ apply (rule cone_ordLeq_cexp)
+ apply (rule cone_ordLeq_Cnotzero)
+ using csum_Cnotzero2 ctwo_Cnotzero apply blast
+ by (rule natLeq_Card_order)
+ finally show ?thesis .
+qed
+
+lemma rcset_to_rcset: "countable A \<Longrightarrow> rcset (the_inv rcset A) = A"
+apply (rule f_the_inv_into_f)
+unfolding inj_on_def rcset_inj using rcset_surj by auto
+
+lemma Collect_Int_Times:
+"{(x, y). R x y} \<inter> A \<times> B = {(x, y). R x y \<and> x \<in> A \<and> y \<in> B}"
+by auto
+
+lemma rcset_natural': "rcset (cIm f x) = f ` rcset x"
+unfolding cIm_def[abs_def] by simp
+
+definition cset_rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a cset \<Rightarrow> 'b cset \<Rightarrow> bool" where
+"cset_rel R a b \<longleftrightarrow>
+ (\<forall>t \<in> rcset a. \<exists>u \<in> rcset b. R t u) \<and>
+ (\<forall>t \<in> rcset b. \<exists>u \<in> rcset a. R u t)"
+
+lemma cset_rel_aux:
+"(\<forall>t \<in> rcset a. \<exists>u \<in> rcset b. R t u) \<and> (\<forall>t \<in> rcset b. \<exists>u \<in> rcset a. R u t) \<longleftrightarrow>
+ (a, b) \<in> (Gr {x. rcset x \<subseteq> {(a, b). R a b}} (cIm fst))\<inverse> O
+ Gr {x. rcset x \<subseteq> {(a, b). R a b}} (cIm snd)" (is "?L = ?R")
+proof
+ assume ?L
+ def R' \<equiv> "the_inv rcset (Collect (split R) \<inter> (rcset a \<times> rcset b))"
+ (is "the_inv rcset ?L'")
+ have "countable ?L'" by auto
+ hence *: "rcset R' = ?L'" unfolding R'_def using fset_to_fset by (intro rcset_to_rcset)
+ show ?R unfolding Gr_def relcomp_unfold converse_unfold
+ proof (intro CollectI prod_caseI exI conjI)
+ have "rcset a = fst ` ({(x, y). R x y} \<inter> rcset a \<times> rcset b)" (is "_ = ?A")
+ using conjunct1[OF `?L`] unfolding image_def by (auto simp add: Collect_Int_Times)
+ hence "a = acset ?A" by (metis acset_rcset)
+ thus "(R', a) = (R', cIm fst R')" unfolding cIm_def * by auto
+ have "rcset b = snd ` ({(x, y). R x y} \<inter> rcset a \<times> rcset b)" (is "_ = ?B")
+ using conjunct2[OF `?L`] unfolding image_def by (auto simp add: Collect_Int_Times)
+ hence "b = acset ?B" by (metis acset_rcset)
+ thus "(R', b) = (R', cIm snd R')" unfolding cIm_def * by auto
+ qed (auto simp add: *)
+next
+ assume ?R thus ?L unfolding Gr_def relcomp_unfold converse_unfold
+ apply (simp add: subset_eq Ball_def)
+ apply (rule conjI)
+ apply (clarsimp, metis (lifting, no_types) rcset_natural' image_iff surjective_pairing)
+ apply (clarsimp)
+ by (metis Domain.intros Range.simps rcset_natural' fst_eq_Domain snd_eq_Range)
+qed
+
+bnf_def cIm [rcset] "\<lambda>_::'a cset. natLeq" ["cEmp"] cset_rel
+proof -
+ show "cIm id = id" unfolding cIm_def[abs_def] id_def by auto
+next
+ fix f g show "cIm (g \<circ> f) = cIm g \<circ> cIm f"
+ unfolding cIm_def[abs_def] apply(rule ext) unfolding comp_def by auto
+next
+ fix C f g assume eq: "\<And>a. a \<in> rcset C \<Longrightarrow> f a = g a"
+ thus "cIm f C = cIm g C"
+ unfolding cIm_def[abs_def] unfolding image_def by auto
+next
+ fix f show "rcset \<circ> cIm f = op ` f \<circ> rcset" unfolding cIm_def[abs_def] by auto
+next
+ show "card_order natLeq" by (rule natLeq_card_order)
+next
+ show "cinfinite natLeq" by (rule natLeq_cinfinite)
+next
+ fix C show "|rcset C| \<le>o natLeq" using rcset unfolding countable_def .
+next
+ fix A :: "'a set"
+ have "|{Z. rcset Z \<subseteq> A}| \<le>o |acset ` {X. X \<subseteq> A \<and> countable X}|"
+ apply(rule card_of_mono1) unfolding Pow_def image_def
+ proof (rule Collect_mono, clarsimp)
+ fix x
+ assume "rcset x \<subseteq> A"
+ hence "rcset x \<subseteq> A \<and> countable (rcset x) \<and> x = acset (rcset x)"
+ using acset_rcset[of x] rcset[of x] by force
+ thus "\<exists>y \<subseteq> A. countable y \<and> x = acset y" by blast
+ qed
+ also have "|acset ` {X. X \<subseteq> A \<and> countable X}| \<le>o |{X. X \<subseteq> A \<and> countable X}|"
+ using card_of_image .
+ also have "|{X. X \<subseteq> A \<and> countable X}| \<le>o ( |A| +c ctwo) ^c natLeq"
+ using card_of_countable_sets .
+ finally show "|{Z. rcset Z \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq" .
+next
+ fix A B1 B2 f1 f2 p1 p2
+ assume wp: "wpull A B1 B2 f1 f2 p1 p2"
+ show "wpull {x. rcset x \<subseteq> A} {x. rcset x \<subseteq> B1} {x. rcset x \<subseteq> B2}
+ (cIm f1) (cIm f2) (cIm p1) (cIm p2)"
+ unfolding wpull_def proof safe
+ fix y1 y2
+ assume Y1: "rcset y1 \<subseteq> B1" and Y2: "rcset y2 \<subseteq> B2"
+ assume "cIm f1 y1 = cIm f2 y2"
+ hence EQ: "f1 ` (rcset y1) = f2 ` (rcset y2)"
+ unfolding cIm_def by auto
+ with Y1 Y2 obtain X where X: "X \<subseteq> A"
+ and Y1: "p1 ` X = rcset y1" and Y2: "p2 ` X = rcset y2"
+ using wpull_image[OF wp] unfolding wpull_def Pow_def
+ unfolding Bex_def mem_Collect_eq apply -
+ apply(erule allE[of _ "rcset y1"], erule allE[of _ "rcset y2"]) by auto
+ have "\<forall> y1' \<in> rcset y1. \<exists> x. x \<in> X \<and> y1' = p1 x" using Y1 by auto
+ then obtain q1 where q1: "\<forall> y1' \<in> rcset y1. q1 y1' \<in> X \<and> y1' = p1 (q1 y1')" by metis
+ have "\<forall> y2' \<in> rcset y2. \<exists> x. x \<in> X \<and> y2' = p2 x" using Y2 by auto
+ then obtain q2 where q2: "\<forall> y2' \<in> rcset y2. q2 y2' \<in> X \<and> y2' = p2 (q2 y2')" by metis
+ def X' \<equiv> "q1 ` (rcset y1) \<union> q2 ` (rcset y2)"
+ have X': "X' \<subseteq> A" and Y1: "p1 ` X' = rcset y1" and Y2: "p2 ` X' = rcset y2"
+ using X Y1 Y2 q1 q2 unfolding X'_def by fast+
+ have fX': "countable X'" unfolding X'_def by simp
+ then obtain x where X'eq: "X' = rcset x" by (metis rcset_acset)
+ show "\<exists>x\<in>{x. rcset x \<subseteq> A}. cIm p1 x = y1 \<and> cIm p2 x = y2"
+ apply(intro bexI[of _ "x"]) using X' Y1 Y2 unfolding X'eq cIm_def by auto
+ qed
+next
+ fix R
+ show "{p. cset_rel (\<lambda>x y. (x, y) \<in> R) (fst p) (snd p)} =
+ (Gr {x. rcset x \<subseteq> R} (cIm fst))\<inverse> O Gr {x. rcset x \<subseteq> R} (cIm snd)"
+ unfolding cset_rel_def cset_rel_aux by simp
+qed (unfold cEmp_def, auto)
+
+
+(* Multisets *)
+
+(* The cardinal of a mutiset: this, and the following basic lemmas about it,
+should eventually go into Multiset.thy *)
+definition "mcard M \<equiv> setsum (count M) {a. count M a > 0}"
+
+lemma mcard_emp[simp]: "mcard {#} = 0"
+unfolding mcard_def by auto
+
+lemma mcard_emp_iff[simp]: "mcard M = 0 \<longleftrightarrow> M = {#}"
+unfolding mcard_def apply safe
+ apply simp_all
+ by (metis multi_count_eq zero_multiset.rep_eq)
+
+lemma mcard_singl[simp]: "mcard {#a#} = Suc 0"
+unfolding mcard_def by auto
+
+lemma mcard_Plus[simp]: "mcard (M + N) = mcard M + mcard N"
+proof-
+ have "setsum (count M) {a. 0 < count M a + count N a} =
+ setsum (count M) {a. a \<in># M}"
+ apply(rule setsum_mono_zero_cong_right) by auto
+ moreover
+ have "setsum (count N) {a. 0 < count M a + count N a} =
+ setsum (count N) {a. a \<in># N}"
+ apply(rule setsum_mono_zero_cong_right) by auto
+ ultimately show ?thesis
+ unfolding mcard_def count_union[THEN ext] comm_monoid_add_class.setsum.F_fun_f by simp
+qed
+
+lemma setsum_gt_0_iff:
+fixes f :: "'a \<Rightarrow> nat" assumes "finite A"
+shows "setsum f A > 0 \<longleftrightarrow> (\<exists> a \<in> A. f a > 0)"
+(is "?L \<longleftrightarrow> ?R")
+proof-
+ have "?L \<longleftrightarrow> \<not> setsum f A = 0" by fast
+ also have "... \<longleftrightarrow> (\<exists> a \<in> A. f a \<noteq> 0)" using assms by simp
+ also have "... \<longleftrightarrow> ?R" by simp
+ finally show ?thesis .
+qed
+
+(* *)
+definition mmap :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> nat) \<Rightarrow> 'b \<Rightarrow> nat" where
+"mmap h f b = setsum f {a. h a = b \<and> f a > 0}"
+
+lemma mmap_id: "mmap id = id"
+proof (rule ext)+
+ fix f a show "mmap id f a = id f a"
+ proof(cases "f a = 0")
+ case False
+ hence 1: "{aa. aa = a \<and> 0 < f aa} = {a}" by auto
+ show ?thesis by (simp add: mmap_def id_apply 1)
+ qed(unfold mmap_def, auto)
+qed
+
+lemma inj_on_setsum_inv:
+assumes f: "f \<in> multiset"
+and 1: "(0::nat) < setsum f {a. h a = b' \<and> 0 < f a}" (is "0 < setsum f ?A'")
+and 2: "{a. h a = b \<and> 0 < f a} = {a. h a = b' \<and> 0 < f a}" (is "?A = ?A'")
+shows "b = b'"
+proof-
+ have "finite ?A'" using f unfolding multiset_def by auto
+ hence "?A' \<noteq> {}" using 1 setsum_gt_0_iff by auto
+ thus ?thesis using 2 by auto
+qed
+
+lemma mmap_comp:
+fixes h1 :: "'a \<Rightarrow> 'b" and h2 :: "'b \<Rightarrow> 'c"
+assumes f: "f \<in> multiset"
+shows "mmap (h2 o h1) f = (mmap h2 o mmap h1) f"
+unfolding mmap_def[abs_def] comp_def proof(rule ext)+
+ fix c :: 'c
+ let ?A = "{a. h2 (h1 a) = c \<and> 0 < f a}"
+ let ?As = "\<lambda> b. {a. h1 a = b \<and> 0 < f a}"
+ let ?B = "{b. h2 b = c \<and> 0 < setsum f (?As b)}"
+ have 0: "{?As b | b. b \<in> ?B} = ?As ` ?B" by auto
+ have "\<And> b. finite (?As b)" using f unfolding multiset_def by simp
+ hence "?B = {b. h2 b = c \<and> ?As b \<noteq> {}}" using setsum_gt_0_iff by auto
+ hence A: "?A = \<Union> {?As b | b. b \<in> ?B}" by auto
+ have "setsum f ?A = setsum (setsum f) {?As b | b. b \<in> ?B}"
+ unfolding A apply(rule setsum_Union_disjoint)
+ using f unfolding multiset_def by auto
+ also have "... = setsum (setsum f) (?As ` ?B)" unfolding 0 ..
+ also have "... = setsum (setsum f o ?As) ?B" apply(rule setsum_reindex)
+ unfolding inj_on_def apply auto using inj_on_setsum_inv[OF f, of h1] by blast
+ also have "... = setsum (\<lambda> b. setsum f (?As b)) ?B" unfolding comp_def ..
+ finally show "setsum f ?A = setsum (\<lambda> b. setsum f (?As b)) ?B" .
+qed
+
+lemma mmap_comp1:
+fixes h1 :: "'a \<Rightarrow> 'b" and h2 :: "'b \<Rightarrow> 'c"
+assumes "f \<in> multiset"
+shows "mmap (\<lambda> a. h2 (h1 a)) f = mmap h2 (mmap h1 f)"
+using mmap_comp[OF assms] unfolding comp_def by auto
+
+lemma mmap:
+assumes "f \<in> multiset"
+shows "mmap h f \<in> multiset"
+using assms unfolding mmap_def[abs_def] multiset_def proof safe
+ assume fin: "finite {a. 0 < f a}" (is "finite ?A")
+ show "finite {b. 0 < setsum f {a. h a = b \<and> 0 < f a}}"
+ (is "finite {b. 0 < setsum f (?As b)}")
+ proof- let ?B = "{b. 0 < setsum f (?As b)}"
+ have "\<And> b. finite (?As b)" using assms unfolding multiset_def by simp
+ hence B: "?B = {b. ?As b \<noteq> {}}" using setsum_gt_0_iff by auto
+ hence "?B \<subseteq> h ` ?A" by auto
+ thus ?thesis using finite_surj[OF fin] by auto
+ qed
+qed
+
+lemma mmap_cong:
+assumes "\<And>a. a \<in># M \<Longrightarrow> f a = g a"
+shows "mmap f (count M) = mmap g (count M)"
+using assms unfolding mmap_def[abs_def] by (intro ext, intro setsum_cong) auto
+
+abbreviation supp where "supp f \<equiv> {a. f a > 0}"
+
+lemma mmap_image_comp:
+assumes f: "f \<in> multiset"
+shows "(supp o mmap h) f = (image h o supp) f"
+unfolding mmap_def[abs_def] comp_def proof-
+ have "\<And> b. finite {a. h a = b \<and> 0 < f a}" (is "\<And> b. finite (?As b)")
+ using f unfolding multiset_def by auto
+ thus "{b. 0 < setsum f (?As b)} = h ` {a. 0 < f a}"
+ using setsum_gt_0_iff by auto
+qed
+
+lemma mmap_image:
+assumes f: "f \<in> multiset"
+shows "supp (mmap h f) = h ` (supp f)"
+using mmap_image_comp[OF assms] unfolding comp_def .
+
+lemma set_of_Abs_multiset:
+assumes f: "f \<in> multiset"
+shows "set_of (Abs_multiset f) = supp f"
+using assms unfolding set_of_def by (auto simp: Abs_multiset_inverse)
+
+lemma supp_count:
+"supp (count M) = set_of M"
+using assms unfolding set_of_def by auto
+
+lemma multiset_of_surj:
+"multiset_of ` {as. set as \<subseteq> A} = {M. set_of M \<subseteq> A}"
+proof safe
+ fix M assume M: "set_of M \<subseteq> A"
+ obtain as where eq: "M = multiset_of as" using surj_multiset_of unfolding surj_def by auto
+ hence "set as \<subseteq> A" using M by auto
+ thus "M \<in> multiset_of ` {as. set as \<subseteq> A}" using eq by auto
+next
+ show "\<And>x xa xb. \<lbrakk>set xa \<subseteq> A; xb \<in> set_of (multiset_of xa)\<rbrakk> \<Longrightarrow> xb \<in> A"
+ by (erule set_mp) (unfold set_of_multiset_of)
+qed
+
+lemma card_of_set_of:
+"|{M. set_of M \<subseteq> A}| \<le>o |{as. set as \<subseteq> A}|"
+apply(rule card_of_ordLeqI2[of _ multiset_of]) using multiset_of_surj by auto
+
+lemma nat_sum_induct:
+assumes "\<And>n1 n2. (\<And> m1 m2. m1 + m2 < n1 + n2 \<Longrightarrow> phi m1 m2) \<Longrightarrow> phi n1 n2"
+shows "phi (n1::nat) (n2::nat)"
+proof-
+ let ?chi = "\<lambda> n1n2 :: nat * nat. phi (fst n1n2) (snd n1n2)"
+ have "?chi (n1,n2)"
+ apply(induct rule: measure_induct[of "\<lambda> n1n2. fst n1n2 + snd n1n2" ?chi])
+ using assms by (metis fstI sndI)
+ thus ?thesis by simp
+qed
+
+lemma matrix_count:
+fixes ct1 ct2 :: "nat \<Rightarrow> nat"
+assumes "setsum ct1 {..<Suc n1} = setsum ct2 {..<Suc n2}"
+shows
+"\<exists> ct. (\<forall> i1 \<le> n1. setsum (\<lambda> i2. ct i1 i2) {..<Suc n2} = ct1 i1) \<and>
+ (\<forall> i2 \<le> n2. setsum (\<lambda> i1. ct i1 i2) {..<Suc n1} = ct2 i2)"
+(is "?phi ct1 ct2 n1 n2")
+proof-
+ have "\<forall> ct1 ct2 :: nat \<Rightarrow> nat.
+ setsum ct1 {..<Suc n1} = setsum ct2 {..<Suc n2} \<longrightarrow> ?phi ct1 ct2 n1 n2"
+ proof(induct rule: nat_sum_induct[of
+"\<lambda> n1 n2. \<forall> ct1 ct2 :: nat \<Rightarrow> nat.
+ setsum ct1 {..<Suc n1} = setsum ct2 {..<Suc n2} \<longrightarrow> ?phi ct1 ct2 n1 n2"],
+ clarify)
+ fix n1 n2 :: nat and ct1 ct2 :: "nat \<Rightarrow> nat"
+ assume IH: "\<And> m1 m2. m1 + m2 < n1 + n2 \<Longrightarrow>
+ \<forall> dt1 dt2 :: nat \<Rightarrow> nat.
+ setsum dt1 {..<Suc m1} = setsum dt2 {..<Suc m2} \<longrightarrow> ?phi dt1 dt2 m1 m2"
+ and ss: "setsum ct1 {..<Suc n1} = setsum ct2 {..<Suc n2}"
+ show "?phi ct1 ct2 n1 n2"
+ proof(cases n1)
+ case 0 note n1 = 0
+ show ?thesis
+ proof(cases n2)
+ case 0 note n2 = 0
+ let ?ct = "\<lambda> i1 i2. ct2 0"
+ show ?thesis apply(rule exI[of _ ?ct]) using n1 n2 ss by simp
+ next
+ case (Suc m2) note n2 = Suc
+ let ?ct = "\<lambda> i1 i2. ct2 i2"
+ show ?thesis apply(rule exI[of _ ?ct]) using n1 n2 ss by auto
+ qed
+ next
+ case (Suc m1) note n1 = Suc
+ show ?thesis
+ proof(cases n2)
+ case 0 note n2 = 0
+ let ?ct = "\<lambda> i1 i2. ct1 i1"
+ show ?thesis apply(rule exI[of _ ?ct]) using n1 n2 ss by auto
+ next
+ case (Suc m2) note n2 = Suc
+ show ?thesis
+ proof(cases "ct1 n1 \<le> ct2 n2")
+ case True
+ def dt2 \<equiv> "\<lambda> i2. if i2 = n2 then ct2 i2 - ct1 n1 else ct2 i2"
+ have "setsum ct1 {..<Suc m1} = setsum dt2 {..<Suc n2}"
+ unfolding dt2_def using ss n1 True by auto
+ hence "?phi ct1 dt2 m1 n2" using IH[of m1 n2] n1 by simp
+ then obtain dt where
+ 1: "\<And> i1. i1 \<le> m1 \<Longrightarrow> setsum (\<lambda> i2. dt i1 i2) {..<Suc n2} = ct1 i1" and
+ 2: "\<And> i2. i2 \<le> n2 \<Longrightarrow> setsum (\<lambda> i1. dt i1 i2) {..<Suc m1} = dt2 i2" by auto
+ let ?ct = "\<lambda> i1 i2. if i1 = n1 then (if i2 = n2 then ct1 n1 else 0)
+ else dt i1 i2"
+ show ?thesis apply(rule exI[of _ ?ct])
+ using n1 n2 1 2 True unfolding dt2_def by simp
+ next
+ case False
+ hence False: "ct2 n2 < ct1 n1" by simp
+ def dt1 \<equiv> "\<lambda> i1. if i1 = n1 then ct1 i1 - ct2 n2 else ct1 i1"
+ have "setsum dt1 {..<Suc n1} = setsum ct2 {..<Suc m2}"
+ unfolding dt1_def using ss n2 False by auto
+ hence "?phi dt1 ct2 n1 m2" using IH[of n1 m2] n2 by simp
+ then obtain dt where
+ 1: "\<And> i1. i1 \<le> n1 \<Longrightarrow> setsum (\<lambda> i2. dt i1 i2) {..<Suc m2} = dt1 i1" and
+ 2: "\<And> i2. i2 \<le> m2 \<Longrightarrow> setsum (\<lambda> i1. dt i1 i2) {..<Suc n1} = ct2 i2" by force
+ let ?ct = "\<lambda> i1 i2. if i2 = n2 then (if i1 = n1 then ct2 n2 else 0)
+ else dt i1 i2"
+ show ?thesis apply(rule exI[of _ ?ct])
+ using n1 n2 1 2 False unfolding dt1_def by simp
+ qed
+ qed
+ qed
+ qed
+ thus ?thesis using assms by auto
+qed
+
+definition
+"inj2 u B1 B2 \<equiv>
+ \<forall> b1 b1' b2 b2'. {b1,b1'} \<subseteq> B1 \<and> {b2,b2'} \<subseteq> B2 \<and> u b1 b2 = u b1' b2'
+ \<longrightarrow> b1 = b1' \<and> b2 = b2'"
+
+lemma matrix_setsum_finite:
+assumes B1: "B1 \<noteq> {}" "finite B1" and B2: "B2 \<noteq> {}" "finite B2" and u: "inj2 u B1 B2"
+and ss: "setsum N1 B1 = setsum N2 B2"
+shows "\<exists> M :: 'a \<Rightarrow> nat.
+ (\<forall> b1 \<in> B1. setsum (\<lambda> b2. M (u b1 b2)) B2 = N1 b1) \<and>
+ (\<forall> b2 \<in> B2. setsum (\<lambda> b1. M (u b1 b2)) B1 = N2 b2)"
+proof-
+ obtain n1 where "card B1 = Suc n1" using B1 by (metis card_insert finite.simps)
+ then obtain e1 where e1: "bij_betw e1 {..<Suc n1} B1"
+ using ex_bij_betw_finite_nat[OF B1(2)] by (metis atLeast0LessThan bij_betw_the_inv_into)
+ hence e1_inj: "inj_on e1 {..<Suc n1}" and e1_surj: "e1 ` {..<Suc n1} = B1"
+ unfolding bij_betw_def by auto
+ def f1 \<equiv> "inv_into {..<Suc n1} e1"
+ have f1: "bij_betw f1 B1 {..<Suc n1}"
+ and f1e1[simp]: "\<And> i1. i1 < Suc n1 \<Longrightarrow> f1 (e1 i1) = i1"
+ and e1f1[simp]: "\<And> b1. b1 \<in> B1 \<Longrightarrow> e1 (f1 b1) = b1" unfolding f1_def
+ apply (metis bij_betw_inv_into e1, metis bij_betw_inv_into_left e1 lessThan_iff)
+ by (metis e1_surj f_inv_into_f)
+ (* *)
+ obtain n2 where "card B2 = Suc n2" using B2 by (metis card_insert finite.simps)
+ then obtain e2 where e2: "bij_betw e2 {..<Suc n2} B2"
+ using ex_bij_betw_finite_nat[OF B2(2)] by (metis atLeast0LessThan bij_betw_the_inv_into)
+ hence e2_inj: "inj_on e2 {..<Suc n2}" and e2_surj: "e2 ` {..<Suc n2} = B2"
+ unfolding bij_betw_def by auto
+ def f2 \<equiv> "inv_into {..<Suc n2} e2"
+ have f2: "bij_betw f2 B2 {..<Suc n2}"
+ and f2e2[simp]: "\<And> i2. i2 < Suc n2 \<Longrightarrow> f2 (e2 i2) = i2"
+ and e2f2[simp]: "\<And> b2. b2 \<in> B2 \<Longrightarrow> e2 (f2 b2) = b2" unfolding f2_def
+ apply (metis bij_betw_inv_into e2, metis bij_betw_inv_into_left e2 lessThan_iff)
+ by (metis e2_surj f_inv_into_f)
+ (* *)
+ let ?ct1 = "N1 o e1" let ?ct2 = "N2 o e2"
+ have ss: "setsum ?ct1 {..<Suc n1} = setsum ?ct2 {..<Suc n2}"
+ unfolding setsum_reindex[OF e1_inj, symmetric] setsum_reindex[OF e2_inj, symmetric]
+ e1_surj e2_surj using ss .
+ obtain ct where
+ ct1: "\<And> i1. i1 \<le> n1 \<Longrightarrow> setsum (\<lambda> i2. ct i1 i2) {..<Suc n2} = ?ct1 i1" and
+ ct2: "\<And> i2. i2 \<le> n2 \<Longrightarrow> setsum (\<lambda> i1. ct i1 i2) {..<Suc n1} = ?ct2 i2"
+ using matrix_count[OF ss] by blast
+ (* *)
+ def A \<equiv> "{u b1 b2 | b1 b2. b1 \<in> B1 \<and> b2 \<in> B2}"
+ have "\<forall> a \<in> A. \<exists> b1b2 \<in> B1 <*> B2. u (fst b1b2) (snd b1b2) = a"
+ unfolding A_def Ball_def mem_Collect_eq by auto
+ then obtain h1h2 where h12:
+ "\<And>a. a \<in> A \<Longrightarrow> u (fst (h1h2 a)) (snd (h1h2 a)) = a \<and> h1h2 a \<in> B1 <*> B2" by metis
+ def h1 \<equiv> "fst o h1h2" def h2 \<equiv> "snd o h1h2"
+ have h12[simp]: "\<And>a. a \<in> A \<Longrightarrow> u (h1 a) (h2 a) = a"
+ "\<And> a. a \<in> A \<Longrightarrow> h1 a \<in> B1" "\<And> a. a \<in> A \<Longrightarrow> h2 a \<in> B2"
+ using h12 unfolding h1_def h2_def by force+
+ {fix b1 b2 assume b1: "b1 \<in> B1" and b2: "b2 \<in> B2"
+ hence inA: "u b1 b2 \<in> A" unfolding A_def by auto
+ hence "u b1 b2 = u (h1 (u b1 b2)) (h2 (u b1 b2))" by auto
+ moreover have "h1 (u b1 b2) \<in> B1" "h2 (u b1 b2) \<in> B2" using inA by auto
+ ultimately have "h1 (u b1 b2) = b1 \<and> h2 (u b1 b2) = b2"
+ using u b1 b2 unfolding inj2_def by fastforce
+ }
+ hence h1[simp]: "\<And> b1 b2. \<lbrakk>b1 \<in> B1; b2 \<in> B2\<rbrakk> \<Longrightarrow> h1 (u b1 b2) = b1" and
+ h2[simp]: "\<And> b1 b2. \<lbrakk>b1 \<in> B1; b2 \<in> B2\<rbrakk> \<Longrightarrow> h2 (u b1 b2) = b2" by auto
+ def M \<equiv> "\<lambda> a. ct (f1 (h1 a)) (f2 (h2 a))"
+ show ?thesis
+ apply(rule exI[of _ M]) proof safe
+ fix b1 assume b1: "b1 \<in> B1"
+ hence f1b1: "f1 b1 \<le> n1" using f1 unfolding bij_betw_def
+ by (metis bij_betwE f1 lessThan_iff less_Suc_eq_le)
+ have "(\<Sum>b2\<in>B2. M (u b1 b2)) = (\<Sum>i2<Suc n2. ct (f1 b1) (f2 (e2 i2)))"
+ unfolding e2_surj[symmetric] setsum_reindex[OF e2_inj]
+ unfolding M_def comp_def apply(intro setsum_cong) apply force
+ by (metis e2_surj b1 h1 h2 imageI)
+ also have "... = N1 b1" using b1 ct1[OF f1b1] by simp
+ finally show "(\<Sum>b2\<in>B2. M (u b1 b2)) = N1 b1" .
+ next
+ fix b2 assume b2: "b2 \<in> B2"
+ hence f2b2: "f2 b2 \<le> n2" using f2 unfolding bij_betw_def
+ by (metis bij_betwE f2 lessThan_iff less_Suc_eq_le)
+ have "(\<Sum>b1\<in>B1. M (u b1 b2)) = (\<Sum>i1<Suc n1. ct (f1 (e1 i1)) (f2 b2))"
+ unfolding e1_surj[symmetric] setsum_reindex[OF e1_inj]
+ unfolding M_def comp_def apply(intro setsum_cong) apply force
+ by (metis e1_surj b2 h1 h2 imageI)
+ also have "... = N2 b2" using b2 ct2[OF f2b2] by simp
+ finally show "(\<Sum>b1\<in>B1. M (u b1 b2)) = N2 b2" .
+ qed
+qed
+
+lemma supp_vimage_mmap:
+assumes "M \<in> multiset"
+shows "supp M \<subseteq> f -` (supp (mmap f M))"
+using assms by (auto simp: mmap_image)
+
+lemma mmap_ge_0:
+assumes "M \<in> multiset"
+shows "0 < mmap f M b \<longleftrightarrow> (\<exists>a. 0 < M a \<and> f a = b)"
+proof-
+ have f: "finite {a. f a = b \<and> 0 < M a}" using assms unfolding multiset_def by auto
+ show ?thesis unfolding mmap_def setsum_gt_0_iff[OF f] by auto
+qed
+
+lemma finite_twosets:
+assumes "finite B1" and "finite B2"
+shows "finite {u b1 b2 |b1 b2. b1 \<in> B1 \<and> b2 \<in> B2}" (is "finite ?A")
+proof-
+ have A: "?A = (\<lambda> b1b2. u (fst b1b2) (snd b1b2)) ` (B1 <*> B2)" by force
+ show ?thesis unfolding A using finite_cartesian_product[OF assms] by auto
+qed
+
+lemma wp_mmap:
+fixes A :: "'a set" and B1 :: "'b1 set" and B2 :: "'b2 set"
+assumes wp: "wpull A B1 B2 f1 f2 p1 p2"
+shows
+"wpull {M. M \<in> multiset \<and> supp M \<subseteq> A}
+ {N1. N1 \<in> multiset \<and> supp N1 \<subseteq> B1} {N2. N2 \<in> multiset \<and> supp N2 \<subseteq> B2}
+ (mmap f1) (mmap f2) (mmap p1) (mmap p2)"
+unfolding wpull_def proof (safe, unfold Bex_def mem_Collect_eq)
+ fix N1 :: "'b1 \<Rightarrow> nat" and N2 :: "'b2 \<Rightarrow> nat"
+ assume mmap': "mmap f1 N1 = mmap f2 N2"
+ and N1[simp]: "N1 \<in> multiset" "supp N1 \<subseteq> B1"
+ and N2[simp]: "N2 \<in> multiset" "supp N2 \<subseteq> B2"
+ have mN1[simp]: "mmap f1 N1 \<in> multiset" using N1 by (auto simp: mmap)
+ have mN2[simp]: "mmap f2 N2 \<in> multiset" using N2 by (auto simp: mmap)
+ def P \<equiv> "mmap f1 N1"
+ have P1: "P = mmap f1 N1" and P2: "P = mmap f2 N2" unfolding P_def using mmap' by auto
+ note P = P1 P2
+ have P_mult[simp]: "P \<in> multiset" unfolding P_def using N1 by auto
+ have fin_N1[simp]: "finite (supp N1)" using N1(1) unfolding multiset_def by auto
+ have fin_N2[simp]: "finite (supp N2)" using N2(1) unfolding multiset_def by auto
+ have fin_P[simp]: "finite (supp P)" using P_mult unfolding multiset_def by auto
+ (* *)
+ def set1 \<equiv> "\<lambda> c. {b1 \<in> supp N1. f1 b1 = c}"
+ have set1[simp]: "\<And> c b1. b1 \<in> set1 c \<Longrightarrow> f1 b1 = c" unfolding set1_def by auto
+ have fin_set1: "\<And> c. c \<in> supp P \<Longrightarrow> finite (set1 c)"
+ using N1(1) unfolding set1_def multiset_def by auto
+ have set1_NE: "\<And> c. c \<in> supp P \<Longrightarrow> set1 c \<noteq> {}"
+ unfolding set1_def P1 mmap_ge_0[OF N1(1)] by auto
+ have supp_N1_set1: "supp N1 = (\<Union> c \<in> supp P. set1 c)"
+ using supp_vimage_mmap[OF N1(1), of f1] unfolding set1_def P1 by auto
+ hence set1_inclN1: "\<And>c. c \<in> supp P \<Longrightarrow> set1 c \<subseteq> supp N1" by auto
+ hence set1_incl: "\<And> c. c \<in> supp P \<Longrightarrow> set1 c \<subseteq> B1" using N1(2) by blast
+ have set1_disj: "\<And> c c'. c \<noteq> c' \<Longrightarrow> set1 c \<inter> set1 c' = {}"
+ unfolding set1_def by auto
+ have setsum_set1: "\<And> c. setsum N1 (set1 c) = P c"
+ unfolding P1 set1_def mmap_def apply(rule setsum_cong) by auto
+ (* *)
+ def set2 \<equiv> "\<lambda> c. {b2 \<in> supp N2. f2 b2 = c}"
+ have set2[simp]: "\<And> c b2. b2 \<in> set2 c \<Longrightarrow> f2 b2 = c" unfolding set2_def by auto
+ have fin_set2: "\<And> c. c \<in> supp P \<Longrightarrow> finite (set2 c)"
+ using N2(1) unfolding set2_def multiset_def by auto
+ have set2_NE: "\<And> c. c \<in> supp P \<Longrightarrow> set2 c \<noteq> {}"
+ unfolding set2_def P2 mmap_ge_0[OF N2(1)] by auto
+ have supp_N2_set2: "supp N2 = (\<Union> c \<in> supp P. set2 c)"
+ using supp_vimage_mmap[OF N2(1), of f2] unfolding set2_def P2 by auto
+ hence set2_inclN2: "\<And>c. c \<in> supp P \<Longrightarrow> set2 c \<subseteq> supp N2" by auto
+ hence set2_incl: "\<And> c. c \<in> supp P \<Longrightarrow> set2 c \<subseteq> B2" using N2(2) by blast
+ have set2_disj: "\<And> c c'. c \<noteq> c' \<Longrightarrow> set2 c \<inter> set2 c' = {}"
+ unfolding set2_def by auto
+ have setsum_set2: "\<And> c. setsum N2 (set2 c) = P c"
+ unfolding P2 set2_def mmap_def apply(rule setsum_cong) by auto
+ (* *)
+ have ss: "\<And> c. c \<in> supp P \<Longrightarrow> setsum N1 (set1 c) = setsum N2 (set2 c)"
+ unfolding setsum_set1 setsum_set2 ..
+ have "\<forall> c \<in> supp P. \<forall> b1b2 \<in> (set1 c) \<times> (set2 c).
+ \<exists> a \<in> A. p1 a = fst b1b2 \<and> p2 a = snd b1b2"
+ using wp set1_incl set2_incl unfolding wpull_def Ball_def mem_Collect_eq
+ by simp (metis set1 set2 set_rev_mp)
+ then obtain uu where uu:
+ "\<forall> c \<in> supp P. \<forall> b1b2 \<in> (set1 c) \<times> (set2 c).
+ uu c b1b2 \<in> A \<and> p1 (uu c b1b2) = fst b1b2 \<and> p2 (uu c b1b2) = snd b1b2" by metis
+ def u \<equiv> "\<lambda> c b1 b2. uu c (b1,b2)"
+ have u[simp]:
+ "\<And> c b1 b2. \<lbrakk>c \<in> supp P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk> \<Longrightarrow> u c b1 b2 \<in> A"
+ "\<And> c b1 b2. \<lbrakk>c \<in> supp P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk> \<Longrightarrow> p1 (u c b1 b2) = b1"
+ "\<And> c b1 b2. \<lbrakk>c \<in> supp P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk> \<Longrightarrow> p2 (u c b1 b2) = b2"
+ using uu unfolding u_def by auto
+ {fix c assume c: "c \<in> supp P"
+ have "inj2 (u c) (set1 c) (set2 c)" unfolding inj2_def proof clarify
+ fix b1 b1' b2 b2'
+ assume "{b1, b1'} \<subseteq> set1 c" "{b2, b2'} \<subseteq> set2 c" and 0: "u c b1 b2 = u c b1' b2'"
+ hence "p1 (u c b1 b2) = b1 \<and> p2 (u c b1 b2) = b2 \<and>
+ p1 (u c b1' b2') = b1' \<and> p2 (u c b1' b2') = b2'"
+ using u(2)[OF c] u(3)[OF c] by simp metis
+ thus "b1 = b1' \<and> b2 = b2'" using 0 by auto
+ qed
+ } note inj = this
+ def sset \<equiv> "\<lambda> c. {u c b1 b2 | b1 b2. b1 \<in> set1 c \<and> b2 \<in> set2 c}"
+ have fin_sset[simp]: "\<And> c. c \<in> supp P \<Longrightarrow> finite (sset c)" unfolding sset_def
+ using fin_set1 fin_set2 finite_twosets by blast
+ have sset_A: "\<And> c. c \<in> supp P \<Longrightarrow> sset c \<subseteq> A" unfolding sset_def by auto
+ {fix c a assume c: "c \<in> supp P" and ac: "a \<in> sset c"
+ then obtain b1 b2 where b1: "b1 \<in> set1 c" and b2: "b2 \<in> set2 c"
+ and a: "a = u c b1 b2" unfolding sset_def by auto
+ have "p1 a \<in> set1 c" and p2a: "p2 a \<in> set2 c"
+ using ac a b1 b2 c u(2) u(3) by simp+
+ hence "u c (p1 a) (p2 a) = a" unfolding a using b1 b2 inj[OF c]
+ unfolding inj2_def by (metis c u(2) u(3))
+ } note u_p12[simp] = this
+ {fix c a assume c: "c \<in> supp P" and ac: "a \<in> sset c"
+ hence "p1 a \<in> set1 c" unfolding sset_def by auto
+ }note p1[simp] = this
+ {fix c a assume c: "c \<in> supp P" and ac: "a \<in> sset c"
+ hence "p2 a \<in> set2 c" unfolding sset_def by auto
+ }note p2[simp] = this
+ (* *)
+ {fix c assume c: "c \<in> supp P"
+ hence "\<exists> M. (\<forall> b1 \<in> set1 c. setsum (\<lambda> b2. M (u c b1 b2)) (set2 c) = N1 b1) \<and>
+ (\<forall> b2 \<in> set2 c. setsum (\<lambda> b1. M (u c b1 b2)) (set1 c) = N2 b2)"
+ unfolding sset_def
+ using matrix_setsum_finite[OF set1_NE[OF c] fin_set1[OF c]
+ set2_NE[OF c] fin_set2[OF c] inj[OF c] ss[OF c]] by auto
+ }
+ then obtain Ms where
+ ss1: "\<And> c b1. \<lbrakk>c \<in> supp P; b1 \<in> set1 c\<rbrakk> \<Longrightarrow>
+ setsum (\<lambda> b2. Ms c (u c b1 b2)) (set2 c) = N1 b1" and
+ ss2: "\<And> c b2. \<lbrakk>c \<in> supp P; b2 \<in> set2 c\<rbrakk> \<Longrightarrow>
+ setsum (\<lambda> b1. Ms c (u c b1 b2)) (set1 c) = N2 b2"
+ by metis
+ def SET \<equiv> "\<Union> c \<in> supp P. sset c"
+ have fin_SET[simp]: "finite SET" unfolding SET_def apply(rule finite_UN_I) by auto
+ have SET_A: "SET \<subseteq> A" unfolding SET_def using sset_A by auto
+ have u_SET[simp]: "\<And> c b1 b2. \<lbrakk>c \<in> supp P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk> \<Longrightarrow> u c b1 b2 \<in> SET"
+ unfolding SET_def sset_def by blast
+ {fix c a assume c: "c \<in> supp P" and a: "a \<in> SET" and p1a: "p1 a \<in> set1 c"
+ then obtain c' where c': "c' \<in> supp P" and ac': "a \<in> sset c'"
+ unfolding SET_def by auto
+ hence "p1 a \<in> set1 c'" unfolding sset_def by auto
+ hence eq: "c = c'" using p1a c c' set1_disj by auto
+ hence "a \<in> sset c" using ac' by simp
+ } note p1_rev = this
+ {fix c a assume c: "c \<in> supp P" and a: "a \<in> SET" and p2a: "p2 a \<in> set2 c"
+ then obtain c' where c': "c' \<in> supp P" and ac': "a \<in> sset c'"
+ unfolding SET_def by auto
+ hence "p2 a \<in> set2 c'" unfolding sset_def by auto
+ hence eq: "c = c'" using p2a c c' set2_disj by auto
+ hence "a \<in> sset c" using ac' by simp
+ } note p2_rev = this
+ (* *)
+ have "\<forall> a \<in> SET. \<exists> c \<in> supp P. a \<in> sset c" unfolding SET_def by auto
+ then obtain h where h: "\<forall> a \<in> SET. h a \<in> supp P \<and> a \<in> sset (h a)" by metis
+ have h_u[simp]: "\<And> c b1 b2. \<lbrakk>c \<in> supp P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk>
+ \<Longrightarrow> h (u c b1 b2) = c"
+ by (metis h p2 set2 u(3) u_SET)
+ have h_u1: "\<And> c b1 b2. \<lbrakk>c \<in> supp P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk>
+ \<Longrightarrow> h (u c b1 b2) = f1 b1"
+ using h unfolding sset_def by auto
+ have h_u2: "\<And> c b1 b2. \<lbrakk>c \<in> supp P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk>
+ \<Longrightarrow> h (u c b1 b2) = f2 b2"
+ using h unfolding sset_def by auto
+ def M \<equiv> "\<lambda> a. if a \<in> SET \<and> p1 a \<in> supp N1 \<and> p2 a \<in> supp N2 then Ms (h a) a else 0"
+ have sM: "supp M \<subseteq> SET" "supp M \<subseteq> p1 -` (supp N1)" "supp M \<subseteq> p2 -` (supp N2)"
+ unfolding M_def by auto
+ show "\<exists>M. (M \<in> multiset \<and> supp M \<subseteq> A) \<and> mmap p1 M = N1 \<and> mmap p2 M = N2"
+ proof(rule exI[of _ M], safe)
+ show "M \<in> multiset"
+ unfolding multiset_def using finite_subset[OF sM(1) fin_SET] by simp
+ next
+ fix a assume "0 < M a"
+ thus "a \<in> A" unfolding M_def using SET_A by (cases "a \<in> SET") auto
+ next
+ show "mmap p1 M = N1"
+ unfolding mmap_def[abs_def] proof(rule ext)
+ fix b1
+ let ?K = "{a. p1 a = b1 \<and> 0 < M a}"
+ show "setsum M ?K = N1 b1"
+ proof(cases "b1 \<in> supp N1")
+ case False
+ hence "?K = {}" using sM(2) by auto
+ thus ?thesis using False by auto
+ next
+ case True
+ def c \<equiv> "f1 b1"
+ have c: "c \<in> supp P" and b1: "b1 \<in> set1 c"
+ unfolding set1_def c_def P1 using True by (auto simp: mmap_image)
+ have "setsum M ?K = setsum M {a. p1 a = b1 \<and> a \<in> SET}"
+ apply(rule setsum_mono_zero_cong_left) unfolding M_def by auto
+ also have "... = setsum M ((\<lambda> b2. u c b1 b2) ` (set2 c))"
+ apply(rule setsum_cong) using c b1 proof safe
+ fix a assume p1a: "p1 a \<in> set1 c" and "0 < P c" and "a \<in> SET"
+ hence ac: "a \<in> sset c" using p1_rev by auto
+ hence "a = u c (p1 a) (p2 a)" using c by auto
+ moreover have "p2 a \<in> set2 c" using ac c by auto
+ ultimately show "a \<in> u c (p1 a) ` set2 c" by auto
+ next
+ fix b2 assume b1: "b1 \<in> set1 c" and b2: "b2 \<in> set2 c"
+ hence "u c b1 b2 \<in> SET" using c by auto
+ qed auto
+ also have "... = setsum (\<lambda> b2. M (u c b1 b2)) (set2 c)"
+ unfolding comp_def[symmetric] apply(rule setsum_reindex)
+ using inj unfolding inj_on_def inj2_def using b1 c u(3) by blast
+ also have "... = N1 b1" unfolding ss1[OF c b1, symmetric]
+ apply(rule setsum_cong[OF refl]) unfolding M_def
+ using True h_u[OF c b1] set2_def u(2,3)[OF c b1] u_SET[OF c b1] by fastforce
+ finally show ?thesis .
+ qed
+ qed
+ next
+ show "mmap p2 M = N2"
+ unfolding mmap_def[abs_def] proof(rule ext)
+ fix b2
+ let ?K = "{a. p2 a = b2 \<and> 0 < M a}"
+ show "setsum M ?K = N2 b2"
+ proof(cases "b2 \<in> supp N2")
+ case False
+ hence "?K = {}" using sM(3) by auto
+ thus ?thesis using False by auto
+ next
+ case True
+ def c \<equiv> "f2 b2"
+ have c: "c \<in> supp P" and b2: "b2 \<in> set2 c"
+ unfolding set2_def c_def P2 using True by (auto simp: mmap_image)
+ have "setsum M ?K = setsum M {a. p2 a = b2 \<and> a \<in> SET}"
+ apply(rule setsum_mono_zero_cong_left) unfolding M_def by auto
+ also have "... = setsum M ((\<lambda> b1. u c b1 b2) ` (set1 c))"
+ apply(rule setsum_cong) using c b2 proof safe
+ fix a assume p2a: "p2 a \<in> set2 c" and "0 < P c" and "a \<in> SET"
+ hence ac: "a \<in> sset c" using p2_rev by auto
+ hence "a = u c (p1 a) (p2 a)" using c by auto
+ moreover have "p1 a \<in> set1 c" using ac c by auto
+ ultimately show "a \<in> (\<lambda>b1. u c b1 (p2 a)) ` set1 c" by auto
+ next
+ fix b2 assume b1: "b1 \<in> set1 c" and b2: "b2 \<in> set2 c"
+ hence "u c b1 b2 \<in> SET" using c by auto
+ qed auto
+ also have "... = setsum (M o (\<lambda> b1. u c b1 b2)) (set1 c)"
+ apply(rule setsum_reindex)
+ using inj unfolding inj_on_def inj2_def using b2 c u(2) by blast
+ also have "... = setsum (\<lambda> b1. M (u c b1 b2)) (set1 c)"
+ unfolding comp_def[symmetric] by simp
+ also have "... = N2 b2" unfolding ss2[OF c b2, symmetric]
+ apply(rule setsum_cong[OF refl]) unfolding M_def set2_def
+ using True h_u1[OF c _ b2] u(2,3)[OF c _ b2] u_SET[OF c _ b2]
+ unfolding set1_def by fastforce
+ finally show ?thesis .
+ qed
+ qed
+ qed
+qed
+
+definition multiset_map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a multiset \<Rightarrow> 'b multiset" where
+"multiset_map h = Abs_multiset \<circ> mmap h \<circ> count"
+
+bnf_def multiset_map [set_of] "\<lambda>_::'a multiset. natLeq" ["{#}"]
+unfolding multiset_map_def
+proof -
+ show "Abs_multiset \<circ> mmap id \<circ> count = id" unfolding mmap_id by (auto simp: count_inverse)
+next
+ fix f g
+ show "Abs_multiset \<circ> mmap (g \<circ> f) \<circ> count =
+ Abs_multiset \<circ> mmap g \<circ> count \<circ> (Abs_multiset \<circ> mmap f \<circ> count)"
+ unfolding comp_def apply(rule ext)
+ by (auto simp: Abs_multiset_inverse count mmap_comp1 mmap)
+next
+ fix M f g assume eq: "\<And>a. a \<in> set_of M \<Longrightarrow> f a = g a"
+ thus "(Abs_multiset \<circ> mmap f \<circ> count) M = (Abs_multiset \<circ> mmap g \<circ> count) M" apply auto
+ unfolding cIm_def[abs_def] image_def
+ by (auto intro!: mmap_cong simp: Abs_multiset_inject count mmap)
+next
+ fix f show "set_of \<circ> (Abs_multiset \<circ> mmap f \<circ> count) = op ` f \<circ> set_of"
+ by (auto simp: count mmap mmap_image set_of_Abs_multiset supp_count)
+next
+ show "card_order natLeq" by (rule natLeq_card_order)
+next
+ show "cinfinite natLeq" by (rule natLeq_cinfinite)
+next
+ fix M show "|set_of M| \<le>o natLeq"
+ apply(rule ordLess_imp_ordLeq)
+ unfolding finite_iff_ordLess_natLeq[symmetric] using finite_set_of .
+next
+ fix A :: "'a set"
+ have "|{M. set_of M \<subseteq> A}| \<le>o |{as. set as \<subseteq> A}|" using card_of_set_of .
+ also have "|{as. set as \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq"
+ by (rule list_in_bd)
+ finally show "|{M. set_of M \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq" .
+next
+ fix A B1 B2 f1 f2 p1 p2
+ let ?map = "\<lambda> f. Abs_multiset \<circ> mmap f \<circ> count"
+ assume wp: "wpull A B1 B2 f1 f2 p1 p2"
+ show "wpull {x. set_of x \<subseteq> A} {x. set_of x \<subseteq> B1} {x. set_of x \<subseteq> B2}
+ (?map f1) (?map f2) (?map p1) (?map p2)"
+ unfolding wpull_def proof safe
+ fix y1 y2
+ assume y1: "set_of y1 \<subseteq> B1" and y2: "set_of y2 \<subseteq> B2"
+ and m: "?map f1 y1 = ?map f2 y2"
+ def N1 \<equiv> "count y1" def N2 \<equiv> "count y2"
+ have "N1 \<in> multiset \<and> supp N1 \<subseteq> B1" and "N2 \<in> multiset \<and> supp N2 \<subseteq> B2"
+ and "mmap f1 N1 = mmap f2 N2"
+ using y1 y2 m unfolding N1_def N2_def
+ by (auto simp: Abs_multiset_inject count mmap)
+ then obtain M where M: "M \<in> multiset \<and> supp M \<subseteq> A"
+ and N1: "mmap p1 M = N1" and N2: "mmap p2 M = N2"
+ using wp_mmap[OF wp] unfolding wpull_def by auto
+ def x \<equiv> "Abs_multiset M"
+ show "\<exists>x\<in>{x. set_of x \<subseteq> A}. ?map p1 x = y1 \<and> ?map p2 x = y2"
+ apply(intro bexI[of _ x]) using M N1 N2 unfolding N1_def N2_def x_def
+ by (auto simp: count_inverse Abs_multiset_inverse)
+ qed
+qed (unfold set_of_empty, auto)
+
+inductive multiset_rel' where
+Zero: "multiset_rel' R {#} {#}"
+|
+Plus: "\<lbrakk>R a b; multiset_rel' R M N\<rbrakk> \<Longrightarrow> multiset_rel' R (M + {#a#}) (N + {#b#})"
+
+lemma multiset_map_Zero_iff[simp]: "multiset_map f M = {#} \<longleftrightarrow> M = {#}"
+by (metis image_is_empty multiset.set_natural' set_of_eq_empty_iff)
+
+lemma multiset_map_Zero[simp]: "multiset_map f {#} = {#}" by simp
+
+lemma multiset_rel_Zero: "multiset_rel R {#} {#}"
+unfolding multiset_rel_def Gr_def relcomp_unfold by auto
+
+declare multiset.count[simp]
+declare mmap[simp]
+declare Abs_multiset_inverse[simp]
+declare multiset.count_inverse[simp]
+declare union_preserves_multiset[simp]
+
+lemma mmap_Plus[simp]:
+assumes "K \<in> multiset" and "L \<in> multiset"
+shows "mmap f (\<lambda>a. K a + L a) a = mmap f K a + mmap f L a"
+proof-
+ have "{aa. f aa = a \<and> (0 < K aa \<or> 0 < L aa)} \<subseteq>
+ {aa. 0 < K aa} \<union> {aa. 0 < L aa}" (is "?C \<subseteq> ?A \<union> ?B") by auto
+ moreover have "finite (?A \<union> ?B)" apply(rule finite_UnI)
+ using assms unfolding multiset_def by auto
+ ultimately have C: "finite ?C" using finite_subset by blast
+ have "setsum K {aa. f aa = a \<and> 0 < K aa} = setsum K {aa. f aa = a \<and> 0 < K aa + L aa}"
+ apply(rule setsum_mono_zero_cong_left) using C by auto
+ moreover
+ have "setsum L {aa. f aa = a \<and> 0 < L aa} = setsum L {aa. f aa = a \<and> 0 < K aa + L aa}"
+ apply(rule setsum_mono_zero_cong_left) using C by auto
+ ultimately show ?thesis
+ unfolding mmap_def unfolding comm_monoid_add_class.setsum.F_fun_f by auto
+qed
+
+lemma multiset_map_Plus[simp]:
+"multiset_map f (M1 + M2) = multiset_map f M1 + multiset_map f M2"
+unfolding multiset_map_def
+apply(subst multiset.count_inject[symmetric])
+unfolding plus_multiset.rep_eq comp_def by auto
+
+lemma multiset_map_singl[simp]: "multiset_map f {#a#} = {#f a#}"
+proof-
+ have 0: "\<And> b. card {aa. a = aa \<and> (a = aa \<longrightarrow> f aa = b)} =
+ (if b = f a then 1 else 0)" by auto
+ thus ?thesis
+ unfolding multiset_map_def comp_def mmap_def[abs_def] map_fun_def
+ by (simp, simp add: single_def)
+qed
+
+lemma multiset_rel_Plus:
+assumes ab: "R a b" and MN: "multiset_rel R M N"
+shows "multiset_rel R (M + {#a#}) (N + {#b#})"
+proof-
+ {fix y assume "R a b" and "set_of y \<subseteq> {(x, y). R x y}"
+ hence "\<exists>ya. multiset_map fst y + {#a#} = multiset_map fst ya \<and>
+ multiset_map snd y + {#b#} = multiset_map snd ya \<and>
+ set_of ya \<subseteq> {(x, y). R x y}"
+ apply(intro exI[of _ "y + {#(a,b)#}"]) by auto
+ }
+ thus ?thesis
+ using assms
+ unfolding multiset_rel_def Gr_def relcomp_unfold by force
+qed
+
+lemma multiset_rel'_imp_multiset_rel:
+"multiset_rel' R M N \<Longrightarrow> multiset_rel R M N"
+apply(induct rule: multiset_rel'.induct)
+using multiset_rel_Zero multiset_rel_Plus by auto
+
+lemma mcard_multiset_map[simp]: "mcard (multiset_map f M) = mcard M"
+proof-
+ def A \<equiv> "\<lambda> b. {a. f a = b \<and> a \<in># M}"
+ let ?B = "{b. 0 < setsum (count M) (A b)}"
+ have "{b. \<exists>a. f a = b \<and> a \<in># M} \<subseteq> f ` {a. a \<in># M}" by auto
+ moreover have "finite (f ` {a. a \<in># M})" apply(rule finite_imageI)
+ using finite_Collect_mem .
+ ultimately have fin: "finite {b. \<exists>a. f a = b \<and> a \<in># M}" by(rule finite_subset)
+ have i: "inj_on A ?B" unfolding inj_on_def A_def apply clarsimp
+ by (metis (lifting, mono_tags) mem_Collect_eq rel_simps(54)
+ setsum_gt_0_iff setsum_infinite)
+ have 0: "\<And> b. 0 < setsum (count M) (A b) \<longleftrightarrow> (\<exists> a \<in> A b. count M a > 0)"
+ apply safe
+ apply (metis less_not_refl setsum_gt_0_iff setsum_infinite)
+ by (metis A_def finite_Collect_conjI finite_Collect_mem setsum_gt_0_iff)
+ hence AB: "A ` ?B = {A b | b. \<exists> a \<in> A b. count M a > 0}" by auto
+
+ have "setsum (\<lambda> x. setsum (count M) (A x)) ?B = setsum (setsum (count M) o A) ?B"
+ unfolding comp_def ..
+ also have "... = (\<Sum>x\<in> A ` ?B. setsum (count M) x)"
+ unfolding comm_monoid_add_class.setsum_reindex[OF i, symmetric] ..
+ also have "... = setsum (count M) (\<Union>x\<in>A ` {b. 0 < setsum (count M) (A b)}. x)"
+ (is "_ = setsum (count M) ?J")
+ apply(rule comm_monoid_add_class.setsum_UN_disjoint[symmetric])
+ using 0 fin unfolding A_def by (auto intro!: finite_imageI)
+ also have "?J = {a. a \<in># M}" unfolding AB unfolding A_def by auto
+ finally have "setsum (\<lambda> x. setsum (count M) (A x)) ?B =
+ setsum (count M) {a. a \<in># M}" .
+ thus ?thesis unfolding A_def mcard_def multiset_map_def by (simp add: mmap_def)
+qed
+
+lemma multiset_rel_mcard:
+assumes "multiset_rel R M N"
+shows "mcard M = mcard N"
+using assms unfolding multiset_rel_def relcomp_unfold Gr_def by auto
+
+lemma multiset_induct2[case_names empty addL addR]:
+assumes empty: "P {#} {#}"
+and addL: "\<And>M N a. P M N \<Longrightarrow> P (M + {#a#}) N"
+and addR: "\<And>M N a. P M N \<Longrightarrow> P M (N + {#a#})"
+shows "P M N"
+apply(induct N rule: multiset_induct)
+ apply(induct M rule: multiset_induct, rule empty, erule addL)
+ apply(induct M rule: multiset_induct, erule addR, erule addR)
+done
+
+lemma multiset_induct2_mcard[consumes 1, case_names empty add]:
+assumes c: "mcard M = mcard N"
+and empty: "P {#} {#}"
+and add: "\<And>M N a b. P M N \<Longrightarrow> P (M + {#a#}) (N + {#b#})"
+shows "P M N"
+using c proof(induct M arbitrary: N rule: measure_induct_rule[of mcard])
+ case (less M) show ?case
+ proof(cases "M = {#}")
+ case True hence "N = {#}" using less.prems by auto
+ thus ?thesis using True empty by auto
+ next
+ case False then obtain M1 a where M: "M = M1 + {#a#}" by (metis multi_nonempty_split)
+ have "N \<noteq> {#}" using False less.prems by auto
+ then obtain N1 b where N: "N = N1 + {#b#}" by (metis multi_nonempty_split)
+ have "mcard M1 = mcard N1" using less.prems unfolding M N by auto
+ thus ?thesis using M N less.hyps add by auto
+ qed
+qed
+
+lemma msed_map_invL:
+assumes "multiset_map f (M + {#a#}) = N"
+shows "\<exists> N1. N = N1 + {#f a#} \<and> multiset_map f M = N1"
+proof-
+ have "f a \<in># N"
+ using assms multiset.set_natural'[of f "M + {#a#}"] by auto
+ then obtain N1 where N: "N = N1 + {#f a#}" using multi_member_split by metis
+ have "multiset_map f M = N1" using assms unfolding N by simp
+ thus ?thesis using N by blast
+qed
+
+lemma msed_map_invR:
+assumes "multiset_map f M = N + {#b#}"
+shows "\<exists> M1 a. M = M1 + {#a#} \<and> f a = b \<and> multiset_map f M1 = N"
+proof-
+ obtain a where a: "a \<in># M" and fa: "f a = b"
+ using multiset.set_natural'[of f M] unfolding assms
+ by (metis image_iff mem_set_of_iff union_single_eq_member)
+ then obtain M1 where M: "M = M1 + {#a#}" using multi_member_split by metis
+ have "multiset_map f M1 = N" using assms unfolding M fa[symmetric] by simp
+ thus ?thesis using M fa by blast
+qed
+
+lemma msed_rel_invL:
+assumes "multiset_rel R (M + {#a#}) N"
+shows "\<exists> N1 b. N = N1 + {#b#} \<and> R a b \<and> multiset_rel R M N1"
+proof-
+ obtain K where KM: "multiset_map fst K = M + {#a#}"
+ and KN: "multiset_map snd K = N" and sK: "set_of K \<subseteq> {(a, b). R a b}"
+ using assms
+ unfolding multiset_rel_def Gr_def relcomp_unfold by auto
+ obtain K1 ab where K: "K = K1 + {#ab#}" and a: "fst ab = a"
+ and K1M: "multiset_map fst K1 = M" using msed_map_invR[OF KM] by auto
+ obtain N1 where N: "N = N1 + {#snd ab#}" and K1N1: "multiset_map snd K1 = N1"
+ using msed_map_invL[OF KN[unfolded K]] by auto
+ have Rab: "R a (snd ab)" using sK a unfolding K by auto
+ have "multiset_rel R M N1" using sK K1M K1N1
+ unfolding K multiset_rel_def Gr_def relcomp_unfold by auto
+ thus ?thesis using N Rab by auto
+qed
+
+lemma msed_rel_invR:
+assumes "multiset_rel R M (N + {#b#})"
+shows "\<exists> M1 a. M = M1 + {#a#} \<and> R a b \<and> multiset_rel R M1 N"
+proof-
+ obtain K where KN: "multiset_map snd K = N + {#b#}"
+ and KM: "multiset_map fst K = M" and sK: "set_of K \<subseteq> {(a, b). R a b}"
+ using assms
+ unfolding multiset_rel_def Gr_def relcomp_unfold by auto
+ obtain K1 ab where K: "K = K1 + {#ab#}" and b: "snd ab = b"
+ and K1N: "multiset_map snd K1 = N" using msed_map_invR[OF KN] by auto
+ obtain M1 where M: "M = M1 + {#fst ab#}" and K1M1: "multiset_map fst K1 = M1"
+ using msed_map_invL[OF KM[unfolded K]] by auto
+ have Rab: "R (fst ab) b" using sK b unfolding K by auto
+ have "multiset_rel R M1 N" using sK K1N K1M1
+ unfolding K multiset_rel_def Gr_def relcomp_unfold by auto
+ thus ?thesis using M Rab by auto
+qed
+
+lemma multiset_rel_imp_multiset_rel':
+assumes "multiset_rel R M N"
+shows "multiset_rel' R M N"
+using assms proof(induct M arbitrary: N rule: measure_induct_rule[of mcard])
+ case (less M)
+ have c: "mcard M = mcard N" using multiset_rel_mcard[OF less.prems] .
+ show ?case
+ proof(cases "M = {#}")
+ case True hence "N = {#}" using c by simp
+ thus ?thesis using True multiset_rel'.Zero by auto
+ next
+ case False then obtain M1 a where M: "M = M1 + {#a#}" by (metis multi_nonempty_split)
+ obtain N1 b where N: "N = N1 + {#b#}" and R: "R a b" and ms: "multiset_rel R M1 N1"
+ using msed_rel_invL[OF less.prems[unfolded M]] by auto
+ have "multiset_rel' R M1 N1" using less.hyps[of M1 N1] ms unfolding M by simp
+ thus ?thesis using multiset_rel'.Plus[of R a b, OF R] unfolding M N by simp
+ qed
+qed
+
+lemma multiset_rel_multiset_rel':
+"multiset_rel R M N = multiset_rel' R M N"
+using multiset_rel_imp_multiset_rel' multiset_rel'_imp_multiset_rel by auto
+
+(* The main end product for multiset_rel: inductive characterization *)
+theorems multiset_rel_induct[case_names empty add, induct pred: multiset_rel] =
+ multiset_rel'.induct[unfolded multiset_rel_multiset_rel'[symmetric]]
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/README.html Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,54 @@
+<!DOCTYPE HTML PUBLIC "-//W3C//DTD HTML 4.01 Transitional//EN" "http://www.w3.org/TR/html4/loose.dtd">
+
+<html>
+
+<head>
+ <meta http-equiv="Content-Type" content="text/html; charset=iso-8859-1">
+ <title>BNF Package</title>
+</head>
+
+<body>
+
+<h3><i>BNF</i>: A (co)datatype package based on bounded natural functors
+(BNFs)</h3>
+
+<p>
+The <i>BNF</i> package provides a fully modular framework for constructing
+inductive and coinductive datatypes in HOL, with support for mixed mutual and
+nested (co)recursion. Mixed (co)recursion enables type definitions involving
+both datatypes and codatatypes, such as the type of finitely branching trees of
+possibly infinite depth. The framework draws heavily from category theory.
+
+<p>
+The package is described in the following paper:
+
+<ul>
+ <li><a href="http://www21.in.tum.de/~traytel/papers/codatatypes/index.html">Foundational, Compositional (Co)datatypes for Higher-Order Logic—Category Theory Applied to Theorem Proving</a>, <br>
+ Dmitriy Traytel, Andrei Popescu, and Jasmin Christian Blanchette.<br>
+ <i>Logic in Computer Science (LICS 2012)</i>, 2012.
+</ul>
+
+<p>
+The main entry point for applications is <tt>BNF.thy</tt>. The <tt>Examples</tt>
+directory contains various examples of (co)datatypes, including the examples
+from the paper.
+
+<p>
+The key notion underlying the package is that of a <i>bounded natural functor</i>
+(<i>BNF</i>)—an enriched type constructor satisfying specific properties
+preserved by interesting categorical operations (composition, least fixed point,
+and greatest fixed point). The <tt>Basic_BNFs.thy</tt> and <tt>More_BNFs.thy</tt>
+files register various basic types, notably for sums, products, function spaces,
+finite sets, multisets, and countable sets. Custom BNFs can be registered as well.
+
+<p>
+<b>Warning:</b> The package is under development. Please contact any nonempty
+subset of
+<a href="mailto:traytel@in.tum.de">the</a>
+<a href="mailto:popescua@in.tum.de">above</a>
+<a href="mailto:blanchette@in.tum.de">authors</a>
+if you have questions or comments.
+
+</body>
+
+</html>
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Tools/bnf_comp.ML Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,726 @@
+(* Title: HOL/BNF/Tools/bnf_comp.ML
+ Author: Dmitriy Traytel, TU Muenchen
+ Author: Jasmin Blanchette, TU Muenchen
+ Copyright 2012
+
+Composition of bounded natural functors.
+*)
+
+signature BNF_COMP =
+sig
+ type unfold_set
+ val empty_unfolds: unfold_set
+ val map_unfolds_of: unfold_set -> thm list
+ val rel_unfolds_of: unfold_set -> thm list
+ val set_unfoldss_of: unfold_set -> thm list list
+ val srel_unfolds_of: unfold_set -> thm list
+
+ val bnf_of_typ: BNF_Def.const_policy -> (binding -> binding) ->
+ ((string * sort) list list -> (string * sort) list) -> typ -> unfold_set * Proof.context ->
+ (BNF_Def.BNF * (typ list * typ list)) * (unfold_set * Proof.context)
+ val default_comp_sort: (string * sort) list list -> (string * sort) list
+ val normalize_bnfs: (int -> binding -> binding) -> ''a list list -> ''a list ->
+ (''a list list -> ''a list) -> BNF_Def.BNF list -> unfold_set -> Proof.context ->
+ (int list list * ''a list) * (BNF_Def.BNF list * (unfold_set * Proof.context))
+ val seal_bnf: unfold_set -> binding -> typ list -> BNF_Def.BNF -> Proof.context ->
+ (BNF_Def.BNF * typ list) * local_theory
+end;
+
+structure BNF_Comp : BNF_COMP =
+struct
+
+open BNF_Def
+open BNF_Util
+open BNF_Tactics
+open BNF_Comp_Tactics
+
+type unfold_set = {
+ map_unfolds: thm list,
+ set_unfoldss: thm list list,
+ rel_unfolds: thm list,
+ srel_unfolds: thm list
+};
+
+val empty_unfolds = {map_unfolds = [], set_unfoldss = [], rel_unfolds = [], srel_unfolds = []};
+
+fun add_to_thms thms new = thms |> not (Thm.is_reflexive new) ? insert Thm.eq_thm new;
+fun adds_to_thms thms news = insert (eq_set Thm.eq_thm) (no_reflexive news) thms;
+
+fun add_to_unfolds map sets rel srel
+ {map_unfolds, set_unfoldss, rel_unfolds, srel_unfolds} =
+ {map_unfolds = add_to_thms map_unfolds map,
+ set_unfoldss = adds_to_thms set_unfoldss sets,
+ rel_unfolds = add_to_thms rel_unfolds rel,
+ srel_unfolds = add_to_thms srel_unfolds srel};
+
+fun add_bnf_to_unfolds bnf =
+ add_to_unfolds (map_def_of_bnf bnf) (set_defs_of_bnf bnf) (rel_def_of_bnf bnf)
+ (srel_def_of_bnf bnf);
+
+val map_unfolds_of = #map_unfolds;
+val set_unfoldss_of = #set_unfoldss;
+val rel_unfolds_of = #rel_unfolds;
+val srel_unfolds_of = #srel_unfolds;
+
+val bdTN = "bdT";
+
+fun mk_killN n = "_kill" ^ string_of_int n;
+fun mk_liftN n = "_lift" ^ string_of_int n;
+fun mk_permuteN src dest =
+ "_permute_" ^ implode (map string_of_int src) ^ "_" ^ implode (map string_of_int dest);
+
+(*copied from Envir.expand_term_free*)
+fun expand_term_const defs =
+ let
+ val eqs = map ((fn ((x, U), u) => (x, (U, u))) o apfst dest_Const) defs;
+ val get = fn Const (x, _) => AList.lookup (op =) eqs x | _ => NONE;
+ in Envir.expand_term get end;
+
+fun clean_compose_bnf const_policy qualify b outer inners (unfold_set, lthy) =
+ let
+ val olive = live_of_bnf outer;
+ val onwits = nwits_of_bnf outer;
+ val odead = dead_of_bnf outer;
+ val inner = hd inners;
+ val ilive = live_of_bnf inner;
+ val ideads = map dead_of_bnf inners;
+ val inwitss = map nwits_of_bnf inners;
+
+ (* TODO: check olive = length inners > 0,
+ forall inner from inners. ilive = live,
+ forall inner from inners. idead = dead *)
+
+ val (oDs, lthy1) = apfst (map TFree)
+ (Variable.invent_types (replicate odead HOLogic.typeS) lthy);
+ val (Dss, lthy2) = apfst (map (map TFree))
+ (fold_map Variable.invent_types (map (fn n => replicate n HOLogic.typeS) ideads) lthy1);
+ val (Ass, lthy3) = apfst (replicate ilive o map TFree)
+ (Variable.invent_types (replicate ilive HOLogic.typeS) lthy2);
+ val As = if ilive > 0 then hd Ass else [];
+ val Ass_repl = replicate olive As;
+ val (Bs, _(*lthy4*)) = apfst (map TFree)
+ (Variable.invent_types (replicate ilive HOLogic.typeS) lthy3);
+ val Bss_repl = replicate olive Bs;
+
+ val ((((fs', Qs'), Asets), xs), _(*names_lthy*)) = lthy
+ |> apfst snd o mk_Frees' "f" (map2 (curry (op -->)) As Bs)
+ ||>> apfst snd o mk_Frees' "Q" (map2 mk_pred2T As Bs)
+ ||>> mk_Frees "A" (map HOLogic.mk_setT As)
+ ||>> mk_Frees "x" As;
+
+ val CAs = map3 mk_T_of_bnf Dss Ass_repl inners;
+ val CCA = mk_T_of_bnf oDs CAs outer;
+ val CBs = map3 mk_T_of_bnf Dss Bss_repl inners;
+ val outer_sets = mk_sets_of_bnf (replicate olive oDs) (replicate olive CAs) outer;
+ val inner_setss = map3 mk_sets_of_bnf (map (replicate ilive) Dss) (replicate olive Ass) inners;
+ val inner_bds = map3 mk_bd_of_bnf Dss Ass_repl inners;
+ val outer_bd = mk_bd_of_bnf oDs CAs outer;
+
+ (*%f1 ... fn. outer.map (inner_1.map f1 ... fn) ... (inner_m.map f1 ... fn)*)
+ val mapx = fold_rev Term.abs fs'
+ (Term.list_comb (mk_map_of_bnf oDs CAs CBs outer,
+ map2 (fn Ds => (fn f => Term.list_comb (f, map Bound (ilive - 1 downto 0))) o
+ mk_map_of_bnf Ds As Bs) Dss inners));
+ (*%Q1 ... Qn. outer.rel (inner_1.rel Q1 ... Qn) ... (inner_m.rel Q1 ... Qn)*)
+ val rel = fold_rev Term.abs Qs'
+ (Term.list_comb (mk_rel_of_bnf oDs CAs CBs outer,
+ map2 (fn Ds => (fn f => Term.list_comb (f, map Bound (ilive - 1 downto 0))) o
+ mk_rel_of_bnf Ds As Bs) Dss inners));
+
+ (*Union o collect {outer.set_1 ... outer.set_m} o outer.map inner_1.set_i ... inner_m.set_i*)
+ (*Union o collect {image inner_1.set_i o outer.set_1 ... image inner_m.set_i o outer.set_m}*)
+ fun mk_set i =
+ let
+ val (setTs, T) = `(replicate olive o HOLogic.mk_setT) (nth As i);
+ val outer_set = mk_collect
+ (mk_sets_of_bnf (replicate olive oDs) (replicate olive setTs) outer)
+ (mk_T_of_bnf oDs setTs outer --> HOLogic.mk_setT T);
+ val inner_sets = map (fn sets => nth sets i) inner_setss;
+ val outer_map = mk_map_of_bnf oDs CAs setTs outer;
+ val map_inner_sets = Term.list_comb (outer_map, inner_sets);
+ val collect_image = mk_collect
+ (map2 (fn f => fn set => HOLogic.mk_comp (mk_image f, set)) inner_sets outer_sets)
+ (CCA --> HOLogic.mk_setT T);
+ in
+ (Library.foldl1 HOLogic.mk_comp [mk_Union T, outer_set, map_inner_sets],
+ HOLogic.mk_comp (mk_Union T, collect_image))
+ end;
+
+ val (sets, sets_alt) = map_split mk_set (0 upto ilive - 1);
+
+ (*(inner_1.bd +c ... +c inner_m.bd) *c outer.bd*)
+ val bd = Term.absdummy CCA (mk_cprod (Library.foldr1 (uncurry mk_csum) inner_bds) outer_bd);
+
+ fun map_id_tac {context = ctxt, ...} =
+ let
+ (*order the theorems by reverse size to prevent bad interaction with nonconfluent rewrite
+ rules*)
+ val thms = (map map_id_of_bnf inners
+ |> map (`(Term.size_of_term o Thm.prop_of))
+ |> sort (rev_order o int_ord o pairself fst)
+ |> map snd) @ [map_id_of_bnf outer];
+ in
+ (EVERY' (map (fn thm => subst_tac ctxt [thm]) thms) THEN' rtac refl) 1
+ end;
+
+ fun map_comp_tac _ =
+ mk_comp_map_comp_tac (map_comp_of_bnf outer) (map_cong_of_bnf outer)
+ (map map_comp_of_bnf inners);
+
+ fun mk_single_set_natural_tac i _ =
+ mk_comp_set_natural_tac (map_comp_of_bnf outer) (map_cong_of_bnf outer)
+ (collect_set_natural_of_bnf outer)
+ (map ((fn thms => nth thms i) o set_natural_of_bnf) inners);
+
+ val set_natural_tacs = map mk_single_set_natural_tac (0 upto ilive - 1);
+
+ fun bd_card_order_tac _ =
+ mk_comp_bd_card_order_tac (map bd_card_order_of_bnf inners) (bd_card_order_of_bnf outer);
+
+ fun bd_cinfinite_tac _ =
+ mk_comp_bd_cinfinite_tac (bd_cinfinite_of_bnf inner) (bd_cinfinite_of_bnf outer);
+
+ val set_alt_thms =
+ if ! quick_and_dirty then
+ []
+ else
+ map (fn goal =>
+ Skip_Proof.prove lthy [] [] goal
+ (fn {context, ...} => (mk_comp_set_alt_tac context (collect_set_natural_of_bnf outer)))
+ |> Thm.close_derivation)
+ (map2 (curry (HOLogic.mk_Trueprop o HOLogic.mk_eq)) sets sets_alt);
+
+ fun map_cong_tac _ =
+ mk_comp_map_cong_tac set_alt_thms (map_cong_of_bnf outer) (map map_cong_of_bnf inners);
+
+ val set_bd_tacs =
+ if ! quick_and_dirty then
+ replicate (length set_alt_thms) (K all_tac)
+ else
+ let
+ val outer_set_bds = set_bd_of_bnf outer;
+ val inner_set_bdss = map set_bd_of_bnf inners;
+ val inner_bd_Card_orders = map bd_Card_order_of_bnf inners;
+ fun single_set_bd_thm i j =
+ @{thm comp_single_set_bd} OF [nth inner_bd_Card_orders j, nth (nth inner_set_bdss j) i,
+ nth outer_set_bds j]
+ val single_set_bd_thmss =
+ map ((fn f => map f (0 upto olive - 1)) o single_set_bd_thm) (0 upto ilive - 1);
+ in
+ map2 (fn set_alt => fn single_set_bds => fn {context, ...} =>
+ mk_comp_set_bd_tac context set_alt single_set_bds)
+ set_alt_thms single_set_bd_thmss
+ end;
+
+ val in_alt_thm =
+ let
+ val inx = mk_in Asets sets CCA;
+ val in_alt = mk_in (map2 (mk_in Asets) inner_setss CAs) outer_sets CCA;
+ val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt));
+ in
+ Skip_Proof.prove lthy [] [] goal
+ (fn {context, ...} => mk_comp_in_alt_tac context set_alt_thms)
+ |> Thm.close_derivation
+ end;
+
+ fun in_bd_tac _ =
+ mk_comp_in_bd_tac in_alt_thm (map in_bd_of_bnf inners) (in_bd_of_bnf outer)
+ (map bd_Cinfinite_of_bnf inners) (bd_Card_order_of_bnf outer);
+
+ fun map_wpull_tac _ =
+ mk_map_wpull_tac in_alt_thm (map map_wpull_of_bnf inners) (map_wpull_of_bnf outer);
+
+ fun srel_O_Gr_tac _ =
+ let
+ val basic_thms = @{thms mem_Collect_eq fst_conv snd_conv}; (*TODO: tune*)
+ val outer_srel_Gr = srel_Gr_of_bnf outer RS sym;
+ val outer_srel_cong = srel_cong_of_bnf outer;
+ val thm =
+ (trans OF [in_alt_thm RS @{thm subst_rel_def},
+ trans OF [@{thm arg_cong2[of _ _ _ _ relcomp]} OF
+ [trans OF [outer_srel_Gr RS @{thm arg_cong[of _ _ converse]},
+ srel_converse_of_bnf outer RS sym], outer_srel_Gr],
+ trans OF [srel_O_of_bnf outer RS sym, outer_srel_cong OF
+ (map (fn bnf => srel_O_Gr_of_bnf bnf RS sym) inners)]]] RS sym)
+ |> unfold_thms lthy (basic_thms @ srel_def_of_bnf outer :: map srel_def_of_bnf inners);
+ in
+ unfold_thms_tac lthy basic_thms THEN rtac thm 1
+ end;
+
+ val tacs = zip_axioms map_id_tac map_comp_tac map_cong_tac set_natural_tacs bd_card_order_tac
+ bd_cinfinite_tac set_bd_tacs in_bd_tac map_wpull_tac srel_O_Gr_tac;
+
+ val outer_wits = mk_wits_of_bnf (replicate onwits oDs) (replicate onwits CAs) outer;
+
+ val inner_witss = map (map (fn (I, wit) => Term.list_comb (wit, map (nth xs) I)))
+ (map3 (fn Ds => fn n => mk_wits_of_bnf (replicate n Ds) (replicate n As))
+ Dss inwitss inners);
+
+ val inner_witsss = map (map (nth inner_witss) o fst) outer_wits;
+
+ val wits = (inner_witsss, (map (single o snd) outer_wits))
+ |-> map2 (fold (map_product (fn iwit => fn owit => owit $ iwit)))
+ |> flat
+ |> map (`(fn t => Term.add_frees t []))
+ |> minimize_wits
+ |> map (fn (frees, t) => fold absfree frees t);
+
+ fun wit_tac {context = ctxt, ...} =
+ mk_comp_wit_tac ctxt (wit_thms_of_bnf outer) (collect_set_natural_of_bnf outer)
+ (maps wit_thms_of_bnf inners);
+
+ val (bnf', lthy') =
+ bnf_def const_policy (K Derive_Few_Facts) qualify tacs wit_tac (SOME (oDs @ flat Dss))
+ (((((b, mapx), sets), bd), wits), SOME rel) lthy;
+ in
+ (bnf', (add_bnf_to_unfolds bnf' unfold_set, lthy'))
+ end;
+
+(* Killing live variables *)
+
+fun kill_bnf qualify n bnf (unfold_set, lthy) = if n = 0 then (bnf, (unfold_set, lthy)) else
+ let
+ val b = Binding.suffix_name (mk_killN n) (name_of_bnf bnf);
+ val live = live_of_bnf bnf;
+ val dead = dead_of_bnf bnf;
+ val nwits = nwits_of_bnf bnf;
+
+ (* TODO: check 0 < n <= live *)
+
+ val (Ds, lthy1) = apfst (map TFree)
+ (Variable.invent_types (replicate dead HOLogic.typeS) lthy);
+ val ((killedAs, As), lthy2) = apfst (`(take n) o map TFree)
+ (Variable.invent_types (replicate live HOLogic.typeS) lthy1);
+ val (Bs, _(*lthy3*)) = apfst (append killedAs o map TFree)
+ (Variable.invent_types (replicate (live - n) HOLogic.typeS) lthy2);
+
+ val ((Asets, lives), _(*names_lthy*)) = lthy
+ |> mk_Frees "A" (map HOLogic.mk_setT (drop n As))
+ ||>> mk_Frees "x" (drop n As);
+ val xs = map (fn T => HOLogic.choice_const T $ absdummy T @{term True}) killedAs @ lives;
+
+ val T = mk_T_of_bnf Ds As bnf;
+
+ (*bnf.map id ... id*)
+ val mapx = Term.list_comb (mk_map_of_bnf Ds As Bs bnf, map HOLogic.id_const killedAs);
+ (*bnf.rel (op =) ... (op =)*)
+ val rel = Term.list_comb (mk_rel_of_bnf Ds As Bs bnf, map HOLogic.eq_const killedAs);
+
+ val bnf_sets = mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf;
+ val sets = drop n bnf_sets;
+
+ (*(|UNIV :: A1 set| +c ... +c |UNIV :: An set|) *c bnf.bd*)
+ val bnf_bd = mk_bd_of_bnf Ds As bnf;
+ val bd = mk_cprod
+ (Library.foldr1 (uncurry mk_csum) (map (mk_card_of o HOLogic.mk_UNIV) killedAs)) bnf_bd;
+
+ fun map_id_tac _ = rtac (map_id_of_bnf bnf) 1;
+ fun map_comp_tac {context, ...} =
+ unfold_thms_tac context ((map_comp_of_bnf bnf RS sym) :: @{thms o_assoc id_o o_id}) THEN
+ rtac refl 1;
+ fun map_cong_tac {context, ...} =
+ mk_kill_map_cong_tac context n (live - n) (map_cong_of_bnf bnf);
+ val set_natural_tacs = map (fn thm => fn _ => rtac thm 1) (drop n (set_natural_of_bnf bnf));
+ fun bd_card_order_tac _ = mk_kill_bd_card_order_tac n (bd_card_order_of_bnf bnf);
+ fun bd_cinfinite_tac _ = mk_kill_bd_cinfinite_tac (bd_Cinfinite_of_bnf bnf);
+ val set_bd_tacs =
+ map (fn thm => fn _ => mk_kill_set_bd_tac (bd_Card_order_of_bnf bnf) thm)
+ (drop n (set_bd_of_bnf bnf));
+
+ val in_alt_thm =
+ let
+ val inx = mk_in Asets sets T;
+ val in_alt = mk_in (map HOLogic.mk_UNIV killedAs @ Asets) bnf_sets T;
+ val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt));
+ in
+ Skip_Proof.prove lthy [] [] goal (K kill_in_alt_tac) |> Thm.close_derivation
+ end;
+
+ fun in_bd_tac _ =
+ mk_kill_in_bd_tac n (live > n) in_alt_thm (in_bd_of_bnf bnf) (bd_Card_order_of_bnf bnf)
+ (bd_Cinfinite_of_bnf bnf) (bd_Cnotzero_of_bnf bnf);
+ fun map_wpull_tac _ = mk_map_wpull_tac in_alt_thm [] (map_wpull_of_bnf bnf);
+
+ fun srel_O_Gr_tac _ =
+ let
+ val srel_Gr = srel_Gr_of_bnf bnf RS sym
+ val thm =
+ (trans OF [in_alt_thm RS @{thm subst_rel_def},
+ trans OF [@{thm arg_cong2[of _ _ _ _ relcomp]} OF
+ [trans OF [srel_Gr RS @{thm arg_cong[of _ _ converse]},
+ srel_converse_of_bnf bnf RS sym], srel_Gr],
+ trans OF [srel_O_of_bnf bnf RS sym, srel_cong_of_bnf bnf OF
+ (replicate n @{thm trans[OF Gr_UNIV_id[OF refl] Id_alt[symmetric]]} @
+ replicate (live - n) @{thm Gr_fst_snd})]]] RS sym)
+ |> unfold_thms lthy (srel_def_of_bnf bnf :: @{thms Id_def' mem_Collect_eq split_conv});
+ in
+ rtac thm 1
+ end;
+
+ val tacs = zip_axioms map_id_tac map_comp_tac map_cong_tac set_natural_tacs bd_card_order_tac
+ bd_cinfinite_tac set_bd_tacs in_bd_tac map_wpull_tac srel_O_Gr_tac;
+
+ val bnf_wits = mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf;
+
+ val wits = map (fn t => fold absfree (Term.add_frees t []) t)
+ (map (fn (I, wit) => Term.list_comb (wit, map (nth xs) I)) bnf_wits);
+
+ fun wit_tac _ = mk_simple_wit_tac (wit_thms_of_bnf bnf);
+
+ val (bnf', lthy') =
+ bnf_def Smart_Inline (K Derive_Few_Facts) qualify tacs wit_tac (SOME (killedAs @ Ds))
+ (((((b, mapx), sets), Term.absdummy T bd), wits), SOME rel) lthy;
+ in
+ (bnf', (add_bnf_to_unfolds bnf' unfold_set, lthy'))
+ end;
+
+(* Adding dummy live variables *)
+
+fun lift_bnf qualify n bnf (unfold_set, lthy) = if n = 0 then (bnf, (unfold_set, lthy)) else
+ let
+ val b = Binding.suffix_name (mk_liftN n) (name_of_bnf bnf);
+ val live = live_of_bnf bnf;
+ val dead = dead_of_bnf bnf;
+ val nwits = nwits_of_bnf bnf;
+
+ (* TODO: check 0 < n *)
+
+ val (Ds, lthy1) = apfst (map TFree)
+ (Variable.invent_types (replicate dead HOLogic.typeS) lthy);
+ val ((newAs, As), lthy2) = apfst (chop n o map TFree)
+ (Variable.invent_types (replicate (n + live) HOLogic.typeS) lthy1);
+ val ((newBs, Bs), _(*lthy3*)) = apfst (chop n o map TFree)
+ (Variable.invent_types (replicate (n + live) HOLogic.typeS) lthy2);
+
+ val (Asets, _(*names_lthy*)) = lthy
+ |> mk_Frees "A" (map HOLogic.mk_setT (newAs @ As));
+
+ val T = mk_T_of_bnf Ds As bnf;
+
+ (*%f1 ... fn. bnf.map*)
+ val mapx =
+ fold_rev Term.absdummy (map2 (curry (op -->)) newAs newBs) (mk_map_of_bnf Ds As Bs bnf);
+ (*%Q1 ... Qn. bnf.rel*)
+ val rel = fold_rev Term.absdummy (map2 mk_pred2T newAs newBs) (mk_rel_of_bnf Ds As Bs bnf);
+
+ val bnf_sets = mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf;
+ val sets = map (fn A => absdummy T (HOLogic.mk_set A [])) newAs @ bnf_sets;
+
+ val bd = mk_bd_of_bnf Ds As bnf;
+
+ fun map_id_tac _ = rtac (map_id_of_bnf bnf) 1;
+ fun map_comp_tac {context, ...} =
+ unfold_thms_tac context ((map_comp_of_bnf bnf RS sym) :: @{thms o_assoc id_o o_id}) THEN
+ rtac refl 1;
+ fun map_cong_tac {context, ...} =
+ rtac (map_cong_of_bnf bnf) 1 THEN REPEAT_DETERM_N live (Goal.assume_rule_tac context 1);
+ val set_natural_tacs =
+ if ! quick_and_dirty then
+ replicate (n + live) (K all_tac)
+ else
+ replicate n (K empty_natural_tac) @
+ map (fn thm => fn _ => rtac thm 1) (set_natural_of_bnf bnf);
+ fun bd_card_order_tac _ = rtac (bd_card_order_of_bnf bnf) 1;
+ fun bd_cinfinite_tac _ = rtac (bd_cinfinite_of_bnf bnf) 1;
+ val set_bd_tacs =
+ if ! quick_and_dirty then
+ replicate (n + live) (K all_tac)
+ else
+ replicate n (K (mk_lift_set_bd_tac (bd_Card_order_of_bnf bnf))) @
+ (map (fn thm => fn _ => rtac thm 1) (set_bd_of_bnf bnf));
+
+ val in_alt_thm =
+ let
+ val inx = mk_in Asets sets T;
+ val in_alt = mk_in (drop n Asets) bnf_sets T;
+ val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt));
+ in
+ Skip_Proof.prove lthy [] [] goal (K lift_in_alt_tac) |> Thm.close_derivation
+ end;
+
+ fun in_bd_tac _ = mk_lift_in_bd_tac n in_alt_thm (in_bd_of_bnf bnf) (bd_Card_order_of_bnf bnf);
+ fun map_wpull_tac _ = mk_map_wpull_tac in_alt_thm [] (map_wpull_of_bnf bnf);
+
+ fun srel_O_Gr_tac _ =
+ mk_simple_srel_O_Gr_tac lthy (srel_def_of_bnf bnf) (srel_O_Gr_of_bnf bnf) in_alt_thm;
+
+ val tacs = zip_axioms map_id_tac map_comp_tac map_cong_tac set_natural_tacs bd_card_order_tac
+ bd_cinfinite_tac set_bd_tacs in_bd_tac map_wpull_tac srel_O_Gr_tac;
+
+ val wits = map snd (mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf);
+
+ fun wit_tac _ = mk_simple_wit_tac (wit_thms_of_bnf bnf);
+
+ val (bnf', lthy') =
+ bnf_def Smart_Inline (K Derive_Few_Facts) qualify tacs wit_tac (SOME Ds)
+ (((((b, mapx), sets), Term.absdummy T bd), wits), SOME rel) lthy;
+
+ in
+ (bnf', (add_bnf_to_unfolds bnf' unfold_set, lthy'))
+ end;
+
+(* Changing the order of live variables *)
+
+fun permute_bnf qualify src dest bnf (unfold_set, lthy) =
+ if src = dest then (bnf, (unfold_set, lthy)) else
+ let
+ val b = Binding.suffix_name (mk_permuteN src dest) (name_of_bnf bnf);
+ val live = live_of_bnf bnf;
+ val dead = dead_of_bnf bnf;
+ val nwits = nwits_of_bnf bnf;
+ fun permute xs = mk_permute src dest xs;
+ fun permute_rev xs = mk_permute dest src xs;
+
+ val (Ds, lthy1) = apfst (map TFree)
+ (Variable.invent_types (replicate dead HOLogic.typeS) lthy);
+ val (As, lthy2) = apfst (map TFree)
+ (Variable.invent_types (replicate live HOLogic.typeS) lthy1);
+ val (Bs, _(*lthy3*)) = apfst (map TFree)
+ (Variable.invent_types (replicate live HOLogic.typeS) lthy2);
+
+ val (Asets, _(*names_lthy*)) = lthy
+ |> mk_Frees "A" (map HOLogic.mk_setT (permute As));
+
+ val T = mk_T_of_bnf Ds As bnf;
+
+ (*%f(1) ... f(n). bnf.map f\<sigma>(1) ... f\<sigma>(n)*)
+ val mapx = fold_rev Term.absdummy (permute (map2 (curry op -->) As Bs))
+ (Term.list_comb (mk_map_of_bnf Ds As Bs bnf, permute_rev (map Bound (live - 1 downto 0))));
+ (*%Q(1) ... Q(n). bnf.rel Q\<sigma>(1) ... Q\<sigma>(n)*)
+ val rel = fold_rev Term.absdummy (permute (map2 mk_pred2T As Bs))
+ (Term.list_comb (mk_rel_of_bnf Ds As Bs bnf, permute_rev (map Bound (live - 1 downto 0))));
+
+ val bnf_sets = mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf;
+ val sets = permute bnf_sets;
+
+ val bd = mk_bd_of_bnf Ds As bnf;
+
+ fun map_id_tac _ = rtac (map_id_of_bnf bnf) 1;
+ fun map_comp_tac _ = rtac (map_comp_of_bnf bnf) 1;
+ fun map_cong_tac {context, ...} =
+ rtac (map_cong_of_bnf bnf) 1 THEN REPEAT_DETERM_N live (Goal.assume_rule_tac context 1);
+ val set_natural_tacs = permute (map (fn thm => fn _ => rtac thm 1) (set_natural_of_bnf bnf));
+ fun bd_card_order_tac _ = rtac (bd_card_order_of_bnf bnf) 1;
+ fun bd_cinfinite_tac _ = rtac (bd_cinfinite_of_bnf bnf) 1;
+ val set_bd_tacs = permute (map (fn thm => fn _ => rtac thm 1) (set_bd_of_bnf bnf));
+
+ val in_alt_thm =
+ let
+ val inx = mk_in Asets sets T;
+ val in_alt = mk_in (permute_rev Asets) bnf_sets T;
+ val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt));
+ in
+ Skip_Proof.prove lthy [] [] goal (K (mk_permute_in_alt_tac src dest))
+ |> Thm.close_derivation
+ end;
+
+ fun in_bd_tac _ =
+ mk_permute_in_bd_tac src dest in_alt_thm (in_bd_of_bnf bnf) (bd_Card_order_of_bnf bnf);
+ fun map_wpull_tac _ = mk_map_wpull_tac in_alt_thm [] (map_wpull_of_bnf bnf);
+
+ fun srel_O_Gr_tac _ =
+ mk_simple_srel_O_Gr_tac lthy (srel_def_of_bnf bnf) (srel_O_Gr_of_bnf bnf) in_alt_thm;
+
+ val tacs = zip_axioms map_id_tac map_comp_tac map_cong_tac set_natural_tacs bd_card_order_tac
+ bd_cinfinite_tac set_bd_tacs in_bd_tac map_wpull_tac srel_O_Gr_tac;
+
+ val wits = map snd (mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf);
+
+ fun wit_tac _ = mk_simple_wit_tac (wit_thms_of_bnf bnf);
+
+ val (bnf', lthy') =
+ bnf_def Smart_Inline (K Derive_Few_Facts) qualify tacs wit_tac (SOME Ds)
+ (((((b, mapx), sets), Term.absdummy T bd), wits), SOME rel) lthy;
+ in
+ (bnf', (add_bnf_to_unfolds bnf' unfold_set, lthy'))
+ end;
+
+(* Composition pipeline *)
+
+fun permute_and_kill qualify n src dest bnf =
+ bnf
+ |> permute_bnf qualify src dest
+ #> uncurry (kill_bnf qualify n);
+
+fun lift_and_permute qualify n src dest bnf =
+ bnf
+ |> lift_bnf qualify n
+ #> uncurry (permute_bnf qualify src dest);
+
+fun normalize_bnfs qualify Ass Ds sort bnfs unfold_set lthy =
+ let
+ val before_kill_src = map (fn As => 0 upto (length As - 1)) Ass;
+ val kill_poss = map (find_indices Ds) Ass;
+ val live_poss = map2 (subtract (op =)) kill_poss before_kill_src;
+ val before_kill_dest = map2 append kill_poss live_poss;
+ val kill_ns = map length kill_poss;
+ val (inners', (unfold_set', lthy')) =
+ fold_map5 (fn i => permute_and_kill (qualify i))
+ (if length bnfs = 1 then [0] else (1 upto length bnfs))
+ kill_ns before_kill_src before_kill_dest bnfs (unfold_set, lthy);
+
+ val Ass' = map2 (map o nth) Ass live_poss;
+ val As = sort Ass';
+ val after_lift_dest = replicate (length Ass') (0 upto (length As - 1));
+ val old_poss = map (map (fn x => find_index (fn y => x = y) As)) Ass';
+ val new_poss = map2 (subtract (op =)) old_poss after_lift_dest;
+ val after_lift_src = map2 append new_poss old_poss;
+ val lift_ns = map (fn xs => length As - length xs) Ass';
+ in
+ ((kill_poss, As), fold_map5 (fn i => lift_and_permute (qualify i))
+ (if length bnfs = 1 then [0] else (1 upto length bnfs))
+ lift_ns after_lift_src after_lift_dest inners' (unfold_set', lthy'))
+ end;
+
+fun default_comp_sort Ass =
+ Library.sort (Term_Ord.typ_ord o pairself TFree) (fold (fold (insert (op =))) Ass []);
+
+fun compose_bnf const_policy qualify sort outer inners oDs Dss tfreess (unfold_set, lthy) =
+ let
+ val b = name_of_bnf outer;
+
+ val Ass = map (map Term.dest_TFree) tfreess;
+ val Ds = fold (fold Term.add_tfreesT) (oDs :: Dss) [];
+
+ val ((kill_poss, As), (inners', (unfold_set', lthy'))) =
+ normalize_bnfs qualify Ass Ds sort inners unfold_set lthy;
+
+ val Ds = oDs @ flat (map3 (append oo map o nth) tfreess kill_poss Dss);
+ val As = map TFree As;
+ in
+ apfst (rpair (Ds, As))
+ (clean_compose_bnf const_policy (qualify 0) b outer inners' (unfold_set', lthy'))
+ end;
+
+(* Hide the type of the bound (optimization) and unfold the definitions (nicer to the user) *)
+
+fun seal_bnf unfold_set b Ds bnf lthy =
+ let
+ val live = live_of_bnf bnf;
+ val nwits = nwits_of_bnf bnf;
+
+ val (As, lthy1) = apfst (map TFree)
+ (Variable.invent_types (replicate live HOLogic.typeS) (fold Variable.declare_typ Ds lthy));
+ val (Bs, _) = apfst (map TFree)
+ (Variable.invent_types (replicate live HOLogic.typeS) lthy1);
+
+ val map_unfolds = map_unfolds_of unfold_set;
+ val set_unfoldss = set_unfoldss_of unfold_set;
+ val rel_unfolds = rel_unfolds_of unfold_set;
+ val srel_unfolds = srel_unfolds_of unfold_set;
+
+ val expand_maps =
+ fold expand_term_const (map (single o Logic.dest_equals o Thm.prop_of) map_unfolds);
+ val expand_sets =
+ fold expand_term_const (map (map (Logic.dest_equals o Thm.prop_of)) set_unfoldss);
+ val expand_rels =
+ fold expand_term_const (map (single o Logic.dest_equals o Thm.prop_of) rel_unfolds);
+ val unfold_maps = fold (unfold_thms lthy o single) map_unfolds;
+ val unfold_sets = fold (unfold_thms lthy) set_unfoldss;
+ val unfold_rels = unfold_thms lthy rel_unfolds;
+ val unfold_srels = unfold_thms lthy srel_unfolds;
+ val unfold_all = unfold_sets o unfold_maps o unfold_rels o unfold_srels;
+ val bnf_map = expand_maps (mk_map_of_bnf Ds As Bs bnf);
+ val bnf_sets = map (expand_maps o expand_sets)
+ (mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf);
+ val bnf_bd = mk_bd_of_bnf Ds As bnf;
+ val bnf_rel = expand_rels (mk_rel_of_bnf Ds As Bs bnf);
+ val T = mk_T_of_bnf Ds As bnf;
+
+ (*bd should only depend on dead type variables!*)
+ val bd_repT = fst (dest_relT (fastype_of bnf_bd));
+ val bdT_bind = Binding.suffix_name ("_" ^ bdTN) b;
+ val params = fold Term.add_tfreesT Ds [];
+ val deads = map TFree params;
+
+ val ((bdT_name, (bdT_glob_info, bdT_loc_info)), lthy) =
+ typedef false NONE (bdT_bind, params, NoSyn)
+ (HOLogic.mk_UNIV bd_repT) NONE (EVERY' [rtac exI, rtac UNIV_I] 1) lthy;
+
+ val bnf_bd' = mk_dir_image bnf_bd
+ (Const (#Abs_name bdT_glob_info, bd_repT --> Type (bdT_name, deads)))
+
+ val Abs_bdT_inj = mk_Abs_inj_thm (#Abs_inject bdT_loc_info);
+ val Abs_bdT_bij = mk_Abs_bij_thm lthy Abs_bdT_inj (#Abs_cases bdT_loc_info);
+
+ val bd_ordIso = @{thm dir_image} OF [Abs_bdT_inj, bd_Card_order_of_bnf bnf];
+ val bd_card_order =
+ @{thm card_order_dir_image} OF [Abs_bdT_bij, bd_card_order_of_bnf bnf];
+ val bd_cinfinite =
+ (@{thm Cinfinite_cong} OF [bd_ordIso, bd_Cinfinite_of_bnf bnf]) RS conjunct1;
+
+ val set_bds =
+ map (fn thm => @{thm ordLeq_ordIso_trans} OF [thm, bd_ordIso]) (set_bd_of_bnf bnf);
+ val in_bd =
+ @{thm ordLeq_ordIso_trans} OF [in_bd_of_bnf bnf,
+ @{thm cexp_cong2_Cnotzero} OF [bd_ordIso, if live = 0 then
+ @{thm ctwo_Cnotzero} else @{thm ctwo_Cnotzero} RS @{thm csum_Cnotzero2},
+ bd_Card_order_of_bnf bnf]];
+
+ fun mk_tac thm {context = ctxt, prems = _} =
+ (rtac (unfold_all thm) THEN'
+ SOLVE o REPEAT_DETERM o (atac ORELSE' Goal.assume_rule_tac ctxt)) 1;
+
+ val tacs = zip_axioms (mk_tac (map_id_of_bnf bnf)) (mk_tac (map_comp_of_bnf bnf))
+ (mk_tac (map_cong_of_bnf bnf)) (map mk_tac (set_natural_of_bnf bnf))
+ (K (rtac bd_card_order 1)) (K (rtac bd_cinfinite 1)) (map mk_tac set_bds) (mk_tac in_bd)
+ (mk_tac (map_wpull_of_bnf bnf))
+ (mk_tac (unfold_thms lthy [srel_def_of_bnf bnf] (srel_O_Gr_of_bnf bnf)));
+
+ val bnf_wits = map snd (mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf);
+
+ fun wit_tac _ = mk_simple_wit_tac (map unfold_all (wit_thms_of_bnf bnf));
+
+ val policy = user_policy Derive_All_Facts;
+
+ val (bnf', lthy') = bnf_def Hardly_Inline policy I tacs wit_tac (SOME deads)
+ (((((b, bnf_map), bnf_sets), Term.absdummy T bnf_bd'), bnf_wits), SOME bnf_rel) lthy;
+ in
+ ((bnf', deads), lthy')
+ end;
+
+val ID_bnf = the (bnf_of @{context} "Basic_BNFs.ID");
+val DEADID_bnf = the (bnf_of @{context} "Basic_BNFs.DEADID");
+
+fun bnf_of_typ _ _ _ (T as TFree _) accum = ((ID_bnf, ([], [T])), accum)
+ | bnf_of_typ _ _ _ (TVar _) _ = error "Unexpected schematic variable"
+ | bnf_of_typ const_policy qualify' sort (T as Type (C, Ts)) (unfold_set, lthy) =
+ let
+ val tfrees = Term.add_tfreesT T [];
+ val bnf_opt = if null tfrees then NONE else bnf_of lthy C;
+ in
+ (case bnf_opt of
+ NONE => ((DEADID_bnf, ([T], [])), (unfold_set, lthy))
+ | SOME bnf =>
+ if forall (can Term.dest_TFree) Ts andalso length Ts = length tfrees then
+ let
+ val T' = T_of_bnf bnf;
+ val deads = deads_of_bnf bnf;
+ val lives = lives_of_bnf bnf;
+ val tvars' = Term.add_tvarsT T' [];
+ val deads_lives =
+ pairself (map (Term.typ_subst_TVars (map fst tvars' ~~ map TFree tfrees)))
+ (deads, lives);
+ in ((bnf, deads_lives), (unfold_set, lthy)) end
+ else
+ let
+ val name = Long_Name.base_name C;
+ fun qualify i =
+ let val namei = name ^ nonzero_string_of_int i;
+ in qualify' o Binding.qualify true namei end;
+ val odead = dead_of_bnf bnf;
+ val olive = live_of_bnf bnf;
+ val oDs_pos = find_indices [TFree ("dead", [])] (snd (Term.dest_Type
+ (mk_T_of_bnf (replicate odead (TFree ("dead", []))) (replicate olive dummyT) bnf)));
+ val oDs = map (nth Ts) oDs_pos;
+ val Ts' = map (nth Ts) (subtract (op =) oDs_pos (0 upto length Ts - 1));
+ val ((inners, (Dss, Ass)), (unfold_set', lthy')) =
+ apfst (apsnd split_list o split_list)
+ (fold_map2 (fn i => bnf_of_typ Smart_Inline (qualify i) sort)
+ (if length Ts' = 1 then [0] else (1 upto length Ts')) Ts' (unfold_set, lthy));
+ in
+ compose_bnf const_policy qualify sort bnf inners oDs Dss Ass (unfold_set', lthy')
+ end)
+ end;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Tools/bnf_comp_tactics.ML Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,416 @@
+(* Title: HOL/BNF/Tools/bnf_comp_tactics.ML
+ Author: Dmitriy Traytel, TU Muenchen
+ Author: Jasmin Blanchette, TU Muenchen
+ Copyright 2012
+
+Tactics for composition of bounded natural functors.
+*)
+
+signature BNF_COMP_TACTICS =
+sig
+ val mk_comp_bd_card_order_tac: thm list -> thm -> tactic
+ val mk_comp_bd_cinfinite_tac: thm -> thm -> tactic
+ val mk_comp_in_alt_tac: Proof.context -> thm list -> tactic
+ val mk_comp_in_bd_tac: thm -> thm list -> thm -> thm list -> thm -> tactic
+ val mk_comp_map_comp_tac: thm -> thm -> thm list -> tactic
+ val mk_comp_map_cong_tac: thm list -> thm -> thm list -> tactic
+ val mk_comp_set_alt_tac: Proof.context -> thm -> tactic
+ val mk_comp_set_bd_tac: Proof.context -> thm -> thm list -> tactic
+ val mk_comp_set_natural_tac: thm -> thm -> thm -> thm list -> tactic
+ val mk_comp_wit_tac: Proof.context -> thm list -> thm -> thm list -> tactic
+
+ val mk_kill_bd_card_order_tac: int -> thm -> tactic
+ val mk_kill_bd_cinfinite_tac: thm -> tactic
+ val kill_in_alt_tac: tactic
+ val mk_kill_in_bd_tac: int -> bool -> thm -> thm -> thm -> thm -> thm -> tactic
+ val mk_kill_map_cong_tac: Proof.context -> int -> int -> thm -> tactic
+ val mk_kill_set_bd_tac: thm -> thm -> tactic
+
+ val empty_natural_tac: tactic
+ val lift_in_alt_tac: tactic
+ val mk_lift_in_bd_tac: int -> thm -> thm -> thm -> tactic
+ val mk_lift_set_bd_tac: thm -> tactic
+
+ val mk_permute_in_alt_tac: ''a list -> ''a list -> tactic
+ val mk_permute_in_bd_tac: ''a list -> ''a list -> thm -> thm -> thm -> tactic
+
+ val mk_map_wpull_tac: thm -> thm list -> thm -> tactic
+ val mk_simple_srel_O_Gr_tac: Proof.context -> thm -> thm -> thm -> tactic
+ val mk_simple_wit_tac: thm list -> tactic
+end;
+
+structure BNF_Comp_Tactics : BNF_COMP_TACTICS =
+struct
+
+open BNF_Util
+open BNF_Tactics
+
+val Card_order_csum = @{thm Card_order_csum};
+val Card_order_ctwo = @{thm Card_order_ctwo};
+val Cnotzero_UNIV = @{thm Cnotzero_UNIV};
+val arg_cong_Union = @{thm arg_cong[of _ _ Union]};
+val card_of_Card_order = @{thm card_of_Card_order};
+val csum_Cnotzero1 = @{thm csum_Cnotzero1};
+val csum_Cnotzero2 = @{thm csum_Cnotzero2};
+val ctwo_Cnotzero = @{thm ctwo_Cnotzero};
+val o_eq_dest_lhs = @{thm o_eq_dest_lhs};
+val ordIso_transitive = @{thm ordIso_transitive};
+val ordLeq_csum2 = @{thm ordLeq_csum2};
+val trans_image_cong_o_apply = @{thm trans[OF image_cong[OF o_apply refl]]};
+val trans_o_apply = @{thm trans[OF o_apply]};
+
+
+
+(* Composition *)
+
+fun mk_comp_set_alt_tac ctxt collect_set_natural =
+ unfold_thms_tac ctxt @{thms sym[OF o_assoc]} THEN
+ unfold_thms_tac ctxt [collect_set_natural RS sym] THEN
+ rtac refl 1;
+
+fun mk_comp_map_comp_tac Gmap_comp Gmap_cong map_comps =
+ EVERY' ([rtac ext, rtac sym, rtac trans_o_apply,
+ rtac (Gmap_comp RS sym RS o_eq_dest_lhs RS trans), rtac Gmap_cong] @
+ map (fn thm => rtac (thm RS sym RS fun_cong)) map_comps) 1;
+
+fun mk_comp_set_natural_tac Gmap_comp Gmap_cong Gset_natural set_naturals =
+ EVERY' ([rtac ext] @
+ replicate 3 (rtac trans_o_apply) @
+ [rtac (arg_cong_Union RS trans),
+ rtac (@{thm arg_cong2[of _ _ _ _ collect, OF refl]} RS trans),
+ rtac (Gmap_comp RS sym RS o_eq_dest_lhs RS trans),
+ rtac Gmap_cong] @
+ map (fn thm => rtac (thm RS fun_cong)) set_naturals @
+ [rtac (Gset_natural RS o_eq_dest_lhs), rtac sym, rtac trans_o_apply,
+ rtac trans_image_cong_o_apply, rtac trans_image_cong_o_apply,
+ rtac (@{thm image_cong} OF [Gset_natural RS o_eq_dest_lhs RS arg_cong_Union, refl] RS trans),
+ rtac @{thm trans[OF pointfreeE[OF Union_natural[symmetric]]]}, rtac arg_cong_Union,
+ rtac @{thm trans[OF o_eq_dest_lhs[OF image_o_collect[symmetric]]]},
+ rtac @{thm fun_cong[OF arg_cong[of _ _ collect]]}] @
+ [REPEAT_DETERM_N (length set_naturals) o EVERY' [rtac @{thm trans[OF image_insert]},
+ rtac @{thm arg_cong2[of _ _ _ _ insert]}, rtac ext, rtac trans_o_apply,
+ rtac trans_image_cong_o_apply, rtac @{thm trans[OF image_image]},
+ rtac @{thm sym[OF trans[OF o_apply]]}, rtac @{thm image_cong[OF refl o_apply]}],
+ rtac @{thm image_empty}]) 1;
+
+fun mk_comp_map_cong_tac comp_set_alts map_cong map_congs =
+ let
+ val n = length comp_set_alts;
+ in
+ (if n = 0 then rtac refl 1
+ else rtac map_cong 1 THEN
+ EVERY' (map_index (fn (i, map_cong) =>
+ rtac map_cong THEN' EVERY' (map_index (fn (k, set_alt) =>
+ EVERY' [select_prem_tac n (dtac @{thm meta_spec}) (k + 1), etac @{thm meta_mp},
+ rtac (equalityD2 RS set_mp), rtac (set_alt RS fun_cong RS trans),
+ rtac trans_o_apply, rtac (@{thm collect_def} RS arg_cong_Union),
+ rtac @{thm UnionI}, rtac @{thm UN_I}, REPEAT_DETERM_N i o rtac @{thm insertI2},
+ rtac @{thm insertI1}, rtac (o_apply RS equalityD2 RS set_mp),
+ etac @{thm imageI}, atac])
+ comp_set_alts))
+ map_congs) 1)
+ end;
+
+fun mk_comp_bd_card_order_tac Fbd_card_orders Gbd_card_order =
+ let
+ val (card_orders, last_card_order) = split_last Fbd_card_orders;
+ fun gen_before thm = rtac @{thm card_order_csum} THEN' rtac thm;
+ in
+ (rtac @{thm card_order_cprod} THEN'
+ WRAP' gen_before (K (K all_tac)) card_orders (rtac last_card_order) THEN'
+ rtac Gbd_card_order) 1
+ end;
+
+fun mk_comp_bd_cinfinite_tac Fbd_cinfinite Gbd_cinfinite =
+ (rtac @{thm cinfinite_cprod} THEN'
+ ((K (TRY ((rtac @{thm cinfinite_csum} THEN' rtac disjI1) 1)) THEN'
+ ((rtac @{thm cinfinite_csum} THEN' rtac disjI1 THEN' rtac Fbd_cinfinite) ORELSE'
+ rtac Fbd_cinfinite)) ORELSE'
+ rtac Fbd_cinfinite) THEN'
+ rtac Gbd_cinfinite) 1;
+
+fun mk_comp_set_bd_tac ctxt comp_set_alt Gset_Fset_bds =
+ let
+ val (bds, last_bd) = split_last Gset_Fset_bds;
+ fun gen_before bd =
+ rtac ctrans THEN' rtac @{thm Un_csum} THEN'
+ rtac ctrans THEN' rtac @{thm csum_mono} THEN'
+ rtac bd;
+ fun gen_after _ = rtac @{thm ordIso_imp_ordLeq} THEN' rtac @{thm cprod_csum_distrib1};
+ in
+ unfold_thms_tac ctxt [comp_set_alt] THEN
+ rtac @{thm comp_set_bd_Union_o_collect} 1 THEN
+ unfold_thms_tac ctxt @{thms Union_image_insert Union_image_empty Union_Un_distrib o_apply} THEN
+ (rtac ctrans THEN'
+ WRAP' gen_before gen_after bds (rtac last_bd) THEN'
+ rtac @{thm ordIso_imp_ordLeq} THEN'
+ rtac @{thm cprod_com}) 1
+ end;
+
+val comp_in_alt_thms = @{thms o_apply collect_def SUP_def image_insert image_empty Union_insert
+ Union_empty Un_empty_right Union_Un_distrib Un_subset_iff conj_subset_def UN_image_subset
+ conj_assoc};
+
+fun mk_comp_in_alt_tac ctxt comp_set_alts =
+ unfold_thms_tac ctxt (comp_set_alts @ comp_in_alt_thms) THEN
+ unfold_thms_tac ctxt @{thms set_eq_subset} THEN
+ rtac conjI 1 THEN
+ REPEAT_DETERM (
+ rtac @{thm subsetI} 1 THEN
+ unfold_thms_tac ctxt @{thms mem_Collect_eq Ball_def} THEN
+ (REPEAT_DETERM (CHANGED (etac conjE 1)) THEN
+ REPEAT_DETERM (CHANGED ((
+ (rtac conjI THEN' (atac ORELSE' rtac subset_UNIV)) ORELSE'
+ atac ORELSE'
+ (rtac subset_UNIV)) 1)) ORELSE rtac subset_UNIV 1));
+
+fun mk_comp_in_bd_tac comp_in_alt Fin_bds Gin_bd Fbd_Cinfs Gbd_Card_order =
+ let
+ val (bds, last_bd) = split_last Fin_bds;
+ val (Cinfs, _) = split_last Fbd_Cinfs;
+ fun gen_before (bd, _) = rtac ctrans THEN' rtac @{thm csum_mono} THEN' rtac bd;
+ fun gen_after (_, (bd_Cinf, next_bd_Cinf)) =
+ TRY o (rtac @{thm csum_cexp} THEN'
+ rtac bd_Cinf THEN'
+ (TRY o (rtac @{thm Cinfinite_csum} THEN' rtac disjI1) THEN' rtac next_bd_Cinf ORELSE'
+ rtac next_bd_Cinf) THEN'
+ ((rtac Card_order_csum THEN' rtac ordLeq_csum2) ORELSE'
+ (rtac Card_order_ctwo THEN' rtac @{thm ordLeq_refl})) THEN'
+ rtac Card_order_ctwo);
+ in
+ (rtac @{thm ordIso_ordLeq_trans} THEN'
+ rtac @{thm card_of_ordIso_subst} THEN'
+ rtac comp_in_alt THEN'
+ rtac ctrans THEN'
+ rtac Gin_bd THEN'
+ rtac @{thm ordLeq_ordIso_trans} THEN'
+ rtac @{thm cexp_mono1} THEN'
+ rtac @{thm ordLeq_ordIso_trans} THEN'
+ rtac @{thm csum_mono1} THEN'
+ WRAP' gen_before gen_after (bds ~~ (Cinfs ~~ tl Fbd_Cinfs)) (rtac last_bd) THEN'
+ rtac @{thm csum_absorb1} THEN'
+ rtac @{thm Cinfinite_cexp} THEN'
+ (rtac ordLeq_csum2 ORELSE' rtac @{thm ordLeq_refl}) THEN'
+ rtac Card_order_ctwo THEN'
+ (TRY o (rtac @{thm Cinfinite_csum} THEN' rtac disjI1) THEN' rtac (hd Fbd_Cinfs) ORELSE'
+ rtac (hd Fbd_Cinfs)) THEN'
+ rtac @{thm ctwo_ordLeq_Cinfinite} THEN'
+ rtac @{thm Cinfinite_cexp} THEN'
+ (rtac ordLeq_csum2 ORELSE' rtac @{thm ordLeq_refl}) THEN'
+ rtac Card_order_ctwo THEN'
+ (TRY o (rtac @{thm Cinfinite_csum} THEN' rtac disjI1) THEN' rtac (hd Fbd_Cinfs) ORELSE'
+ rtac (hd Fbd_Cinfs)) THEN'
+ rtac disjI1 THEN'
+ TRY o rtac csum_Cnotzero2 THEN'
+ rtac ctwo_Cnotzero THEN'
+ rtac Gbd_Card_order THEN'
+ rtac @{thm cexp_cprod} THEN'
+ TRY o rtac csum_Cnotzero2 THEN'
+ rtac ctwo_Cnotzero) 1
+ end;
+
+val comp_wit_thms = @{thms Union_empty_conv o_apply collect_def SUP_def
+ Union_image_insert Union_image_empty};
+
+fun mk_comp_wit_tac ctxt Gwit_thms collect_set_natural Fwit_thms =
+ ALLGOALS (dtac @{thm in_Union_o_assoc}) THEN
+ unfold_thms_tac ctxt (collect_set_natural :: comp_wit_thms) THEN
+ REPEAT_DETERM (
+ atac 1 ORELSE
+ REPEAT_DETERM (eresolve_tac @{thms UnionE UnE imageE} 1) THEN
+ (TRY o dresolve_tac Gwit_thms THEN'
+ (etac FalseE ORELSE'
+ hyp_subst_tac THEN'
+ dresolve_tac Fwit_thms THEN'
+ (etac FalseE ORELSE' atac))) 1);
+
+
+
+(* Kill operation *)
+
+fun mk_kill_map_cong_tac ctxt n m map_cong =
+ (rtac map_cong THEN' EVERY' (replicate n (rtac refl)) THEN'
+ EVERY' (replicate m (Goal.assume_rule_tac ctxt))) 1;
+
+fun mk_kill_bd_card_order_tac n bd_card_order =
+ (rtac @{thm card_order_cprod} THEN'
+ K (REPEAT_DETERM_N (n - 1)
+ ((rtac @{thm card_order_csum} THEN'
+ rtac @{thm card_of_card_order_on}) 1)) THEN'
+ rtac @{thm card_of_card_order_on} THEN'
+ rtac bd_card_order) 1;
+
+fun mk_kill_bd_cinfinite_tac bd_Cinfinite =
+ (rtac @{thm cinfinite_cprod2} THEN'
+ TRY o rtac csum_Cnotzero1 THEN'
+ rtac Cnotzero_UNIV THEN'
+ rtac bd_Cinfinite) 1;
+
+fun mk_kill_set_bd_tac bd_Card_order set_bd =
+ (rtac ctrans THEN'
+ rtac set_bd THEN'
+ rtac @{thm ordLeq_cprod2} THEN'
+ TRY o rtac csum_Cnotzero1 THEN'
+ rtac Cnotzero_UNIV THEN'
+ rtac bd_Card_order) 1
+
+val kill_in_alt_tac =
+ ((rtac @{thm Collect_cong} THEN' rtac iffI) 1 THEN
+ REPEAT_DETERM (CHANGED (etac conjE 1)) THEN
+ REPEAT_DETERM (CHANGED ((etac conjI ORELSE'
+ rtac conjI THEN' rtac subset_UNIV) 1)) THEN
+ (rtac subset_UNIV ORELSE' atac) 1 THEN
+ REPEAT_DETERM (CHANGED (etac conjE 1)) THEN
+ REPEAT_DETERM (CHANGED ((etac conjI ORELSE' atac) 1))) ORELSE
+ ((rtac @{thm UNIV_eq_I} THEN' rtac CollectI) 1 THEN
+ REPEAT_DETERM (TRY (rtac conjI 1) THEN rtac subset_UNIV 1));
+
+fun mk_kill_in_bd_tac n nontrivial_kill_in in_alt in_bd bd_Card_order bd_Cinfinite bd_Cnotzero =
+ (rtac @{thm ordIso_ordLeq_trans} THEN'
+ rtac @{thm card_of_ordIso_subst} THEN'
+ rtac in_alt THEN'
+ rtac ctrans THEN'
+ rtac in_bd THEN'
+ rtac @{thm ordIso_ordLeq_trans} THEN'
+ rtac @{thm cexp_cong1}) 1 THEN
+ (if nontrivial_kill_in then
+ rtac ordIso_transitive 1 THEN
+ REPEAT_DETERM_N (n - 1)
+ ((rtac @{thm csum_cong1} THEN'
+ rtac @{thm ordIso_symmetric} THEN'
+ rtac @{thm csum_assoc} THEN'
+ rtac ordIso_transitive) 1) THEN
+ (rtac @{thm ordIso_refl} THEN'
+ rtac Card_order_csum THEN'
+ rtac ordIso_transitive THEN'
+ rtac @{thm csum_assoc} THEN'
+ rtac ordIso_transitive THEN'
+ rtac @{thm csum_cong1} THEN'
+ K (mk_flatten_assoc_tac
+ (rtac @{thm ordIso_refl} THEN'
+ FIRST' [rtac card_of_Card_order, rtac Card_order_csum])
+ ordIso_transitive @{thm csum_assoc} @{thm csum_cong}) THEN'
+ rtac @{thm ordIso_refl} THEN'
+ (rtac card_of_Card_order ORELSE' rtac Card_order_csum)) 1
+ else all_tac) THEN
+ (rtac @{thm csum_com} THEN'
+ rtac bd_Card_order THEN'
+ rtac disjI1 THEN'
+ rtac csum_Cnotzero2 THEN'
+ rtac ctwo_Cnotzero THEN'
+ rtac disjI1 THEN'
+ rtac csum_Cnotzero2 THEN'
+ TRY o rtac csum_Cnotzero1 THEN'
+ rtac Cnotzero_UNIV THEN'
+ rtac @{thm ordLeq_ordIso_trans} THEN'
+ rtac @{thm cexp_mono1} THEN'
+ rtac ctrans THEN'
+ rtac @{thm csum_mono2} THEN'
+ rtac @{thm ordLeq_cprod1} THEN'
+ (rtac card_of_Card_order ORELSE' rtac Card_order_csum) THEN'
+ rtac bd_Cnotzero THEN'
+ rtac @{thm csum_cexp'} THEN'
+ rtac @{thm Cinfinite_cprod2} THEN'
+ TRY o rtac csum_Cnotzero1 THEN'
+ rtac Cnotzero_UNIV THEN'
+ rtac bd_Cinfinite THEN'
+ ((rtac Card_order_ctwo THEN' rtac @{thm ordLeq_refl}) ORELSE'
+ (rtac Card_order_csum THEN' rtac ordLeq_csum2)) THEN'
+ rtac Card_order_ctwo THEN'
+ rtac disjI1 THEN'
+ rtac csum_Cnotzero2 THEN'
+ TRY o rtac csum_Cnotzero1 THEN'
+ rtac Cnotzero_UNIV THEN'
+ rtac bd_Card_order THEN'
+ rtac @{thm cexp_cprod_ordLeq} THEN'
+ TRY o rtac csum_Cnotzero2 THEN'
+ rtac ctwo_Cnotzero THEN'
+ rtac @{thm Cinfinite_cprod2} THEN'
+ TRY o rtac csum_Cnotzero1 THEN'
+ rtac Cnotzero_UNIV THEN'
+ rtac bd_Cinfinite THEN'
+ rtac bd_Cnotzero THEN'
+ rtac @{thm ordLeq_cprod2} THEN'
+ TRY o rtac csum_Cnotzero1 THEN'
+ rtac Cnotzero_UNIV THEN'
+ rtac bd_Card_order) 1;
+
+
+
+(* Lift operation *)
+
+val empty_natural_tac = rtac @{thm empty_natural} 1;
+
+fun mk_lift_set_bd_tac bd_Card_order = (rtac @{thm Card_order_empty} THEN' rtac bd_Card_order) 1;
+
+val lift_in_alt_tac =
+ ((rtac @{thm Collect_cong} THEN' rtac iffI) 1 THEN
+ REPEAT_DETERM (CHANGED (etac conjE 1)) THEN
+ REPEAT_DETERM (CHANGED ((etac conjI ORELSE' atac) 1)) THEN
+ REPEAT_DETERM (CHANGED (etac conjE 1)) THEN
+ REPEAT_DETERM (CHANGED ((etac conjI ORELSE'
+ rtac conjI THEN' rtac @{thm empty_subsetI}) 1)) THEN
+ (rtac @{thm empty_subsetI} ORELSE' atac) 1) ORELSE
+ ((rtac sym THEN' rtac @{thm UNIV_eq_I} THEN' rtac CollectI) 1 THEN
+ REPEAT_DETERM (TRY (rtac conjI 1) THEN rtac @{thm empty_subsetI} 1));
+
+fun mk_lift_in_bd_tac n in_alt in_bd bd_Card_order =
+ (rtac @{thm ordIso_ordLeq_trans} THEN'
+ rtac @{thm card_of_ordIso_subst} THEN'
+ rtac in_alt THEN'
+ rtac ctrans THEN'
+ rtac in_bd THEN'
+ rtac @{thm cexp_mono1}) 1 THEN
+ ((rtac @{thm csum_mono1} 1 THEN
+ REPEAT_DETERM_N (n - 1)
+ ((rtac ctrans THEN'
+ rtac ordLeq_csum2 THEN'
+ (rtac Card_order_csum ORELSE' rtac card_of_Card_order)) 1) THEN
+ (rtac ordLeq_csum2 THEN'
+ (rtac Card_order_csum ORELSE' rtac card_of_Card_order)) 1) ORELSE
+ (rtac ordLeq_csum2 THEN' rtac Card_order_ctwo) 1) THEN
+ (rtac disjI1 THEN' TRY o rtac csum_Cnotzero2 THEN' rtac ctwo_Cnotzero
+ THEN' rtac bd_Card_order) 1;
+
+
+
+(* Permute operation *)
+
+fun mk_permute_in_alt_tac src dest =
+ (rtac @{thm Collect_cong} THEN'
+ mk_rotate_eq_tac (rtac refl) trans @{thm conj_assoc} @{thm conj_commute} @{thm conj_cong}
+ dest src) 1;
+
+fun mk_permute_in_bd_tac src dest in_alt in_bd bd_Card_order =
+ (rtac @{thm ordIso_ordLeq_trans} THEN'
+ rtac @{thm card_of_ordIso_subst} THEN'
+ rtac in_alt THEN'
+ rtac @{thm ordLeq_ordIso_trans} THEN'
+ rtac in_bd THEN'
+ rtac @{thm cexp_cong1} THEN'
+ rtac @{thm csum_cong1} THEN'
+ mk_rotate_eq_tac
+ (rtac @{thm ordIso_refl} THEN'
+ FIRST' [rtac card_of_Card_order, rtac Card_order_csum])
+ ordIso_transitive @{thm csum_assoc} @{thm csum_com} @{thm csum_cong}
+ src dest THEN'
+ rtac bd_Card_order THEN'
+ rtac disjI1 THEN'
+ TRY o rtac csum_Cnotzero2 THEN'
+ rtac ctwo_Cnotzero THEN'
+ rtac disjI1 THEN'
+ TRY o rtac csum_Cnotzero2 THEN'
+ rtac ctwo_Cnotzero) 1;
+
+fun mk_map_wpull_tac comp_in_alt inner_map_wpulls outer_map_wpull =
+ (rtac (@{thm wpull_cong} OF (replicate 3 comp_in_alt)) THEN' rtac outer_map_wpull) 1 THEN
+ WRAP (fn thm => rtac thm 1 THEN REPEAT_DETERM (atac 1)) (K all_tac) inner_map_wpulls all_tac THEN
+ TRY (REPEAT_DETERM (atac 1 ORELSE rtac @{thm wpull_id} 1));
+
+fun mk_simple_srel_O_Gr_tac ctxt srel_def srel_O_Gr in_alt_thm =
+ rtac (unfold_thms ctxt [srel_def]
+ (trans OF [srel_O_Gr, in_alt_thm RS @{thm subst_rel_def} RS sym])) 1;
+
+fun mk_simple_wit_tac wit_thms = ALLGOALS (atac ORELSE' eresolve_tac (@{thm emptyE} :: wit_thms));
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Tools/bnf_def.ML Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,1238 @@
+(* Title: HOL/BNF/Tools/bnf_def.ML
+ Author: Dmitriy Traytel, TU Muenchen
+ Author: Jasmin Blanchette, TU Muenchen
+ Copyright 2012
+
+Definition of bounded natural functors.
+*)
+
+signature BNF_DEF =
+sig
+ type BNF
+ type nonemptiness_witness = {I: int list, wit: term, prop: thm list}
+
+ val bnf_of: Proof.context -> string -> BNF option
+ val register_bnf: string -> (BNF * local_theory) -> (BNF * local_theory)
+
+ val name_of_bnf: BNF -> binding
+ val T_of_bnf: BNF -> typ
+ val live_of_bnf: BNF -> int
+ val lives_of_bnf: BNF -> typ list
+ val dead_of_bnf: BNF -> int
+ val deads_of_bnf: BNF -> typ list
+ val nwits_of_bnf: BNF -> int
+
+ val mapN: string
+ val relN: string
+ val setN: string
+ val mk_setN: int -> string
+ val srelN: string
+
+ val map_of_bnf: BNF -> term
+
+ val mk_T_of_bnf: typ list -> typ list -> BNF -> typ
+ val mk_bd_of_bnf: typ list -> typ list -> BNF -> term
+ val mk_map_of_bnf: typ list -> typ list -> typ list -> BNF -> term
+ val mk_rel_of_bnf: typ list -> typ list -> typ list -> BNF -> term
+ val mk_sets_of_bnf: typ list list -> typ list list -> BNF -> term list
+ val mk_srel_of_bnf: typ list -> typ list -> typ list -> BNF -> term
+ val mk_wits_of_bnf: typ list list -> typ list list -> BNF -> (int list * term) list
+
+ val bd_Card_order_of_bnf: BNF -> thm
+ val bd_Cinfinite_of_bnf: BNF -> thm
+ val bd_Cnotzero_of_bnf: BNF -> thm
+ val bd_card_order_of_bnf: BNF -> thm
+ val bd_cinfinite_of_bnf: BNF -> thm
+ val collect_set_natural_of_bnf: BNF -> thm
+ val in_bd_of_bnf: BNF -> thm
+ val in_cong_of_bnf: BNF -> thm
+ val in_mono_of_bnf: BNF -> thm
+ val in_srel_of_bnf: BNF -> thm
+ val map_comp'_of_bnf: BNF -> thm
+ val map_comp_of_bnf: BNF -> thm
+ val map_cong_of_bnf: BNF -> thm
+ val map_def_of_bnf: BNF -> thm
+ val map_id'_of_bnf: BNF -> thm
+ val map_id_of_bnf: BNF -> thm
+ val map_wppull_of_bnf: BNF -> thm
+ val map_wpull_of_bnf: BNF -> thm
+ val rel_def_of_bnf: BNF -> thm
+ val set_bd_of_bnf: BNF -> thm list
+ val set_defs_of_bnf: BNF -> thm list
+ val set_natural'_of_bnf: BNF -> thm list
+ val set_natural_of_bnf: BNF -> thm list
+ val sets_of_bnf: BNF -> term list
+ val srel_def_of_bnf: BNF -> thm
+ val srel_Gr_of_bnf: BNF -> thm
+ val srel_Id_of_bnf: BNF -> thm
+ val srel_O_of_bnf: BNF -> thm
+ val srel_O_Gr_of_bnf: BNF -> thm
+ val srel_cong_of_bnf: BNF -> thm
+ val srel_converse_of_bnf: BNF -> thm
+ val srel_mono_of_bnf: BNF -> thm
+ val wit_thms_of_bnf: BNF -> thm list
+ val wit_thmss_of_bnf: BNF -> thm list list
+
+ val mk_witness: int list * term -> thm list -> nonemptiness_witness
+ val minimize_wits: (''a list * 'b) list -> (''a list * 'b) list
+ val wits_of_bnf: BNF -> nonemptiness_witness list
+
+ val zip_axioms: 'a -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a -> 'a list
+
+ datatype const_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline
+ datatype fact_policy =
+ Derive_Few_Facts | Derive_All_Facts | Derive_All_Facts_Note_Most | Note_All_Facts_and_Axioms
+ val bnf_note_all: bool Config.T
+ val user_policy: fact_policy -> Proof.context -> fact_policy
+
+ val print_bnfs: Proof.context -> unit
+ val bnf_def: const_policy -> (Proof.context -> fact_policy) -> (binding -> binding) ->
+ ({prems: thm list, context: Proof.context} -> tactic) list ->
+ ({prems: thm list, context: Proof.context} -> tactic) -> typ list option ->
+ ((((binding * term) * term list) * term) * term list) * term option -> local_theory ->
+ BNF * local_theory
+end;
+
+structure BNF_Def : BNF_DEF =
+struct
+
+open BNF_Util
+open BNF_Tactics
+open BNF_Def_Tactics
+
+type axioms = {
+ map_id: thm,
+ map_comp: thm,
+ map_cong: thm,
+ set_natural: thm list,
+ bd_card_order: thm,
+ bd_cinfinite: thm,
+ set_bd: thm list,
+ in_bd: thm,
+ map_wpull: thm,
+ srel_O_Gr: thm
+};
+
+fun mk_axioms' (((((((((id, comp), cong), nat), c_o), cinf), set_bd), in_bd), wpull), srel) =
+ {map_id = id, map_comp = comp, map_cong = cong, set_natural = nat, bd_card_order = c_o,
+ bd_cinfinite = cinf, set_bd = set_bd, in_bd = in_bd, map_wpull = wpull, srel_O_Gr = srel};
+
+fun dest_cons [] = raise Empty
+ | dest_cons (x :: xs) = (x, xs);
+
+fun mk_axioms n thms = thms
+ |> map the_single
+ |> dest_cons
+ ||>> dest_cons
+ ||>> dest_cons
+ ||>> chop n
+ ||>> dest_cons
+ ||>> dest_cons
+ ||>> chop n
+ ||>> dest_cons
+ ||>> dest_cons
+ ||> the_single
+ |> mk_axioms';
+
+fun zip_axioms mid mcomp mcong snat bdco bdinf sbd inbd wpull srel =
+ [mid, mcomp, mcong] @ snat @ [bdco, bdinf] @ sbd @ [inbd, wpull, srel];
+
+fun dest_axioms {map_id, map_comp, map_cong, set_natural, bd_card_order, bd_cinfinite, set_bd,
+ in_bd, map_wpull, srel_O_Gr} =
+ zip_axioms map_id map_comp map_cong set_natural bd_card_order bd_cinfinite set_bd in_bd map_wpull
+ srel_O_Gr;
+
+fun map_axioms f {map_id, map_comp, map_cong, set_natural, bd_card_order, bd_cinfinite, set_bd,
+ in_bd, map_wpull, srel_O_Gr} =
+ {map_id = f map_id,
+ map_comp = f map_comp,
+ map_cong = f map_cong,
+ set_natural = map f set_natural,
+ bd_card_order = f bd_card_order,
+ bd_cinfinite = f bd_cinfinite,
+ set_bd = map f set_bd,
+ in_bd = f in_bd,
+ map_wpull = f map_wpull,
+ srel_O_Gr = f srel_O_Gr};
+
+val morph_axioms = map_axioms o Morphism.thm;
+
+type defs = {
+ map_def: thm,
+ set_defs: thm list,
+ rel_def: thm,
+ srel_def: thm
+}
+
+fun mk_defs map sets rel srel = {map_def = map, set_defs = sets, rel_def = rel, srel_def = srel};
+
+fun map_defs f {map_def, set_defs, rel_def, srel_def} =
+ {map_def = f map_def, set_defs = map f set_defs, rel_def = f rel_def, srel_def = f srel_def};
+
+val morph_defs = map_defs o Morphism.thm;
+
+type facts = {
+ bd_Card_order: thm,
+ bd_Cinfinite: thm,
+ bd_Cnotzero: thm,
+ collect_set_natural: thm lazy,
+ in_cong: thm lazy,
+ in_mono: thm lazy,
+ in_srel: thm lazy,
+ map_comp': thm lazy,
+ map_id': thm lazy,
+ map_wppull: thm lazy,
+ set_natural': thm lazy list,
+ srel_cong: thm lazy,
+ srel_mono: thm lazy,
+ srel_Id: thm lazy,
+ srel_Gr: thm lazy,
+ srel_converse: thm lazy,
+ srel_O: thm lazy
+};
+
+fun mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_natural in_cong in_mono in_srel
+ map_comp' map_id' map_wppull set_natural' srel_cong srel_mono srel_Id srel_Gr srel_converse
+ srel_O = {
+ bd_Card_order = bd_Card_order,
+ bd_Cinfinite = bd_Cinfinite,
+ bd_Cnotzero = bd_Cnotzero,
+ collect_set_natural = collect_set_natural,
+ in_cong = in_cong,
+ in_mono = in_mono,
+ in_srel = in_srel,
+ map_comp' = map_comp',
+ map_id' = map_id',
+ map_wppull = map_wppull,
+ set_natural' = set_natural',
+ srel_cong = srel_cong,
+ srel_mono = srel_mono,
+ srel_Id = srel_Id,
+ srel_Gr = srel_Gr,
+ srel_converse = srel_converse,
+ srel_O = srel_O};
+
+fun map_facts f {
+ bd_Card_order,
+ bd_Cinfinite,
+ bd_Cnotzero,
+ collect_set_natural,
+ in_cong,
+ in_mono,
+ in_srel,
+ map_comp',
+ map_id',
+ map_wppull,
+ set_natural',
+ srel_cong,
+ srel_mono,
+ srel_Id,
+ srel_Gr,
+ srel_converse,
+ srel_O} =
+ {bd_Card_order = f bd_Card_order,
+ bd_Cinfinite = f bd_Cinfinite,
+ bd_Cnotzero = f bd_Cnotzero,
+ collect_set_natural = Lazy.map f collect_set_natural,
+ in_cong = Lazy.map f in_cong,
+ in_mono = Lazy.map f in_mono,
+ in_srel = Lazy.map f in_srel,
+ map_comp' = Lazy.map f map_comp',
+ map_id' = Lazy.map f map_id',
+ map_wppull = Lazy.map f map_wppull,
+ set_natural' = map (Lazy.map f) set_natural',
+ srel_cong = Lazy.map f srel_cong,
+ srel_mono = Lazy.map f srel_mono,
+ srel_Id = Lazy.map f srel_Id,
+ srel_Gr = Lazy.map f srel_Gr,
+ srel_converse = Lazy.map f srel_converse,
+ srel_O = Lazy.map f srel_O};
+
+val morph_facts = map_facts o Morphism.thm;
+
+type nonemptiness_witness = {
+ I: int list,
+ wit: term,
+ prop: thm list
+};
+
+fun mk_witness (I, wit) prop = {I = I, wit = wit, prop = prop};
+fun map_witness f g {I, wit, prop} = {I = I, wit = f wit, prop = map g prop};
+fun morph_witness phi = map_witness (Morphism.term phi) (Morphism.thm phi);
+
+datatype BNF = BNF of {
+ name: binding,
+ T: typ,
+ live: int,
+ lives: typ list, (*source type variables of map, only for composition*)
+ lives': typ list, (*target type variables of map, only for composition*)
+ dead: int,
+ deads: typ list, (*only for composition*)
+ map: term,
+ sets: term list,
+ bd: term,
+ axioms: axioms,
+ defs: defs,
+ facts: facts,
+ nwits: int,
+ wits: nonemptiness_witness list,
+ rel: term,
+ srel: term
+};
+
+(* getters *)
+
+fun rep_bnf (BNF bnf) = bnf;
+val name_of_bnf = #name o rep_bnf;
+val T_of_bnf = #T o rep_bnf;
+fun mk_T_of_bnf Ds Ts bnf =
+ let val bnf_rep = rep_bnf bnf
+ in Term.typ_subst_atomic ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#T bnf_rep) end;
+val live_of_bnf = #live o rep_bnf;
+val lives_of_bnf = #lives o rep_bnf;
+val dead_of_bnf = #dead o rep_bnf;
+val deads_of_bnf = #deads o rep_bnf;
+val axioms_of_bnf = #axioms o rep_bnf;
+val facts_of_bnf = #facts o rep_bnf;
+val nwits_of_bnf = #nwits o rep_bnf;
+val wits_of_bnf = #wits o rep_bnf;
+
+(*terms*)
+val map_of_bnf = #map o rep_bnf;
+val sets_of_bnf = #sets o rep_bnf;
+fun mk_map_of_bnf Ds Ts Us bnf =
+ let val bnf_rep = rep_bnf bnf;
+ in
+ Term.subst_atomic_types
+ ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#map bnf_rep)
+ end;
+fun mk_sets_of_bnf Dss Tss bnf =
+ let val bnf_rep = rep_bnf bnf;
+ in
+ map2 (fn (Ds, Ts) => Term.subst_atomic_types
+ ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts))) (Dss ~~ Tss) (#sets bnf_rep)
+ end;
+val bd_of_bnf = #bd o rep_bnf;
+fun mk_bd_of_bnf Ds Ts bnf =
+ let val bnf_rep = rep_bnf bnf;
+ in Term.subst_atomic_types ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#bd bnf_rep) end;
+fun mk_wits_of_bnf Dss Tss bnf =
+ let
+ val bnf_rep = rep_bnf bnf;
+ val wits = map (fn x => (#I x, #wit x)) (#wits bnf_rep);
+ in
+ map2 (fn (Ds, Ts) => apsnd (Term.subst_atomic_types
+ ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)))) (Dss ~~ Tss) wits
+ end;
+val rel_of_bnf = #rel o rep_bnf;
+fun mk_rel_of_bnf Ds Ts Us bnf =
+ let val bnf_rep = rep_bnf bnf;
+ in
+ Term.subst_atomic_types
+ ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#rel bnf_rep)
+ end;
+val srel_of_bnf = #srel o rep_bnf;
+fun mk_srel_of_bnf Ds Ts Us bnf =
+ let val bnf_rep = rep_bnf bnf;
+ in
+ Term.subst_atomic_types
+ ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#srel bnf_rep)
+ end;
+
+(*thms*)
+val bd_card_order_of_bnf = #bd_card_order o #axioms o rep_bnf;
+val bd_cinfinite_of_bnf = #bd_cinfinite o #axioms o rep_bnf;
+val bd_Card_order_of_bnf = #bd_Card_order o #facts o rep_bnf;
+val bd_Cinfinite_of_bnf = #bd_Cinfinite o #facts o rep_bnf;
+val bd_Cnotzero_of_bnf = #bd_Cnotzero o #facts o rep_bnf;
+val collect_set_natural_of_bnf = Lazy.force o #collect_set_natural o #facts o rep_bnf;
+val in_bd_of_bnf = #in_bd o #axioms o rep_bnf;
+val in_cong_of_bnf = Lazy.force o #in_cong o #facts o rep_bnf;
+val in_mono_of_bnf = Lazy.force o #in_mono o #facts o rep_bnf;
+val in_srel_of_bnf = Lazy.force o #in_srel o #facts o rep_bnf;
+val map_def_of_bnf = #map_def o #defs o rep_bnf;
+val map_id_of_bnf = #map_id o #axioms o rep_bnf;
+val map_id'_of_bnf = Lazy.force o #map_id' o #facts o rep_bnf;
+val map_comp_of_bnf = #map_comp o #axioms o rep_bnf;
+val map_comp'_of_bnf = Lazy.force o #map_comp' o #facts o rep_bnf;
+val map_cong_of_bnf = #map_cong o #axioms o rep_bnf;
+val map_wppull_of_bnf = Lazy.force o #map_wppull o #facts o rep_bnf;
+val map_wpull_of_bnf = #map_wpull o #axioms o rep_bnf;
+val rel_def_of_bnf = #rel_def o #defs o rep_bnf;
+val set_bd_of_bnf = #set_bd o #axioms o rep_bnf;
+val set_defs_of_bnf = #set_defs o #defs o rep_bnf;
+val set_natural_of_bnf = #set_natural o #axioms o rep_bnf;
+val set_natural'_of_bnf = map Lazy.force o #set_natural' o #facts o rep_bnf;
+val srel_cong_of_bnf = Lazy.force o #srel_cong o #facts o rep_bnf;
+val srel_mono_of_bnf = Lazy.force o #srel_mono o #facts o rep_bnf;
+val srel_def_of_bnf = #srel_def o #defs o rep_bnf;
+val srel_Id_of_bnf = Lazy.force o #srel_Id o #facts o rep_bnf;
+val srel_Gr_of_bnf = Lazy.force o #srel_Gr o #facts o rep_bnf;
+val srel_converse_of_bnf = Lazy.force o #srel_converse o #facts o rep_bnf;
+val srel_O_of_bnf = Lazy.force o #srel_O o #facts o rep_bnf;
+val srel_O_Gr_of_bnf = #srel_O_Gr o #axioms o rep_bnf;
+val wit_thms_of_bnf = maps #prop o wits_of_bnf;
+val wit_thmss_of_bnf = map #prop o wits_of_bnf;
+
+fun mk_bnf name T live lives lives' dead deads map sets bd axioms defs facts wits rel srel =
+ BNF {name = name, T = T,
+ live = live, lives = lives, lives' = lives', dead = dead, deads = deads,
+ map = map, sets = sets, bd = bd,
+ axioms = axioms, defs = defs, facts = facts,
+ nwits = length wits, wits = wits, rel = rel, srel = srel};
+
+fun morph_bnf phi (BNF {name = name, T = T, live = live, lives = lives, lives' = lives',
+ dead = dead, deads = deads, map = map, sets = sets, bd = bd,
+ axioms = axioms, defs = defs, facts = facts,
+ nwits = nwits, wits = wits, rel = rel, srel = srel}) =
+ BNF {name = Morphism.binding phi name, T = Morphism.typ phi T,
+ live = live, lives = List.map (Morphism.typ phi) lives,
+ lives' = List.map (Morphism.typ phi) lives',
+ dead = dead, deads = List.map (Morphism.typ phi) deads,
+ map = Morphism.term phi map, sets = List.map (Morphism.term phi) sets,
+ bd = Morphism.term phi bd,
+ axioms = morph_axioms phi axioms,
+ defs = morph_defs phi defs,
+ facts = morph_facts phi facts,
+ nwits = nwits,
+ wits = List.map (morph_witness phi) wits,
+ rel = Morphism.term phi rel, srel = Morphism.term phi srel};
+
+fun eq_bnf (BNF {T = T1, live = live1, dead = dead1, ...},
+ BNF {T = T2, live = live2, dead = dead2, ...}) =
+ Type.could_unify (T1, T2) andalso live1 = live2 andalso dead1 = dead2;
+
+structure Data = Generic_Data
+(
+ type T = BNF Symtab.table;
+ val empty = Symtab.empty;
+ val extend = I;
+ val merge = Symtab.merge eq_bnf;
+);
+
+val bnf_of = Symtab.lookup o Data.get o Context.Proof;
+
+
+
+(* Utilities *)
+
+fun normalize_set insts instA set =
+ let
+ val (T, T') = dest_funT (fastype_of set);
+ val A = fst (Term.dest_TVar (HOLogic.dest_setT T'));
+ val params = Term.add_tvar_namesT T [];
+ in Term.subst_TVars ((A :: params) ~~ (instA :: insts)) set end;
+
+fun normalize_rel ctxt instTs instA instB rel =
+ let
+ val thy = Proof_Context.theory_of ctxt;
+ val tyenv =
+ Sign.typ_match thy (fastype_of rel, Library.foldr (op -->) (instTs, mk_pred2T instA instB))
+ Vartab.empty;
+ in Envir.subst_term (tyenv, Vartab.empty) rel end
+ handle Type.TYPE_MATCH => error "Bad predicator";
+
+fun normalize_srel ctxt instTs instA instB srel =
+ let
+ val thy = Proof_Context.theory_of ctxt;
+ val tyenv =
+ Sign.typ_match thy (fastype_of srel, Library.foldr (op -->) (instTs, mk_relT (instA, instB)))
+ Vartab.empty;
+ in Envir.subst_term (tyenv, Vartab.empty) srel end
+ handle Type.TYPE_MATCH => error "Bad relator";
+
+fun normalize_wit insts CA As wit =
+ let
+ fun strip_param (Ts, T as Type (@{type_name fun}, [T1, T2])) =
+ if Type.raw_instance (CA, T) then (Ts, T) else strip_param (T1 :: Ts, T2)
+ | strip_param x = x;
+ val (Ts, T) = strip_param ([], fastype_of wit);
+ val subst = Term.add_tvar_namesT T [] ~~ insts;
+ fun find y = find_index (fn x => x = y) As;
+ in
+ (map (find o Term.typ_subst_TVars subst) (rev Ts), Term.subst_TVars subst wit)
+ end;
+
+fun minimize_wits wits =
+ let
+ fun minimize done [] = done
+ | minimize done ((I, wit) :: todo) =
+ if exists (fn (J, _) => subset (op =) (J, I)) (done @ todo)
+ then minimize done todo
+ else minimize ((I, wit) :: done) todo;
+ in minimize [] wits end;
+
+
+
+(* Names *)
+
+val mapN = "map";
+val setN = "set";
+fun mk_setN i = setN ^ nonzero_string_of_int i;
+val bdN = "bd";
+val witN = "wit";
+fun mk_witN i = witN ^ nonzero_string_of_int i;
+val relN = "rel";
+val srelN = "srel";
+
+val bd_card_orderN = "bd_card_order";
+val bd_cinfiniteN = "bd_cinfinite";
+val bd_Card_orderN = "bd_Card_order";
+val bd_CinfiniteN = "bd_Cinfinite";
+val bd_CnotzeroN = "bd_Cnotzero";
+val collect_set_naturalN = "collect_set_natural";
+val in_bdN = "in_bd";
+val in_monoN = "in_mono";
+val in_srelN = "in_srel";
+val map_idN = "map_id";
+val map_id'N = "map_id'";
+val map_compN = "map_comp";
+val map_comp'N = "map_comp'";
+val map_congN = "map_cong";
+val map_wpullN = "map_wpull";
+val srel_IdN = "srel_Id";
+val srel_GrN = "srel_Gr";
+val srel_converseN = "srel_converse";
+val srel_monoN = "srel_mono"
+val srel_ON = "srel_comp";
+val srel_O_GrN = "srel_comp_Gr";
+val set_naturalN = "set_natural";
+val set_natural'N = "set_natural'";
+val set_bdN = "set_bd";
+
+datatype const_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline;
+
+datatype fact_policy =
+ Derive_Few_Facts | Derive_All_Facts | Derive_All_Facts_Note_Most | Note_All_Facts_and_Axioms;
+
+val bnf_note_all = Attrib.setup_config_bool @{binding bnf_note_all} (K false);
+
+fun user_policy policy ctxt =
+ if Config.get ctxt bnf_note_all then Note_All_Facts_and_Axioms else policy;
+
+val smart_max_inline_size = 25; (*FUDGE*)
+
+
+(* Define new BNFs *)
+
+fun prepare_def const_policy mk_fact_policy qualify prep_term Ds_opt
+ (((((raw_b, raw_map), raw_sets), raw_bd_Abs), raw_wits), raw_rel_opt) no_defs_lthy =
+ let
+ val fact_policy = mk_fact_policy no_defs_lthy;
+ val b = qualify raw_b;
+ val live = length raw_sets;
+ val nwits = length raw_wits;
+
+ val map_rhs = prep_term no_defs_lthy raw_map;
+ val set_rhss = map (prep_term no_defs_lthy) raw_sets;
+ val (bd_rhsT, bd_rhs) = (case prep_term no_defs_lthy raw_bd_Abs of
+ Abs (_, T, t) => (T, t)
+ | _ => error "Bad bound constant");
+ val wit_rhss = map (prep_term no_defs_lthy) raw_wits;
+
+ fun err T =
+ error ("Trying to register the type " ^ quote (Syntax.string_of_typ no_defs_lthy T) ^
+ " as unnamed BNF");
+
+ val (b, key) =
+ if Binding.eq_name (b, Binding.empty) then
+ (case bd_rhsT of
+ Type (C, Ts) => if forall (is_some o try dest_TFree) Ts
+ then (Binding.qualified_name C, C) else err bd_rhsT
+ | T => err T)
+ else (b, Local_Theory.full_name no_defs_lthy b);
+
+ fun maybe_define user_specified (b, rhs) lthy =
+ let
+ val inline =
+ (user_specified orelse fact_policy = Derive_Few_Facts) andalso
+ (case const_policy of
+ Dont_Inline => false
+ | Hardly_Inline => Term.is_Free rhs orelse Term.is_Const rhs
+ | Smart_Inline => Term.size_of_term rhs <= smart_max_inline_size
+ | Do_Inline => true)
+ in
+ if inline then
+ ((rhs, Drule.reflexive_thm), lthy)
+ else
+ let val b = b () in
+ apfst (apsnd snd) (Local_Theory.define ((b, NoSyn), ((Thm.def_binding b, []), rhs))
+ lthy)
+ end
+ end;
+
+ fun maybe_restore lthy_old lthy =
+ lthy |> not (pointer_eq (lthy_old, lthy)) ? Local_Theory.restore;
+
+ val map_bind_def = (fn () => Binding.suffix_name ("_" ^ mapN) b, map_rhs);
+ val set_binds_defs =
+ let
+ val bs = if live = 1 then [fn () => Binding.suffix_name ("_" ^ setN) b]
+ else map (fn i => fn () => Binding.suffix_name ("_" ^ mk_setN i) b) (1 upto live)
+ in map2 pair bs set_rhss end;
+ val bd_bind_def = (fn () => Binding.suffix_name ("_" ^ bdN) b, bd_rhs);
+ val wit_binds_defs =
+ let
+ val bs = if nwits = 1 then [fn () => Binding.suffix_name ("_" ^ witN) b]
+ else map (fn i => fn () => Binding.suffix_name ("_" ^ mk_witN i) b) (1 upto nwits);
+ in map2 pair bs wit_rhss end;
+
+ val (((((bnf_map_term, raw_map_def),
+ (bnf_set_terms, raw_set_defs)),
+ (bnf_bd_term, raw_bd_def)),
+ (bnf_wit_terms, raw_wit_defs)), (lthy, lthy_old)) =
+ no_defs_lthy
+ |> maybe_define true map_bind_def
+ ||>> apfst split_list o fold_map (maybe_define true) set_binds_defs
+ ||>> maybe_define true bd_bind_def
+ ||>> apfst split_list o fold_map (maybe_define true) wit_binds_defs
+ ||> `(maybe_restore no_defs_lthy);
+
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+
+ val bnf_map_def = Morphism.thm phi raw_map_def;
+ val bnf_set_defs = map (Morphism.thm phi) raw_set_defs;
+ val bnf_bd_def = Morphism.thm phi raw_bd_def;
+ val bnf_wit_defs = map (Morphism.thm phi) raw_wit_defs;
+
+ val bnf_map = Morphism.term phi bnf_map_term;
+
+ (*TODO: handle errors*)
+ (*simple shape analysis of a map function*)
+ val ((alphas, betas), (CA, _)) =
+ fastype_of bnf_map
+ |> strip_typeN live
+ |>> map_split dest_funT
+ ||> dest_funT
+ handle TYPE _ => error "Bad map function";
+
+ val CA_params = map TVar (Term.add_tvarsT CA []);
+
+ val bnf_sets = map2 (normalize_set CA_params) alphas (map (Morphism.term phi) bnf_set_terms);
+ val bdT = Morphism.typ phi bd_rhsT;
+ val bnf_bd =
+ Term.subst_TVars (Term.add_tvar_namesT bdT [] ~~ CA_params) (Morphism.term phi bnf_bd_term);
+ val bnf_wits = map (normalize_wit CA_params CA alphas o Morphism.term phi) bnf_wit_terms;
+
+ (*TODO: assert Ds = (TVars of bnf_map) \ (alphas @ betas) as sets*)
+ val deads = (case Ds_opt of
+ NONE => subtract (op =) (alphas @ betas) (map TVar (Term.add_tvars bnf_map []))
+ | SOME Ds => map (Morphism.typ phi) Ds);
+ val dead = length deads;
+
+ (*TODO: further checks of type of bnf_map*)
+ (*TODO: check types of bnf_sets*)
+ (*TODO: check type of bnf_bd*)
+ (*TODO: check type of bnf_rel*)
+
+ val ((((((((((As', Bs'), Cs), Ds), B1Ts), B2Ts), domTs), ranTs), ranTs'), ranTs''),
+ (Ts, T)) = lthy
+ |> mk_TFrees live
+ ||>> mk_TFrees live
+ ||>> mk_TFrees live
+ ||>> mk_TFrees dead
+ ||>> mk_TFrees live
+ ||>> mk_TFrees live
+ ||>> mk_TFrees live
+ ||>> mk_TFrees live
+ ||>> mk_TFrees live
+ ||>> mk_TFrees live
+ ||> fst o mk_TFrees 1
+ ||> the_single
+ ||> `(replicate live);
+
+ fun mk_bnf_map As' Bs' =
+ Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As') @ (betas ~~ Bs')) bnf_map;
+ fun mk_bnf_t As' = Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As'));
+ fun mk_bnf_T As' = Term.typ_subst_atomic ((deads ~~ Ds) @ (alphas ~~ As'));
+
+ val (setRTs, RTs) = map_split (`HOLogic.mk_setT o HOLogic.mk_prodT) (As' ~~ Bs');
+ val setRTsAsCs = map (HOLogic.mk_setT o HOLogic.mk_prodT) (As' ~~ Cs);
+ val setRTsBsCs = map (HOLogic.mk_setT o HOLogic.mk_prodT) (Bs' ~~ Cs);
+ val setRT's = map (HOLogic.mk_setT o HOLogic.mk_prodT) (Bs' ~~ As');
+ val self_setRTs = map (HOLogic.mk_setT o HOLogic.mk_prodT) (As' ~~ As');
+ val QTs = map2 mk_pred2T As' Bs';
+
+ val CA' = mk_bnf_T As' CA;
+ val CB' = mk_bnf_T Bs' CA;
+ val CC' = mk_bnf_T Cs CA;
+ val CRs' = mk_bnf_T RTs CA;
+ val CA'CB' = HOLogic.mk_prodT (CA', CB');
+
+ val bnf_map_AsAs = mk_bnf_map As' As';
+ val bnf_map_AsBs = mk_bnf_map As' Bs';
+ val bnf_map_AsCs = mk_bnf_map As' Cs;
+ val bnf_map_BsCs = mk_bnf_map Bs' Cs;
+ val bnf_sets_As = map (mk_bnf_t As') bnf_sets;
+ val bnf_sets_Bs = map (mk_bnf_t Bs') bnf_sets;
+ val bnf_bd_As = mk_bnf_t As' bnf_bd;
+ val bnf_wit_As = map (apsnd (mk_bnf_t As')) bnf_wits;
+
+ val (((((((((((((((((((((((((fs, fs_copy), gs), hs), p), (x, x')), (y, y')), (z, z')), zs), As),
+ As_copy), Xs), B1s), B2s), f1s), f2s), e1s), e2s), p1s), p2s), bs), (Rs, Rs')), Rs_copy), Ss),
+ (Qs, Qs')), _) = lthy
+ |> mk_Frees "f" (map2 (curry (op -->)) As' Bs')
+ ||>> mk_Frees "f" (map2 (curry (op -->)) As' Bs')
+ ||>> mk_Frees "g" (map2 (curry (op -->)) Bs' Cs)
+ ||>> mk_Frees "h" (map2 (curry (op -->)) As' Ts)
+ ||>> yield_singleton (mk_Frees "p") CA'CB'
+ ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "x") CA'
+ ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "y") CB'
+ ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "z") CRs'
+ ||>> mk_Frees "z" As'
+ ||>> mk_Frees "A" (map HOLogic.mk_setT As')
+ ||>> mk_Frees "A" (map HOLogic.mk_setT As')
+ ||>> mk_Frees "A" (map HOLogic.mk_setT domTs)
+ ||>> mk_Frees "B1" (map HOLogic.mk_setT B1Ts)
+ ||>> mk_Frees "B2" (map HOLogic.mk_setT B2Ts)
+ ||>> mk_Frees "f1" (map2 (curry (op -->)) B1Ts ranTs)
+ ||>> mk_Frees "f2" (map2 (curry (op -->)) B2Ts ranTs)
+ ||>> mk_Frees "e1" (map2 (curry (op -->)) B1Ts ranTs')
+ ||>> mk_Frees "e2" (map2 (curry (op -->)) B2Ts ranTs'')
+ ||>> mk_Frees "p1" (map2 (curry (op -->)) domTs B1Ts)
+ ||>> mk_Frees "p2" (map2 (curry (op -->)) domTs B2Ts)
+ ||>> mk_Frees "b" As'
+ ||>> mk_Frees' "R" setRTs
+ ||>> mk_Frees "R" setRTs
+ ||>> mk_Frees "S" setRTsBsCs
+ ||>> mk_Frees' "Q" QTs;
+
+ (*Gr (in R1 .. Rn) (map fst .. fst)^-1 O Gr (in R1 .. Rn) (map snd .. snd)*)
+ val O_Gr =
+ let
+ val map1 = Term.list_comb (mk_bnf_map RTs As', map fst_const RTs);
+ val map2 = Term.list_comb (mk_bnf_map RTs Bs', map snd_const RTs);
+ val bnf_in = mk_in (map Free Rs') (map (mk_bnf_t RTs) bnf_sets) CRs';
+ in
+ mk_rel_comp (mk_converse (mk_Gr bnf_in map1), mk_Gr bnf_in map2)
+ end;
+
+ fun mk_predicate_of_set x_name y_name t =
+ let
+ val (T, U) = HOLogic.dest_prodT (HOLogic.dest_setT (fastype_of t));
+ val x = Free (x_name, T);
+ val y = Free (y_name, U);
+ in fold_rev Term.lambda [x, y] (HOLogic.mk_mem (HOLogic.mk_prod (x, y), t)) end;
+
+ val rel_rhs = (case raw_rel_opt of
+ NONE =>
+ fold_rev absfree Qs' (mk_predicate_of_set (fst x') (fst y')
+ (Term.list_comb (fold_rev Term.absfree Rs' O_Gr, map3 (fn Q => fn T => fn U =>
+ HOLogic.Collect_const (HOLogic.mk_prodT (T, U)) $ HOLogic.mk_split Q) Qs As' Bs')))
+ | SOME raw_rel => prep_term no_defs_lthy raw_rel);
+
+ val rel_bind_def = (fn () => Binding.suffix_name ("_" ^ relN) b, rel_rhs);
+
+ val ((bnf_rel_term, raw_rel_def), (lthy, lthy_old)) =
+ lthy
+ |> maybe_define (is_some raw_rel_opt) rel_bind_def
+ ||> `(maybe_restore lthy);
+
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+ val bnf_rel_def = Morphism.thm phi raw_rel_def;
+ val bnf_rel = Morphism.term phi bnf_rel_term;
+
+ fun mk_bnf_rel QTs CA' CB' = normalize_rel lthy QTs CA' CB' bnf_rel;
+
+ val rel = mk_bnf_rel QTs CA' CB';
+
+ val srel_rhs =
+ fold_rev Term.absfree Rs' (HOLogic.Collect_const CA'CB' $
+ Term.lambda p (Term.list_comb (rel, map (mk_predicate_of_set (fst x') (fst y')) Rs) $
+ HOLogic.mk_fst p $ HOLogic.mk_snd p));
+
+ val srel_bind_def = (fn () => Binding.suffix_name ("_" ^ srelN) b, srel_rhs);
+
+ val ((bnf_srel_term, raw_srel_def), (lthy, lthy_old)) =
+ lthy
+ |> maybe_define false srel_bind_def
+ ||> `(maybe_restore lthy);
+
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+ val bnf_srel_def = Morphism.thm phi raw_srel_def;
+ val bnf_srel = Morphism.term phi bnf_srel_term;
+
+ fun mk_bnf_srel setRTs CA' CB' = normalize_srel lthy setRTs CA' CB' bnf_srel;
+
+ val srel = mk_bnf_srel setRTs CA' CB';
+
+ val _ = case no_reflexive (raw_map_def :: raw_set_defs @ [raw_bd_def] @
+ raw_wit_defs @ [raw_rel_def, raw_srel_def]) of
+ [] => ()
+ | defs => Proof_Display.print_consts true lthy_old (K false)
+ (map (dest_Free o fst o Logic.dest_equals o prop_of) defs);
+
+ val map_id_goal =
+ let
+ val bnf_map_app_id = Term.list_comb (bnf_map_AsAs, map HOLogic.id_const As');
+ in
+ HOLogic.mk_Trueprop
+ (HOLogic.mk_eq (bnf_map_app_id, HOLogic.id_const CA'))
+ end;
+
+ val map_comp_goal =
+ let
+ val bnf_map_app_comp = Term.list_comb (bnf_map_AsCs, map2 (curry HOLogic.mk_comp) gs fs);
+ val comp_bnf_map_app = HOLogic.mk_comp
+ (Term.list_comb (bnf_map_BsCs, gs),
+ Term.list_comb (bnf_map_AsBs, fs));
+ in
+ fold_rev Logic.all (fs @ gs) (mk_Trueprop_eq (bnf_map_app_comp, comp_bnf_map_app))
+ end;
+
+ val map_cong_goal =
+ let
+ fun mk_prem z set f f_copy =
+ Logic.all z (Logic.mk_implies
+ (HOLogic.mk_Trueprop (HOLogic.mk_mem (z, set $ x)),
+ mk_Trueprop_eq (f $ z, f_copy $ z)));
+ val prems = map4 mk_prem zs bnf_sets_As fs fs_copy;
+ val eq = HOLogic.mk_eq (Term.list_comb (bnf_map_AsBs, fs) $ x,
+ Term.list_comb (bnf_map_AsBs, fs_copy) $ x);
+ in
+ fold_rev Logic.all (x :: fs @ fs_copy)
+ (Logic.list_implies (prems, HOLogic.mk_Trueprop eq))
+ end;
+
+ val set_naturals_goal =
+ let
+ fun mk_goal setA setB f =
+ let
+ val set_comp_map =
+ HOLogic.mk_comp (setB, Term.list_comb (bnf_map_AsBs, fs));
+ val image_comp_set = HOLogic.mk_comp (mk_image f, setA);
+ in
+ fold_rev Logic.all fs (mk_Trueprop_eq (set_comp_map, image_comp_set))
+ end;
+ in
+ map3 mk_goal bnf_sets_As bnf_sets_Bs fs
+ end;
+
+ val card_order_bd_goal = HOLogic.mk_Trueprop (mk_card_order bnf_bd_As);
+
+ val cinfinite_bd_goal = HOLogic.mk_Trueprop (mk_cinfinite bnf_bd_As);
+
+ val set_bds_goal =
+ let
+ fun mk_goal set =
+ Logic.all x (HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of (set $ x)) bnf_bd_As));
+ in
+ map mk_goal bnf_sets_As
+ end;
+
+ val in_bd_goal =
+ let
+ val bd = mk_cexp
+ (if live = 0 then ctwo
+ else mk_csum (Library.foldr1 (uncurry mk_csum) (map mk_card_of As)) ctwo)
+ bnf_bd_As;
+ in
+ fold_rev Logic.all As
+ (HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of (mk_in As bnf_sets_As CA')) bd))
+ end;
+
+ val map_wpull_goal =
+ let
+ val prems = map HOLogic.mk_Trueprop
+ (map8 mk_wpull Xs B1s B2s f1s f2s (replicate live NONE) p1s p2s);
+ val CX = mk_bnf_T domTs CA;
+ val CB1 = mk_bnf_T B1Ts CA;
+ val CB2 = mk_bnf_T B2Ts CA;
+ val bnf_sets_CX = map2 (normalize_set (map (mk_bnf_T domTs) CA_params)) domTs bnf_sets;
+ val bnf_sets_CB1 = map2 (normalize_set (map (mk_bnf_T B1Ts) CA_params)) B1Ts bnf_sets;
+ val bnf_sets_CB2 = map2 (normalize_set (map (mk_bnf_T B2Ts) CA_params)) B2Ts bnf_sets;
+ val bnf_map_app_f1 = Term.list_comb (mk_bnf_map B1Ts ranTs, f1s);
+ val bnf_map_app_f2 = Term.list_comb (mk_bnf_map B2Ts ranTs, f2s);
+ val bnf_map_app_p1 = Term.list_comb (mk_bnf_map domTs B1Ts, p1s);
+ val bnf_map_app_p2 = Term.list_comb (mk_bnf_map domTs B2Ts, p2s);
+
+ val map_wpull = mk_wpull (mk_in Xs bnf_sets_CX CX)
+ (mk_in B1s bnf_sets_CB1 CB1) (mk_in B2s bnf_sets_CB2 CB2)
+ bnf_map_app_f1 bnf_map_app_f2 NONE bnf_map_app_p1 bnf_map_app_p2;
+ in
+ fold_rev Logic.all (Xs @ B1s @ B2s @ f1s @ f2s @ p1s @ p2s)
+ (Logic.list_implies (prems, HOLogic.mk_Trueprop map_wpull))
+ end;
+
+ val srel_O_Gr_goal = fold_rev Logic.all Rs (mk_Trueprop_eq (Term.list_comb (srel, Rs), O_Gr));
+
+ val goals = zip_axioms map_id_goal map_comp_goal map_cong_goal set_naturals_goal
+ card_order_bd_goal cinfinite_bd_goal set_bds_goal in_bd_goal map_wpull_goal srel_O_Gr_goal;
+
+ fun mk_wit_goals (I, wit) =
+ let
+ val xs = map (nth bs) I;
+ fun wit_goal i =
+ let
+ val z = nth zs i;
+ val set_wit = nth bnf_sets_As i $ Term.list_comb (wit, xs);
+ val concl = HOLogic.mk_Trueprop
+ (if member (op =) I i then HOLogic.mk_eq (z, nth bs i)
+ else @{term False});
+ in
+ fold_rev Logic.all (z :: xs)
+ (Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (z, set_wit)), concl))
+ end;
+ in
+ map wit_goal (0 upto live - 1)
+ end;
+
+ val wit_goalss = map mk_wit_goals bnf_wit_As;
+
+ fun after_qed thms lthy =
+ let
+ val (axioms, wit_thms) = apfst (mk_axioms live) (chop (length goals) thms);
+
+ val bd_Card_order = #bd_card_order axioms RS @{thm conjunct2[OF card_order_on_Card_order]};
+ val bd_Cinfinite = @{thm conjI} OF [#bd_cinfinite axioms, bd_Card_order];
+ val bd_Cnotzero = bd_Cinfinite RS @{thm Cinfinite_Cnotzero};
+
+ fun mk_lazy f = if fact_policy <> Derive_Few_Facts then Lazy.value (f ()) else Lazy.lazy f;
+
+ fun mk_collect_set_natural () =
+ let
+ val defT = mk_bnf_T Ts CA --> HOLogic.mk_setT T;
+ val collect_map = HOLogic.mk_comp
+ (mk_collect (map (mk_bnf_t Ts) bnf_sets) defT,
+ Term.list_comb (mk_bnf_map As' Ts, hs));
+ val image_collect = mk_collect
+ (map2 (fn h => fn set => HOLogic.mk_comp (mk_image h, set)) hs bnf_sets_As)
+ defT;
+ (*collect {set1 ... setm} o map f1 ... fm = collect {f1` o set1 ... fm` o setm}*)
+ val goal = fold_rev Logic.all hs (mk_Trueprop_eq (collect_map, image_collect));
+ in
+ Skip_Proof.prove lthy [] [] goal
+ (fn {context = ctxt, ...} => mk_collect_set_natural_tac ctxt (#set_natural axioms))
+ |> Thm.close_derivation
+ end;
+
+ val collect_set_natural = mk_lazy mk_collect_set_natural;
+
+ fun mk_in_mono () =
+ let
+ val prems_mono = map2 (HOLogic.mk_Trueprop oo mk_subset) As As_copy;
+ val in_mono_goal =
+ fold_rev Logic.all (As @ As_copy)
+ (Logic.list_implies (prems_mono, HOLogic.mk_Trueprop
+ (mk_subset (mk_in As bnf_sets_As CA') (mk_in As_copy bnf_sets_As CA'))));
+ in
+ Skip_Proof.prove lthy [] [] in_mono_goal (K (mk_in_mono_tac live))
+ |> Thm.close_derivation
+ end;
+
+ val in_mono = mk_lazy mk_in_mono;
+
+ fun mk_in_cong () =
+ let
+ val prems_cong = map2 (HOLogic.mk_Trueprop oo curry HOLogic.mk_eq) As As_copy;
+ val in_cong_goal =
+ fold_rev Logic.all (As @ As_copy)
+ (Logic.list_implies (prems_cong, HOLogic.mk_Trueprop
+ (HOLogic.mk_eq (mk_in As bnf_sets_As CA', mk_in As_copy bnf_sets_As CA'))));
+ in
+ Skip_Proof.prove lthy [] [] in_cong_goal (K ((TRY o hyp_subst_tac THEN' rtac refl) 1))
+ |> Thm.close_derivation
+ end;
+
+ val in_cong = mk_lazy mk_in_cong;
+
+ val map_id' = mk_lazy (fn () => mk_id' (#map_id axioms));
+ val map_comp' = mk_lazy (fn () => mk_comp' (#map_comp axioms));
+
+ val set_natural' =
+ map (fn thm => mk_lazy (fn () => mk_set_natural' thm)) (#set_natural axioms);
+
+ fun mk_map_wppull () =
+ let
+ val prems = if live = 0 then [] else
+ [HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+ (map8 mk_wpull Xs B1s B2s f1s f2s (map SOME (e1s ~~ e2s)) p1s p2s))];
+ val CX = mk_bnf_T domTs CA;
+ val CB1 = mk_bnf_T B1Ts CA;
+ val CB2 = mk_bnf_T B2Ts CA;
+ val bnf_sets_CX =
+ map2 (normalize_set (map (mk_bnf_T domTs) CA_params)) domTs bnf_sets;
+ val bnf_sets_CB1 =
+ map2 (normalize_set (map (mk_bnf_T B1Ts) CA_params)) B1Ts bnf_sets;
+ val bnf_sets_CB2 =
+ map2 (normalize_set (map (mk_bnf_T B2Ts) CA_params)) B2Ts bnf_sets;
+ val bnf_map_app_f1 = Term.list_comb (mk_bnf_map B1Ts ranTs, f1s);
+ val bnf_map_app_f2 = Term.list_comb (mk_bnf_map B2Ts ranTs, f2s);
+ val bnf_map_app_e1 = Term.list_comb (mk_bnf_map B1Ts ranTs', e1s);
+ val bnf_map_app_e2 = Term.list_comb (mk_bnf_map B2Ts ranTs'', e2s);
+ val bnf_map_app_p1 = Term.list_comb (mk_bnf_map domTs B1Ts, p1s);
+ val bnf_map_app_p2 = Term.list_comb (mk_bnf_map domTs B2Ts, p2s);
+
+ val concl = mk_wpull (mk_in Xs bnf_sets_CX CX)
+ (mk_in B1s bnf_sets_CB1 CB1) (mk_in B2s bnf_sets_CB2 CB2)
+ bnf_map_app_f1 bnf_map_app_f2 (SOME (bnf_map_app_e1, bnf_map_app_e2))
+ bnf_map_app_p1 bnf_map_app_p2;
+
+ val goal =
+ fold_rev Logic.all (Xs @ B1s @ B2s @ f1s @ f2s @ e1s @ e2s @ p1s @ p2s)
+ (Logic.list_implies (prems, HOLogic.mk_Trueprop concl))
+ in
+ Skip_Proof.prove lthy [] [] goal
+ (fn _ => mk_map_wppull_tac (#map_id axioms) (#map_cong axioms)
+ (#map_wpull axioms) (Lazy.force map_comp') (map Lazy.force set_natural'))
+ |> Thm.close_derivation
+ end;
+
+ val srel_O_Grs = no_refl [#srel_O_Gr axioms];
+
+ val map_wppull = mk_lazy mk_map_wppull;
+
+ fun mk_srel_Gr () =
+ let
+ val lhs = Term.list_comb (srel, map2 mk_Gr As fs);
+ val rhs = mk_Gr (mk_in As bnf_sets_As CA') (Term.list_comb (bnf_map_AsBs, fs));
+ val goal = fold_rev Logic.all (As @ fs) (mk_Trueprop_eq (lhs, rhs));
+ in
+ Skip_Proof.prove lthy [] [] goal
+ (mk_srel_Gr_tac srel_O_Grs (#map_id axioms) (#map_cong axioms) (Lazy.force map_id')
+ (Lazy.force map_comp') (map Lazy.force set_natural'))
+ |> Thm.close_derivation
+ end;
+
+ val srel_Gr = mk_lazy mk_srel_Gr;
+
+ fun mk_srel_prems f = map2 (HOLogic.mk_Trueprop oo f) Rs Rs_copy
+ fun mk_srel_concl f = HOLogic.mk_Trueprop
+ (f (Term.list_comb (srel, Rs), Term.list_comb (srel, Rs_copy)));
+
+ fun mk_srel_mono () =
+ let
+ val mono_prems = mk_srel_prems mk_subset;
+ val mono_concl = mk_srel_concl (uncurry mk_subset);
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (mono_prems, mono_concl)))
+ (mk_srel_mono_tac srel_O_Grs (Lazy.force in_mono))
+ |> Thm.close_derivation
+ end;
+
+ fun mk_srel_cong () =
+ let
+ val cong_prems = mk_srel_prems (curry HOLogic.mk_eq);
+ val cong_concl = mk_srel_concl HOLogic.mk_eq;
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (cong_prems, cong_concl)))
+ (fn _ => (TRY o hyp_subst_tac THEN' rtac refl) 1)
+ |> Thm.close_derivation
+ end;
+
+ val srel_mono = mk_lazy mk_srel_mono;
+ val srel_cong = mk_lazy mk_srel_cong;
+
+ fun mk_srel_Id () =
+ let val relAsAs = mk_bnf_srel self_setRTs CA' CA' in
+ Skip_Proof.prove lthy [] []
+ (HOLogic.mk_Trueprop
+ (HOLogic.mk_eq (Term.list_comb (relAsAs, map Id_const As'), Id_const CA')))
+ (mk_srel_Id_tac live (Lazy.force srel_Gr) (#map_id axioms))
+ |> Thm.close_derivation
+ end;
+
+ val srel_Id = mk_lazy mk_srel_Id;
+
+ fun mk_srel_converse () =
+ let
+ val relBsAs = mk_bnf_srel setRT's CB' CA';
+ val lhs = Term.list_comb (relBsAs, map mk_converse Rs);
+ val rhs = mk_converse (Term.list_comb (srel, Rs));
+ val le_goal = fold_rev Logic.all Rs (HOLogic.mk_Trueprop (mk_subset lhs rhs));
+ val le_thm = Skip_Proof.prove lthy [] [] le_goal
+ (mk_srel_converse_le_tac srel_O_Grs (Lazy.force srel_Id) (#map_cong axioms)
+ (Lazy.force map_comp') (map Lazy.force set_natural'))
+ |> Thm.close_derivation
+ val goal = fold_rev Logic.all Rs (mk_Trueprop_eq (lhs, rhs));
+ in
+ Skip_Proof.prove lthy [] [] goal (fn _ => mk_srel_converse_tac le_thm)
+ |> Thm.close_derivation
+ end;
+
+ val srel_converse = mk_lazy mk_srel_converse;
+
+ fun mk_srel_O () =
+ let
+ val relAsCs = mk_bnf_srel setRTsAsCs CA' CC';
+ val relBsCs = mk_bnf_srel setRTsBsCs CB' CC';
+ val lhs = Term.list_comb (relAsCs, map2 (curry mk_rel_comp) Rs Ss);
+ val rhs = mk_rel_comp (Term.list_comb (srel, Rs), Term.list_comb (relBsCs, Ss));
+ val goal = fold_rev Logic.all (Rs @ Ss) (mk_Trueprop_eq (lhs, rhs));
+ in
+ Skip_Proof.prove lthy [] [] goal
+ (mk_srel_O_tac srel_O_Grs (Lazy.force srel_Id) (#map_cong axioms)
+ (Lazy.force map_wppull) (Lazy.force map_comp') (map Lazy.force set_natural'))
+ |> Thm.close_derivation
+ end;
+
+ val srel_O = mk_lazy mk_srel_O;
+
+ fun mk_in_srel () =
+ let
+ val bnf_in = mk_in Rs (map (mk_bnf_t RTs) bnf_sets) CRs';
+ val map1 = Term.list_comb (mk_bnf_map RTs As', map fst_const RTs);
+ val map2 = Term.list_comb (mk_bnf_map RTs Bs', map snd_const RTs);
+ val map_fst_eq = HOLogic.mk_eq (map1 $ z, x);
+ val map_snd_eq = HOLogic.mk_eq (map2 $ z, y);
+ val lhs = HOLogic.mk_mem (HOLogic.mk_prod (x, y), Term.list_comb (srel, Rs));
+ val rhs =
+ HOLogic.mk_exists (fst z', snd z', HOLogic.mk_conj (HOLogic.mk_mem (z, bnf_in),
+ HOLogic.mk_conj (map_fst_eq, map_snd_eq)));
+ val goal =
+ fold_rev Logic.all (x :: y :: Rs) (mk_Trueprop_eq (lhs, rhs));
+ in
+ Skip_Proof.prove lthy [] [] goal (mk_in_srel_tac srel_O_Grs (length bnf_sets))
+ |> Thm.close_derivation
+ end;
+
+ val in_srel = mk_lazy mk_in_srel;
+
+ val defs = mk_defs bnf_map_def bnf_set_defs bnf_rel_def bnf_srel_def;
+
+ val facts = mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_natural in_cong
+ in_mono in_srel map_comp' map_id' map_wppull set_natural' srel_cong srel_mono srel_Id
+ srel_Gr srel_converse srel_O;
+
+ val wits = map2 mk_witness bnf_wits wit_thms;
+
+ val bnf_rel =
+ Term.subst_atomic_types ((Ds ~~ deads) @ (As' ~~ alphas) @ (Bs' ~~ betas)) rel;
+ val bnf_srel =
+ Term.subst_atomic_types ((Ds ~~ deads) @ (As' ~~ alphas) @ (Bs' ~~ betas)) srel;
+
+ val bnf = mk_bnf b CA live alphas betas dead deads bnf_map bnf_sets bnf_bd axioms defs facts
+ wits bnf_rel bnf_srel;
+ in
+ (bnf, lthy
+ |> (if fact_policy = Note_All_Facts_and_Axioms then
+ let
+ val witNs = if length wits = 1 then [witN] else map mk_witN (1 upto length wits);
+ val notes =
+ [(bd_card_orderN, [#bd_card_order axioms]),
+ (bd_cinfiniteN, [#bd_cinfinite axioms]),
+ (bd_Card_orderN, [#bd_Card_order facts]),
+ (bd_CinfiniteN, [#bd_Cinfinite facts]),
+ (bd_CnotzeroN, [#bd_Cnotzero facts]),
+ (collect_set_naturalN, [Lazy.force (#collect_set_natural facts)]),
+ (in_bdN, [#in_bd axioms]),
+ (in_monoN, [Lazy.force (#in_mono facts)]),
+ (in_srelN, [Lazy.force (#in_srel facts)]),
+ (map_compN, [#map_comp axioms]),
+ (map_idN, [#map_id axioms]),
+ (map_wpullN, [#map_wpull axioms]),
+ (set_naturalN, #set_natural axioms),
+ (set_bdN, #set_bd axioms)] @
+ map2 pair witNs wit_thms
+ |> map (fn (thmN, thms) =>
+ ((qualify (Binding.qualify true (Binding.name_of b) (Binding.name thmN)), []),
+ [(thms, [])]));
+ in
+ Local_Theory.notes notes #> snd
+ end
+ else
+ I)
+ |> (if fact_policy = Note_All_Facts_and_Axioms orelse
+ fact_policy = Derive_All_Facts_Note_Most then
+ let
+ val notes =
+ [(map_comp'N, [Lazy.force (#map_comp' facts)]),
+ (map_congN, [#map_cong axioms]),
+ (map_id'N, [Lazy.force (#map_id' facts)]),
+ (set_natural'N, map Lazy.force (#set_natural' facts)),
+ (srel_O_GrN, srel_O_Grs),
+ (srel_IdN, [Lazy.force (#srel_Id facts)]),
+ (srel_GrN, [Lazy.force (#srel_Gr facts)]),
+ (srel_converseN, [Lazy.force (#srel_converse facts)]),
+ (srel_monoN, [Lazy.force (#srel_mono facts)]),
+ (srel_ON, [Lazy.force (#srel_O facts)])]
+ |> filter_out (null o #2)
+ |> map (fn (thmN, thms) =>
+ ((qualify (Binding.qualify true (Binding.name_of b) (Binding.name thmN)), []),
+ [(thms, [])]));
+ in
+ Local_Theory.notes notes #> snd
+ end
+ else
+ I))
+ end;
+
+ val one_step_defs =
+ no_reflexive (bnf_map_def :: bnf_bd_def :: bnf_set_defs @ bnf_wit_defs @ [bnf_rel_def,
+ bnf_srel_def]);
+ in
+ (key, goals, wit_goalss, after_qed, lthy, one_step_defs)
+ end;
+
+fun register_bnf key (bnf, lthy) =
+ (bnf, Local_Theory.declaration {syntax = false, pervasive = true}
+ (fn phi => Data.map (Symtab.update_new (key, morph_bnf phi bnf))) lthy);
+
+(* TODO: Once the invariant "nwits > 0" holds, remove "mk_conjunction_balanced'" and "rtac TrueI"
+ below *)
+fun mk_conjunction_balanced' [] = @{prop True}
+ | mk_conjunction_balanced' ts = Logic.mk_conjunction_balanced ts;
+
+fun bnf_def const_policy fact_policy qualify tacs wit_tac Ds =
+ (fn (_, goals, wit_goalss, after_qed, lthy, one_step_defs) =>
+ let
+ val wits_tac =
+ K (TRYALL Goal.conjunction_tac) THEN' K (TRYALL (rtac TrueI)) THEN'
+ mk_unfold_thms_then_tac lthy one_step_defs wit_tac;
+ val wit_goals = map mk_conjunction_balanced' wit_goalss;
+ val wit_thms =
+ Skip_Proof.prove lthy [] [] (mk_conjunction_balanced' wit_goals) wits_tac
+ |> Conjunction.elim_balanced (length wit_goals)
+ |> map2 (Conjunction.elim_balanced o length) wit_goalss
+ |> map (map (Thm.close_derivation o Thm.forall_elim_vars 0));
+ in
+ map2 (Thm.close_derivation oo Skip_Proof.prove lthy [] [])
+ goals (map (mk_unfold_thms_then_tac lthy one_step_defs) tacs)
+ |> (fn thms => after_qed (map single thms @ wit_thms) lthy)
+ end) oo prepare_def const_policy fact_policy qualify (K I) Ds;
+
+val bnf_def_cmd = (fn (key, goals, wit_goals, after_qed, lthy, defs) =>
+ Proof.unfolding ([[(defs, [])]])
+ (Proof.theorem NONE (snd o register_bnf key oo after_qed)
+ (map (single o rpair []) goals @ map (map (rpair [])) wit_goals) lthy)) oo
+ prepare_def Do_Inline (user_policy Derive_All_Facts_Note_Most) I Syntax.read_term NONE;
+
+fun print_bnfs ctxt =
+ let
+ fun pretty_set sets i = Pretty.block
+ [Pretty.str (mk_setN (i + 1) ^ ":"), Pretty.brk 1,
+ Pretty.quote (Syntax.pretty_term ctxt (nth sets i))];
+
+ fun pretty_bnf (key, BNF {T = T, map = map, sets = sets, bd = bd,
+ live = live, lives = lives, dead = dead, deads = deads, ...}) =
+ Pretty.big_list
+ (Pretty.string_of (Pretty.block [Pretty.str key, Pretty.str ":", Pretty.brk 1,
+ Pretty.quote (Syntax.pretty_typ ctxt T)]))
+ ([Pretty.block [Pretty.str "live:", Pretty.brk 1, Pretty.str (string_of_int live),
+ Pretty.brk 3, Pretty.list "[" "]" (List.map (Syntax.pretty_typ ctxt) lives)],
+ Pretty.block [Pretty.str "dead:", Pretty.brk 1, Pretty.str (string_of_int dead),
+ Pretty.brk 3, Pretty.list "[" "]" (List.map (Syntax.pretty_typ ctxt) deads)],
+ Pretty.block [Pretty.str (mapN ^ ":"), Pretty.brk 1,
+ Pretty.quote (Syntax.pretty_term ctxt map)]] @
+ List.map (pretty_set sets) (0 upto length sets - 1) @
+ [Pretty.block [Pretty.str (bdN ^ ":"), Pretty.brk 1,
+ Pretty.quote (Syntax.pretty_term ctxt bd)]]);
+ in
+ Pretty.big_list "BNFs:" (map pretty_bnf (Symtab.dest (Data.get (Context.Proof ctxt))))
+ |> Pretty.writeln
+ end;
+
+val _ =
+ Outer_Syntax.improper_command @{command_spec "print_bnfs"} "print all BNFs"
+ (Scan.succeed (Toplevel.keep (print_bnfs o Toplevel.context_of)));
+
+val _ =
+ Outer_Syntax.local_theory_to_proof @{command_spec "bnf_def"} "define a BNF for an existing type"
+ ((parse_opt_binding_colon -- Parse.term --
+ (@{keyword "["} |-- Parse.list Parse.term --| @{keyword "]"}) -- Parse.term --
+ (@{keyword "["} |-- Parse.list Parse.term --| @{keyword "]"}) -- Scan.option Parse.term)
+ >> bnf_def_cmd);
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Tools/bnf_def_tactics.ML Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,209 @@
+(* Title: HOL/BNF/Tools/bnf_def_tactics.ML
+ Author: Dmitriy Traytel, TU Muenchen
+ Author: Jasmin Blanchette, TU Muenchen
+ Copyright 2012
+
+Tactics for definition of bounded natural functors.
+*)
+
+signature BNF_DEF_TACTICS =
+sig
+ val mk_collect_set_natural_tac: Proof.context -> thm list -> tactic
+ val mk_id': thm -> thm
+ val mk_comp': thm -> thm
+ val mk_in_mono_tac: int -> tactic
+ val mk_map_wppull_tac: thm -> thm -> thm -> thm -> thm list -> tactic
+ val mk_set_natural': thm -> thm
+
+ val mk_srel_Gr_tac: thm list -> thm -> thm -> thm -> thm -> thm list ->
+ {prems: thm list, context: Proof.context} -> tactic
+ val mk_srel_Id_tac: int -> thm -> thm -> {prems: 'a, context: Proof.context} -> tactic
+ val mk_srel_O_tac: thm list -> thm -> thm -> thm -> thm -> thm list ->
+ {prems: thm list, context: Proof.context} -> tactic
+ val mk_in_srel_tac: thm list -> int -> {prems: 'b, context: Proof.context} -> tactic
+ val mk_srel_converse_tac: thm -> tactic
+ val mk_srel_converse_le_tac: thm list -> thm -> thm -> thm -> thm list ->
+ {prems: thm list, context: Proof.context} -> tactic
+ val mk_srel_mono_tac: thm list -> thm -> {prems: 'a, context: Proof.context} -> tactic
+end;
+
+structure BNF_Def_Tactics : BNF_DEF_TACTICS =
+struct
+
+open BNF_Util
+open BNF_Tactics
+
+fun mk_id' id = mk_trans (fun_cong OF [id]) @{thm id_apply};
+fun mk_comp' comp = @{thm o_eq_dest_lhs} OF [mk_sym comp];
+fun mk_set_natural' set_natural = set_natural RS @{thm pointfreeE};
+fun mk_in_mono_tac n = if n = 0 then rtac subset_UNIV 1
+ else (rtac subsetI THEN'
+ rtac CollectI) 1 THEN
+ REPEAT_DETERM (eresolve_tac [CollectE, conjE] 1) THEN
+ REPEAT_DETERM_N (n - 1)
+ ((rtac conjI THEN' etac subset_trans THEN' atac) 1) THEN
+ (etac subset_trans THEN' atac) 1;
+
+fun mk_collect_set_natural_tac ctxt set_naturals =
+ substs_tac ctxt (@{thms collect_o image_insert image_empty} @ set_naturals) 1 THEN rtac refl 1;
+
+fun mk_map_wppull_tac map_id map_cong map_wpull map_comp set_naturals =
+ if null set_naturals then
+ EVERY' [rtac @{thm wppull_id}, rtac map_wpull, rtac map_id, rtac map_id] 1
+ else EVERY' [REPEAT_DETERM o etac conjE, REPEAT_DETERM o dtac @{thm wppull_thePull},
+ REPEAT_DETERM o etac exE, rtac @{thm wpull_wppull}, rtac map_wpull,
+ REPEAT_DETERM o rtac @{thm wpull_thePull}, rtac ballI,
+ REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac conjI, rtac CollectI,
+ CONJ_WRAP' (fn thm => EVERY' [rtac (thm RS @{thm ord_eq_le_trans}),
+ rtac @{thm image_subsetI}, rtac conjunct1, etac bspec, etac set_mp, atac])
+ set_naturals,
+ CONJ_WRAP' (fn thm => EVERY' [rtac (map_comp RS trans), rtac (map_comp RS trans),
+ rtac (map_comp RS trans RS sym), rtac map_cong,
+ REPEAT_DETERM_N (length set_naturals) o EVERY' [rtac (o_apply RS trans),
+ rtac (o_apply RS trans RS sym), rtac (o_apply RS trans), rtac thm,
+ rtac conjunct2, etac bspec, etac set_mp, atac]]) [conjunct1, conjunct2]] 1;
+
+fun mk_srel_Gr_tac srel_O_Grs map_id map_cong map_id' map_comp set_naturals
+ {context = ctxt, prems = _} =
+ let
+ val n = length set_naturals;
+ in
+ if null set_naturals then
+ unfold_thms_tac ctxt srel_O_Grs THEN EVERY' [rtac @{thm Gr_UNIV_id}, rtac map_id] 1
+ else unfold_thms_tac ctxt (@{thm Gr_def} :: srel_O_Grs) THEN
+ EVERY' [rtac equalityI, rtac subsetI,
+ REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE, @{thm relcompE}, @{thm converseE}],
+ REPEAT_DETERM o dtac Pair_eqD,
+ REPEAT_DETERM o etac conjE, hyp_subst_tac,
+ rtac CollectI, rtac exI, rtac conjI, rtac Pair_eqI, rtac conjI, rtac refl,
+ rtac sym, rtac trans, rtac map_comp, rtac map_cong,
+ REPEAT_DETERM_N n o EVERY' [dtac @{thm set_rev_mp}, atac,
+ REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac,
+ rtac (o_apply RS trans), rtac (@{thm fst_conv} RS arg_cong RS trans),
+ rtac (@{thm snd_conv} RS sym)],
+ rtac CollectI,
+ CONJ_WRAP' (fn thm => EVERY' [rtac (thm RS @{thm ord_eq_le_trans}),
+ rtac @{thm image_subsetI}, dtac @{thm set_rev_mp}, atac,
+ REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac,
+ stac @{thm fst_conv}, atac]) set_naturals,
+ rtac @{thm subrelI}, etac CollectE, REPEAT_DETERM o eresolve_tac [exE, conjE],
+ REPEAT_DETERM o dtac Pair_eqD,
+ REPEAT_DETERM o etac conjE, hyp_subst_tac,
+ rtac @{thm relcompI}, rtac @{thm converseI},
+ EVERY' (map2 (fn convol => fn map_id =>
+ EVERY' [rtac CollectI, rtac exI, rtac conjI,
+ rtac Pair_eqI, rtac conjI, rtac refl, rtac sym,
+ rtac (box_equals OF [map_cong, map_comp RS sym, map_id]),
+ REPEAT_DETERM_N n o rtac (convol RS fun_cong),
+ REPEAT_DETERM o eresolve_tac [CollectE, conjE],
+ rtac CollectI,
+ CONJ_WRAP' (fn thm =>
+ EVERY' [rtac @{thm ord_eq_le_trans}, rtac thm, rtac @{thm image_subsetI},
+ rtac @{thm convol_memI[of id _ "%x. x", OF id_apply refl]}, etac set_mp, atac])
+ set_naturals])
+ @{thms fst_convol snd_convol} [map_id', refl])] 1
+ end;
+
+fun mk_srel_Id_tac n srel_Gr map_id {context = ctxt, prems = _} =
+ unfold_thms_tac ctxt [srel_Gr, @{thm Id_alt}] THEN
+ subst_tac ctxt [map_id] 1 THEN
+ (if n = 0 then rtac refl 1
+ else EVERY' [rtac @{thm arg_cong2[of _ _ _ _ Gr]},
+ rtac equalityI, rtac subset_UNIV, rtac subsetI, rtac CollectI,
+ CONJ_WRAP' (K (rtac subset_UNIV)) (1 upto n), rtac refl] 1);
+
+fun mk_srel_mono_tac srel_O_Grs in_mono {context = ctxt, prems = _} =
+ unfold_thms_tac ctxt srel_O_Grs THEN
+ EVERY' [rtac @{thm relcomp_mono}, rtac @{thm iffD2[OF converse_mono]},
+ rtac @{thm Gr_mono}, rtac in_mono, REPEAT_DETERM o atac,
+ rtac @{thm Gr_mono}, rtac in_mono, REPEAT_DETERM o atac] 1;
+
+fun mk_srel_converse_le_tac srel_O_Grs srel_Id map_cong map_comp set_naturals
+ {context = ctxt, prems = _} =
+ let
+ val n = length set_naturals;
+ in
+ if null set_naturals then
+ unfold_thms_tac ctxt [srel_Id] THEN rtac equalityD2 1 THEN rtac @{thm converse_Id} 1
+ else unfold_thms_tac ctxt (@{thm Gr_def} :: srel_O_Grs) THEN
+ EVERY' [rtac @{thm subrelI},
+ REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE, @{thm relcompE}, @{thm converseE}],
+ REPEAT_DETERM o dtac Pair_eqD,
+ REPEAT_DETERM o etac conjE, hyp_subst_tac, rtac @{thm converseI},
+ rtac @{thm relcompI}, rtac @{thm converseI},
+ EVERY' (map (fn thm => EVERY' [rtac CollectI, rtac exI,
+ rtac conjI, rtac Pair_eqI, rtac conjI, rtac refl, rtac trans,
+ rtac map_cong, REPEAT_DETERM_N n o rtac thm,
+ rtac (map_comp RS sym), rtac CollectI,
+ CONJ_WRAP' (fn thm => EVERY' [rtac (thm RS @{thm ord_eq_le_trans}),
+ etac @{thm flip_rel}]) set_naturals]) [@{thm snd_fst_flip}, @{thm fst_snd_flip}])] 1
+ end;
+
+fun mk_srel_converse_tac le_converse =
+ EVERY' [rtac equalityI, rtac le_converse, rtac @{thm xt1(6)}, rtac @{thm converse_shift},
+ rtac le_converse, REPEAT_DETERM o stac @{thm converse_converse}, rtac subset_refl] 1;
+
+fun mk_srel_O_tac srel_O_Grs srel_Id map_cong map_wppull map_comp set_naturals
+ {context = ctxt, prems = _} =
+ let
+ val n = length set_naturals;
+ fun in_tac nthO_in = rtac CollectI THEN'
+ CONJ_WRAP' (fn thm => EVERY' [rtac (thm RS @{thm ord_eq_le_trans}),
+ rtac @{thm image_subsetI}, rtac nthO_in, etac set_mp, atac]) set_naturals;
+ in
+ if null set_naturals then unfold_thms_tac ctxt [srel_Id] THEN rtac (@{thm Id_O_R} RS sym) 1
+ else unfold_thms_tac ctxt (@{thm Gr_def} :: srel_O_Grs) THEN
+ EVERY' [rtac equalityI, rtac @{thm subrelI},
+ REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE, @{thm relcompE}, @{thm converseE}],
+ REPEAT_DETERM o dtac Pair_eqD,
+ REPEAT_DETERM o etac conjE, hyp_subst_tac,
+ rtac @{thm relcompI}, rtac @{thm relcompI}, rtac @{thm converseI},
+ rtac CollectI, rtac exI, rtac conjI, rtac Pair_eqI, rtac conjI, rtac refl,
+ rtac sym, rtac trans, rtac map_comp, rtac sym, rtac map_cong,
+ REPEAT_DETERM_N n o rtac @{thm fst_fstO},
+ in_tac @{thm fstO_in},
+ rtac CollectI, rtac exI, rtac conjI, rtac Pair_eqI, rtac conjI, rtac refl,
+ rtac sym, rtac trans, rtac map_comp, rtac map_cong,
+ REPEAT_DETERM_N n o EVERY' [rtac trans, rtac o_apply, rtac ballE, rtac subst,
+ rtac @{thm csquare_def}, rtac @{thm csquare_fstO_sndO}, atac, etac notE,
+ etac set_mp, atac],
+ in_tac @{thm fstO_in},
+ rtac @{thm relcompI}, rtac @{thm converseI},
+ rtac CollectI, rtac exI, rtac conjI, rtac Pair_eqI, rtac conjI, rtac refl,
+ rtac sym, rtac trans, rtac map_comp, rtac map_cong,
+ REPEAT_DETERM_N n o rtac o_apply,
+ in_tac @{thm sndO_in},
+ rtac CollectI, rtac exI, rtac conjI, rtac Pair_eqI, rtac conjI, rtac refl,
+ rtac sym, rtac trans, rtac map_comp, rtac sym, rtac map_cong,
+ REPEAT_DETERM_N n o rtac @{thm snd_sndO},
+ in_tac @{thm sndO_in},
+ rtac @{thm subrelI},
+ REPEAT_DETERM o eresolve_tac [CollectE, @{thm relcompE}, @{thm converseE}],
+ REPEAT_DETERM o eresolve_tac [exE, conjE],
+ REPEAT_DETERM o dtac Pair_eqD,
+ REPEAT_DETERM o etac conjE, hyp_subst_tac,
+ rtac allE, rtac subst, rtac @{thm wppull_def}, rtac map_wppull,
+ CONJ_WRAP' (K (rtac @{thm wppull_fstO_sndO})) set_naturals,
+ etac allE, etac impE, etac conjI, etac conjI, atac,
+ REPEAT_DETERM o eresolve_tac [bexE, conjE],
+ rtac @{thm relcompI}, rtac @{thm converseI},
+ EVERY' (map (fn thm => EVERY' [rtac CollectI, rtac exI,
+ rtac conjI, rtac Pair_eqI, rtac conjI, rtac refl, rtac sym, rtac trans,
+ rtac trans, rtac map_cong, REPEAT_DETERM_N n o rtac thm,
+ rtac (map_comp RS sym), atac, atac]) [@{thm fst_fstO}, @{thm snd_sndO}])] 1
+ end;
+
+fun mk_in_srel_tac srel_O_Grs m {context = ctxt, prems = _} =
+ let
+ val ls' = replicate (Int.max (1, m)) ();
+ in
+ unfold_thms_tac ctxt (srel_O_Grs @
+ @{thms Gr_def converse_unfold relcomp_unfold mem_Collect_eq prod.cases Pair_eq}) THEN
+ EVERY' [rtac iffI, REPEAT_DETERM o eresolve_tac [exE, conjE], hyp_subst_tac, rtac exI,
+ rtac conjI, CONJ_WRAP' (K atac) ls', rtac conjI, rtac refl, rtac refl,
+ REPEAT_DETERM o eresolve_tac [exE, conjE], rtac exI, rtac conjI,
+ REPEAT_DETERM_N 2 o EVERY' [rtac exI, rtac conjI, etac @{thm conjI[OF refl sym]},
+ CONJ_WRAP' (K atac) ls']] 1
+ end;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Tools/bnf_fp.ML Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,442 @@
+(* Title: HOL/BNF/Tools/bnf_fp.ML
+ Author: Dmitriy Traytel, TU Muenchen
+ Copyright 2012
+
+Shared library for the datatype and codatatype constructions.
+*)
+
+signature BNF_FP =
+sig
+ val time: Timer.real_timer -> string -> Timer.real_timer
+
+ 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 corecN: string
+ val corecsN: string
+ val ctorN: string
+ val ctor_dtorN: string
+ val ctor_dtor_unfoldsN: string
+ val ctor_dtor_corecsN: string
+ val ctor_exhaustN: string
+ val ctor_induct2N: string
+ val ctor_inductN: string
+ val ctor_injectN: string
+ val ctor_foldN: string
+ val ctor_fold_uniqueN: string
+ val ctor_foldsN: string
+ val ctor_recN: string
+ val ctor_recsN: string
+ val disc_unfold_iffN: string
+ val disc_unfoldsN: string
+ val disc_corec_iffN: string
+ val disc_corecsN: string
+ val dtorN: string
+ val dtor_coinductN: string
+ val dtor_unfoldN: string
+ val dtor_unfold_uniqueN: string
+ val dtor_unfoldsN: string
+ val dtor_corecN: string
+ val dtor_corecsN: string
+ val dtor_exhaustN: string
+ val dtor_ctorN: string
+ val dtor_injectN: string
+ val dtor_strong_coinductN: string
+ val exhaustN: string
+ val foldN: string
+ val foldsN: string
+ val hsetN: string
+ val hset_recN: string
+ val inductN: string
+ val injectN: string
+ val isNodeN: string
+ val lsbisN: string
+ val map_simpsN: string
+ val map_uniqueN: string
+ val min_algN: string
+ val morN: string
+ val nchotomyN: string
+ val recN: string
+ val recsN: string
+ val rel_coinductN: string
+ val rel_simpN: string
+ val rel_strong_coinductN: string
+ val rvN: string
+ val sel_unfoldsN: string
+ val sel_corecsN: string
+ val set_inclN: string
+ val set_set_inclN: string
+ val simpsN: string
+ val srel_coinductN: string
+ val srel_simpN: string
+ val srel_strong_coinductN: string
+ val strTN: string
+ val str_initN: string
+ val strongN: string
+ val sum_bdN: string
+ val sum_bdTN: string
+ val unfoldN: string
+ val unfoldsN: string
+ val uniqueN: string
+
+ val mk_exhaustN: string -> string
+ val mk_injectN: string -> string
+ val mk_nchotomyN: string -> string
+ val mk_set_simpsN: int -> string
+ val mk_set_minimalN: int -> string
+ val mk_set_inductN: int -> string
+
+ val mk_common_name: string list -> string
+
+ val split_conj_thm: thm -> thm list
+ val split_conj_prems: int -> thm -> thm
+
+ val retype_free: typ -> term -> term
+
+ val mk_sumTN: typ list -> typ
+ val mk_sumTN_balanced: typ list -> typ
+
+ val id_const: typ -> term
+ val id_abs: typ -> term
+
+ val Inl_const: typ -> typ -> term
+ val Inr_const: typ -> typ -> term
+
+ val mk_Inl: typ -> term -> term
+ val mk_Inr: typ -> term -> term
+ val mk_InN: typ list -> term -> int -> term
+ val mk_InN_balanced: typ -> int -> term -> int -> term
+ val mk_sum_case: term * term -> term
+ val mk_sum_caseN: term list -> term
+ val mk_sum_caseN_balanced: term list -> term
+
+ val dest_sumT: typ -> typ * typ
+ val dest_sumTN: int -> typ -> typ list
+ val dest_sumTN_balanced: int -> typ -> typ list
+ val dest_tupleT: int -> typ -> typ list
+
+ val mk_Field: term -> term
+ val mk_If: term -> term -> term -> term
+ val mk_union: term * term -> term
+
+ val mk_sumEN: int -> thm
+ val mk_sumEN_balanced: int -> thm
+ val mk_sumEN_tupled_balanced: int list -> thm
+ val mk_sum_casesN: int -> int -> thm
+ val mk_sum_casesN_balanced: int -> int -> thm
+
+ val fixpoint: ('a * 'a -> bool) -> ('a list -> 'a list) -> 'a list -> 'a list
+
+ val fp_bnf: (mixfix list -> (string * sort) list option -> binding list ->
+ typ list * typ list list -> BNF_Def.BNF list -> local_theory -> 'a) ->
+ binding list -> mixfix list -> (string * sort) list -> ((string * sort) * typ) list ->
+ local_theory -> BNF_Def.BNF list * 'a
+ val fp_bnf_cmd: (mixfix list -> (string * sort) list option -> binding list ->
+ typ list * typ list list -> BNF_Def.BNF list -> local_theory -> 'a) ->
+ binding list * (string list * string list) -> local_theory -> 'a
+end;
+
+structure BNF_FP : BNF_FP =
+struct
+
+open BNF_Comp
+open BNF_Def
+open BNF_Util
+
+val timing = true;
+fun time timer msg = (if timing
+ then warning (msg ^ ": " ^ ATP_Util.string_from_time (Timer.checkRealTimer timer))
+ 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 foldsN = foldN ^ "s"
+val unfoldN = unN ^ foldN
+val unfoldsN = unfoldN ^ "s"
+val uniqueN = "_unique"
+val simpsN = "simps"
+val ctorN = "ctor"
+val dtorN = "dtor"
+val ctor_foldN = ctorN ^ "_" ^ foldN
+val ctor_foldsN = ctor_foldN ^ "s"
+val dtor_unfoldN = dtorN ^ "_" ^ unfoldN
+val dtor_unfoldsN = dtor_unfoldN ^ "s"
+val ctor_fold_uniqueN = ctor_foldN ^ uniqueN
+val dtor_unfold_uniqueN = dtor_unfoldN ^ uniqueN
+val ctor_dtor_unfoldsN = ctorN ^ "_" ^ dtor_unfoldN ^ "s"
+val map_simpsN = mapN ^ "_" ^ simpsN
+val map_uniqueN = mapN ^ 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"
+fun mk_set_simpsN i = mk_setN i ^ "_" ^ simpsN
+fun mk_set_minimalN i = mk_setN i ^ "_minimal"
+fun mk_set_inductN i = mk_setN i ^ "_induct"
+
+val str_initN = "str_init"
+val recN = "rec"
+val recsN = recN ^ "s"
+val corecN = coN ^ recN
+val corecsN = corecN ^ "s"
+val ctor_recN = ctorN ^ "_" ^ recN
+val ctor_recsN = ctor_recN ^ "s"
+val dtor_corecN = dtorN ^ "_" ^ corecN
+val dtor_corecsN = dtor_corecN ^ "s"
+val ctor_dtor_corecsN = ctorN ^ "_" ^ dtor_corecN ^ "s"
+
+val ctor_dtorN = ctorN ^ "_" ^ dtorN
+val dtor_ctorN = dtorN ^ "_" ^ ctorN
+val nchotomyN = "nchotomy"
+fun mk_nchotomyN s = s ^ "_" ^ nchotomyN
+val injectN = "inject"
+fun mk_injectN s = s ^ "_" ^ injectN
+val exhaustN = "exhaust"
+fun mk_exhaustN s = s ^ "_" ^ exhaustN
+val ctor_injectN = mk_injectN ctorN
+val ctor_exhaustN = mk_exhaustN ctorN
+val dtor_injectN = mk_injectN dtorN
+val dtor_exhaustN = mk_exhaustN dtorN
+val inductN = "induct"
+val coinductN = coN ^ inductN
+val ctor_inductN = ctorN ^ "_" ^ inductN
+val ctor_induct2N = ctor_inductN ^ "2"
+val dtor_coinductN = dtorN ^ "_" ^ coinductN
+val rel_coinductN = relN ^ "_" ^ coinductN
+val srel_coinductN = srelN ^ "_" ^ coinductN
+val simpN = "_simp";
+val srel_simpN = srelN ^ simpN;
+val rel_simpN = relN ^ simpN;
+val strongN = "strong_"
+val dtor_strong_coinductN = dtorN ^ "_" ^ strongN ^ coinductN
+val rel_strong_coinductN = relN ^ "_" ^ strongN ^ coinductN
+val srel_strong_coinductN = srelN ^ "_" ^ strongN ^ coinductN
+val hsetN = "Hset"
+val hset_recN = hsetN ^ "_rec"
+val set_inclN = "set_incl"
+val set_set_inclN = "set_set_incl"
+
+val caseN = "case"
+val discN = "disc"
+val disc_unfoldsN = discN ^ "_" ^ unfoldsN
+val disc_corecsN = discN ^ "_" ^ corecsN
+val iffN = "_iff"
+val disc_unfold_iffN = discN ^ "_" ^ unfoldN ^ iffN
+val disc_corec_iffN = discN ^ "_" ^ corecN ^ iffN
+val selN = "sel"
+val sel_unfoldsN = selN ^ "_" ^ unfoldsN
+val sel_corecsN = selN ^ "_" ^ corecsN
+
+val mk_common_name = space_implode "_";
+
+fun retype_free T (Free (s, _)) = Free (s, T);
+
+fun dest_sumT (Type (@{type_name sum}, [T, T'])) = (T, T');
+
+fun dest_sumTN 1 T = [T]
+ | dest_sumTN n (Type (@{type_name sum}, [T, T'])) = T :: dest_sumTN (n - 1) T';
+
+val dest_sumTN_balanced = Balanced_Tree.dest dest_sumT;
+
+(* TODO: move something like this to "HOLogic"? *)
+fun dest_tupleT 0 @{typ unit} = []
+ | dest_tupleT 1 T = [T]
+ | dest_tupleT n (Type (@{type_name prod}, [T, T'])) = T :: dest_tupleT (n - 1) T';
+
+val mk_sumTN = Library.foldr1 mk_sumT;
+val mk_sumTN_balanced = Balanced_Tree.make mk_sumT;
+
+fun id_const T = Const (@{const_name id}, T --> T);
+fun id_abs T = Abs (Name.uu, T, Bound 0);
+
+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_InN [_] t 1 = t
+ | mk_InN (_ :: Ts) t 1 = mk_Inl (mk_sumTN Ts) t
+ | mk_InN (LT :: Ts) t m = mk_Inr LT (mk_InN Ts t (m - 1))
+ | mk_InN Ts t _ = raise (TYPE ("mk_InN", Ts, [t]));
+
+fun mk_InN_balanced sum_T n t k =
+ let
+ fun repair_types T (Const (s as @{const_name Inl}, _) $ t) = repair_inj_types T s fst t
+ | repair_types T (Const (s as @{const_name Inr}, _) $ t) = repair_inj_types T s snd t
+ | repair_types _ t = t
+ and repair_inj_types T s get t =
+ let val T' = get (dest_sumT T) in
+ Const (s, T' --> T) $ repair_types T' t
+ end;
+ in
+ Balanced_Tree.access {left = mk_Inl dummyT, right = mk_Inr dummyT, init = t} n k
+ |> repair_types sum_T
+ end;
+
+fun mk_sum_case (f, g) =
+ let
+ val fT = fastype_of f;
+ val gT = fastype_of g;
+ in
+ Const (@{const_name sum_case},
+ fT --> gT --> mk_sumT (domain_type fT, domain_type gT) --> range_type fT) $ f $ g
+ end;
+
+val mk_sum_caseN = Library.foldr1 mk_sum_case;
+val mk_sum_caseN_balanced = Balanced_Tree.make mk_sum_case;
+
+fun mk_If p t f =
+ let val T = fastype_of t;
+ in Const (@{const_name If}, HOLogic.boolT --> T --> T --> 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;
+
+val mk_union = HOLogic.mk_binop @{const_name sup};
+
+(*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_sumEN 1 = @{thm one_pointE}
+ | mk_sumEN 2 = @{thm sumE}
+ | mk_sumEN n =
+ (fold (fn i => fn thm => @{thm obj_sum_step} RSN (i, thm)) (2 upto n - 1) @{thm obj_sumE}) OF
+ replicate n (impI RS allI);
+
+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_sumEN_balanced' n all_impIs = mk_obj_sumEN_balanced n OF all_impIs RS @{thm obj_one_pointE};
+
+fun mk_sumEN_balanced 1 = @{thm one_pointE} (*optimization*)
+ | mk_sumEN_balanced 2 = @{thm sumE} (*optimization*)
+ | mk_sumEN_balanced n = mk_sumEN_balanced' n (replicate n (impI RS allI));
+
+fun mk_tupled_allIN 0 = @{thm unit_all_impI}
+ | mk_tupled_allIN 1 = @{thm impI[THEN allI]}
+ | mk_tupled_allIN 2 = @{thm prod_all_impI} (*optimization*)
+ | mk_tupled_allIN n = mk_tupled_allIN (n - 1) RS @{thm prod_all_impI_step};
+
+fun mk_sumEN_tupled_balanced ms =
+ let val n = length ms in
+ if forall (curry (op =) 1) ms then mk_sumEN_balanced n
+ else mk_sumEN_balanced' n (map mk_tupled_allIN ms)
+ end;
+
+fun mk_sum_casesN 1 1 = refl
+ | mk_sum_casesN _ 1 = @{thm sum.cases(1)}
+ | mk_sum_casesN 2 2 = @{thm sum.cases(2)}
+ | mk_sum_casesN n k = trans OF [@{thm sum_case_step(2)}, mk_sum_casesN (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_casesN_balanced 1 1 = refl
+ | mk_sum_casesN_balanced n k =
+ Balanced_Tree.access {left = mk_sum_step @{thm sum.cases(1)} @{thm sum_case_step(1)},
+ right = mk_sum_step @{thm sum.cases(2)} @{thm sum_case_step(2)}, init = refl} n k;
+
+(* FIXME: because of "@ lhss", the output could contain type variables that are not in the input;
+ also, "fp_sort" should put the "resBs" first and in the order in which they appear *)
+fun fp_sort lhss NONE Ass = Library.sort (Term_Ord.typ_ord o pairself TFree)
+ (subtract (op =) lhss (fold (fold (insert (op =))) Ass [])) @ lhss
+ | fp_sort lhss (SOME resBs) Ass =
+ (subtract (op =) lhss (filter (fn T => exists (fn Ts => member (op =) Ts T) Ass) resBs)) @ lhss;
+
+fun mk_fp_bnf timer construct resBs bs sort lhss bnfs deadss livess unfold_set lthy =
+ let
+ val name = mk_common_name (map Binding.name_of bs);
+ fun qualify i =
+ let val namei = name ^ nonzero_string_of_int i;
+ in Binding.qualify true namei end;
+
+ val Ass = map (map dest_TFree) livess;
+ val resDs = (case resBs of NONE => [] | SOME Ts => fold (subtract (op =)) Ass Ts);
+ val Ds = fold (fold Term.add_tfreesT) deadss [];
+
+ val _ = (case Library.inter (op =) Ds lhss of [] => ()
+ | A :: _ => error ("Nonadmissible type recursion (cannot take fixed point of dead type \
+ \variable " ^ quote (Syntax.string_of_typ lthy (TFree A)) ^ ")"));
+
+ val timer = time (timer "Construction of BNFs");
+
+ val ((kill_poss, _), (bnfs', (unfold_set', lthy'))) =
+ normalize_bnfs qualify Ass Ds sort bnfs unfold_set lthy;
+
+ val Dss = map3 (append oo map o nth) livess kill_poss deadss;
+
+ val ((bnfs'', deadss), lthy'') =
+ fold_map3 (seal_bnf unfold_set') (map (Binding.prefix_name preN) bs) Dss bnfs' lthy'
+ |>> split_list;
+
+ val timer = time (timer "Normalization & sealing of BNFs");
+
+ val res = construct resBs bs (map TFree resDs, deadss) bnfs'' lthy'';
+
+ val timer = time (timer "FP construction in total");
+ in
+ timer; (bnfs'', res)
+ end;
+
+fun fp_bnf construct bs mixfixes resBs eqs lthy =
+ let
+ val timer = time (Timer.startRealTimer ());
+ val (lhss, rhss) = split_list eqs;
+ val sort = fp_sort lhss (SOME resBs);
+ fun qualify b = Binding.qualify true (Binding.name_of (Binding.prefix_name rawN b));
+ val ((bnfs, (Dss, Ass)), (unfold_set, lthy')) = apfst (apsnd split_list o split_list)
+ (fold_map2 (fn b => bnf_of_typ Smart_Inline (qualify b) sort) bs rhss
+ (empty_unfolds, lthy));
+ in
+ mk_fp_bnf timer (construct mixfixes) (SOME resBs) bs sort lhss bnfs Dss Ass unfold_set lthy'
+ end;
+
+fun fp_bnf_cmd construct (bs, (raw_lhss, raw_bnfs)) lthy =
+ let
+ val timer = time (Timer.startRealTimer ());
+ val lhss = map (dest_TFree o Syntax.read_typ lthy) raw_lhss;
+ val sort = fp_sort lhss NONE;
+ fun qualify b = Binding.qualify true (Binding.name_of (Binding.prefix_name rawN b));
+ val ((bnfs, (Dss, Ass)), (unfold_set, lthy')) = apfst (apsnd split_list o split_list)
+ (fold_map2 (fn b => fn rawT =>
+ (bnf_of_typ Smart_Inline (qualify b) sort (Syntax.read_typ lthy rawT)))
+ bs raw_bnfs (empty_unfolds, lthy));
+ in
+ snd (mk_fp_bnf timer
+ (construct (map (K NoSyn) bs)) NONE bs sort lhss bnfs Dss Ass unfold_set lthy')
+ end;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Tools/bnf_fp_sugar.ML Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,911 @@
+(* Title: HOL/BNF/Tools/bnf_fp_sugar.ML
+ Author: Jasmin Blanchette, TU Muenchen
+ Copyright 2012
+
+Sugared datatype and codatatype constructions.
+*)
+
+signature BNF_FP_SUGAR =
+sig
+ val datatyp: bool ->
+ (mixfix list -> (string * sort) list option -> binding list -> typ list * typ list list ->
+ BNF_Def.BNF list -> local_theory ->
+ (term list * term list * term list * term list * thm * thm list * thm list * thm list *
+ thm list * thm list) * local_theory) ->
+ bool * ((((typ * sort) list * binding) * mixfix) * ((((binding * binding) *
+ (binding * typ) list) * (binding * term) list) * mixfix) list) list ->
+ local_theory -> local_theory
+ val parse_datatype_cmd: bool ->
+ (mixfix list -> (string * sort) list option -> binding list -> typ list * typ list list ->
+ BNF_Def.BNF list -> local_theory ->
+ (term list * term list * term list * term list * thm * thm list * thm list * thm list *
+ thm list * thm list) * local_theory) ->
+ (local_theory -> local_theory) parser
+end;
+
+structure BNF_FP_Sugar : BNF_FP_SUGAR =
+struct
+
+open BNF_Util
+open BNF_Wrap
+open BNF_Def
+open BNF_FP
+open BNF_FP_Sugar_Tactics
+
+val simp_attrs = @{attributes [simp]};
+
+fun split_list8 xs =
+ (map #1 xs, map #2 xs, map #3 xs, map #4 xs, map #5 xs, map #6 xs, map #7 xs, map #8 xs);
+
+fun resort_tfree S (TFree (s, _)) = TFree (s, S);
+
+fun typ_subst inst (T as Type (s, Ts)) =
+ (case AList.lookup (op =) inst T of
+ NONE => Type (s, map (typ_subst inst) Ts)
+ | SOME T' => T')
+ | typ_subst inst T = the_default T (AList.lookup (op =) inst T);
+
+val lists_bmoc = fold (fn xs => fn t => Term.list_comb (t, xs));
+
+fun mk_tupled_fun x f xs = HOLogic.tupled_lambda x (Term.list_comb (f, xs));
+fun mk_uncurried_fun f xs = mk_tupled_fun (HOLogic.mk_tuple xs) f xs;
+fun mk_uncurried2_fun f xss =
+ mk_tupled_fun (HOLogic.mk_tuple (map HOLogic.mk_tuple xss)) f (flat xss);
+
+fun tick u f = Term.lambda u (HOLogic.mk_prod (u, f $ u));
+
+fun tack z_name (c, u) f =
+ let val z = Free (z_name, mk_sumT (fastype_of u, fastype_of c)) in
+ Term.lambda z (mk_sum_case (Term.lambda u u, Term.lambda c (f $ c)) $ z)
+ end;
+
+fun cannot_merge_types () = error "Mutually recursive types must have the same type parameters";
+
+fun merge_type_arg T T' = if T = T' then T else cannot_merge_types ();
+
+fun merge_type_args (As, As') =
+ if length As = length As' then map2 merge_type_arg As As' else cannot_merge_types ();
+
+fun is_triv_implies thm =
+ op aconv (Logic.dest_implies (Thm.prop_of thm))
+ handle TERM _ => false;
+
+fun type_args_constrained_of (((cAs, _), _), _) = cAs;
+fun type_binding_of (((_, b), _), _) = b;
+fun mixfix_of ((_, mx), _) = mx;
+fun ctr_specs_of (_, ctr_specs) = ctr_specs;
+
+fun disc_of ((((disc, _), _), _), _) = disc;
+fun ctr_of ((((_, ctr), _), _), _) = ctr;
+fun args_of (((_, args), _), _) = args;
+fun defaults_of ((_, ds), _) = ds;
+fun ctr_mixfix_of (_, mx) = mx;
+
+fun define_datatype prepare_constraint prepare_typ prepare_term lfp construct (no_dests, specs)
+ no_defs_lthy0 =
+ let
+ (* TODO: sanity checks on arguments *)
+ (* TODO: integration with function package ("size") *)
+
+ val _ = if not lfp andalso no_dests then error "Cannot define destructor-less codatatypes"
+ else ();
+
+ val nn = length specs;
+ val fp_bs = map type_binding_of specs;
+ val fp_b_names = map Binding.name_of fp_bs;
+ val fp_common_name = mk_common_name fp_b_names;
+
+ fun prepare_type_arg (ty, c) =
+ let val TFree (s, _) = prepare_typ no_defs_lthy0 ty in
+ TFree (s, prepare_constraint no_defs_lthy0 c)
+ end;
+
+ val Ass0 = map (map prepare_type_arg o type_args_constrained_of) specs;
+ val unsorted_Ass0 = map (map (resort_tfree HOLogic.typeS)) Ass0;
+ val unsorted_As = Library.foldr1 merge_type_args unsorted_Ass0;
+
+ val ((Bs, Cs), no_defs_lthy) =
+ no_defs_lthy0
+ |> fold (Variable.declare_typ o resort_tfree dummyS) unsorted_As
+ |> mk_TFrees nn
+ ||>> mk_TFrees nn;
+
+ (* TODO: cleaner handling of fake contexts, without "background_theory" *)
+ (*the "perhaps o try" below helps gracefully handles the case where the new type is defined in a
+ locale and shadows an existing global type*)
+ val fake_thy =
+ Theory.copy #> fold (fn spec => perhaps (try (Sign.add_type no_defs_lthy
+ (type_binding_of spec, length (type_args_constrained_of spec), mixfix_of spec)))) specs;
+ val fake_lthy = Proof_Context.background_theory fake_thy no_defs_lthy;
+
+ fun mk_fake_T b =
+ Type (fst (Term.dest_Type (Proof_Context.read_type_name fake_lthy true (Binding.name_of b))),
+ unsorted_As);
+
+ val fake_Ts = map mk_fake_T fp_bs;
+
+ val mixfixes = map mixfix_of specs;
+
+ val _ = (case duplicates Binding.eq_name fp_bs of [] => ()
+ | b :: _ => error ("Duplicate type name declaration " ^ quote (Binding.name_of b)));
+
+ val ctr_specss = map ctr_specs_of specs;
+
+ val disc_bindingss = map (map disc_of) ctr_specss;
+ val ctr_bindingss =
+ map2 (fn fp_b_name => map (Binding.qualify false fp_b_name o ctr_of)) fp_b_names ctr_specss;
+ val ctr_argsss = map (map args_of) ctr_specss;
+ val ctr_mixfixess = map (map ctr_mixfix_of) ctr_specss;
+
+ val sel_bindingsss = map (map (map fst)) ctr_argsss;
+ val fake_ctr_Tsss0 = map (map (map (prepare_typ fake_lthy o snd))) ctr_argsss;
+ val raw_sel_defaultsss = map (map defaults_of) ctr_specss;
+
+ val (As :: _) :: fake_ctr_Tsss =
+ burrow (burrow (Syntax.check_typs fake_lthy)) (Ass0 :: fake_ctr_Tsss0);
+
+ val _ = (case duplicates (op =) unsorted_As of [] => ()
+ | A :: _ => error ("Duplicate type parameter " ^
+ quote (Syntax.string_of_typ no_defs_lthy A)));
+
+ val rhs_As' = fold (fold (fold Term.add_tfreesT)) fake_ctr_Tsss [];
+ val _ = (case subtract (op =) (map dest_TFree As) rhs_As' of
+ [] => ()
+ | A' :: _ => error ("Extra type variable on right-hand side: " ^
+ quote (Syntax.string_of_typ no_defs_lthy (TFree A'))));
+
+ fun eq_fpT (T as Type (s, Us)) (Type (s', Us')) =
+ s = s' andalso (Us = Us' orelse error ("Illegal occurrence of recursive type " ^
+ quote (Syntax.string_of_typ fake_lthy T)))
+ | eq_fpT _ _ = false;
+
+ fun freeze_fp (T as Type (s, Us)) =
+ (case find_index (eq_fpT T) fake_Ts of ~1 => Type (s, map freeze_fp Us) | j => nth Bs j)
+ | freeze_fp T = T;
+
+ val ctr_TsssBs = map (map (map freeze_fp)) fake_ctr_Tsss;
+ val ctr_sum_prod_TsBs = map (mk_sumTN_balanced o map HOLogic.mk_tupleT) ctr_TsssBs;
+
+ val fp_eqs =
+ map dest_TFree Bs ~~ map (Term.typ_subst_atomic (As ~~ unsorted_As)) ctr_sum_prod_TsBs;
+
+ val (pre_bnfs, ((dtors0, ctors0, fp_folds0, fp_recs0, fp_induct, dtor_ctors, ctor_dtors,
+ ctor_injects, fp_fold_thms, fp_rec_thms), lthy)) =
+ fp_bnf construct fp_bs mixfixes (map dest_TFree unsorted_As) fp_eqs no_defs_lthy0;
+
+ fun add_nesty_bnf_names Us =
+ let
+ fun add (Type (s, Ts)) ss =
+ let val (needs, ss') = fold_map add Ts ss in
+ if exists I needs then (true, insert (op =) s ss') else (false, ss')
+ end
+ | add T ss = (member (op =) Us T, ss);
+ in snd oo add end;
+
+ fun nesty_bnfs Us =
+ map_filter (bnf_of lthy) (fold (fold (fold (add_nesty_bnf_names Us))) ctr_TsssBs []);
+
+ val nesting_bnfs = nesty_bnfs As;
+ val nested_bnfs = nesty_bnfs Bs;
+
+ val timer = time (Timer.startRealTimer ());
+
+ fun mk_ctor_or_dtor get_T Ts t =
+ let val Type (_, Ts0) = get_T (fastype_of t) in
+ Term.subst_atomic_types (Ts0 ~~ Ts) t
+ end;
+
+ val mk_ctor = mk_ctor_or_dtor range_type;
+ val mk_dtor = mk_ctor_or_dtor domain_type;
+
+ val ctors = map (mk_ctor As) ctors0;
+ val dtors = map (mk_dtor As) dtors0;
+
+ val fpTs = map (domain_type o fastype_of) dtors;
+
+ val exists_fp_subtype = exists_subtype (member (op =) fpTs);
+
+ val ctr_Tsss = map (map (map (Term.typ_subst_atomic (Bs ~~ fpTs)))) ctr_TsssBs;
+ val ns = map length ctr_Tsss;
+ val kss = map (fn n => 1 upto n) ns;
+ val mss = map (map length) ctr_Tsss;
+ val Css = map2 replicate ns Cs;
+
+ fun mk_rec_like Ts Us t =
+ let
+ val (bindings, body) = strip_type (fastype_of t);
+ val (f_Us, prebody) = split_last bindings;
+ val Type (_, Ts0) = if lfp then prebody else body;
+ val Us0 = distinct (op =) (map (if lfp then body_type else domain_type) f_Us);
+ in
+ Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t
+ end;
+
+ val fp_folds as fp_fold1 :: _ = map (mk_rec_like As Cs) fp_folds0;
+ val fp_recs as fp_rec1 :: _ = map (mk_rec_like As Cs) fp_recs0;
+
+ val fp_fold_fun_Ts = fst (split_last (binder_types (fastype_of fp_fold1)));
+ val fp_rec_fun_Ts = fst (split_last (binder_types (fastype_of fp_rec1)));
+
+ val (((fold_only as (gss, _, _), rec_only as (hss, _, _)),
+ (zs, cs, cpss, unfold_only as ((pgss, crgsss), _), corec_only as ((phss, cshsss), _))),
+ names_lthy) =
+ if lfp then
+ let
+ val y_Tsss =
+ map3 (fn n => fn ms => map2 dest_tupleT ms o dest_sumTN_balanced n o domain_type)
+ ns mss fp_fold_fun_Ts;
+ val g_Tss = map2 (map2 (curry (op --->))) y_Tsss Css;
+
+ val ((gss, ysss), lthy) =
+ lthy
+ |> mk_Freess "f" g_Tss
+ ||>> mk_Freesss "x" y_Tsss;
+ val yssss = map (map (map single)) ysss;
+
+ fun dest_rec_prodT (T as Type (@{type_name prod}, Us as [_, U])) =
+ if member (op =) Cs U then Us else [T]
+ | dest_rec_prodT T = [T];
+
+ val z_Tssss =
+ map3 (fn n => fn ms => map2 (map dest_rec_prodT oo dest_tupleT) ms o
+ dest_sumTN_balanced n o domain_type) ns mss fp_rec_fun_Ts;
+ val h_Tss = map2 (map2 (fold_rev (curry (op --->)))) z_Tssss Css;
+
+ val hss = map2 (map2 retype_free) h_Tss gss;
+ val zssss_hd = map2 (map2 (map2 (retype_free o hd))) z_Tssss ysss;
+ val (zssss_tl, lthy) =
+ lthy
+ |> mk_Freessss "y" (map (map (map tl)) z_Tssss);
+ val zssss = map2 (map2 (map2 cons)) zssss_hd zssss_tl;
+ in
+ ((((gss, g_Tss, yssss), (hss, h_Tss, zssss)),
+ ([], [], [], (([], []), ([], [])), (([], []), ([], [])))), lthy)
+ end
+ else
+ let
+ (*avoid "'a itself" arguments in coiterators and corecursors*)
+ val mss' = map (fn [0] => [1] | ms => ms) mss;
+
+ val p_Tss = map2 (fn n => replicate (Int.max (0, n - 1)) o mk_pred1T) ns Cs;
+
+ fun zip_predss_getterss qss fss = maps (op @) (qss ~~ fss);
+
+ fun zip_preds_predsss_gettersss [] [qss] [fss] = zip_predss_getterss qss fss
+ | zip_preds_predsss_gettersss (p :: ps) (qss :: qsss) (fss :: fsss) =
+ p :: zip_predss_getterss qss fss @ zip_preds_predsss_gettersss ps qsss fsss;
+
+ fun mk_types maybe_dest_sumT fun_Ts =
+ let
+ val f_sum_prod_Ts = map range_type fun_Ts;
+ val f_prod_Tss = map2 dest_sumTN_balanced ns f_sum_prod_Ts;
+ val f_Tssss =
+ map3 (fn C => map2 (map (map (curry (op -->) C) o maybe_dest_sumT) oo dest_tupleT))
+ Cs mss' f_prod_Tss;
+ val q_Tssss =
+ map (map (map (fn [_] => [] | [_, C] => [mk_pred1T (domain_type C)]))) f_Tssss;
+ val pf_Tss = map3 zip_preds_predsss_gettersss p_Tss q_Tssss f_Tssss;
+ in (q_Tssss, f_sum_prod_Ts, f_Tssss, pf_Tss) end;
+
+ val (r_Tssss, g_sum_prod_Ts, g_Tssss, pg_Tss) = mk_types single fp_fold_fun_Ts;
+
+ val ((((Free (z, _), cs), pss), gssss), lthy) =
+ lthy
+ |> yield_singleton (mk_Frees "z") dummyT
+ ||>> mk_Frees "a" Cs
+ ||>> mk_Freess "p" p_Tss
+ ||>> mk_Freessss "g" g_Tssss;
+ val rssss = map (map (map (fn [] => []))) r_Tssss;
+
+ fun dest_corec_sumT (T as Type (@{type_name sum}, Us as [_, U])) =
+ if member (op =) Cs U then Us else [T]
+ | dest_corec_sumT T = [T];
+
+ val (s_Tssss, h_sum_prod_Ts, h_Tssss, ph_Tss) = mk_types dest_corec_sumT fp_rec_fun_Ts;
+
+ val hssss_hd = map2 (map2 (map2 (fn T :: _ => fn [g] => retype_free T g))) h_Tssss gssss;
+ val ((sssss, hssss_tl), lthy) =
+ lthy
+ |> mk_Freessss "q" s_Tssss
+ ||>> mk_Freessss "h" (map (map (map tl)) h_Tssss);
+ val hssss = map2 (map2 (map2 cons)) hssss_hd hssss_tl;
+
+ val cpss = map2 (fn c => map (fn p => p $ c)) cs pss;
+
+ fun mk_preds_getters_join [] [cf] = cf
+ | mk_preds_getters_join [cq] [cf, cf'] =
+ mk_If cq (mk_Inl (fastype_of cf') cf) (mk_Inr (fastype_of cf) cf');
+
+ fun mk_terms qssss fssss =
+ let
+ val pfss = map3 zip_preds_predsss_gettersss pss qssss fssss;
+ val cqssss = map2 (fn c => map (map (map (fn f => f $ c)))) cs qssss;
+ val cfssss = map2 (fn c => map (map (map (fn f => f $ c)))) cs fssss;
+ val cqfsss = map2 (map2 (map2 mk_preds_getters_join)) cqssss cfssss;
+ in (pfss, cqfsss) end;
+ in
+ (((([], [], []), ([], [], [])),
+ ([z], cs, cpss, (mk_terms rssss gssss, (g_sum_prod_Ts, pg_Tss)),
+ (mk_terms sssss hssss, (h_sum_prod_Ts, ph_Tss)))), lthy)
+ end;
+
+ fun define_ctrs_case_for_type ((((((((((((((((((fp_b, fpT), C), ctor), dtor), fp_fold), fp_rec),
+ ctor_dtor), dtor_ctor), ctor_inject), n), ks), ms), ctr_bindings), ctr_mixfixes), ctr_Tss),
+ disc_bindings), sel_bindingss), raw_sel_defaultss) no_defs_lthy =
+ let
+ val fp_b_name = Binding.name_of fp_b;
+
+ val dtorT = domain_type (fastype_of ctor);
+ val ctr_prod_Ts = map HOLogic.mk_tupleT ctr_Tss;
+ val ctr_sum_prod_T = mk_sumTN_balanced ctr_prod_Ts;
+ val case_Ts = map (fn Ts => Ts ---> C) ctr_Tss;
+
+ val ((((w, fs), xss), u'), _) =
+ no_defs_lthy
+ |> yield_singleton (mk_Frees "w") dtorT
+ ||>> mk_Frees "f" case_Ts
+ ||>> mk_Freess "x" ctr_Tss
+ ||>> yield_singleton Variable.variant_fixes fp_b_name;
+
+ val u = Free (u', fpT);
+
+ val ctr_rhss =
+ map2 (fn k => fn xs => fold_rev Term.lambda xs (ctor $
+ mk_InN_balanced ctr_sum_prod_T n (HOLogic.mk_tuple xs) k)) ks xss;
+
+ val case_binding = Binding.suffix_name ("_" ^ caseN) fp_b;
+
+ val case_rhs =
+ fold_rev Term.lambda (fs @ [u])
+ (mk_sum_caseN_balanced (map2 mk_uncurried_fun fs xss) $ (dtor $ u));
+
+ val ((raw_case :: raw_ctrs, raw_case_def :: raw_ctr_defs), (lthy', lthy)) = no_defs_lthy
+ |> apfst split_list o fold_map3 (fn b => fn mx => fn rhs =>
+ Local_Theory.define ((b, mx), ((Thm.def_binding b, []), rhs)) #>> apsnd snd)
+ (case_binding :: ctr_bindings) (NoSyn :: ctr_mixfixes) (case_rhs :: ctr_rhss)
+ ||> `Local_Theory.restore;
+
+ val phi = Proof_Context.export_morphism lthy lthy';
+
+ val ctr_defs = map (Morphism.thm phi) raw_ctr_defs;
+ val case_def = Morphism.thm phi raw_case_def;
+
+ val ctrs0 = map (Morphism.term phi) raw_ctrs;
+ val casex0 = Morphism.term phi raw_case;
+
+ val ctrs = map (mk_ctr As) ctrs0;
+
+ fun exhaust_tac {context = ctxt, ...} =
+ let
+ val ctor_iff_dtor_thm =
+ let
+ val goal =
+ fold_rev Logic.all [w, u]
+ (mk_Trueprop_eq (HOLogic.mk_eq (u, ctor $ w), HOLogic.mk_eq (dtor $ u, w)));
+ in
+ Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
+ mk_ctor_iff_dtor_tac ctxt (map (SOME o certifyT lthy) [dtorT, fpT])
+ (certify lthy ctor) (certify lthy dtor) ctor_dtor dtor_ctor)
+ |> Thm.close_derivation
+ |> Morphism.thm phi
+ end;
+
+ val sumEN_thm' =
+ unfold_thms lthy @{thms all_unit_eq}
+ (Drule.instantiate' (map (SOME o certifyT lthy) ctr_prod_Ts) []
+ (mk_sumEN_balanced n))
+ |> Morphism.thm phi;
+ in
+ mk_exhaust_tac ctxt n ctr_defs ctor_iff_dtor_thm sumEN_thm'
+ end;
+
+ val inject_tacss =
+ map2 (fn 0 => K [] | _ => fn ctr_def => [fn {context = ctxt, ...} =>
+ mk_inject_tac ctxt ctr_def ctor_inject]) ms ctr_defs;
+
+ val half_distinct_tacss =
+ map (map (fn (def, def') => fn {context = ctxt, ...} =>
+ mk_half_distinct_tac ctxt ctor_inject [def, def'])) (mk_half_pairss ctr_defs);
+
+ val case_tacs =
+ map3 (fn k => fn m => fn ctr_def => fn {context = ctxt, ...} =>
+ mk_case_tac ctxt n k m case_def ctr_def dtor_ctor) ks ms ctr_defs;
+
+ val tacss = [exhaust_tac] :: inject_tacss @ half_distinct_tacss @ [case_tacs];
+
+ fun define_fold_rec (wrap_res, no_defs_lthy) =
+ let
+ val fpT_to_C = fpT --> C;
+
+ fun generate_rec_like (suf, fp_rec_like, (fss, f_Tss, xssss)) =
+ let
+ val res_T = fold_rev (curry (op --->)) f_Tss fpT_to_C;
+ val binding = Binding.suffix_name ("_" ^ suf) fp_b;
+ val spec =
+ mk_Trueprop_eq (lists_bmoc fss (Free (Binding.name_of binding, res_T)),
+ Term.list_comb (fp_rec_like,
+ map2 (mk_sum_caseN_balanced oo map2 mk_uncurried2_fun) fss xssss));
+ in (binding, spec) end;
+
+ val rec_like_infos =
+ [(foldN, fp_fold, fold_only),
+ (recN, fp_rec, rec_only)];
+
+ val (bindings, specs) = map generate_rec_like rec_like_infos |> split_list;
+
+ val ((csts, defs), (lthy', lthy)) = no_defs_lthy
+ |> apfst split_list o fold_map2 (fn b => fn spec =>
+ Specification.definition (SOME (b, NONE, NoSyn), ((Thm.def_binding b, []), spec))
+ #>> apsnd snd) bindings specs
+ ||> `Local_Theory.restore;
+
+ val phi = Proof_Context.export_morphism lthy lthy';
+
+ val [fold_def, rec_def] = map (Morphism.thm phi) defs;
+
+ val [foldx, recx] = map (mk_rec_like As Cs o Morphism.term phi) csts;
+ in
+ ((wrap_res, ctrs, foldx, recx, xss, ctr_defs, fold_def, rec_def), lthy)
+ end;
+
+ fun define_unfold_corec (wrap_res, no_defs_lthy) =
+ let
+ val B_to_fpT = C --> fpT;
+
+ fun mk_preds_getterss_join c n cps sum_prod_T cqfss =
+ Term.lambda c (mk_IfN sum_prod_T cps
+ (map2 (mk_InN_balanced sum_prod_T n) (map HOLogic.mk_tuple cqfss) (1 upto n)));
+
+ fun generate_corec_like (suf, fp_rec_like, ((pfss, cqfsss), (f_sum_prod_Ts,
+ pf_Tss))) =
+ let
+ val res_T = fold_rev (curry (op --->)) pf_Tss B_to_fpT;
+ val binding = Binding.suffix_name ("_" ^ suf) fp_b;
+ val spec =
+ mk_Trueprop_eq (lists_bmoc pfss (Free (Binding.name_of binding, res_T)),
+ Term.list_comb (fp_rec_like,
+ map5 mk_preds_getterss_join cs ns cpss f_sum_prod_Ts cqfsss));
+ in (binding, spec) end;
+
+ val corec_like_infos =
+ [(unfoldN, fp_fold, unfold_only),
+ (corecN, fp_rec, corec_only)];
+
+ val (bindings, specs) = map generate_corec_like corec_like_infos |> split_list;
+
+ val ((csts, defs), (lthy', lthy)) = no_defs_lthy
+ |> apfst split_list o fold_map2 (fn b => fn spec =>
+ Specification.definition (SOME (b, NONE, NoSyn), ((Thm.def_binding b, []), spec))
+ #>> apsnd snd) bindings specs
+ ||> `Local_Theory.restore;
+
+ val phi = Proof_Context.export_morphism lthy lthy';
+
+ val [unfold_def, corec_def] = map (Morphism.thm phi) defs;
+
+ val [unfold, corec] = map (mk_rec_like As Cs o Morphism.term phi) csts;
+ in
+ ((wrap_res, ctrs, unfold, corec, xss, ctr_defs, unfold_def, corec_def), lthy)
+ end;
+
+ fun wrap lthy =
+ let val sel_defaultss = map (map (apsnd (prepare_term lthy))) raw_sel_defaultss in
+ wrap_datatype tacss (((no_dests, ctrs0), casex0), (disc_bindings, (sel_bindingss,
+ sel_defaultss))) lthy
+ end;
+
+ val define_rec_likes = if lfp then define_fold_rec else define_unfold_corec;
+ in
+ ((wrap, define_rec_likes), lthy')
+ end;
+
+ val pre_map_defs = map map_def_of_bnf pre_bnfs;
+ val pre_set_defss = map set_defs_of_bnf pre_bnfs;
+ val nested_set_natural's = maps set_natural'_of_bnf nested_bnfs;
+ val nesting_map_ids = map map_id_of_bnf nesting_bnfs;
+
+ fun mk_map live Ts Us t =
+ let
+ val (Type (_, Ts0), Type (_, Us0)) = strip_typeN (live + 1) (fastype_of t) |>> List.last
+ in
+ Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t
+ end;
+
+ fun build_map build_arg (Type (s, Ts)) (Type (_, Us)) =
+ let
+ val bnf = the (bnf_of lthy s);
+ val live = live_of_bnf bnf;
+ val mapx = mk_map live Ts Us (map_of_bnf bnf);
+ val TUs = map dest_funT (fst (strip_typeN live (fastype_of mapx)));
+ val args = map build_arg TUs;
+ in Term.list_comb (mapx, args) end;
+
+ val mk_simp_thmss =
+ map3 (fn (_, _, injects, distincts, cases, _, _, _) => fn rec_likes => fn fold_likes =>
+ injects @ distincts @ cases @ rec_likes @ fold_likes);
+
+ fun derive_induct_fold_rec_thms_for_types ((wrap_ress, ctrss, folds, recs, xsss, ctr_defss,
+ fold_defs, rec_defs), lthy) =
+ let
+ val (((phis, phis'), us'), names_lthy) =
+ lthy
+ |> mk_Frees' "P" (map mk_pred1T fpTs)
+ ||>> Variable.variant_fixes fp_b_names;
+
+ val us = map2 (curry Free) us' fpTs;
+
+ fun mk_sets_nested bnf =
+ let
+ val Type (T_name, Us) = T_of_bnf bnf;
+ val lives = lives_of_bnf bnf;
+ val sets = sets_of_bnf bnf;
+ fun mk_set U =
+ (case find_index (curry (op =) U) lives of
+ ~1 => Term.dummy
+ | i => nth sets i);
+ in
+ (T_name, map mk_set Us)
+ end;
+
+ val setss_nested = map mk_sets_nested nested_bnfs;
+
+ val (induct_thms, induct_thm) =
+ let
+ fun mk_set Ts t =
+ let val Type (_, Ts0) = domain_type (fastype_of t) in
+ Term.subst_atomic_types (Ts0 ~~ Ts) t
+ end;
+
+ fun mk_raw_prem_prems names_lthy (x as Free (s, T as Type (T_name, Ts0))) =
+ (case find_index (curry (op =) T) fpTs of
+ ~1 =>
+ (case AList.lookup (op =) setss_nested T_name of
+ NONE => []
+ | SOME raw_sets0 =>
+ let
+ val (Ts, raw_sets) =
+ split_list (filter (exists_fp_subtype o fst) (Ts0 ~~ raw_sets0));
+ val sets = map (mk_set Ts0) raw_sets;
+ val (ys, names_lthy') = names_lthy |> mk_Frees s Ts;
+ val xysets = map (pair x) (ys ~~ sets);
+ val ppremss = map (mk_raw_prem_prems names_lthy') ys;
+ in
+ flat (map2 (map o apfst o cons) xysets ppremss)
+ end)
+ | i => [([], (i + 1, x))])
+ | mk_raw_prem_prems _ _ = [];
+
+ fun close_prem_prem xs t =
+ fold_rev Logic.all (map Free (drop (nn + length xs)
+ (rev (Term.add_frees t (map dest_Free xs @ phis'))))) t;
+
+ fun mk_prem_prem xs (xysets, (j, x)) =
+ close_prem_prem xs (Logic.list_implies (map (fn (x', (y, set)) =>
+ HOLogic.mk_Trueprop (HOLogic.mk_mem (y, set $ x'))) xysets,
+ HOLogic.mk_Trueprop (nth phis (j - 1) $ x)));
+
+ fun mk_raw_prem phi ctr ctr_Ts =
+ let
+ val (xs, names_lthy') = names_lthy |> mk_Frees "x" ctr_Ts;
+ val pprems = maps (mk_raw_prem_prems names_lthy') xs;
+ in (xs, pprems, HOLogic.mk_Trueprop (phi $ Term.list_comb (ctr, xs))) end;
+
+ fun mk_prem (xs, raw_pprems, concl) =
+ fold_rev Logic.all xs (Logic.list_implies (map (mk_prem_prem xs) raw_pprems, concl));
+
+ val raw_premss = map3 (map2 o mk_raw_prem) phis ctrss ctr_Tsss;
+
+ val goal =
+ Library.foldr (Logic.list_implies o apfst (map mk_prem)) (raw_premss,
+ HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (curry (op $)) phis us)));
+
+ val kksss = map (map (map (fst o snd) o #2)) raw_premss;
+
+ val ctor_induct' = fp_induct OF (map mk_sumEN_tupled_balanced mss);
+
+ val induct_thm =
+ Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
+ mk_induct_tac ctxt ns mss kksss (flat ctr_defss) ctor_induct'
+ nested_set_natural's pre_set_defss)
+ |> singleton (Proof_Context.export names_lthy lthy)
+ in
+ `(conj_dests nn) induct_thm
+ end;
+
+ (* TODO: Generate nicer names in case of clashes *)
+ val induct_cases = Datatype_Prop.indexify_names (maps (map base_name_of_ctr) ctrss);
+
+ val (fold_thmss, rec_thmss) =
+ let
+ val xctrss = map2 (map2 (curry Term.list_comb)) ctrss xsss;
+ val gfolds = map (lists_bmoc gss) folds;
+ val hrecs = map (lists_bmoc hss) recs;
+
+ fun mk_goal fss frec_like xctr f xs fxs =
+ fold_rev (fold_rev Logic.all) (xs :: fss)
+ (mk_Trueprop_eq (frec_like $ xctr, Term.list_comb (f, fxs)));
+
+ fun build_call frec_likes maybe_tick (T, U) =
+ if T = U then
+ id_const T
+ else
+ (case find_index (curry (op =) T) fpTs of
+ ~1 => build_map (build_call frec_likes maybe_tick) T U
+ | j => maybe_tick (nth us j) (nth frec_likes j));
+
+ fun mk_U maybe_mk_prodT =
+ typ_subst (map2 (fn fpT => fn C => (fpT, maybe_mk_prodT fpT C)) fpTs Cs);
+
+ fun intr_calls frec_likes maybe_cons maybe_tick maybe_mk_prodT (x as Free (_, T)) =
+ if member (op =) fpTs T then
+ maybe_cons x [build_call frec_likes (K I) (T, mk_U (K I) T) $ x]
+ else if exists_fp_subtype T then
+ [build_call frec_likes maybe_tick (T, mk_U maybe_mk_prodT T) $ x]
+ else
+ [x];
+
+ val gxsss = map (map (maps (intr_calls gfolds (K I) (K I) (K I)))) xsss;
+ val hxsss = map (map (maps (intr_calls hrecs cons tick (curry HOLogic.mk_prodT)))) xsss;
+
+ val fold_goalss = map5 (map4 o mk_goal gss) gfolds xctrss gss xsss gxsss;
+ val rec_goalss = map5 (map4 o mk_goal hss) hrecs xctrss hss xsss hxsss;
+
+ val fold_tacss =
+ map2 (map o mk_rec_like_tac pre_map_defs nesting_map_ids fold_defs) fp_fold_thms
+ ctr_defss;
+ val rec_tacss =
+ map2 (map o mk_rec_like_tac pre_map_defs nesting_map_ids rec_defs) fp_rec_thms
+ ctr_defss;
+
+ fun prove goal tac = Skip_Proof.prove lthy [] [] goal (tac o #context);
+ in
+ (map2 (map2 prove) fold_goalss fold_tacss,
+ map2 (map2 prove) rec_goalss rec_tacss)
+ end;
+
+ val simp_thmss = mk_simp_thmss wrap_ress rec_thmss fold_thmss;
+
+ val induct_case_names_attr = Attrib.internal (K (Rule_Cases.case_names induct_cases));
+ fun induct_type_attr T_name = Attrib.internal (K (Induct.induct_type T_name));
+
+ (* TODO: Also note "recs", "simps", and "splits" if "nn > 1" (for compatibility with the
+ old package)? And for codatatypes as well? *)
+ val common_notes =
+ (if nn > 1 then [(inductN, [induct_thm], [induct_case_names_attr])] else [])
+ |> map (fn (thmN, thms, attrs) =>
+ ((Binding.qualify true fp_common_name (Binding.name thmN), attrs), [(thms, [])]));
+
+ val notes =
+ [(inductN, map single induct_thms,
+ fn T_name => [induct_case_names_attr, induct_type_attr T_name]),
+ (foldsN, fold_thmss, K (Code.add_default_eqn_attrib :: simp_attrs)),
+ (recsN, rec_thmss, K (Code.add_default_eqn_attrib :: simp_attrs)),
+ (simpsN, simp_thmss, K [])]
+ |> maps (fn (thmN, thmss, attrs) =>
+ map3 (fn b => fn Type (T_name, _) => fn thms =>
+ ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), attrs T_name),
+ [(thms, [])])) fp_bs fpTs thmss);
+ in
+ lthy |> Local_Theory.notes (common_notes @ notes) |> snd
+ end;
+
+ fun derive_coinduct_unfold_corec_thms_for_types ((wrap_ress, ctrss, unfolds, corecs, _,
+ ctr_defss, unfold_defs, corec_defs), lthy) =
+ let
+ val discss = map (map (mk_disc_or_sel As) o #1) wrap_ress;
+ val selsss = map #2 wrap_ress;
+ val disc_thmsss = map #6 wrap_ress;
+ val discIss = map #7 wrap_ress;
+ val sel_thmsss = map #8 wrap_ress;
+
+ val (us', _) =
+ lthy
+ |> Variable.variant_fixes fp_b_names;
+
+ val us = map2 (curry Free) us' fpTs;
+
+ val (coinduct_thms, coinduct_thm) =
+ let
+ val coinduct_thm = fp_induct;
+ in
+ `(conj_dests nn) coinduct_thm
+ end;
+
+ fun mk_maybe_not pos = not pos ? HOLogic.mk_not;
+
+ val z = the_single zs;
+ val gunfolds = map (lists_bmoc pgss) unfolds;
+ val hcorecs = map (lists_bmoc phss) corecs;
+
+ val (unfold_thmss, corec_thmss, safe_unfold_thmss, safe_corec_thmss) =
+ let
+ fun mk_goal pfss c cps fcorec_like n k ctr m cfs' =
+ fold_rev (fold_rev Logic.all) ([c] :: pfss)
+ (Logic.list_implies (seq_conds (HOLogic.mk_Trueprop oo mk_maybe_not) n k cps,
+ mk_Trueprop_eq (fcorec_like $ c, Term.list_comb (ctr, take m cfs'))));
+
+ fun build_call frec_likes maybe_tack (T, U) =
+ if T = U then
+ id_const T
+ else
+ (case find_index (curry (op =) U) fpTs of
+ ~1 => build_map (build_call frec_likes maybe_tack) T U
+ | j => maybe_tack (nth cs j, nth us j) (nth frec_likes j));
+
+ fun mk_U maybe_mk_sumT =
+ typ_subst (map2 (fn C => fn fpT => (maybe_mk_sumT fpT C, fpT)) Cs fpTs);
+
+ fun intr_calls frec_likes maybe_mk_sumT maybe_tack cqf =
+ let val T = fastype_of cqf in
+ if exists_subtype (member (op =) Cs) T then
+ build_call frec_likes maybe_tack (T, mk_U maybe_mk_sumT T) $ cqf
+ else
+ cqf
+ end;
+
+ val crgsss' = map (map (map (intr_calls gunfolds (K I) (K I)))) crgsss;
+ val cshsss' = map (map (map (intr_calls hcorecs (curry mk_sumT) (tack z)))) cshsss;
+
+ val unfold_goalss =
+ map8 (map4 oooo mk_goal pgss) cs cpss gunfolds ns kss ctrss mss crgsss';
+ val corec_goalss =
+ map8 (map4 oooo mk_goal phss) cs cpss hcorecs ns kss ctrss mss cshsss';
+
+ val unfold_tacss =
+ map3 (map oo mk_corec_like_tac unfold_defs nesting_map_ids) fp_fold_thms pre_map_defs
+ ctr_defss;
+ val corec_tacss =
+ map3 (map oo mk_corec_like_tac corec_defs nesting_map_ids) fp_rec_thms pre_map_defs
+ ctr_defss;
+
+ fun prove goal tac =
+ Skip_Proof.prove lthy [] [] goal (tac o #context) |> Thm.close_derivation;
+
+ val unfold_thmss = map2 (map2 prove) unfold_goalss unfold_tacss;
+ val corec_thmss =
+ map2 (map2 prove) corec_goalss corec_tacss
+ |> map (map (unfold_thms lthy @{thms sum_case_if}));
+
+ val unfold_safesss = map2 (map2 (map2 (curry (op =)))) crgsss' crgsss;
+ val corec_safesss = map2 (map2 (map2 (curry (op =)))) cshsss' cshsss;
+
+ val filter_safesss =
+ map2 (map_filter (fn (safes, thm) => if forall I safes then SOME thm else NONE) oo
+ curry (op ~~));
+
+ val safe_unfold_thmss = filter_safesss unfold_safesss unfold_thmss;
+ val safe_corec_thmss = filter_safesss corec_safesss corec_thmss;
+ in
+ (unfold_thmss, corec_thmss, safe_unfold_thmss, safe_corec_thmss)
+ end;
+
+ val (disc_unfold_iff_thmss, disc_corec_iff_thmss) =
+ let
+ fun mk_goal c cps fcorec_like n k disc =
+ mk_Trueprop_eq (disc $ (fcorec_like $ c),
+ if n = 1 then @{const True}
+ else Library.foldr1 HOLogic.mk_conj (seq_conds mk_maybe_not n k cps));
+
+ val unfold_goalss = map6 (map2 oooo mk_goal) cs cpss gunfolds ns kss discss;
+ val corec_goalss = map6 (map2 oooo mk_goal) cs cpss hcorecs ns kss discss;
+
+ fun mk_case_split' cp =
+ Drule.instantiate' [] [SOME (certify lthy cp)] @{thm case_split};
+
+ val case_splitss' = map (map mk_case_split') cpss;
+
+ val unfold_tacss =
+ map3 (map oo mk_disc_corec_like_iff_tac) case_splitss' unfold_thmss disc_thmsss;
+ val corec_tacss =
+ map3 (map oo mk_disc_corec_like_iff_tac) case_splitss' corec_thmss disc_thmsss;
+
+ fun prove goal tac =
+ Skip_Proof.prove lthy [] [] goal (tac o #context)
+ |> Thm.close_derivation
+ |> singleton (Proof_Context.export names_lthy no_defs_lthy);
+
+ fun proves [_] [_] = []
+ | proves goals tacs = map2 prove goals tacs;
+ in
+ (map2 proves unfold_goalss unfold_tacss,
+ map2 proves corec_goalss corec_tacss)
+ end;
+
+ fun mk_disc_corec_like_thms corec_likes discIs =
+ map (op RS) (filter_out (is_triv_implies o snd) (corec_likes ~~ discIs));
+
+ val disc_unfold_thmss = map2 mk_disc_corec_like_thms unfold_thmss discIss;
+ val disc_corec_thmss = map2 mk_disc_corec_like_thms corec_thmss discIss;
+
+ fun mk_sel_corec_like_thm corec_like_thm sel sel_thm =
+ let
+ val (domT, ranT) = dest_funT (fastype_of sel);
+ val arg_cong' =
+ Drule.instantiate' (map (SOME o certifyT lthy) [domT, ranT])
+ [NONE, NONE, SOME (certify lthy sel)] arg_cong
+ |> Thm.varifyT_global;
+ val sel_thm' = sel_thm RSN (2, trans);
+ in
+ corec_like_thm RS arg_cong' RS sel_thm'
+ end;
+
+ fun mk_sel_corec_like_thms corec_likess =
+ map3 (map3 (map2 o mk_sel_corec_like_thm)) corec_likess selsss sel_thmsss |> map flat;
+
+ val sel_unfold_thmss = mk_sel_corec_like_thms unfold_thmss;
+ val sel_corec_thmss = mk_sel_corec_like_thms corec_thmss;
+
+ fun zip_corec_like_thms corec_likes disc_corec_likes sel_corec_likes =
+ corec_likes @ disc_corec_likes @ sel_corec_likes;
+
+ val simp_thmss =
+ mk_simp_thmss wrap_ress
+ (map3 zip_corec_like_thms safe_corec_thmss disc_corec_thmss sel_corec_thmss)
+ (map3 zip_corec_like_thms safe_unfold_thmss disc_unfold_thmss sel_unfold_thmss);
+
+ val anonymous_notes =
+ [(flat safe_unfold_thmss @ flat safe_corec_thmss, simp_attrs)]
+ |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])]));
+
+ val common_notes =
+ (if nn > 1 then [(coinductN, [coinduct_thm], [])] (* FIXME: attribs *) else [])
+ |> map (fn (thmN, thms, attrs) =>
+ ((Binding.qualify true fp_common_name (Binding.name thmN), attrs), [(thms, [])]));
+
+ val notes =
+ [(coinductN, map single coinduct_thms, []), (* FIXME: attribs *)
+ (unfoldsN, unfold_thmss, []),
+ (corecsN, corec_thmss, []),
+ (disc_unfold_iffN, disc_unfold_iff_thmss, simp_attrs),
+ (disc_unfoldsN, disc_unfold_thmss, simp_attrs),
+ (disc_corec_iffN, disc_corec_iff_thmss, simp_attrs),
+ (disc_corecsN, disc_corec_thmss, simp_attrs),
+ (sel_unfoldsN, sel_unfold_thmss, simp_attrs),
+ (sel_corecsN, sel_corec_thmss, simp_attrs),
+ (simpsN, simp_thmss, [])]
+ |> maps (fn (thmN, thmss, attrs) =>
+ map_filter (fn (_, []) => NONE | (b, thms) =>
+ SOME ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), attrs),
+ [(thms, [])])) (fp_bs ~~ thmss));
+ in
+ lthy |> Local_Theory.notes (anonymous_notes @ common_notes @ notes) |> snd
+ end;
+
+ fun wrap_types_and_define_rec_likes ((wraps, define_rec_likess), lthy) =
+ fold_map2 (curry (op o)) define_rec_likess wraps lthy |>> split_list8
+
+ val lthy' = lthy
+ |> fold_map define_ctrs_case_for_type (fp_bs ~~ fpTs ~~ Cs ~~ ctors ~~ dtors ~~ fp_folds ~~
+ fp_recs ~~ ctor_dtors ~~ dtor_ctors ~~ ctor_injects ~~ ns ~~ kss ~~ mss ~~ ctr_bindingss ~~
+ ctr_mixfixess ~~ ctr_Tsss ~~ disc_bindingss ~~ sel_bindingsss ~~ raw_sel_defaultsss)
+ |>> split_list |> wrap_types_and_define_rec_likes
+ |> (if lfp then derive_induct_fold_rec_thms_for_types
+ else derive_coinduct_unfold_corec_thms_for_types);
+
+ val timer = time (timer ("Constructors, discriminators, selectors, etc., for the new " ^
+ (if lfp then "" else "co") ^ "datatype"));
+ in
+ timer; lthy'
+ end;
+
+val datatyp = define_datatype (K I) (K I) (K I);
+
+val datatype_cmd = define_datatype Typedecl.read_constraint Syntax.parse_typ Syntax.read_term;
+
+val parse_ctr_arg =
+ @{keyword "("} |-- parse_binding_colon -- Parse.typ --| @{keyword ")"} ||
+ (Parse.typ >> pair Binding.empty);
+
+val parse_defaults =
+ @{keyword "("} |-- @{keyword "defaults"} |-- Scan.repeat parse_bound_term --| @{keyword ")"};
+
+val parse_single_spec =
+ Parse.type_args_constrained -- Parse.binding -- Parse.opt_mixfix --
+ (@{keyword "="} |-- Parse.enum1 "|" (parse_opt_binding_colon -- Parse.binding --
+ Scan.repeat parse_ctr_arg -- Scan.optional parse_defaults [] -- Parse.opt_mixfix));
+
+val parse_datatype = parse_wrap_options -- Parse.and_list1 parse_single_spec;
+
+fun parse_datatype_cmd lfp construct = parse_datatype >> datatype_cmd lfp construct;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Tools/bnf_fp_sugar_tactics.ML Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,133 @@
+(* Title: HOL/BNF/Tools/bnf_fp_sugar_tactics.ML
+ Author: Jasmin Blanchette, TU Muenchen
+ Copyright 2012
+
+Tactics for datatype and codatatype sugar.
+*)
+
+signature BNF_FP_SUGAR_TACTICS =
+sig
+ val mk_case_tac: Proof.context -> int -> int -> int -> thm -> thm -> thm -> tactic
+ val mk_corec_like_tac: thm list -> thm list -> thm -> thm -> thm -> Proof.context -> tactic
+ val mk_ctor_iff_dtor_tac: Proof.context -> ctyp option list -> cterm -> cterm -> thm -> thm ->
+ tactic
+ val mk_disc_corec_like_iff_tac: thm list -> thm list -> thm list -> Proof.context -> tactic
+ val mk_exhaust_tac: Proof.context -> int -> thm list -> thm -> thm -> tactic
+ val mk_half_distinct_tac: Proof.context -> thm -> thm list -> tactic
+ val mk_induct_tac: Proof.context -> int list -> int list list -> int list list list -> thm list ->
+ thm -> thm list -> thm list list -> tactic
+ val mk_inject_tac: Proof.context -> thm -> thm -> tactic
+ val mk_rec_like_tac: thm list -> thm list -> thm list -> thm -> thm -> Proof.context -> tactic
+end;
+
+structure BNF_FP_Sugar_Tactics : BNF_FP_SUGAR_TACTICS =
+struct
+
+open BNF_Tactics
+open BNF_Util
+open BNF_FP
+
+val meta_mp = @{thm meta_mp};
+val meta_spec = @{thm meta_spec};
+
+fun inst_spurious_fs lthy thm =
+ let
+ val fs =
+ Term.add_vars (prop_of thm) []
+ |> filter (fn (_, Type (@{type_name fun}, [_, T'])) => T' <> HOLogic.boolT | _ => false);
+ val cfs =
+ map (fn f as (_, T) => (certify lthy (Var f), certify lthy (id_abs (domain_type T)))) fs;
+ in
+ Drule.cterm_instantiate cfs thm
+ end;
+
+val inst_spurious_fs_tac = PRIMITIVE o inst_spurious_fs;
+
+fun mk_case_tac ctxt n k m case_def ctr_def dtor_ctor =
+ unfold_thms_tac ctxt [case_def, ctr_def, dtor_ctor] THEN
+ (rtac (mk_sum_casesN_balanced n k RS ssubst) THEN'
+ REPEAT_DETERM_N (Int.max (0, m - 1)) o rtac (@{thm split} RS ssubst) THEN'
+ rtac refl) 1;
+
+fun mk_exhaust_tac ctxt n ctr_defs ctor_iff_dtor sumEN' =
+ unfold_thms_tac ctxt (ctor_iff_dtor :: ctr_defs) THEN rtac sumEN' 1 THEN
+ unfold_thms_tac ctxt @{thms all_prod_eq} THEN
+ EVERY' (maps (fn k => [select_prem_tac n (rotate_tac 1) k, REPEAT_DETERM o dtac meta_spec,
+ etac meta_mp, atac]) (1 upto n)) 1;
+
+fun mk_ctor_iff_dtor_tac ctxt cTs cctor cdtor ctor_dtor dtor_ctor =
+ (rtac iffI THEN'
+ EVERY' (map3 (fn cTs => fn cx => fn th =>
+ dtac (Drule.instantiate' cTs [NONE, NONE, SOME cx] arg_cong) THEN'
+ SELECT_GOAL (unfold_thms_tac ctxt [th]) THEN'
+ atac) [rev cTs, cTs] [cdtor, cctor] [dtor_ctor, ctor_dtor])) 1;
+
+fun mk_half_distinct_tac ctxt ctor_inject ctr_defs =
+ unfold_thms_tac ctxt (ctor_inject :: @{thms sum.inject} @ ctr_defs) THEN
+ rtac @{thm sum.distinct(1)} 1;
+
+fun mk_inject_tac ctxt ctr_def ctor_inject =
+ unfold_thms_tac ctxt [ctr_def] THEN rtac (ctor_inject RS ssubst) 1 THEN
+ unfold_thms_tac ctxt @{thms sum.inject Pair_eq conj_assoc} THEN rtac refl 1;
+
+val rec_like_unfold_thms =
+ @{thms case_unit comp_def convol_def id_apply map_pair_def sum.simps(5,6) sum_map.simps
+ split_conv};
+
+fun mk_rec_like_tac pre_map_defs map_ids rec_like_defs ctor_rec_like ctr_def ctxt =
+ unfold_thms_tac ctxt (ctr_def :: ctor_rec_like :: rec_like_defs @ pre_map_defs @ map_ids @
+ rec_like_unfold_thms) THEN unfold_thms_tac ctxt @{thms id_def} THEN rtac refl 1;
+
+val corec_like_ss = ss_only @{thms if_True if_False};
+val corec_like_unfold_thms = @{thms id_apply map_pair_def sum_map.simps prod.cases};
+
+fun mk_corec_like_tac corec_like_defs map_ids ctor_dtor_corec_like pre_map_def ctr_def ctxt =
+ unfold_thms_tac ctxt (ctr_def :: corec_like_defs) THEN
+ subst_tac ctxt [ctor_dtor_corec_like] 1 THEN asm_simp_tac corec_like_ss 1 THEN
+ unfold_thms_tac ctxt (pre_map_def :: corec_like_unfold_thms @ map_ids) THEN
+ unfold_thms_tac ctxt @{thms id_def} THEN
+ TRY ((rtac refl ORELSE' subst_tac ctxt @{thms unit_eq} THEN' rtac refl) 1);
+
+fun mk_disc_corec_like_iff_tac case_splits' corec_likes discs ctxt =
+ EVERY (map3 (fn case_split_tac => fn corec_like_thm => fn disc =>
+ case_split_tac 1 THEN unfold_thms_tac ctxt [corec_like_thm] THEN
+ asm_simp_tac (ss_only @{thms simp_thms(7,8,12,14,22,24)}) 1 THEN
+ (if is_refl disc then all_tac else rtac disc 1))
+ (map rtac case_splits' @ [K all_tac]) corec_likes discs);
+
+val solve_prem_prem_tac =
+ REPEAT o (eresolve_tac @{thms bexE rev_bexI} ORELSE' rtac @{thm rev_bexI[OF UNIV_I]} ORELSE'
+ hyp_subst_tac ORELSE' resolve_tac @{thms disjI1 disjI2}) THEN'
+ (rtac refl ORELSE' atac ORELSE' rtac @{thm singletonI});
+
+val induct_prem_prem_thms =
+ @{thms SUP_empty Sup_empty Sup_insert UN_insert Un_empty_left Un_empty_right Un_iff
+ Union_Un_distrib collect_def[abs_def] image_def o_apply map_pair_simp
+ mem_Collect_eq mem_UN_compreh_eq prod_set_simps sum_map.simps sum_set_simps};
+
+fun mk_induct_leverage_prem_prems_tac ctxt nn kks set_natural's pre_set_defs =
+ EVERY' (maps (fn kk => [select_prem_tac nn (dtac meta_spec) kk, etac meta_mp,
+ SELECT_GOAL (unfold_thms_tac ctxt (pre_set_defs @ set_natural's @ induct_prem_prem_thms)),
+ solve_prem_prem_tac]) (rev kks)) 1;
+
+fun mk_induct_discharge_prem_tac ctxt nn n set_natural's pre_set_defs m k kks =
+ let val r = length kks in
+ EVERY' [select_prem_tac n (rotate_tac 1) k, rotate_tac ~1, hyp_subst_tac,
+ REPEAT_DETERM_N m o (dtac meta_spec THEN' rotate_tac ~1)] 1 THEN
+ EVERY [REPEAT_DETERM_N r
+ (rotate_tac ~1 1 THEN dtac meta_mp 1 THEN rotate_tac 1 1 THEN prefer_tac 2),
+ if r > 0 then PRIMITIVE Raw_Simplifier.norm_hhf else all_tac, atac 1,
+ mk_induct_leverage_prem_prems_tac ctxt nn kks set_natural's pre_set_defs]
+ end;
+
+fun mk_induct_tac ctxt ns mss kkss ctr_defs ctor_induct' set_natural's pre_set_defss =
+ let
+ val nn = length ns;
+ val n = Integer.sum ns;
+ in
+ unfold_thms_tac ctxt ctr_defs THEN rtac ctor_induct' 1 THEN inst_spurious_fs_tac ctxt THEN
+ EVERY (map4 (EVERY oooo map3 o mk_induct_discharge_prem_tac ctxt nn n set_natural's)
+ pre_set_defss mss (unflat mss (1 upto n)) kkss)
+ end;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Tools/bnf_gfp.ML Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,3002 @@
+(* Title: HOL/BNF/Tools/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 bnf_gfp: mixfix list -> (string * sort) list option -> binding list ->
+ typ list * typ list list -> BNF_Def.BNF list -> local_theory ->
+ (term list * term list * term list * term list * thm * thm list * thm list * thm list *
+ thm list * thm list) * local_theory
+end;
+
+structure BNF_GFP : BNF_GFP =
+struct
+
+open BNF_Def
+open BNF_Util
+open BNF_Tactics
+open BNF_FP
+open BNF_FP_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 bnf_gfp mixfixes resBs bs (resDs, Dss) bnfs lthy =
+ let
+ 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 b = Binding.name (mk_common_name (map Binding.name_of bs));
+
+ (* TODO: check if m, n, etc., are sane *)
+
+ val deads = fold (union (op =)) Dss resDs;
+ val names_lthy = fold Variable.declare_typ deads lthy;
+
+ (* tvars *)
+ val ((((((((passiveAs, activeAs), allAs)), (passiveBs, activeBs)),
+ (passiveCs, activeCs)), passiveXs), passiveYs), idxT) = names_lthy
+ |> mk_TFrees live
+ |> apfst (`(chop m))
+ ||> mk_TFrees live
+ ||>> apfst (chop m)
+ ||> mk_TFrees live
+ ||>> apfst (chop m)
+ ||>> mk_TFrees m
+ ||>> mk_TFrees m
+ ||> fst o mk_TFrees 1
+ ||> the_single;
+
+ 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';
+
+ (* typs *)
+ val dead_poss =
+ (case resBs of
+ NONE => map SOME deads @ replicate m NONE
+ | SOME Ts => map (fn T => if member (op =) deads (TFree T) then SOME (TFree T) else NONE) Ts);
+ 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 all_fTs = map2 (fn T => fn U => T --> U) allAs allBs;
+ 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);
+ val sumBsAs = map2 (curry mk_sumT) activeBs activeAs;
+ val sumFTs = mk_FTs (passiveAs @ sumBsAs);
+ val sum_sTs = map2 (fn T => fn U => T --> U) activeAs sumFTs;
+
+ (* terms *)
+ val mapsAsAs = map4 mk_map_of_bnf Dss Ass Ass bnfs;
+ val mapsAsBs = map4 mk_map_of_bnf Dss Ass Bss bnfs;
+ val mapsBsCs' = map4 mk_map_of_bnf Dss Bss Css' bnfs;
+ val mapsAsCs' = map4 mk_map_of_bnf Dss Ass Css' bnfs;
+ val map_Inls = map4 mk_map_of_bnf Dss Bss (replicate n (passiveAs @ sumBsAs)) bnfs;
+ val map_Inls_rev = map4 mk_map_of_bnf Dss (replicate n (passiveAs @ sumBsAs)) Bss bnfs;
+ val map_fsts = map4 mk_map_of_bnf Dss (replicate n (passiveAs @ RTs)) Ass bnfs;
+ val map_snds = map4 mk_map_of_bnf Dss (replicate n (passiveAs @ RTs)) Bss bnfs;
+ fun mk_setss Ts = map3 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 = map4 mk_srel_of_bnf Dss Ass Bss bnfs;
+ val bds = map3 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 emptys = map (fn T => HOLogic.mk_set T []) passiveAs;
+ val Zeros = map (fn empty =>
+ HOLogic.mk_tuple (map (fn U => absdummy U empty) activeAs)) emptys;
+ val hrecTs = map fastype_of Zeros;
+ val hsetTs = map (fn hrecT => Library.foldr (op -->) (sTs, HOLogic.natT --> hrecT)) hrecTs;
+
+ val (((((((((((((((((((((((((((((((((((zs, zs'), zs_copy), zs_copy2),
+ z's), As), As_copy), Bs), Bs_copy), B's), B''s), ss), sum_ss), s's), s''s), fs), fs_copy),
+ self_fs), all_fs), gs), all_gs), xFs), xFs_copy), RFs), (Rtuple, Rtuple')), (hrecs, hrecs')),
+ (nat, nat')), Rs), Rs_copy), R's), sRs), (idx, idx')), Idx), Ris), Kss),
+ names_lthy) = lthy
+ |> mk_Frees' "b" activeAs
+ ||>> mk_Frees "b" activeAs
+ ||>> mk_Frees "b" activeAs
+ ||>> mk_Frees "b" activeBs
+ ||>> mk_Frees "A" ATs
+ ||>> mk_Frees "A" ATs
+ ||>> mk_Frees "B" BTs
+ ||>> mk_Frees "B" BTs
+ ||>> mk_Frees "B'" B'Ts
+ ||>> mk_Frees "B''" B''Ts
+ ||>> mk_Frees "s" sTs
+ ||>> mk_Frees "sums" sum_sTs
+ ||>> mk_Frees "s'" s'Ts
+ ||>> mk_Frees "s''" s''Ts
+ ||>> mk_Frees "f" fTs
+ ||>> mk_Frees "f" fTs
+ ||>> mk_Frees "f" self_fTs
+ ||>> mk_Frees "f" all_fTs
+ ||>> mk_Frees "g" gTs
+ ||>> mk_Frees "g" all_gTs
+ ||>> mk_Frees "x" FTsAs
+ ||>> mk_Frees "x" FTsAs
+ ||>> mk_Frees "x" FRTs
+ ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "Rtuple") all_sbisT
+ ||>> mk_Frees' "rec" hrecTs
+ ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "n") HOLogic.natT
+ ||>> 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)
+ ||>> mk_Freess "K" (map (fn AT => map (fn T => T --> AT) activeAs) ATs);
+
+ val passive_UNIVs = map HOLogic.mk_UNIV passiveAs;
+ val passive_diags = map mk_diag As;
+ val active_UNIVs = map HOLogic.mk_UNIV activeAs;
+ val sum_UNIVs = map HOLogic.mk_UNIV sumBsAs;
+ val passive_ids = map HOLogic.id_const passiveAs;
+ val active_ids = map HOLogic.id_const activeAs;
+ val Inls = map2 Inl_const activeBs 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 bd_Cnotzeros = map bd_Cnotzero_of_bnf bnfs;
+ val bd_Cnotzero = hd bd_Cnotzeros;
+ val in_bds = map in_bd_of_bnf bnfs;
+ val in_monos = map in_mono_of_bnf bnfs;
+ val map_comps = map map_comp_of_bnf bnfs;
+ val map_comp's = map map_comp'_of_bnf bnfs;
+ val map_congs = map map_cong_of_bnf bnfs;
+ val map_id's = map map_id'_of_bnf bnfs;
+ val map_wpulls = map map_wpull_of_bnf bnfs;
+ val set_bdss = map set_bd_of_bnf bnfs;
+ val set_natural'ss = map set_natural'_of_bnf bnfs;
+ val srel_congs = map srel_cong_of_bnf bnfs;
+ val srel_converses = map srel_converse_of_bnf bnfs;
+ val srel_defs = map srel_def_of_bnf bnfs;
+ val srel_Grs = map srel_Gr_of_bnf bnfs;
+ val srel_Ids = map srel_Id_of_bnf bnfs;
+ val srel_monos = map srel_mono_of_bnf bnfs;
+ val srel_Os = map srel_O_of_bnf bnfs;
+ val srel_O_Grs = map srel_O_Gr_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_comp =
+ 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;
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (x :: fs @ all_gs) (mk_Trueprop_eq (lhs, rhs)))
+ (K (mk_map_comp_id_tac map_comp))
+ |> Thm.close_derivation
+ end;
+
+ val map_comp_id_thms = map5 mk_map_comp_id xFs mapsAsBs mapsBsCs' mapsAsCs' map_comp's;
+
+ (*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_congL x mapAsAs sets map_cong 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 = map4 mk_prem (drop m sets) self_fs zs zs';
+ val goal = mk_Trueprop_eq (Term.list_comb (mapAsAs, passive_ids @ self_fs) $ x, x);
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (x :: self_fs) (Logic.list_implies (prems, goal)))
+ (K (mk_map_congL_tac m map_cong map_id'))
+ |> Thm.close_derivation
+ end;
+
+ val map_congL_thms = map5 mk_map_congL xFs mapsAsAs setssAs map_congs map_id's;
+ 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) xFs xFs_copy;
+ val maps = map (fn mapx => Term.list_comb (mapx, all_fs)) mapsAsBs;
+ val concls =
+ map3 (fn x => fn y => fn mapx => mk_Trueprop_eq (mapx $ x, mapx $ y)) xFs xFs_copy maps;
+ val goals =
+ map4 (fn prem => fn concl => fn x => fn y =>
+ fold_rev Logic.all (x :: y :: all_fs) (Logic.mk_implies (prem, concl)))
+ prems concls xFs xFs_copy;
+ in
+ map (fn goal => Skip_Proof.prove lthy [] [] goal
+ (K ((hyp_subst_tac THEN' rtac refl) 1)) |> Thm.close_derivation) goals
+ end;
+
+ val timer = time (timer "Derived simple theorems");
+
+ (* coalgebra *)
+
+ val coalg_bind = Binding.suffix_name ("_" ^ coN ^ algN) b;
+ val coalg_name = Binding.name_of coalg_bind;
+ val coalg_def_bind = (Thm.def_binding coalg_bind, []);
+
+ (*forall i = 1 ... n: (\<forall>x \<in> Bi. si \<in> Fi_in A1 .. Am B1 ... Bn)*)
+ val coalg_spec =
+ let
+ val coalgT = Library.foldr (op -->) (ATs @ BTs @ sTs, HOLogic.boolT);
+
+ val ins = map3 mk_in (replicate n (As @ 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 lhs = Term.list_comb (Free (coalg_name, coalgT), As @ Bs @ ss);
+ val rhs = Library.foldr1 HOLogic.mk_conj (map5 mk_coalg_conjunct Bs ss ins zs zs')
+ in
+ mk_Trueprop_eq (lhs, rhs)
+ end;
+
+ val ((coalg_free, (_, coalg_def_free)), (lthy, lthy_old)) =
+ lthy
+ |> Specification.definition (SOME (coalg_bind, NONE, NoSyn), (coalg_def_bind, coalg_spec))
+ ||> `Local_Theory.restore;
+
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+ val coalg = fst (Term.dest_Const (Morphism.term phi coalg_free));
+ val coalg_def = Morphism.thm phi coalg_def_free;
+
+ fun mk_coalg As Bs ss =
+ let
+ val args = As @ 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 coalg_prem = HOLogic.mk_Trueprop (mk_coalg As Bs ss);
+
+ val coalg_in_thms = map (fn i =>
+ coalg_def RS @{thm subst[of _ _ "%x. x"]} RS mk_conjunctN n i RS bspec) ks
+
+ val coalg_set_thmss =
+ let
+ val coalg_prem = HOLogic.mk_Trueprop (mk_coalg As Bs ss);
+ fun mk_prem x B = HOLogic.mk_Trueprop (HOLogic.mk_mem (x, B));
+ fun mk_concl s x B set = HOLogic.mk_Trueprop (mk_subset (set $ (s $ x)) B);
+ val prems = map2 mk_prem zs Bs;
+ val conclss = map3 (fn s => fn x => fn sets => map2 (mk_concl s x) (As @ Bs) sets)
+ ss zs setssAs;
+ val goalss = map3 (fn x => fn prem => fn concls => map (fn concl =>
+ fold_rev Logic.all (x :: As @ Bs @ ss)
+ (Logic.list_implies (coalg_prem :: [prem], concl))) concls) zs prems conclss;
+ in
+ map (fn goals => map (fn goal => Skip_Proof.prove lthy [] [] goal
+ (K (mk_coalg_set_tac coalg_def)) |> Thm.close_derivation) goals) goalss
+ end;
+
+ val coalg_set_thmss' = transpose coalg_set_thmss;
+
+ fun mk_tcoalg ATs BTs = mk_coalg (map HOLogic.mk_UNIV ATs) (map HOLogic.mk_UNIV BTs);
+
+ val tcoalg_thm =
+ let
+ val goal = fold_rev Logic.all ss
+ (HOLogic.mk_Trueprop (mk_tcoalg passiveAs activeAs ss))
+ in
+ Skip_Proof.prove lthy [] [] goal
+ (K (stac coalg_def 1 THEN CONJ_WRAP
+ (K (EVERY' [rtac ballI, rtac CollectI,
+ CONJ_WRAP' (K (EVERY' [rtac @{thm subset_UNIV}])) allAs] 1)) ss))
+ |> Thm.close_derivation
+ end;
+
+ val timer = time (timer "Coalgebra definition & thms");
+
+ (* morphism *)
+
+ val mor_bind = Binding.suffix_name ("_" ^ morN) b;
+ val mor_name = Binding.name_of mor_bind;
+ val mor_def_bind = (Thm.def_binding mor_bind, []);
+
+ (*fbetw) forall i = 1 ... n: (\<forall>x \<in> Bi. fi x \<in> B'i)*)
+ (*mor) forall i = 1 ... n: (\<forall>x \<in> Bi.
+ Fi_map id ... id f1 ... fn (si x) = si' (fi x)*)
+ val mor_spec =
+ let
+ val morT = Library.foldr (op -->) (BTs @ sTs @ B'Ts @ s'Ts @ fTs, HOLogic.boolT);
+
+ 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 lhs = Term.list_comb (Free (mor_name, morT), Bs @ ss @ B's @ s's @ fs);
+ val rhs = HOLogic.mk_conj
+ (Library.foldr1 HOLogic.mk_conj (map5 mk_fbetw fs Bs B's zs zs'),
+ Library.foldr1 HOLogic.mk_conj (map7 mk_mor Bs mapsAsBs fs ss s's zs zs'))
+ in
+ mk_Trueprop_eq (lhs, rhs)
+ end;
+
+ val ((mor_free, (_, mor_def_free)), (lthy, lthy_old)) =
+ lthy
+ |> Specification.definition (SOME (mor_bind, NONE, NoSyn), (mor_def_bind, mor_spec))
+ ||> `Local_Theory.restore;
+
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+ val mor = fst (Term.dest_Const (Morphism.term phi mor_free));
+ val mor_def = 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 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 = fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs)
+ (Logic.mk_implies (prem, HOLogic.mk_Trueprop (mk_subset (mk_image f $ B1) B2)));
+ val image_goals = map3 mk_image_goal fs Bs B's;
+ fun mk_elim_goal B mapAsBs f s s' x =
+ fold_rev Logic.all (x :: Bs @ ss @ B's @ s's @ fs)
+ (Logic.list_implies ([prem, HOLogic.mk_Trueprop (HOLogic.mk_mem (x, B))],
+ mk_Trueprop_eq (Term.list_comb (mapAsBs, passive_ids @ fs @ [s $ x]), s' $ (f $ x))));
+ val elim_goals = map6 mk_elim_goal Bs mapsAsBs fs ss s's zs;
+ fun prove goal =
+ Skip_Proof.prove lthy [] [] goal (K (mk_mor_elim_tac mor_def))
+ |> Thm.close_derivation;
+ 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_subset) Bs Bs_copy;
+ val concl = HOLogic.mk_Trueprop (mk_mor Bs ss Bs_copy ss active_ids);
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (Bs @ ss @ Bs_copy) (Logic.list_implies (prems, concl)))
+ (K (mk_mor_incl_tac mor_def map_id's))
+ |> Thm.close_derivation
+ 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));
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (Bs @ ss @ B's @ s's @ B''s @ s''s @ fs @ gs)
+ (Logic.list_implies (prems, concl)))
+ (K (mk_mor_comp_tac mor_def mor_image'_thms morE_thms map_comp_id_thms))
+ |> Thm.close_derivation
+ 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);
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs @ fs_copy)
+ (Logic.list_implies (prems, concl)))
+ (K ((hyp_subst_tac THEN' atac) 1))
+ |> Thm.close_derivation
+ 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 (map4 mk_conjunct mapsAsBs fs ss s's);
+ in
+ Skip_Proof.prove lthy [] [] (fold_rev Logic.all (ss @ s's @ fs) (mk_Trueprop_eq (lhs, rhs)))
+ (K (mk_mor_UNIV_tac morE_thms mor_def))
+ |> Thm.close_derivation
+ 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;
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all ss (HOLogic.mk_Trueprop
+ (mk_mor active_UNIVs ss (map HOLogic.mk_UNIV FTsAs) maps ss)))
+ (K (mk_mor_str_tac ks mor_UNIV_thm))
+ |> Thm.close_derivation
+ end;
+
+ val mor_sum_case_thm =
+ let
+ val maps = map3 (fn s => fn sum_s => fn mapx =>
+ mk_sum_case (HOLogic.mk_comp (Term.list_comb (mapx, passive_ids @ Inls), s), sum_s))
+ s's sum_ss map_Inls;
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (s's @ sum_ss) (HOLogic.mk_Trueprop
+ (mk_mor (map HOLogic.mk_UNIV activeBs) s's sum_UNIVs maps Inls)))
+ (K (mk_mor_sum_case_tac ks mor_UNIV_thm))
+ |> Thm.close_derivation
+ end;
+
+ val timer = time (timer "Morphism definition & thms");
+
+ fun hset_rec_bind j = Binding.suffix_name ("_" ^ hset_recN ^ (if m = 1 then "" else
+ string_of_int j)) b;
+ val hset_rec_name = Binding.name_of o hset_rec_bind;
+ val hset_rec_def_bind = rpair [] o Thm.def_binding o hset_rec_bind;
+
+ fun hset_rec_spec j Zero hsetT hrec hrec' =
+ let
+ fun mk_Suc s setsAs z z' =
+ let
+ val (set, sets) = apfst (fn xs => nth xs (j - 1)) (chop m setsAs);
+ fun mk_UN set k = mk_UNION (set $ (s $ z)) (mk_nthN n hrec k);
+ in
+ Term.absfree z'
+ (mk_union (set $ (s $ z), Library.foldl1 mk_union (map2 mk_UN sets ks)))
+ end;
+
+ val Suc = Term.absdummy HOLogic.natT (Term.absfree hrec'
+ (HOLogic.mk_tuple (map4 mk_Suc ss setssAs zs zs')));
+
+ val lhs = Term.list_comb (Free (hset_rec_name j, hsetT), ss);
+ val rhs = mk_nat_rec Zero Suc;
+ in
+ mk_Trueprop_eq (lhs, rhs)
+ end;
+
+ val ((hset_rec_frees, (_, hset_rec_def_frees)), (lthy, lthy_old)) =
+ lthy
+ |> fold_map5 (fn j => fn Zero => fn hsetT => fn hrec => fn hrec' => Specification.definition
+ (SOME (hset_rec_bind j, NONE, NoSyn),
+ (hset_rec_def_bind j, hset_rec_spec j Zero hsetT hrec hrec')))
+ ls Zeros hsetTs hrecs hrecs'
+ |>> apsnd split_list o split_list
+ ||> `Local_Theory.restore;
+
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+
+ val hset_rec_defs = map (Morphism.thm phi) hset_rec_def_frees;
+ val hset_recs = map (fst o Term.dest_Const o Morphism.term phi) hset_rec_frees;
+
+ fun mk_hset_rec ss nat i j T =
+ let
+ val args = ss @ [nat];
+ val Ts = map fastype_of ss;
+ val bTs = map domain_type Ts;
+ val hrecT = HOLogic.mk_tupleT (map (fn U => U --> HOLogic.mk_setT T) bTs)
+ val hset_recT = Library.foldr (op -->) (Ts, HOLogic.natT --> hrecT);
+ in
+ mk_nthN n (Term.list_comb (Const (nth hset_recs (j - 1), hset_recT), args)) i
+ end;
+
+ val hset_rec_0ss = mk_rec_simps n @{thm nat_rec_0} hset_rec_defs;
+ val hset_rec_Sucss = mk_rec_simps n @{thm nat_rec_Suc} hset_rec_defs;
+ val hset_rec_0ss' = transpose hset_rec_0ss;
+ val hset_rec_Sucss' = transpose hset_rec_Sucss;
+
+ fun hset_bind i j = Binding.suffix_name ("_" ^ hsetN ^
+ (if m = 1 then "" else string_of_int j)) (nth bs (i - 1));
+ val hset_name = Binding.name_of oo hset_bind;
+ val hset_def_bind = rpair [] o Thm.def_binding oo hset_bind;
+
+ fun hset_spec i j =
+ let
+ val U = nth activeAs (i - 1);
+ val z = nth zs (i - 1);
+ val T = nth passiveAs (j - 1);
+ val setT = HOLogic.mk_setT T;
+ val hsetT = Library.foldr (op -->) (sTs, U --> setT);
+
+ val lhs = Term.list_comb (Free (hset_name i j, hsetT), ss @ [z]);
+ val rhs = mk_UNION (HOLogic.mk_UNIV HOLogic.natT)
+ (Term.absfree nat' (mk_hset_rec ss nat i j T $ z));
+ in
+ mk_Trueprop_eq (lhs, rhs)
+ end;
+
+ val ((hset_frees, (_, hset_def_frees)), (lthy, lthy_old)) =
+ lthy
+ |> fold_map (fn i => fold_map (fn j => Specification.definition
+ (SOME (hset_bind i j, NONE, NoSyn), (hset_def_bind i j, hset_spec i j))) ls) ks
+ |>> map (apsnd split_list o split_list)
+ |>> apsnd split_list o split_list
+ ||> `Local_Theory.restore;
+
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+
+ val hset_defss = map (map (Morphism.thm phi)) hset_def_frees;
+ val hset_defss' = transpose hset_defss;
+ val hset_namess = map (map (fst o Term.dest_Const o Morphism.term phi)) hset_frees;
+
+ fun mk_hset ss i j T =
+ let
+ val Ts = map fastype_of ss;
+ val bTs = map domain_type Ts;
+ val hsetT = Library.foldr (op -->) (Ts, nth bTs (i - 1) --> HOLogic.mk_setT T);
+ in
+ Term.list_comb (Const (nth (nth hset_namess (i - 1)) (j - 1), hsetT), ss)
+ end;
+
+ val hsetssAs = map (fn i => map2 (mk_hset ss i) ls passiveAs) ks;
+
+ val (set_incl_hset_thmss, set_hset_incl_hset_thmsss) =
+ let
+ fun mk_set_incl_hset s x set hset = fold_rev Logic.all (x :: ss)
+ (HOLogic.mk_Trueprop (mk_subset (set $ (s $ x)) (hset $ x)));
+
+ fun mk_set_hset_incl_hset s x y set hset1 hset2 =
+ fold_rev Logic.all (x :: y :: ss)
+ (Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (x, set $ (s $ y))),
+ HOLogic.mk_Trueprop (mk_subset (hset1 $ x) (hset2 $ y))));
+
+ val set_incl_hset_goalss =
+ map4 (fn s => fn x => fn sets => fn hsets =>
+ map2 (mk_set_incl_hset s x) (take m sets) hsets)
+ ss zs setssAs hsetssAs;
+
+ (*xk : F(i)set(m+k) (si yi) ==> F(k)_hset(j) s1 ... sn xk <= F(i)_hset(j) s1 ... sn yi*)
+ val set_hset_incl_hset_goalsss =
+ map4 (fn si => fn yi => fn sets => fn hsetsi =>
+ map3 (fn xk => fn set => fn hsetsk =>
+ map2 (mk_set_hset_incl_hset si xk yi set) hsetsk hsetsi)
+ zs_copy (drop m sets) hsetssAs)
+ ss zs setssAs hsetssAs;
+ in
+ (map3 (fn goals => fn defs => fn rec_Sucs =>
+ map3 (fn goal => fn def => fn rec_Suc =>
+ Skip_Proof.prove lthy [] [] goal (K (mk_set_incl_hset_tac def rec_Suc))
+ |> Thm.close_derivation)
+ goals defs rec_Sucs)
+ set_incl_hset_goalss hset_defss hset_rec_Sucss,
+ map3 (fn goalss => fn defsi => fn rec_Sucs =>
+ map3 (fn k => fn goals => fn defsk =>
+ map4 (fn goal => fn defk => fn defi => fn rec_Suc =>
+ Skip_Proof.prove lthy [] [] goal
+ (K (mk_set_hset_incl_hset_tac n [defk, defi] rec_Suc k))
+ |> Thm.close_derivation)
+ goals defsk defsi rec_Sucs)
+ ks goalss hset_defss)
+ set_hset_incl_hset_goalsss hset_defss hset_rec_Sucss)
+ end;
+
+ val set_incl_hset_thmss' = transpose set_incl_hset_thmss;
+ val set_hset_incl_hset_thmsss' = transpose (map transpose set_hset_incl_hset_thmsss);
+ val set_hset_incl_hset_thmsss'' = map transpose set_hset_incl_hset_thmsss';
+ val set_hset_thmss = map (map (fn thm => thm RS @{thm set_mp})) set_incl_hset_thmss;
+ val set_hset_hset_thmsss = map (map (map (fn thm => thm RS @{thm set_mp})))
+ set_hset_incl_hset_thmsss;
+ val set_hset_thmss' = transpose set_hset_thmss;
+ val set_hset_hset_thmsss' = transpose (map transpose set_hset_hset_thmsss);
+
+ val set_incl_hin_thmss =
+ let
+ fun mk_set_incl_hin s x hsets1 set hsets2 T =
+ fold_rev Logic.all (x :: ss @ As)
+ (Logic.list_implies
+ (map2 (fn hset => fn A => HOLogic.mk_Trueprop (mk_subset (hset $ x) A)) hsets1 As,
+ HOLogic.mk_Trueprop (mk_subset (set $ (s $ x)) (mk_in As hsets2 T))));
+
+ val set_incl_hin_goalss =
+ map4 (fn s => fn x => fn sets => fn hsets =>
+ map3 (mk_set_incl_hin s x hsets) (drop m sets) hsetssAs activeAs)
+ ss zs setssAs hsetssAs;
+ in
+ map2 (map2 (fn goal => fn thms =>
+ Skip_Proof.prove lthy [] [] goal (K (mk_set_incl_hin_tac thms))
+ |> Thm.close_derivation))
+ set_incl_hin_goalss set_hset_incl_hset_thmsss
+ end;
+
+ val hset_minimal_thms =
+ let
+ fun mk_passive_prem set s x K =
+ Logic.all x (HOLogic.mk_Trueprop (mk_subset (set $ (s $ x)) (K $ x)));
+
+ fun mk_active_prem s x1 K1 set x2 K2 =
+ fold_rev Logic.all [x1, x2]
+ (Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (x2, set $ (s $ x1))),
+ HOLogic.mk_Trueprop (mk_subset (K2 $ x2) (K1 $ x1))));
+
+ val premss = map2 (fn j => fn Ks =>
+ map4 mk_passive_prem (map (fn xs => nth xs (j - 1)) setssAs) ss zs Ks @
+ flat (map4 (fn sets => fn s => fn x1 => fn K1 =>
+ map3 (mk_active_prem s x1 K1) (drop m sets) zs_copy Ks) setssAs ss zs Ks))
+ ls Kss;
+
+ val hset_rec_minimal_thms =
+ let
+ fun mk_conjunct j T i K x = mk_subset (mk_hset_rec ss nat i j T $ x) (K $ x);
+ fun mk_concl j T Ks = list_all_free zs
+ (Library.foldr1 HOLogic.mk_conj (map3 (mk_conjunct j T) ks Ks zs));
+ val concls = map3 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 certify lthy) [Term.absfree nat' phi, nat]) concls;
+ in
+ map4 (fn goal => fn cts => fn hset_rec_0s => fn hset_rec_Sucs =>
+ singleton (Proof_Context.export names_lthy lthy)
+ (Skip_Proof.prove lthy [] [] goal
+ (mk_hset_rec_minimal_tac m cts hset_rec_0s hset_rec_Sucs))
+ |> Thm.close_derivation)
+ goals ctss hset_rec_0ss' hset_rec_Sucss'
+ end;
+
+ fun mk_conjunct j T i K x = mk_subset (mk_hset ss i j T $ x) (K $ x);
+ fun mk_concl j T Ks = Library.foldr1 HOLogic.mk_conj (map3 (mk_conjunct j T) ks Ks zs);
+ val concls = map3 mk_concl ls passiveAs Kss;
+
+ val goals = map3 (fn Ks => fn prems => fn concl =>
+ fold_rev Logic.all (Ks @ ss @ zs)
+ (Logic.list_implies (prems, HOLogic.mk_Trueprop concl))) Kss premss concls;
+ in
+ map3 (fn goal => fn hset_defs => fn hset_rec_minimal =>
+ Skip_Proof.prove lthy [] [] goal
+ (mk_hset_minimal_tac n hset_defs hset_rec_minimal)
+ |> Thm.close_derivation)
+ goals hset_defss' hset_rec_minimal_thms
+ end;
+
+ val mor_hset_thmss =
+ let
+ val mor_hset_rec_thms =
+ let
+ fun mk_conjunct j T i f x B =
+ HOLogic.mk_imp (HOLogic.mk_mem (x, B), HOLogic.mk_eq
+ (mk_hset_rec s's nat i j T $ (f $ x), mk_hset_rec ss nat i j T $ x));
+
+ fun mk_concl j T = list_all_free zs
+ (Library.foldr1 HOLogic.mk_conj (map4 (mk_conjunct j T) ks fs zs Bs));
+ val concls = map2 mk_concl ls passiveAs;
+
+ val ctss =
+ map (fn phi => map (SOME o certify lthy) [Term.absfree nat' phi, nat]) concls;
+
+ val goals = map (fn concl =>
+ Logic.list_implies ([coalg_prem, mor_prem], HOLogic.mk_Trueprop concl)) concls;
+ in
+ map5 (fn j => fn goal => fn cts => fn hset_rec_0s => fn hset_rec_Sucs =>
+ singleton (Proof_Context.export names_lthy lthy)
+ (Skip_Proof.prove lthy [] [] goal
+ (K (mk_mor_hset_rec_tac m n cts j hset_rec_0s hset_rec_Sucs
+ morE_thms set_natural'ss coalg_set_thmss)))
+ |> Thm.close_derivation)
+ ls goals ctss hset_rec_0ss' hset_rec_Sucss'
+ end;
+
+ val mor_hset_rec_thmss = map (fn thm => map (fn i =>
+ mk_specN n thm RS mk_conjunctN n i RS mp) ks) mor_hset_rec_thms;
+
+ fun mk_prem x B = HOLogic.mk_Trueprop (HOLogic.mk_mem (x, B));
+
+ fun mk_concl j T i f x =
+ mk_Trueprop_eq (mk_hset s's i j T $ (f $ x), mk_hset ss i j T $ x);
+
+ val goalss = map2 (fn j => fn T => map4 (fn i => fn f => fn x => fn B =>
+ fold_rev Logic.all (x :: As @ Bs @ ss @ B's @ s's @ fs)
+ (Logic.list_implies ([coalg_prem, mor_prem,
+ mk_prem x B], mk_concl j T i f x))) ks fs zs Bs) ls passiveAs;
+ in
+ map3 (map3 (fn goal => fn hset_def => fn mor_hset_rec =>
+ Skip_Proof.prove lthy [] [] goal
+ (K (mk_mor_hset_tac hset_def mor_hset_rec))
+ |> Thm.close_derivation))
+ goalss hset_defss' mor_hset_rec_thmss
+ end;
+
+ val timer = time (timer "Hereditary sets");
+
+ (* bisimulation *)
+
+ val bis_bind = Binding.suffix_name ("_" ^ bisN) b;
+ val bis_name = Binding.name_of bis_bind;
+ val bis_def_bind = (Thm.def_binding bis_bind, []);
+
+ fun mk_bis_le_conjunct R B1 B2 = mk_subset R (mk_Times (B1, B2));
+ val bis_le = Library.foldr1 HOLogic.mk_conj (map3 mk_bis_le_conjunct Rs Bs B's)
+
+ val bis_spec =
+ let
+ val bisT = Library.foldr (op -->) (ATs @ BTs @ sTs @ B'Ts @ s'Ts @ setRTs, HOLogic.boolT);
+
+ 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 (As @ 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 lhs = Term.list_comb (Free (bis_name, bisT), As @ Bs @ ss @ B's @ s's @ Rs);
+ val rhs = HOLogic.mk_conj
+ (bis_le, Library.foldr1 HOLogic.mk_conj
+ (map9 mk_bis Rs ss s's zs z's RFs map_fsts map_snds bis_setss))
+ in
+ mk_Trueprop_eq (lhs, rhs)
+ end;
+
+ val ((bis_free, (_, bis_def_free)), (lthy, lthy_old)) =
+ lthy
+ |> Specification.definition (SOME (bis_bind, NONE, NoSyn), (bis_def_bind, bis_spec))
+ ||> `Local_Theory.restore;
+
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+ val bis = fst (Term.dest_Const (Morphism.term phi bis_free));
+ val bis_def = Morphism.thm phi bis_def_free;
+
+ fun mk_bis As Bs1 ss1 Bs2 ss2 Rs =
+ let
+ val args = As @ 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 bis_cong_thm =
+ let
+ val prems = map HOLogic.mk_Trueprop
+ (mk_bis As Bs ss B's s's Rs :: map2 (curry HOLogic.mk_eq) Rs_copy Rs)
+ val concl = HOLogic.mk_Trueprop (mk_bis As Bs ss B's s's Rs_copy);
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (As @ Bs @ ss @ B's @ s's @ Rs @ Rs_copy)
+ (Logic.list_implies (prems, concl)))
+ (K ((hyp_subst_tac THEN' atac) 1))
+ |> Thm.close_derivation
+ end;
+
+ val bis_srel_thm =
+ let
+ fun mk_conjunct R s s' b1 b2 srel =
+ list_all_free [b1, b2] (HOLogic.mk_imp
+ (HOLogic.mk_mem (HOLogic.mk_prod (b1, b2), R),
+ HOLogic.mk_mem (HOLogic.mk_prod (s $ b1, s' $ b2),
+ Term.list_comb (srel, passive_diags @ Rs))));
+
+ val rhs = HOLogic.mk_conj
+ (bis_le, Library.foldr1 HOLogic.mk_conj
+ (map6 mk_conjunct Rs ss s's zs z's relsAsBs))
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (As @ Bs @ ss @ B's @ s's @ Rs)
+ (mk_Trueprop_eq (mk_bis As Bs ss B's s's Rs, rhs)))
+ (K (mk_bis_srel_tac m bis_def srel_O_Grs map_comp's map_congs set_natural'ss))
+ |> Thm.close_derivation
+ end;
+
+ val bis_converse_thm =
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (As @ Bs @ ss @ B's @ s's @ Rs)
+ (Logic.mk_implies
+ (HOLogic.mk_Trueprop (mk_bis As Bs ss B's s's Rs),
+ HOLogic.mk_Trueprop (mk_bis As B's s's Bs ss (map mk_converse Rs)))))
+ (K (mk_bis_converse_tac m bis_srel_thm srel_congs srel_converses))
+ |> Thm.close_derivation;
+
+ val bis_O_thm =
+ let
+ val prems =
+ [HOLogic.mk_Trueprop (mk_bis As Bs ss B's s's Rs),
+ HOLogic.mk_Trueprop (mk_bis As B's s's B''s s''s R's)];
+ val concl =
+ HOLogic.mk_Trueprop (mk_bis As Bs ss B''s s''s (map2 (curry mk_rel_comp) Rs R's));
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (As @ Bs @ ss @ B's @ s's @ B''s @ s''s @ Rs @ R's)
+ (Logic.list_implies (prems, concl)))
+ (K (mk_bis_O_tac m bis_srel_thm srel_congs srel_Os))
+ |> Thm.close_derivation
+ end;
+
+ val bis_Gr_thm =
+ let
+ val concl =
+ HOLogic.mk_Trueprop (mk_bis As Bs ss B's s's (map2 mk_Gr Bs fs));
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (As @ Bs @ ss @ B's @ s's @ fs)
+ (Logic.list_implies ([coalg_prem, mor_prem], concl)))
+ (mk_bis_Gr_tac bis_srel_thm srel_Grs mor_image_thms morE_thms coalg_in_thms)
+ |> Thm.close_derivation
+ 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_diag_thm = bis_cong_thm OF ((mor_id_thm RSN (2, bis_Gr_thm)) ::
+ replicate n @{thm diag_Gr});
+
+ val bis_Union_thm =
+ let
+ val prem =
+ HOLogic.mk_Trueprop (mk_Ball Idx
+ (Term.absfree idx' (mk_bis As Bs ss B's s's (map (fn R => R $ idx) Ris))));
+ val concl =
+ HOLogic.mk_Trueprop (mk_bis As Bs ss B's s's (map (mk_UNION Idx) Ris));
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (Idx :: As @ Bs @ ss @ B's @ s's @ Ris)
+ (Logic.mk_implies (prem, concl)))
+ (mk_bis_Union_tac bis_def in_mono'_thms)
+ |> Thm.close_derivation
+ end;
+
+ (* self-bisimulation *)
+
+ fun mk_sbis As Bs ss Rs = mk_bis As Bs ss Bs ss Rs;
+
+ val sbis_prem = HOLogic.mk_Trueprop (mk_sbis As Bs ss sRs);
+
+ (* largest self-bisimulation *)
+
+ fun lsbis_bind i = Binding.suffix_name ("_" ^ lsbisN ^ (if n = 1 then "" else
+ string_of_int i)) b;
+ val lsbis_name = Binding.name_of o lsbis_bind;
+ 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 As Bs ss sRs)));
+
+ fun lsbis_spec i RT =
+ let
+ fun mk_lsbisT RT =
+ Library.foldr (op -->) (map fastype_of (As @ Bs @ ss), RT);
+ val lhs = Term.list_comb (Free (lsbis_name i, mk_lsbisT RT), As @ Bs @ ss);
+ val rhs = mk_UNION all_sbis (Term.absfree Rtuple' (mk_nthN n Rtuple i));
+ in
+ mk_Trueprop_eq (lhs, rhs)
+ end;
+
+ val ((lsbis_frees, (_, lsbis_def_frees)), (lthy, lthy_old)) =
+ lthy
+ |> fold_map2 (fn i => fn RT => Specification.definition
+ (SOME (lsbis_bind i, NONE, NoSyn), (lsbis_def_bind i, lsbis_spec i RT))) ks setsRTs
+ |>> apsnd split_list o split_list
+ ||> `Local_Theory.restore;
+
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+
+ val lsbis_defs = map (Morphism.thm phi) lsbis_def_frees;
+ val lsbiss = map (fst o Term.dest_Const o Morphism.term phi) lsbis_frees;
+
+ fun mk_lsbis As Bs ss i =
+ let
+ val args = As @ 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 sbis_lsbis_thm =
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (As @ Bs @ ss)
+ (HOLogic.mk_Trueprop (mk_sbis As Bs ss (map (mk_lsbis As Bs ss) ks))))
+ (K (mk_sbis_lsbis_tac lsbis_defs bis_Union_thm bis_cong_thm))
+ |> Thm.close_derivation;
+
+ val lsbis_incl_thms = map (fn i => sbis_lsbis_thm RS
+ (bis_def RS @{thm subst[of _ _ "%x. x"]} RS conjunct1 RS mk_conjunctN n i)) ks;
+ val lsbisE_thms = map (fn i => (mk_specN 2 (sbis_lsbis_thm RS
+ (bis_def RS @{thm subst[of _ _ "%x. x"]} RS conjunct2 RS mk_conjunctN n i))) RS mp) ks;
+
+ val incl_lsbis_thms =
+ let
+ fun mk_concl i R = HOLogic.mk_Trueprop (mk_subset R (mk_lsbis As Bs ss i));
+ val goals = map2 (fn i => fn R => fold_rev Logic.all (As @ Bs @ ss @ sRs)
+ (Logic.mk_implies (sbis_prem, mk_concl i R))) ks sRs;
+ in
+ map3 (fn goal => fn i => fn def => Skip_Proof.prove lthy [] [] goal
+ (K (mk_incl_lsbis_tac n i def)) |> Thm.close_derivation) goals ks lsbis_defs
+ end;
+
+ val equiv_lsbis_thms =
+ let
+ fun mk_concl i B = HOLogic.mk_Trueprop (mk_equiv B (mk_lsbis As Bs ss i));
+ val goals = map2 (fn i => fn B => fold_rev Logic.all (As @ Bs @ ss)
+ (Logic.mk_implies (coalg_prem, mk_concl i B))) ks Bs;
+ in
+ map3 (fn goal => fn l_incl => fn incl_l =>
+ Skip_Proof.prove lthy [] [] goal
+ (K (mk_equiv_lsbis_tac sbis_lsbis_thm l_incl incl_l
+ bis_diag_thm bis_converse_thm bis_O_thm))
+ |> Thm.close_derivation)
+ goals lsbis_incl_thms incl_lsbis_thms
+ end;
+
+ val timer = time (timer "Bisimulations");
+
+ (* bounds *)
+
+ val (lthy, sbd, sbdT,
+ sbd_card_order, sbd_Cinfinite, sbd_Cnotzero, sbd_Card_order, set_sbdss, in_sbds) =
+ if n = 1
+ then (lthy, sum_bd, sum_bdT,
+ bd_card_order, bd_Cinfinite, bd_Cnotzero, bd_Card_order, set_bdss, in_bds)
+ else
+ let
+ val sbdT_bind = Binding.suffix_name ("_" ^ sum_bdTN) b;
+
+ val ((sbdT_name, (sbdT_glob_info, sbdT_loc_info)), lthy) =
+ typedef false NONE (sbdT_bind, params, NoSyn)
+ (HOLogic.mk_UNIV sum_bdT) NONE (EVERY' [rtac exI, rtac UNIV_I] 1) lthy;
+
+ val sbdT = Type (sbdT_name, params');
+ val Abs_sbdT = Const (#Abs_name sbdT_glob_info, sum_bdT --> sbdT);
+
+ val sbd_bind = Binding.suffix_name ("_" ^ sum_bdN) b;
+ val sbd_name = Binding.name_of sbd_bind;
+ val sbd_def_bind = (Thm.def_binding sbd_bind, []);
+
+ val sbd_spec = HOLogic.mk_Trueprop
+ (HOLogic.mk_eq (Free (sbd_name, mk_relT (`I sbdT)), mk_dir_image sum_bd Abs_sbdT));
+
+ val ((sbd_free, (_, sbd_def_free)), (lthy, lthy_old)) =
+ lthy
+ |> Specification.definition (SOME (sbd_bind, NONE, NoSyn), (sbd_def_bind, sbd_spec))
+ ||> `Local_Theory.restore;
+
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+
+ val sbd_def = 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);
+
+ fun mk_sum_Cinfinite [thm] = thm
+ | mk_sum_Cinfinite (thm :: thms) =
+ @{thm Cinfinite_csum_strong} OF [thm, mk_sum_Cinfinite thms];
+
+ val sum_Cinfinite = mk_sum_Cinfinite bd_Cinfinites;
+ val sum_Card_order = sum_Cinfinite RS conjunct2;
+
+ fun mk_sum_card_order [thm] = thm
+ | mk_sum_card_order (thm :: thms) =
+ @{thm card_order_csum} OF [thm, mk_sum_card_order thms];
+
+ val sum_card_order = mk_sum_card_order bd_card_orders;
+
+ val sbd_ordIso = fold_thms lthy [sbd_def]
+ (@{thm dir_image} OF [Abs_sbdT_inj, sum_Card_order]);
+ val sbd_card_order = fold_thms lthy [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_Cnotzero = sbd_Cinfinite RS @{thm Cinfinite_Cnotzero};
+ 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 = map3 mk_set_sbd ks bd_Card_orders set_bdss;
+
+ fun mk_in_sbd 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 _ csum_Cnotzero2[OF ctwo_Cnotzero]]]}));
+ val in_sbds = map4 mk_in_sbd ks bd_Card_orders bd_Cnotzeros in_bds;
+ in
+ (lthy, sbd, sbdT,
+ sbd_card_order, sbd_Cinfinite, sbd_Cnotzero, sbd_Card_order, set_sbdss, in_sbds)
+ end;
+
+ fun mk_sbd_sbd 1 = sbd_Card_order RS @{thm ordIso_refl}
+ | mk_sbd_sbd n = @{thm csum_absorb1} OF
+ [sbd_Cinfinite, mk_sbd_sbd (n - 1) RS @{thm ordIso_imp_ordLeq}];
+
+ val sbd_sbd_thm = mk_sbd_sbd n;
+
+ val sbdTs = replicate n sbdT;
+ val sum_sbd = Library.foldr1 (uncurry mk_csum) (replicate n sbd);
+ 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 = map4 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 = map4 mk_map_of_bnf Dss (replicate n bdTs) (replicate n treeTs) bnfs;
+ val final_maps = map4 mk_map_of_bnf Dss (replicate n treeTs) (replicate n treeQTs) bnfs;
+ val tree_setss = mk_setss treeTs;
+ 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 LevT = Library.foldr (op -->) (sTs, HOLogic.natT --> Lev_recT);
+
+ val Nil = HOLogic.mk_tuple (map3 (fn i => fn z => fn z'=>
+ Term.absfree z' (mk_InN activeAs z i)) ks zs zs');
+ val rv_recT = fastype_of Nil;
+ val rvT = Library.foldr (op -->) (sTs, sum_sbd_listT --> rv_recT);
+
+ val (((((((((((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')),
+ names_lthy) = names_lthy
+ |> 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 *)
+
+ fun isNode_bind i = Binding.suffix_name ("_" ^ isNodeN ^ (if n = 1 then "" else
+ string_of_int i)) b;
+ val isNode_name = Binding.name_of o isNode_bind;
+ val isNode_def_bind = rpair [] o Thm.def_binding o isNode_bind;
+
+ val isNodeT =
+ Library.foldr (op -->) (map fastype_of (As @ [Kl, lab, kl]), HOLogic.boolT);
+
+ val Succs = map3 (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 (passive_sets, active_sets) = chop m (map (fn set => set $ x) sets);
+ val lhs = Term.list_comb (Free (isNode_name i, isNodeT), As @ [Kl, lab, kl]);
+ val rhs = list_exists_free [x]
+ (Library.foldr1 HOLogic.mk_conj (HOLogic.mk_eq (lab $ kl, mk_InN bdFTs x i) ::
+ map2 mk_subset passive_sets As @ map2 (curry HOLogic.mk_eq) active_sets Succs));
+ in
+ mk_Trueprop_eq (lhs, rhs)
+ end;
+
+ val ((isNode_frees, (_, isNode_def_frees)), (lthy, lthy_old)) =
+ lthy
+ |> fold_map3 (fn i => fn x => fn sets => Specification.definition
+ (SOME (isNode_bind i, NONE, NoSyn), (isNode_def_bind i, isNode_spec sets x i)))
+ ks xs isNode_setss
+ |>> apsnd split_list o split_list
+ ||> `Local_Theory.restore;
+
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+
+ val isNode_defs = map (Morphism.thm phi) isNode_def_frees;
+ val isNodes = map (fst o Term.dest_Const o Morphism.term phi) isNode_frees;
+
+ fun mk_isNode As kl i =
+ Term.list_comb (Const (nth isNodes (i - 1), isNodeT), As @ [Kl, lab, kl]);
+
+ val isTree =
+ let
+ val empty = HOLogic.mk_mem (HOLogic.mk_list sum_sbdT [], Kl);
+ val Field = mk_subset Kl (mk_Field (mk_clists sum_sbd));
+ val prefCl = mk_prefCl Kl;
+
+ val tree = mk_Ball Kl (Term.absfree kl'
+ (HOLogic.mk_conj
+ (Library.foldr1 HOLogic.mk_disj (map (mk_isNode As kl) ks),
+ Library.foldr1 HOLogic.mk_conj (map4 (fn Succ => fn i => fn k => fn k' =>
+ mk_Ball Succ (Term.absfree k' (mk_isNode As
+ (mk_append (kl, HOLogic.mk_list sum_sbdT [mk_InN sbdTs k i])) i)))
+ Succs ks kks kks'))));
+
+ val undef = list_all_free [kl] (HOLogic.mk_imp
+ (HOLogic.mk_not (HOLogic.mk_mem (kl, Kl)),
+ HOLogic.mk_eq (lab $ kl, mk_undefined sbdFT)));
+ in
+ Library.foldr1 HOLogic.mk_conj [empty, Field, prefCl, tree, undef]
+ end;
+
+ fun carT_bind i = Binding.suffix_name ("_" ^ carTN ^ (if n = 1 then "" else
+ string_of_int i)) b;
+ val carT_name = Binding.name_of o carT_bind;
+ val carT_def_bind = rpair [] o Thm.def_binding o carT_bind;
+
+ fun carT_spec i =
+ let
+ val carTT = Library.foldr (op -->) (ATs, HOLogic.mk_setT treeT);
+
+ val lhs = Term.list_comb (Free (carT_name i, carTT), As);
+ val rhs = 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 As (HOLogic.mk_list sum_sbdT []) i))));
+ in
+ mk_Trueprop_eq (lhs, rhs)
+ end;
+
+ val ((carT_frees, (_, carT_def_frees)), (lthy, lthy_old)) =
+ lthy
+ |> fold_map (fn i => Specification.definition
+ (SOME (carT_bind i, NONE, NoSyn), (carT_def_bind i, carT_spec i))) ks
+ |>> apsnd split_list o split_list
+ ||> `Local_Theory.restore;
+
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+
+ val carT_defs = map (Morphism.thm phi) carT_def_frees;
+ val carTs = map (fst o Term.dest_Const o Morphism.term phi) carT_frees;
+
+ fun mk_carT As i = Term.list_comb
+ (Const (nth carTs (i - 1),
+ Library.foldr (op -->) (map fastype_of As, HOLogic.mk_setT treeT)), As);
+
+ fun strT_bind i = Binding.suffix_name ("_" ^ strTN ^ (if n = 1 then "" else
+ string_of_int i)) b;
+ val strT_name = Binding.name_of o strT_bind;
+ val strT_def_bind = rpair [] o Thm.def_binding o strT_bind;
+
+ fun strT_spec mapFT FT i =
+ let
+ val strTT = treeT --> FT;
+
+ 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 @ map3 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);
+ val lhs = Free (strT_name i, strTT);
+ val rhs = HOLogic.mk_split (Term.absfree Kl' (Term.absfree lab'
+ (mk_sum_caseN fs $ (lab $ HOLogic.mk_list sum_sbdT []))));
+ in
+ mk_Trueprop_eq (lhs, rhs)
+ end;
+
+ val ((strT_frees, (_, strT_def_frees)), (lthy, lthy_old)) =
+ lthy
+ |> fold_map3 (fn i => fn mapFT => fn FT => Specification.definition
+ (SOME (strT_bind i, NONE, NoSyn), (strT_def_bind i, strT_spec mapFT FT i)))
+ ks tree_maps treeFTs
+ |>> apsnd split_list o split_list
+ ||> `Local_Theory.restore;
+
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+
+ val strT_defs = map ((fn def => trans OF [def RS fun_cong, @{thm prod.cases}]) o
+ Morphism.thm phi) 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 As) ks;
+ val carTAs_copy = map (mk_carT As_copy) ks;
+ val strTAs = map2 mk_strT treeFTs ks;
+ val hset_strTss = map (fn i => map2 (mk_hset strTAs i) ls passiveAs) ks;
+
+ val coalgT_thm =
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all As (HOLogic.mk_Trueprop (mk_coalg As carTAs strTAs)))
+ (mk_coalgT_tac m (coalg_def :: isNode_defs @ carT_defs) strT_defs set_natural'ss)
+ |> Thm.close_derivation;
+
+ val card_of_carT_thms =
+ let
+ val lhs = mk_card_of
+ (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)), isTree))));
+ val rhs = mk_cexp
+ (if m = 0 then ctwo else
+ (mk_csum (Library.foldr1 (uncurry mk_csum) (map mk_card_of As)) ctwo))
+ (mk_cexp sbd sbd);
+ val card_of_carT =
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all As (HOLogic.mk_Trueprop (mk_ordLeq lhs rhs)))
+ (K (mk_card_of_carT_tac m isNode_defs sbd_sbd_thm
+ sbd_card_order sbd_Card_order sbd_Cinfinite sbd_Cnotzero in_sbds))
+ |> Thm.close_derivation
+ in
+ map (fn def => @{thm ordLeq_transitive[OF
+ card_of_mono1[OF ord_eq_le_trans[OF _ Collect_restrict']]]} OF [def, card_of_carT])
+ carT_defs
+ end;
+
+ val carT_set_thmss =
+ let
+ val Kl_lab = HOLogic.mk_prod (Kl, lab);
+ fun mk_goal carT strT set k i =
+ fold_rev Logic.all (sumx :: Kl :: lab :: k :: kl :: As)
+ (Logic.list_implies (map HOLogic.mk_Trueprop
+ [HOLogic.mk_mem (Kl_lab, carT), HOLogic.mk_mem (mk_Cons sumx kl, Kl),
+ HOLogic.mk_eq (sumx, mk_InN sbdTs k i)],
+ HOLogic.mk_Trueprop (HOLogic.mk_mem
+ (HOLogic.mk_prod (mk_Shift Kl sumx, mk_shift lab sumx),
+ set $ (strT $ Kl_lab)))));
+
+ val goalss = map3 (fn carT => fn strT => fn sets =>
+ map3 (mk_goal carT strT) (drop m sets) kks ks) carTAs strTAs tree_setss;
+ in
+ map6 (fn i => fn goals =>
+ fn carT_def => fn strT_def => fn isNode_def => fn set_naturals =>
+ map2 (fn goal => fn set_natural =>
+ Skip_Proof.prove lthy [] [] goal
+ (mk_carT_set_tac n i carT_def strT_def isNode_def set_natural)
+ |> Thm.close_derivation)
+ goals (drop m set_naturals))
+ ks goalss carT_defs strT_defs isNode_defs set_natural'ss
+ end;
+
+ val carT_set_thmss' = transpose carT_set_thmss;
+
+ val isNode_hset_thmss =
+ let
+ val Kl_lab = HOLogic.mk_prod (Kl, lab);
+ fun mk_Kl_lab carT = HOLogic.mk_mem (Kl_lab, carT);
+
+ val strT_hset_thmsss =
+ let
+ val strT_hset_thms =
+ let
+ fun mk_lab_kl i x = HOLogic.mk_eq (lab $ kl, mk_InN bdFTs x i);
+
+ fun mk_inner_conjunct j T i x set i' carT =
+ HOLogic.mk_imp (HOLogic.mk_conj (mk_Kl_lab carT, mk_lab_kl i x),
+ mk_subset (set $ x) (mk_hset strTAs i' j T $ Kl_lab));
+
+ fun mk_conjunct j T i x set =
+ Library.foldr1 HOLogic.mk_conj (map2 (mk_inner_conjunct j T i x set) ks carTAs);
+
+ fun mk_concl j T = list_all_free (Kl :: lab :: xs @ As)
+ (HOLogic.mk_imp (HOLogic.mk_mem (kl, Kl),
+ Library.foldr1 HOLogic.mk_conj (map3 (mk_conjunct j T)
+ ks xs (map (fn xs => nth xs (j - 1)) isNode_setss))));
+ val concls = map2 mk_concl ls passiveAs;
+
+ val cTs = [SOME (certifyT lthy sum_sbdT)];
+ val arg_cong_cTs = map (SOME o certifyT lthy) treeFTs;
+ val ctss =
+ map (fn phi => map (SOME o certify lthy) [Term.absfree kl' phi, kl]) concls;
+
+ val goals = map HOLogic.mk_Trueprop concls;
+ in
+ map5 (fn j => fn goal => fn cts => fn set_incl_hsets => fn set_hset_incl_hsetss =>
+ singleton (Proof_Context.export names_lthy lthy)
+ (Skip_Proof.prove lthy [] [] goal
+ (K (mk_strT_hset_tac n m j arg_cong_cTs cTs cts
+ carT_defs strT_defs isNode_defs
+ set_incl_hsets set_hset_incl_hsetss coalg_set_thmss' carT_set_thmss'
+ coalgT_thm set_natural'ss)))
+ |> Thm.close_derivation)
+ ls goals ctss set_incl_hset_thmss' set_hset_incl_hset_thmsss''
+ end;
+
+ val strT_hset'_thms = map (fn thm => mk_specN (2 + n + m) thm RS mp) strT_hset_thms;
+ in
+ map (fn thm => map (fn i => map (fn i' =>
+ thm RS mk_conjunctN n i RS mk_conjunctN n i' RS mp) ks) ks) strT_hset'_thms
+ end;
+
+ val carT_prems = map (fn carT =>
+ HOLogic.mk_Trueprop (HOLogic.mk_mem (Kl_lab, carT))) carTAs_copy;
+ val prem = HOLogic.mk_Trueprop (HOLogic.mk_mem (kl, Kl));
+ val in_prems = map (fn hsets =>
+ HOLogic.mk_Trueprop (HOLogic.mk_mem (Kl_lab, mk_in As hsets treeT))) hset_strTss;
+ val isNode_premss = replicate n (map (HOLogic.mk_Trueprop o mk_isNode As_copy kl) ks);
+ val conclss = replicate n (map (HOLogic.mk_Trueprop o mk_isNode As kl) ks);
+ in
+ map5 (fn carT_prem => fn isNode_prems => fn in_prem => fn concls => fn strT_hset_thmss =>
+ map4 (fn isNode_prem => fn concl => fn isNode_def => fn strT_hset_thms =>
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (Kl :: lab :: kl :: As @ As_copy)
+ (Logic.list_implies ([carT_prem, prem, isNode_prem, in_prem], concl)))
+ (mk_isNode_hset_tac n isNode_def strT_hset_thms)
+ |> Thm.close_derivation)
+ isNode_prems concls isNode_defs
+ (if m = 0 then replicate n [] else transpose strT_hset_thmss))
+ carT_prems isNode_premss in_prems conclss
+ (if m = 0 then replicate n [] else transpose (map transpose strT_hset_thmsss))
+ end;
+
+ 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 to_sbd_thmss = mk_to_sbd_thmss @{thm toCard};
+ val from_to_sbd_thmss = mk_to_sbd_thmss @{thm fromCard_toCard};
+
+ val Lev_bind = Binding.suffix_name ("_" ^ LevN) b;
+ val Lev_name = Binding.name_of Lev_bind;
+ 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 (map3 mk_set ks sets zs_copy))
+ end;
+
+ val Suc = Term.absdummy HOLogic.natT (Term.absfree Lev_rec'
+ (HOLogic.mk_tuple (map5 mk_Suc ks ss setssAs zs zs')));
+
+ val lhs = Term.list_comb (Free (Lev_name, LevT), ss);
+ val rhs = mk_nat_rec Zero Suc;
+ in
+ mk_Trueprop_eq (lhs, rhs)
+ end;
+
+ val ((Lev_free, (_, Lev_def_free)), (lthy, lthy_old)) =
+ lthy
+ |> Specification.definition (SOME (Lev_bind, NONE, NoSyn), (Lev_def_bind, Lev_spec))
+ ||> `Local_Theory.restore;
+
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+
+ val Lev_def = 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 nat_rec_0} [Lev_def]);
+ val Lev_Sucs = flat (mk_rec_simps n @{thm nat_rec_Suc} [Lev_def]);
+
+ val rv_bind = Binding.suffix_name ("_" ^ rvN) b;
+ val rv_name = Binding.name_of rv_bind;
+ 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_sum_caseN (map mk_case ks) $ sumx)
+ end;
+
+ val Cons = Term.absfree sumx' (Term.absdummy sum_sbd_listT (Term.absfree rv_rec'
+ (HOLogic.mk_tuple (map4 mk_Cons ks ss zs zs'))));
+
+ val lhs = Term.list_comb (Free (rv_name, rvT), ss);
+ val rhs = mk_list_rec Nil Cons;
+ in
+ mk_Trueprop_eq (lhs, rhs)
+ end;
+
+ val ((rv_free, (_, rv_def_free)), (lthy, lthy_old)) =
+ lthy
+ |> Specification.definition (SOME (rv_bind, NONE, NoSyn), (rv_def_bind, rv_spec))
+ ||> `Local_Theory.restore;
+
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+
+ val rv_def = 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 list_rec_Nil} [rv_def]);
+ val rv_Conss = flat (mk_rec_simps n @{thm list_rec_Cons} [rv_def]);
+
+ fun beh_bind i = Binding.suffix_name ("_" ^ behN ^ (if n = 1 then "" else
+ string_of_int i)) b;
+ val beh_name = Binding.name_of o beh_bind;
+ val beh_def_bind = rpair [] o Thm.def_binding o beh_bind;
+
+ fun beh_spec i z =
+ let
+ val mk_behT = Library.foldr (op -->) (map fastype_of (ss @ [z]), treeT);
+
+ 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_If
+ (HOLogic.mk_mem (kl, mk_Lev ss (mk_size kl) i $ z))
+ (mk_sum_caseN (map5 mk_case ks to_sbd_maps ss zs zs') $ (mk_rv ss kl i $ z))
+ (mk_undefined sbdFT));
+
+ val lhs = Term.list_comb (Free (beh_name i, mk_behT), ss) $ z;
+ val rhs = HOLogic.mk_prod (mk_UNION (HOLogic.mk_UNIV HOLogic.natT)
+ (Term.absfree nat' (mk_Lev ss nat i $ z)), Lab);
+ in
+ mk_Trueprop_eq (lhs, rhs)
+ end;
+
+ val ((beh_frees, (_, beh_def_frees)), (lthy, lthy_old)) =
+ lthy
+ |> fold_map2 (fn i => fn z => Specification.definition
+ (SOME (beh_bind i, NONE, NoSyn), (beh_def_bind i, beh_spec i z))) ks zs
+ |>> apsnd split_list o split_list
+ ||> `Local_Theory.restore;
+
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+
+ val beh_defs = map (Morphism.thm phi) 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 Lev_sbd_thms =
+ let
+ fun mk_conjunct i z = mk_subset (mk_Lev ss nat i $ z) (mk_Field (mk_clists sum_sbd));
+ val goal = list_all_free zs
+ (Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct ks zs));
+
+ val cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
+
+ val Lev_sbd = singleton (Proof_Context.export names_lthy lthy)
+ (Skip_Proof.prove lthy [] [] (HOLogic.mk_Trueprop goal)
+ (K (mk_Lev_sbd_tac cts Lev_0s Lev_Sucs to_sbd_thmss))
+ |> Thm.close_derivation);
+
+ val Lev_sbd' = mk_specN n Lev_sbd;
+ in
+ map (fn i => Lev_sbd' RS mk_conjunctN n i) ks
+ end;
+
+ 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 cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
+
+ val length_Lev = singleton (Proof_Context.export names_lthy lthy)
+ (Skip_Proof.prove lthy [] [] (HOLogic.mk_Trueprop goal)
+ (K (mk_length_Lev_tac cts Lev_0s Lev_Sucs))
+ |> Thm.close_derivation);
+
+ 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 = fold_rev Logic.all (z :: kl :: nat :: ss) (Logic.mk_implies
+ (HOLogic.mk_Trueprop (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z)),
+ HOLogic.mk_Trueprop (HOLogic.mk_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 =>
+ Skip_Proof.prove lthy [] [] goal (K (mk_length_Lev'_tac length_Lev))
+ |> Thm.close_derivation) goals length_Levs;
+ in
+ (length_Levs, length_Levs')
+ end;
+
+ val prefCl_Lev_thms =
+ let
+ fun mk_conjunct i z = HOLogic.mk_imp
+ (HOLogic.mk_conj (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z), mk_subset kl_copy kl),
+ HOLogic.mk_mem (kl_copy, mk_Lev ss (mk_size kl_copy) i $ z));
+ val goal = list_all_free (kl :: kl_copy :: zs)
+ (Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct ks zs));
+
+ val cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
+
+ val prefCl_Lev = singleton (Proof_Context.export names_lthy lthy)
+ (Skip_Proof.prove lthy [] [] (HOLogic.mk_Trueprop goal)
+ (K (mk_prefCl_Lev_tac cts Lev_0s Lev_Sucs)))
+ |> Thm.close_derivation;
+
+ val prefCl_Lev' = mk_specN (n + 2) prefCl_Lev;
+ in
+ map (fn i => prefCl_Lev' RS mk_conjunctN n i RS mp) ks
+ 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 cTs = [SOME (certifyT lthy sum_sbdT)];
+ val cts = map (SOME o certify lthy) [Term.absfree kl' goal, kl];
+
+ val rv_last = singleton (Proof_Context.export names_lthy lthy)
+ (Skip_Proof.prove lthy [] [] (HOLogic.mk_Trueprop goal)
+ (K (mk_rv_last_tac cTs cts rv_Nils rv_Conss)))
+ |> Thm.close_derivation;
+
+ 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_rv_Lev_thmsss = if m = 0 then replicate n (replicate n []) else
+ let
+ fun mk_case s sets z z_free = Term.absfree z_free (Library.foldr1 HOLogic.mk_conj
+ (map2 (fn set => fn A => mk_subset (set $ (s $ z)) A) (take m sets) As));
+
+ fun mk_conjunct i z B = HOLogic.mk_imp
+ (HOLogic.mk_conj (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z), HOLogic.mk_mem (z, B)),
+ mk_sum_caseN (map4 mk_case ss setssAs zs zs') $ (mk_rv ss kl i $ z));
+
+ val goal = list_all_free (kl :: zs)
+ (Library.foldr1 HOLogic.mk_conj (map3 mk_conjunct ks zs Bs));
+
+ val cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
+
+ val set_rv_Lev = singleton (Proof_Context.export names_lthy lthy)
+ (Skip_Proof.prove lthy [] []
+ (Logic.mk_implies (coalg_prem, HOLogic.mk_Trueprop goal))
+ (K (mk_set_rv_Lev_tac m cts Lev_0s Lev_Sucs rv_Nils rv_Conss
+ coalg_set_thmss from_to_sbd_thmss)))
+ |> Thm.close_derivation;
+
+ val set_rv_Lev' = mk_specN (n + 1) set_rv_Lev;
+ in
+ map (fn i => map (fn i' =>
+ split_conj_thm (if n = 1 then set_rv_Lev' RS mk_conjunctN n i RS mp
+ else set_rv_Lev' RS mk_conjunctN n i RS mp RSN
+ (2, @{thm sum_case_weak_cong} RS @{thm subst[of _ _ "%x. x"]}) RS
+ (mk_sum_casesN n i' RS @{thm subst[of _ _ "%x. x"]}))) 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 (map3 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 (map4 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 cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
+
+ val set_Lev = singleton (Proof_Context.export names_lthy lthy)
+ (Skip_Proof.prove lthy [] [] (HOLogic.mk_Trueprop goal)
+ (K (mk_set_Lev_tac cts Lev_0s Lev_Sucs rv_Nils rv_Conss from_to_sbd_thmss)))
+ |> Thm.close_derivation;
+
+ 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 (map4 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 cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
+
+ val set_image_Lev = singleton (Proof_Context.export names_lthy lthy)
+ (Skip_Proof.prove lthy [] [] (HOLogic.mk_Trueprop goal)
+ (K (mk_set_image_Lev_tac cts Lev_0s Lev_Sucs rv_Nils rv_Conss
+ from_to_sbd_thmss to_sbd_inj_thmss)))
+ |> Thm.close_derivation;
+
+ 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 =
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (As @ Bs @ ss) (Logic.mk_implies (coalg_prem,
+ HOLogic.mk_Trueprop (mk_mor Bs ss carTAs strTAs (map (mk_beh ss) ks)))))
+ (mk_mor_beh_tac 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 Lev_sbd_thms
+ length_Lev_thms length_Lev'_thms prefCl_Lev_thms rv_last_thmss
+ set_rv_Lev_thmsss set_Lev_thmsss set_image_Lev_thmsss
+ set_natural'ss coalg_set_thmss map_comp_id_thms map_congs map_arg_cong_thms)
+ |> Thm.close_derivation;
+
+ val timer = time (timer "Behavioral morphism");
+
+ fun mk_LSBIS As i = mk_lsbis As (map (mk_carT As) ks) strTAs i;
+ fun mk_car_final As i =
+ mk_quotient (mk_carT As i) (mk_LSBIS As i);
+ fun mk_str_final As i =
+ mk_univ (HOLogic.mk_comp (Term.list_comb (nth final_maps (i - 1),
+ passive_ids @ map (mk_proj o mk_LSBIS As) ks), nth strTAs (i - 1)));
+
+ val car_finalAs = map (mk_car_final As) ks;
+ val str_finalAs = map (mk_str_final As) ks;
+ val car_finals = map (mk_car_final passive_UNIVs) ks;
+ val str_finals = map (mk_str_final passive_UNIVs) 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 =
+ fold_rev Logic.all As (HOLogic.mk_Trueprop
+ (mk_congruent R (HOLogic.mk_comp
+ (Term.list_comb (final_map, passive_ids @ map (mk_proj o mk_LSBIS As) ks), strT))));
+
+ val goals = map3 mk_goal (map (mk_LSBIS As) ks) final_maps strTAs;
+ in
+ map4 (fn goal => fn lsbisE => fn map_comp_id => fn map_cong =>
+ Skip_Proof.prove lthy [] [] goal
+ (K (mk_congruent_str_final_tac m lsbisE map_comp_id map_cong equiv_LSBIS_thms))
+ |> Thm.close_derivation)
+ goals lsbisE_thms map_comp_id_thms map_congs
+ end;
+
+ val coalg_final_thm = Skip_Proof.prove lthy [] [] (fold_rev Logic.all As
+ (HOLogic.mk_Trueprop (mk_coalg As car_finalAs str_finalAs)))
+ (K (mk_coalg_final_tac m coalg_def congruent_str_final_thms equiv_LSBIS_thms
+ set_natural'ss coalgT_set_thmss))
+ |> Thm.close_derivation;
+
+ val mor_T_final_thm = Skip_Proof.prove lthy [] [] (fold_rev Logic.all As
+ (HOLogic.mk_Trueprop (mk_mor carTAs strTAs car_finalAs str_finalAs
+ (map (mk_proj o mk_LSBIS As) ks))))
+ (K (mk_mor_T_final_tac mor_def congruent_str_final_thms equiv_LSBIS_thms))
+ |> Thm.close_derivation;
+
+ val mor_final_thm = mor_comp_thm OF [mor_beh_thm, mor_T_final_thm];
+ val in_car_final_thms = map (fn mor_image' => mor_image' OF
+ [tcoalg_thm RS 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_map4 (fn b => fn mx => fn car_final => fn in_car_final =>
+ typedef false NONE (b, params, mx) car_final NONE
+ (EVERY' [rtac exI, rtac 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 Ts'' = mk_Ts passiveCs;
+ 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 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);
+ val prodTs = map (HOLogic.mk_prodT o `I) Ts;
+ val prodFTs = mk_FTs (passiveAs @ prodTs);
+ val FTs_setss = mk_setss (passiveAs @ Ts);
+ val prodFT_setss = mk_setss (passiveAs @ prodTs);
+ val map_FTs = map2 (fn Ds => mk_map_of_bnf Ds treeQTs (passiveAs @ Ts)) Dss bnfs;
+ val map_FT_nths = map2 (fn Ds =>
+ mk_map_of_bnf Ds (passiveAs @ prodTs) (passiveAs @ Ts)) Dss bnfs;
+ val fstsTs = map fst_const prodTs;
+ val sndsTs = map snd_const prodTs;
+ val dtorTs = map2 (curry (op -->)) Ts FTs;
+ val ctorTs = map2 (curry (op -->)) FTs Ts;
+ val unfold_fTs = map2 (curry op -->) activeAs Ts;
+ val corec_sTs = map (Term.typ_subst_atomic (activeBs ~~ Ts)) sum_sTs;
+ val corec_maps = map (Term.subst_atomic_types (activeBs ~~ Ts)) map_Inls;
+ val corec_maps_rev = map (Term.subst_atomic_types (activeBs ~~ Ts)) map_Inls_rev;
+ val corec_Inls = map (Term.subst_atomic_types (activeBs ~~ Ts)) Inls;
+
+ val (((((((((((((Jzs, Jzs'), (Jz's, Jz's')), Jzs_copy), Jzs1), Jzs2), Jpairs),
+ FJzs), TRs), unfold_fs), unfold_fs_copy), corec_ss), phis), names_lthy) = names_lthy
+ |> mk_Frees' "z" Ts
+ ||>> mk_Frees' "z" Ts'
+ ||>> mk_Frees "z" Ts
+ ||>> mk_Frees "z1" Ts
+ ||>> mk_Frees "z2" Ts
+ ||>> mk_Frees "j" (map2 (curry HOLogic.mk_prodT) Ts Ts')
+ ||>> mk_Frees "x" prodFTs
+ ||>> mk_Frees "R" (map (mk_relT o `I) Ts)
+ ||>> mk_Frees "f" unfold_fTs
+ ||>> mk_Frees "g" unfold_fTs
+ ||>> mk_Frees "s" corec_sTs
+ ||>> mk_Frees "P" (map2 mk_pred2T Ts Ts);
+
+ fun dtor_bind i = Binding.suffix_name ("_" ^ dtorN) (nth bs (i - 1));
+ val dtor_name = Binding.name_of o dtor_bind;
+ val dtor_def_bind = rpair [] o Thm.def_binding o dtor_bind;
+
+ fun dtor_spec i rep str map_FT dtorT Jz Jz' =
+ let
+ val lhs = Free (dtor_name i, dtorT);
+ val rhs = Term.absfree Jz'
+ (Term.list_comb (map_FT, map HOLogic.id_const passiveAs @ Abs_Ts) $
+ (str $ (rep $ Jz)));
+ in
+ mk_Trueprop_eq (lhs, rhs)
+ end;
+
+ val ((dtor_frees, (_, dtor_def_frees)), (lthy, lthy_old)) =
+ lthy
+ |> fold_map7 (fn i => fn rep => fn str => fn mapx => fn dtorT => fn Jz => fn Jz' =>
+ Specification.definition (SOME (dtor_bind i, NONE, NoSyn),
+ (dtor_def_bind i, dtor_spec i rep str mapx dtorT Jz Jz')))
+ ks Rep_Ts str_finals map_FTs dtorTs Jzs Jzs'
+ |>> apsnd split_list o split_list
+ ||> `Local_Theory.restore;
+
+ 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 thm => thm RS fun_cong) o Morphism.thm phi) 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 =
+ Skip_Proof.prove lthy [] []
+ (HOLogic.mk_Trueprop (mk_mor UNIVs dtors car_finals str_finals Rep_Ts))
+ (mk_mor_Rep_tac m (mor_def :: dtor_defs) Reps Abs_inverses coalg_final_set_thmss
+ map_comp_id_thms map_congL_thms)
+ |> Thm.close_derivation;
+
+ val mor_Abs =
+ Skip_Proof.prove lthy [] []
+ (HOLogic.mk_Trueprop (mk_mor car_finals str_finals UNIVs dtors Abs_Ts))
+ (mk_mor_Abs_tac (mor_def :: dtor_defs) Abs_inverses)
+ |> Thm.close_derivation;
+ in
+ (mor_Rep, mor_Abs)
+ end;
+
+ val timer = time (timer "dtor definitions & thms");
+
+ fun unfold_bind i = Binding.suffix_name ("_" ^ dtor_unfoldN) (nth bs (i - 1));
+ val unfold_name = Binding.name_of o unfold_bind;
+ val unfold_def_bind = rpair [] o Thm.def_binding o unfold_bind;
+
+ fun unfold_spec i T AT abs f z z' =
+ let
+ val unfoldT = Library.foldr (op -->) (sTs, AT --> T);
+
+ val lhs = Term.list_comb (Free (unfold_name i, unfoldT), ss);
+ val rhs = Term.absfree z' (abs $ (f $ z));
+ in
+ mk_Trueprop_eq (lhs, rhs)
+ end;
+
+ val ((unfold_frees, (_, unfold_def_frees)), (lthy, lthy_old)) =
+ lthy
+ |> fold_map7 (fn i => fn T => fn AT => fn abs => fn f => fn z => fn z' =>
+ Specification.definition
+ (SOME (unfold_bind i, NONE, NoSyn), (unfold_def_bind i, unfold_spec i T AT abs f z z')))
+ ks Ts activeAs Abs_Ts (map (fn i => HOLogic.mk_comp
+ (mk_proj (mk_LSBIS passive_UNIVs i), mk_beh ss i)) ks) zs zs'
+ |>> apsnd split_list o split_list
+ ||> `Local_Theory.restore;
+
+ 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_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 thm => thm RS fun_cong) o Morphism.thm phi) unfold_def_frees;
+
+ 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 [tcoalg_thm RS mor_final_thm, UNIV_I]) RS sym) morE_thms;
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all ss
+ (HOLogic.mk_Trueprop (mk_mor active_UNIVs ss UNIVs dtors (map (mk_unfold Ts ss) ks))))
+ (K (mk_mor_unfold_tac m mor_UNIV_thm dtor_defs unfold_defs Abs_inverses' morEs'
+ map_comp_id_thms map_congs))
+ |> Thm.close_derivation
+ 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 passive_UNIVs UNIVs dtors TRs);
+ val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+ (map2 (fn R => fn T => mk_subset R (Id_const T)) TRs Ts));
+ val goal = fold_rev Logic.all TRs (Logic.mk_implies (prem, concl));
+ in
+ `split_conj_thm (Skip_Proof.prove lthy [] [] goal
+ (K (mk_raw_coind_tac 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)
+ end;
+
+ val unique_mor_thms =
+ let
+ val prems = [HOLogic.mk_Trueprop (mk_coalg passive_UNIVs Bs ss), HOLogic.mk_Trueprop
+ (HOLogic.mk_conj (mk_mor Bs ss UNIVs dtors unfold_fs,
+ mk_mor Bs ss UNIVs dtors unfold_fs_copy))];
+ fun mk_fun_eq B f g z = HOLogic.mk_imp
+ (HOLogic.mk_mem (z, B), HOLogic.mk_eq (f $ z, g $ z));
+ val unique = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+ (map4 mk_fun_eq Bs unfold_fs unfold_fs_copy zs));
+
+ val unique_mor = Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (Bs @ ss @ unfold_fs @ unfold_fs_copy @ zs)
+ (Logic.list_implies (prems, unique)))
+ (K (mk_unique_mor_tac raw_coind_thms bis_image2_thm))
+ |> Thm.close_derivation;
+ in
+ map (fn thm => conjI RSN (2, thm RS mp)) (split_conj_thm unique_mor)
+ 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 bis_thm = tcoalg_thm RSN (2, tcoalg_thm RS bis_image2_thm);
+ val mor_thm = mor_comp_thm OF [tcoalg_thm RS mor_final_thm, mor_Abs_thm];
+
+ val unique_mor = Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (ss @ unfold_fs) (Logic.mk_implies (prem, unique)))
+ (K (mk_unfold_unique_mor_tac raw_coind_thms bis_thm mor_thm unfold_defs))
+ |> Thm.close_derivation;
+ 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 @{thm ssubst[of _ _ "%x. x"]} 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 = Binding.suffix_name ("_" ^ ctorN) (nth bs (i - 1));
+ val ctor_name = Binding.name_of o ctor_bind;
+ val ctor_def_bind = rpair [] o Thm.def_binding o ctor_bind;
+
+ fun ctor_spec i ctorT =
+ let
+ val lhs = Free (ctor_name i, ctorT);
+ val rhs = mk_unfold Ts map_dtors i;
+ in
+ mk_Trueprop_eq (lhs, rhs)
+ end;
+
+ val ((ctor_frees, (_, ctor_def_frees)), (lthy, lthy_old)) =
+ lthy
+ |> fold_map2 (fn i => fn ctorT =>
+ Specification.definition
+ (SOME (ctor_bind i, NONE, NoSyn), (ctor_def_bind i, ctor_spec i ctorT))) ks ctorTs
+ |>> apsnd split_list o split_list
+ ||> `Local_Theory.restore;
+
+ 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 (Morphism.thm phi) ctor_def_frees;
+
+ val ctor_o_dtor_thms = map2 (fold_thms 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 = map3 mk_goal dtors ctors FTs;
+ in
+ map5 (fn goal => fn ctor_def => fn unfold => fn map_comp_id => fn map_congL =>
+ Skip_Proof.prove lthy [] [] goal
+ (mk_dtor_o_ctor_tac ctor_def unfold map_comp_id map_congL unfold_o_dtor_thms)
+ |> Thm.close_derivation)
+ goals ctor_defs dtor_unfold_thms map_comp_id_thms map_congL_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;
+
+ fun mk_ctor_dtor_unfold_like_thm dtor_inject dtor_ctor unfold =
+ iffD1 OF [dtor_inject, trans OF [unfold, dtor_ctor RS sym]];
+
+ val ctor_dtor_unfold_thms =
+ map3 mk_ctor_dtor_unfold_like_thm dtor_inject_thms dtor_ctor_thms dtor_unfold_thms;
+
+ val timer = time (timer "ctor definitions & thms");
+
+ val corec_Inl_sum_thms =
+ let
+ val mor = mor_comp_thm OF [mor_sum_case_thm, mor_unfold_thm];
+ in
+ map2 (fn unique => fn unfold_dtor =>
+ trans OF [mor RS unique, unfold_dtor]) unfold_unique_mor_thms unfold_dtor_thms
+ end;
+
+ fun corec_bind i = Binding.suffix_name ("_" ^ dtor_corecN) (nth bs (i - 1));
+ val corec_name = Binding.name_of o corec_bind;
+ val corec_def_bind = rpair [] o Thm.def_binding o corec_bind;
+
+ fun corec_spec i T AT =
+ let
+ val corecT = Library.foldr (op -->) (corec_sTs, AT --> T);
+ val maps = map3 (fn dtor => fn sum_s => fn mapx => mk_sum_case
+ (HOLogic.mk_comp (Term.list_comb (mapx, passive_ids @ corec_Inls), dtor), sum_s))
+ dtors corec_ss corec_maps;
+
+ val lhs = Term.list_comb (Free (corec_name i, corecT), corec_ss);
+ val rhs = HOLogic.mk_comp (mk_unfold Ts maps i, Inr_const T AT);
+ in
+ mk_Trueprop_eq (lhs, rhs)
+ end;
+
+ val ((corec_frees, (_, corec_def_frees)), (lthy, lthy_old)) =
+ lthy
+ |> fold_map3 (fn i => fn T => fn AT =>
+ Specification.definition
+ (SOME (corec_bind i, NONE, NoSyn), (corec_def_bind i, corec_spec i T AT)))
+ ks Ts activeAs
+ |>> apsnd split_list o split_list
+ ||> `Local_Theory.restore;
+
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+ val corecs = map (Morphism.term phi) corec_frees;
+ val corec_names = map (fst o dest_Const) corecs;
+ fun mk_corec ss i = Term.list_comb (Const (nth corec_names (i - 1), Library.foldr (op -->)
+ (map fastype_of ss, domain_type (fastype_of (nth ss (i - 1))) --> nth Ts (i - 1))), ss);
+ val corec_defs = map (Morphism.thm phi) corec_def_frees;
+
+ val sum_cases =
+ map2 (fn T => fn i => mk_sum_case (HOLogic.id_const T, mk_corec corec_ss i)) Ts ks;
+ val dtor_corec_thms =
+ let
+ fun mk_goal i corec_s corec_map dtor z =
+ let
+ val lhs = dtor $ (mk_corec corec_ss i $ z);
+ val rhs = Term.list_comb (corec_map, passive_ids @ sum_cases) $ (corec_s $ z);
+ in
+ fold_rev Logic.all (z :: corec_ss) (mk_Trueprop_eq (lhs, rhs))
+ end;
+ val goals = map5 mk_goal ks corec_ss corec_maps_rev dtors zs;
+ in
+ map3 (fn goal => fn unfold => fn map_cong =>
+ Skip_Proof.prove lthy [] [] goal
+ (mk_corec_tac m corec_defs unfold map_cong corec_Inl_sum_thms)
+ |> Thm.close_derivation)
+ goals dtor_unfold_thms map_congs
+ end;
+
+ val ctor_dtor_corec_thms =
+ map3 mk_ctor_dtor_unfold_like_thm dtor_inject_thms dtor_ctor_thms dtor_corec_thms;
+
+ val timer = time (timer "corec definitions & thms");
+
+ val (dtor_coinduct_thm, coinduct_params, srel_coinduct_thm, rel_coinduct_thm,
+ dtor_strong_coinduct_thm, srel_strong_coinduct_thm, rel_strong_coinduct_thm) =
+ let
+ val zs = Jzs1 @ Jzs2;
+ val frees = phis @ zs;
+
+ fun mk_Ids Id = if Id then map Id_const passiveAs else map mk_diag passive_UNIVs;
+
+ fun mk_phi upto_eq phi z1 z2 = if upto_eq
+ then Term.absfree (dest_Free z1) (Term.absfree (dest_Free z2)
+ (HOLogic.mk_disj (phi $ z1 $ z2, HOLogic.mk_eq (z1, z2))))
+ else phi;
+
+ fun phi_srels upto_eq = map4 (fn phi => fn T => fn z1 => fn z2 =>
+ HOLogic.Collect_const (HOLogic.mk_prodT (T, T)) $
+ HOLogic.mk_split (mk_phi upto_eq phi z1 z2)) phis Ts Jzs1 Jzs2;
+
+ val srels = 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
+ (map3 mk_concl phis Jzs1 Jzs2));
+
+ fun mk_srel_prem upto_eq phi dtor srel Jz Jz_copy =
+ let
+ val concl = HOLogic.mk_mem (HOLogic.mk_tuple [dtor $ Jz, dtor $ Jz_copy],
+ Term.list_comb (srel, mk_Ids upto_eq @ phi_srels upto_eq));
+ in
+ HOLogic.mk_Trueprop
+ (list_all_free [Jz, Jz_copy] (HOLogic.mk_imp (phi $ Jz $ Jz_copy, concl)))
+ end;
+
+ val srel_prems = map5 (mk_srel_prem false) phis dtors srels Jzs Jzs_copy;
+ val srel_upto_prems = map5 (mk_srel_prem true) phis dtors srels Jzs Jzs_copy;
+
+ val srel_coinduct_goal = fold_rev Logic.all frees (Logic.list_implies (srel_prems, concl));
+ val coinduct_params = rev (Term.add_tfrees srel_coinduct_goal []);
+
+ val srel_coinduct = unfold_thms lthy @{thms diag_UNIV}
+ (Skip_Proof.prove lthy [] [] srel_coinduct_goal
+ (K (mk_srel_coinduct_tac ks raw_coind_thm bis_srel_thm))
+ |> Thm.close_derivation);
+
+ fun mk_dtor_prem upto_eq phi dtor map_nth sets Jz Jz_copy FJz =
+ let
+ val xs = [Jz, Jz_copy];
+
+ fun mk_map_conjunct nths x =
+ HOLogic.mk_eq (Term.list_comb (map_nth, passive_ids @ nths) $ FJz, dtor $ x);
+
+ fun mk_set_conjunct set phi z1 z2 =
+ list_all_free [z1, z2]
+ (HOLogic.mk_imp (HOLogic.mk_mem (HOLogic.mk_prod (z1, z2), set $ FJz),
+ mk_phi upto_eq phi z1 z2 $ z1 $ z2));
+
+ val concl = list_exists_free [FJz] (HOLogic.mk_conj
+ (Library.foldr1 HOLogic.mk_conj (map2 mk_map_conjunct [fstsTs, sndsTs] xs),
+ Library.foldr1 HOLogic.mk_conj
+ (map4 mk_set_conjunct (drop m sets) phis Jzs1 Jzs2)));
+ in
+ fold_rev Logic.all xs (Logic.mk_implies
+ (HOLogic.mk_Trueprop (Term.list_comb (phi, xs)), HOLogic.mk_Trueprop concl))
+ end;
+
+ fun mk_dtor_prems upto_eq =
+ map7 (mk_dtor_prem upto_eq) phis dtors map_FT_nths prodFT_setss Jzs Jzs_copy FJzs;
+
+ val dtor_prems = mk_dtor_prems false;
+ val dtor_upto_prems = mk_dtor_prems true;
+
+ val dtor_coinduct_goal = fold_rev Logic.all frees (Logic.list_implies (dtor_prems, concl));
+ val dtor_coinduct = Skip_Proof.prove lthy [] [] dtor_coinduct_goal
+ (K (mk_dtor_coinduct_tac m ks raw_coind_thm bis_def))
+ |> Thm.close_derivation;
+
+ val cTs = map (SOME o certifyT lthy o TFree) coinduct_params;
+ val cts = map3 (SOME o certify lthy ooo mk_phi true) phis Jzs1 Jzs2;
+
+ val srel_strong_coinduct = singleton (Proof_Context.export names_lthy lthy)
+ (Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all zs (Logic.list_implies (srel_upto_prems, concl)))
+ (K (mk_srel_strong_coinduct_tac m cTs cts srel_coinduct srel_monos srel_Ids)))
+ |> Thm.close_derivation;
+
+ val dtor_strong_coinduct = singleton (Proof_Context.export names_lthy lthy)
+ (Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all zs (Logic.list_implies (dtor_upto_prems, concl)))
+ (K (mk_dtor_strong_coinduct_tac ks cTs cts dtor_coinduct bis_def
+ (tcoalg_thm RS bis_diag_thm))))
+ |> Thm.close_derivation;
+
+ val rel_of_srel_thms =
+ srel_defs @ @{thms Id_def' mem_Collect_eq fst_conv snd_conv split_conv};
+
+ val rel_coinduct = unfold_thms lthy rel_of_srel_thms srel_coinduct;
+ val rel_strong_coinduct = unfold_thms lthy rel_of_srel_thms srel_strong_coinduct;
+ in
+ (dtor_coinduct, rev (Term.add_tfrees dtor_coinduct_goal []), srel_coinduct, rel_coinduct,
+ dtor_strong_coinduct, srel_strong_coinduct, rel_strong_coinduct)
+ end;
+
+ val timer = time (timer "coinduction");
+
+ (*register new codatatypes as BNFs*)
+ val lthy = if m = 0 then lthy else
+ let
+ val fTs = map2 (curry op -->) passiveAs passiveBs;
+ val gTs = map2 (curry op -->) passiveBs passiveCs;
+ val f1Ts = map2 (curry op -->) passiveAs passiveYs;
+ val f2Ts = map2 (curry op -->) passiveBs passiveYs;
+ val p1Ts = map2 (curry op -->) passiveXs passiveAs;
+ val p2Ts = map2 (curry op -->) passiveXs passiveBs;
+ val pTs = map2 (curry op -->) passiveXs passiveCs;
+ val uTs = map2 (curry op -->) Ts Ts';
+ val JRTs = map2 (curry mk_relT) passiveAs passiveBs;
+ val JphiTs = map2 mk_pred2T passiveAs passiveBs;
+ val prodTs = map2 (curry HOLogic.mk_prodT) Ts Ts';
+ val B1Ts = map HOLogic.mk_setT passiveAs;
+ val B2Ts = map HOLogic.mk_setT passiveBs;
+ val AXTs = map HOLogic.mk_setT passiveXs;
+ val XTs = mk_Ts passiveXs;
+ val YTs = mk_Ts passiveYs;
+
+ val ((((((((((((((((((((fs, fs'), fs_copy), gs), us),
+ (Jys, Jys')), (Jys_copy, Jys'_copy)), set_induct_phiss), JRs), Jphis),
+ B1s), B2s), AXs), f1s), f2s), p1s), p2s), ps), (ys, ys')), (ys_copy, ys'_copy)),
+ names_lthy) = names_lthy
+ |> 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_Freess "P" (map (fn A => map (mk_pred2T A) Ts) passiveAs)
+ ||>> mk_Frees "R" JRTs
+ ||>> mk_Frees "P" JphiTs
+ ||>> mk_Frees "B1" B1Ts
+ ||>> mk_Frees "B2" B2Ts
+ ||>> mk_Frees "A" AXTs
+ ||>> mk_Frees "f1" f1Ts
+ ||>> mk_Frees "f2" f2Ts
+ ||>> mk_Frees "p1" p1Ts
+ ||>> mk_Frees "p2" p2Ts
+ ||>> mk_Frees "p" pTs
+ ||>> mk_Frees' "y" passiveAs
+ ||>> mk_Frees' "y" passiveAs;
+
+ 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 mk_mapsBC = mk_maps passiveBs passiveCs;
+ val mk_mapsAC = mk_maps passiveAs passiveCs;
+ val mk_mapsAY = mk_maps passiveAs passiveYs;
+ val mk_mapsBY = mk_maps passiveBs passiveYs;
+ val mk_mapsXA = mk_maps passiveXs passiveAs;
+ val mk_mapsXB = mk_maps passiveXs passiveBs;
+ val mk_mapsXC = mk_maps passiveXs passiveCs;
+ val fs_maps = map (mk_map_id Ts fs Ts' dtors mk_mapsAB) ks;
+ val fs_copy_maps = map (mk_map_id Ts fs_copy Ts' dtors mk_mapsAB) ks;
+ val gs_maps = map (mk_map_id Ts' gs Ts'' dtor's mk_mapsBC) ks;
+ val fgs_maps =
+ map (mk_map_id Ts (map2 (curry HOLogic.mk_comp) gs fs) Ts'' dtors mk_mapsAC) ks;
+ val Xdtors = mk_dtors passiveXs;
+ val UNIV's = map HOLogic.mk_UNIV Ts';
+ val CUNIVs = map HOLogic.mk_UNIV passiveCs;
+ val UNIV''s = map HOLogic.mk_UNIV Ts'';
+ val fstsTsTs' = map fst_const prodTs;
+ val sndsTsTs' = map snd_const prodTs;
+ val dtor''s = mk_dtors passiveCs;
+ val f1s_maps = map (mk_map_id Ts f1s YTs dtors mk_mapsAY) ks;
+ val f2s_maps = map (mk_map_id Ts' f2s YTs dtor's mk_mapsBY) ks;
+ val pid_maps = map (mk_map_id XTs ps Ts'' Xdtors mk_mapsXC) ks;
+ val pfst_Fmaps =
+ map (mk_Fmap fst_const p1s prodTs) (mk_mapsXA prodTs (fst o HOLogic.dest_prodT));
+ val psnd_Fmaps =
+ map (mk_Fmap snd_const p2s prodTs) (mk_mapsXB prodTs (snd o HOLogic.dest_prodT));
+ val p1id_Fmaps = map (mk_Fmap HOLogic.id_const p1s prodTs) (mk_mapsXA prodTs I);
+ val p2id_Fmaps = map (mk_Fmap HOLogic.id_const p2s prodTs) (mk_mapsXB prodTs I);
+ val pid_Fmaps = map (mk_Fmap HOLogic.id_const ps prodTs) (mk_mapsXC prodTs I);
+
+ val (map_simp_thms, map_thms) =
+ let
+ fun mk_goal fs_map map dtor dtor' = fold_rev Logic.all fs
+ (mk_Trueprop_eq (HOLogic.mk_comp (dtor', fs_map),
+ HOLogic.mk_comp (Term.list_comb (map, fs @ fs_maps), dtor)));
+ val goals = map4 mk_goal fs_maps map_FTFT's dtors dtor's;
+ val cTs = map (SOME o certifyT lthy) FTs';
+ val maps =
+ map5 (fn goal => fn cT => fn unfold => fn map_comp' => fn map_cong =>
+ Skip_Proof.prove lthy [] [] goal
+ (K (mk_map_tac m n cT unfold map_comp' map_cong))
+ |> Thm.close_derivation)
+ goals cTs dtor_unfold_thms map_comp's map_congs;
+ in
+ map_split (fn thm => (thm RS @{thm pointfreeE}, thm)) maps
+ end;
+
+ val map_comp_thms =
+ let
+ val goal = fold_rev Logic.all (fs @ gs)
+ (HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+ (map3 (fn fmap => fn gmap => fn fgmap =>
+ HOLogic.mk_eq (HOLogic.mk_comp (gmap, fmap), fgmap))
+ fs_maps gs_maps fgs_maps)))
+ in
+ split_conj_thm (Skip_Proof.prove lthy [] [] goal
+ (K (mk_map_comp_tac m n map_thms map_comps map_congs dtor_unfold_unique_thm))
+ |> Thm.close_derivation)
+ end;
+
+ val map_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 = map4 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_maps));
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (us @ fs) (Logic.list_implies (prems, goal)))
+ (mk_map_unique_tac dtor_unfold_unique_thm map_comps)
+ |> Thm.close_derivation
+ end;
+
+ val timer = time (timer "map functions for the new codatatypes");
+
+ val bd = mk_ccexp sbd sbd;
+
+ val timer = time (timer "bounds for the new codatatypes");
+
+ val setss_by_bnf = map (fn i => map2 (mk_hset dtors i) ls passiveAs) ks;
+ val setss_by_bnf' = map (fn i => map2 (mk_hset dtor's i) ls passiveBs) ks;
+ val setss_by_range = transpose setss_by_bnf;
+
+ val set_simp_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 =>
+ map4 (fn Fsets =>
+ mk_simp_goal eq (nth Fsets (i - 1)) (drop m Fsets) sets)
+ FTs_setss dtors Jzs sets)
+ ls setss_by_range;
+
+ val le_goals = map
+ (fold_rev Logic.all Jzs o HOLogic.mk_Trueprop o Library.foldr1 HOLogic.mk_conj)
+ (mk_goals (uncurry mk_subset));
+ val set_le_thmss = map split_conj_thm
+ (map4 (fn goal => fn hset_minimal => fn set_hsets => fn set_hset_hsetss =>
+ Skip_Proof.prove lthy [] [] goal
+ (K (mk_set_le_tac n hset_minimal set_hsets set_hset_hsetss))
+ |> Thm.close_derivation)
+ le_goals hset_minimal_thms set_hset_thmss' set_hset_hset_thmsss');
+
+ val simp_goalss = map (map2 (fn z => fn goal =>
+ Logic.all z (HOLogic.mk_Trueprop goal)) Jzs)
+ (mk_goals HOLogic.mk_eq);
+ in
+ map4 (map4 (fn goal => fn set_le => fn set_incl_hset => fn set_hset_incl_hsets =>
+ Skip_Proof.prove lthy [] [] goal
+ (K (mk_set_simp_tac n set_le set_incl_hset set_hset_incl_hsets))
+ |> Thm.close_derivation))
+ simp_goalss set_le_thmss set_incl_hset_thmss' set_hset_incl_hset_thmsss'
+ end;
+
+ val timer = time (timer "set functions for the new codatatypes");
+
+ val colss = map2 (fn j => fn T =>
+ map (fn i => mk_hset_rec dtors nat i j T) ks) ls passiveAs;
+ val colss' = map2 (fn j => fn T =>
+ map (fn i => mk_hset_rec dtor's nat i j T) ks) ls passiveBs;
+ val Xcolss = map2 (fn j => fn T =>
+ map (fn i => mk_hset_rec Xdtors nat i j T) ks) ls passiveXs;
+
+ 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
+ (map4 (mk_col_natural f) fs_maps Jzs cols cols'));
+
+ val goals = map3 mk_goal fs colss colss';
+
+ val ctss =
+ map (fn phi => map (SOME o certify lthy) [Term.absfree nat' phi, nat]) goals;
+
+ val thms =
+ map4 (fn goal => fn cts => fn rec_0s => fn rec_Sucs =>
+ singleton (Proof_Context.export names_lthy lthy)
+ (Skip_Proof.prove lthy [] [] (HOLogic.mk_Trueprop goal)
+ (mk_col_natural_tac cts rec_0s rec_Sucs map_simp_thms set_natural'ss))
+ |> Thm.close_derivation)
+ goals ctss hset_rec_0ss' hset_rec_Sucss';
+ in
+ map (split_conj_thm o mk_specN n) thms
+ end;
+
+ val col_bd_thmss =
+ let
+ fun mk_col_bd z col = mk_ordLeq (mk_card_of (col $ z)) sbd;
+
+ fun mk_goal cols = list_all_free Jzs (Library.foldr1 HOLogic.mk_conj
+ (map2 mk_col_bd Jzs cols));
+
+ val goals = map mk_goal colss;
+
+ val ctss =
+ map (fn phi => map (SOME o certify lthy) [Term.absfree nat' phi, nat]) goals;
+
+ val thms =
+ map5 (fn j => fn goal => fn cts => fn rec_0s => fn rec_Sucs =>
+ singleton (Proof_Context.export names_lthy lthy)
+ (Skip_Proof.prove lthy [] [] (HOLogic.mk_Trueprop goal)
+ (K (mk_col_bd_tac m j cts rec_0s rec_Sucs
+ sbd_Card_order sbd_Cinfinite set_sbdss)))
+ |> Thm.close_derivation)
+ ls goals ctss hset_rec_0ss' hset_rec_Sucss';
+ in
+ map (split_conj_thm o mk_specN n) thms
+ end;
+
+ val map_cong_thms =
+ let
+ val cTs = map (SOME o certifyT 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 (map5 (mk_prem z) sets fs fs_copy ys ys')
+
+ fun mk_map_cong 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'
+ |> certify lthy;
+
+ val cphis =
+ map9 mk_cphi setss_by_bnf Jzs' Jzs fs_maps fs_copy_maps Jys' Jys Jys'_copy Jys_copy;
+
+ val coinduct = Drule.instantiate' cTs (map SOME cphis) dtor_coinduct_thm;
+
+ val goal =
+ HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+ (map4 mk_map_cong setss_by_bnf Jzs fs_maps fs_copy_maps));
+
+ val thm = singleton (Proof_Context.export names_lthy lthy)
+ (Skip_Proof.prove lthy [] [] goal
+ (K (mk_mcong_tac m (rtac coinduct) map_comp's map_simp_thms map_congs set_natural'ss
+ set_hset_thmss set_hset_hset_thmsss)))
+ |> Thm.close_derivation
+ in
+ split_conj_thm thm
+ end;
+
+ val B1_ins = map2 (mk_in B1s) setss_by_bnf Ts;
+ val B2_ins = map2 (mk_in B2s) setss_by_bnf' Ts';
+ val thePulls = map4 mk_thePull B1_ins B2_ins f1s_maps f2s_maps;
+ val thePullTs = passiveXs @ map2 (curry HOLogic.mk_prodT) Ts Ts';
+ val thePull_ins = map2 (mk_in (AXs @ thePulls)) (mk_setss thePullTs) (mk_FTs thePullTs);
+ val pickFs = map5 mk_pickWP thePull_ins pfst_Fmaps psnd_Fmaps
+ (map2 (curry (op $)) dtors Jzs) (map2 (curry (op $)) dtor's Jz's);
+ val pickF_ss = map3 (fn pickF => fn z => fn z' =>
+ HOLogic.mk_split (Term.absfree z (Term.absfree z' pickF))) pickFs Jzs' Jz's';
+ val picks = map (mk_unfold XTs pickF_ss) ks;
+
+ val wpull_prem = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+ (map8 mk_wpull AXs B1s B2s f1s f2s (replicate m NONE) p1s p2s));
+
+ val map_eq_thms = map2 (fn simp => fn diff => box_equals OF [diff RS iffD2, simp, simp])
+ map_simp_thms dtor_inject_thms;
+ val map_wpull_thms = map (fn thm => thm OF
+ (replicate m asm_rl @ replicate n @{thm wpull_thePull})) map_wpulls;
+ val pickWP_assms_tacs =
+ map3 mk_pickWP_assms_tac set_incl_hset_thmss set_incl_hin_thmss map_eq_thms;
+
+ val coalg_thePull_thm =
+ let
+ val coalg = HOLogic.mk_Trueprop
+ (mk_coalg CUNIVs thePulls (map2 (curry HOLogic.mk_comp) pid_Fmaps pickF_ss));
+ val goal = fold_rev Logic.all (AXs @ B1s @ B2s @ f1s @ f2s @ p1s @ p2s @ ps)
+ (Logic.mk_implies (wpull_prem, coalg));
+ in
+ Skip_Proof.prove lthy [] [] goal (mk_coalg_thePull_tac m coalg_def map_wpull_thms
+ set_natural'ss pickWP_assms_tacs)
+ |> Thm.close_derivation
+ end;
+
+ val (mor_thePull_fst_thm, mor_thePull_snd_thm, mor_thePull_pick_thm) =
+ let
+ val mor_fst = HOLogic.mk_Trueprop
+ (mk_mor thePulls (map2 (curry HOLogic.mk_comp) p1id_Fmaps pickF_ss)
+ UNIVs dtors fstsTsTs');
+ val mor_snd = HOLogic.mk_Trueprop
+ (mk_mor thePulls (map2 (curry HOLogic.mk_comp) p2id_Fmaps pickF_ss)
+ UNIV's dtor's sndsTsTs');
+ val mor_pick = HOLogic.mk_Trueprop
+ (mk_mor thePulls (map2 (curry HOLogic.mk_comp) pid_Fmaps pickF_ss)
+ UNIV''s dtor''s (map2 (curry HOLogic.mk_comp) pid_maps picks));
+
+ val fst_goal = fold_rev Logic.all (AXs @ B1s @ B2s @ f1s @ f2s @ p1s @ p2s)
+ (Logic.mk_implies (wpull_prem, mor_fst));
+ val snd_goal = fold_rev Logic.all (AXs @ B1s @ B2s @ f1s @ f2s @ p1s @ p2s)
+ (Logic.mk_implies (wpull_prem, mor_snd));
+ val pick_goal = fold_rev Logic.all (AXs @ B1s @ B2s @ f1s @ f2s @ p1s @ p2s @ ps)
+ (Logic.mk_implies (wpull_prem, mor_pick));
+ in
+ (Skip_Proof.prove lthy [] [] fst_goal (mk_mor_thePull_fst_tac m mor_def map_wpull_thms
+ map_comp's pickWP_assms_tacs) |> Thm.close_derivation,
+ Skip_Proof.prove lthy [] [] snd_goal (mk_mor_thePull_snd_tac m mor_def map_wpull_thms
+ map_comp's pickWP_assms_tacs) |> Thm.close_derivation,
+ Skip_Proof.prove lthy [] [] pick_goal (mk_mor_thePull_pick_tac mor_def dtor_unfold_thms
+ map_comp's) |> Thm.close_derivation)
+ end;
+
+ val pick_col_thmss =
+ let
+ fun mk_conjunct AX Jpair pick thePull col =
+ HOLogic.mk_imp (HOLogic.mk_mem (Jpair, thePull), mk_subset (col $ (pick $ Jpair)) AX);
+
+ fun mk_concl AX cols =
+ list_all_free Jpairs (Library.foldr1 HOLogic.mk_conj
+ (map4 (mk_conjunct AX) Jpairs picks thePulls cols));
+
+ val concls = map2 mk_concl AXs Xcolss;
+
+ val ctss =
+ map (fn phi => map (SOME o certify lthy) [Term.absfree nat' phi, nat]) concls;
+
+ val goals =
+ map (fn concl => Logic.mk_implies (wpull_prem, HOLogic.mk_Trueprop concl)) concls;
+
+ val thms =
+ map5 (fn j => fn goal => fn cts => fn rec_0s => fn rec_Sucs =>
+ singleton (Proof_Context.export names_lthy lthy) (Skip_Proof.prove lthy [] [] goal
+ (mk_pick_col_tac m j cts rec_0s rec_Sucs dtor_unfold_thms set_natural'ss
+ map_wpull_thms pickWP_assms_tacs))
+ |> Thm.close_derivation)
+ ls goals ctss hset_rec_0ss' hset_rec_Sucss';
+ in
+ map (map (fn thm => thm RS mp) o split_conj_thm o mk_specN n) thms
+ end;
+
+ val timer = time (timer "helpers for BNF properties");
+
+ val map_id_tacs =
+ map2 (K oo mk_map_id_tac map_thms) dtor_unfold_unique_thms unfold_dtor_thms;
+ val map_comp_tacs = map (fn thm => K (rtac (thm RS sym) 1)) map_comp_thms;
+ val map_cong_tacs = map (mk_map_cong_tac m) map_cong_thms;
+ val set_nat_tacss =
+ map2 (map2 (K oo mk_set_natural_tac)) hset_defss (transpose col_natural_thmss);
+
+ val bd_co_tacs = replicate n (K (mk_bd_card_order_tac sbd_card_order));
+ val bd_cinf_tacs = replicate n (K (mk_bd_cinfinite_tac sbd_Cinfinite));
+
+ val set_bd_tacss =
+ map2 (map2 (K oo mk_set_bd_tac sbd_Cinfinite)) hset_defss (transpose col_bd_thmss);
+
+ val in_bd_tacs = map7 (fn i => fn isNode_hsets => fn carT_def =>
+ fn card_of_carT => fn mor_image => fn Rep_inverse => fn mor_hsets =>
+ K (mk_in_bd_tac (nth isNode_hsets (i - 1)) isNode_hsets carT_def
+ card_of_carT mor_image Rep_inverse mor_hsets
+ sbd_Cnotzero sbd_Card_order mor_Rep_thm coalgT_thm mor_T_final_thm tcoalg_thm))
+ ks isNode_hset_thmss carT_defs card_of_carT_thms
+ mor_image'_thms Rep_inverses (transpose mor_hset_thmss);
+
+ val map_wpull_tacs =
+ map3 (K ooo mk_wpull_tac m coalg_thePull_thm mor_thePull_fst_thm mor_thePull_snd_thm
+ mor_thePull_pick_thm) unique_mor_thms (transpose pick_col_thmss) hset_defss;
+
+ val srel_O_Gr_tacs = replicate n (simple_srel_O_Gr_tac o #context);
+
+ val tacss = map10 zip_axioms map_id_tacs map_comp_tacs map_cong_tacs set_nat_tacss
+ bd_co_tacs bd_cinf_tacs set_bd_tacss in_bd_tacs map_wpull_tacs srel_O_Gr_tacs;
+
+ val (hset_dtor_incl_thmss, hset_hset_dtor_incl_thmsss, hset_induct_thms) =
+ let
+ fun tinst_of dtor =
+ map (SOME o certify lthy) (dtor :: remove (op =) dtor dtors);
+ fun tinst_of' dtor = case tinst_of dtor of t :: ts => t :: NONE :: ts;
+ val Tinst = map (pairself (certifyT lthy))
+ (map Logic.varifyT_global (deads @ allAs) ~~ (deads @ passiveAs @ Ts));
+ val set_incl_thmss =
+ map2 (fn dtor => map (singleton (Proof_Context.export names_lthy lthy) o
+ Drule.instantiate' [] (tinst_of' dtor) o
+ Thm.instantiate (Tinst, []) o Drule.zero_var_indexes))
+ dtors set_incl_hset_thmss;
+
+ val tinst = interleave (map (SOME o certify lthy) dtors) (replicate n NONE)
+ val set_minimal_thms =
+ map (Drule.instantiate' [] tinst o Thm.instantiate (Tinst, []) o
+ Drule.zero_var_indexes)
+ hset_minimal_thms;
+
+ val set_set_incl_thmsss =
+ map2 (fn dtor => map (map (singleton (Proof_Context.export names_lthy lthy) o
+ Drule.instantiate' [] (NONE :: tinst_of' dtor) o
+ Thm.instantiate (Tinst, []) o Drule.zero_var_indexes)))
+ dtors set_hset_incl_hset_thmsss;
+
+ val set_set_incl_thmsss' = transpose (map transpose set_set_incl_thmsss);
+
+ val incls =
+ maps (map (fn thm => thm RS @{thm subset_Collect_iff})) set_incl_thmss @
+ @{thms subset_Collect_iff[OF subset_refl]};
+
+ fun mk_induct_tinst phis jsets y y' =
+ map4 (fn phi => fn jset => fn Jz => fn Jz' =>
+ SOME (certify 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';
+ val set_induct_thms =
+ map6 (fn set_minimal => fn set_set_inclss => fn jsets => fn y => fn y' => fn phis =>
+ ((set_minimal
+ |> Drule.instantiate' [] (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 asm_rl)))
+ set_minimal_thms set_set_incl_thmsss' setss_by_range ys ys' set_induct_phiss
+ in
+ (set_incl_thmss, set_set_incl_thmsss, set_induct_thms)
+ end;
+
+ 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;
+
+ 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
+ fold_rev Logic.all (map (nth ys) I @ Jzs) (HOLogic.mk_Trueprop
+ (Library.foldr1 HOLogic.mk_conj (map4 mk_conjunct sets Jzs dummys wits)))
+ end;
+ val goals = map5 mk_goal setss_by_range ys ys_copy ys'_copy ls;
+ in
+ map2 (fn goal => fn induct =>
+ Skip_Proof.prove lthy [] [] goal
+ (mk_coind_wit_tac induct dtor_unfold_thms (flat set_natural'ss) wit_thms)
+ |> Thm.close_derivation)
+ goals hset_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 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;
+
+ val 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));
+
+ val wit_tac = mk_wit_tac n dtor_ctor_thms (flat set_simp_thmss) (maps wit_thms_of_bnf bnfs);
+
+ val policy = user_policy Derive_All_Facts_Note_Most;
+
+ val (Jbnfs, lthy) =
+ fold_map6 (fn tacs => fn b => fn mapx => fn sets => fn T => fn (thms, wits) => fn lthy =>
+ bnf_def Dont_Inline policy I tacs (wit_tac thms) (SOME deads)
+ (((((b, fold_rev Term.absfree fs' mapx), sets), absdummy T bd), wits), NONE) lthy
+ |> register_bnf (Local_Theory.full_name lthy b))
+ tacss bs fs_maps setss_by_bnf Ts all_witss lthy;
+
+ val fold_maps = fold_thms lthy (map (fn bnf =>
+ mk_unabs_def m (map_def_of_bnf bnf RS @{thm meta_eq_to_obj_eq})) Jbnfs);
+
+ val fold_sets = fold_thms lthy (maps (fn bnf =>
+ map (fn thm => thm RS @{thm meta_eq_to_obj_eq}) (set_defs_of_bnf bnf)) Jbnfs);
+
+ val timer = time (timer "registered new codatatypes as BNFs");
+
+ val set_incl_thmss = map (map fold_sets) hset_dtor_incl_thmss;
+ val set_set_incl_thmsss = map (map (map fold_sets)) hset_hset_dtor_incl_thmsss;
+ val set_induct_thms = map fold_sets hset_induct_thms;
+
+ val srels = map2 (fn Ds => mk_srel_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs;
+ val Jsrels = map (mk_srel_of_bnf deads passiveAs passiveBs) Jbnfs;
+ val rels = map2 (fn Ds => mk_rel_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs;
+ val Jrels = map (mk_rel_of_bnf deads passiveAs passiveBs) Jbnfs;
+
+ val JrelRs = map (fn Jsrel => Term.list_comb (Jsrel, JRs)) Jsrels;
+ val relRs = map (fn srel => Term.list_comb (srel, JRs @ JrelRs)) srels;
+ val Jpredphis = map (fn Jsrel => Term.list_comb (Jsrel, Jphis)) Jrels;
+ val predphis = map (fn srel => Term.list_comb (srel, Jphis @ Jpredphis)) rels;
+
+ val in_srels = map in_srel_of_bnf bnfs;
+ val in_Jsrels = map in_srel_of_bnf Jbnfs;
+ val Jsrel_defs = map srel_def_of_bnf Jbnfs;
+ val Jrel_defs = map rel_def_of_bnf Jbnfs;
+
+ val folded_map_simp_thms = map fold_maps map_simp_thms;
+ val folded_set_simp_thmss = map (map fold_sets) set_simp_thmss;
+ val folded_set_simp_thmss' = transpose folded_set_simp_thmss;
+
+ val Jsrel_simp_thms =
+ let
+ fun mk_goal Jz Jz' dtor dtor' JrelR relR = fold_rev Logic.all (Jz :: Jz' :: JRs)
+ (mk_Trueprop_eq (HOLogic.mk_mem (HOLogic.mk_prod (Jz, Jz'), JrelR),
+ HOLogic.mk_mem (HOLogic.mk_prod (dtor $ Jz, dtor' $ Jz'), relR)));
+ val goals = map6 mk_goal Jzs Jz's dtors dtor's JrelRs relRs;
+ in
+ map12 (fn i => fn goal => fn in_srel => fn map_comp => fn map_cong =>
+ fn map_simp => fn set_simps => fn dtor_inject => fn dtor_ctor =>
+ fn set_naturals => fn set_incls => fn set_set_inclss =>
+ Skip_Proof.prove lthy [] [] goal
+ (K (mk_srel_simp_tac in_Jsrels i in_srel map_comp map_cong map_simp set_simps
+ dtor_inject dtor_ctor set_naturals set_incls set_set_inclss))
+ |> Thm.close_derivation)
+ ks goals in_srels map_comp's map_congs folded_map_simp_thms folded_set_simp_thmss'
+ dtor_inject_thms dtor_ctor_thms set_natural'ss set_incl_thmss set_set_incl_thmsss
+ end;
+
+ val Jrel_simp_thms =
+ let
+ fun mk_goal Jz Jz' dtor dtor' Jpredphi predphi = fold_rev Logic.all (Jz :: Jz' :: Jphis)
+ (mk_Trueprop_eq (Jpredphi $ Jz $ Jz', predphi $ (dtor $ Jz) $ (dtor' $ Jz')));
+ val goals = map6 mk_goal Jzs Jz's dtors dtor's Jpredphis predphis;
+ in
+ map3 (fn goal => fn srel_def => fn Jsrel_simp =>
+ Skip_Proof.prove lthy [] [] goal
+ (mk_rel_simp_tac srel_def Jrel_defs Jsrel_defs Jsrel_simp)
+ |> Thm.close_derivation)
+ goals srel_defs Jsrel_simp_thms
+ end;
+
+ val timer = time (timer "additional properties");
+
+ val ls' = if m = 1 then [0] else ls;
+
+ val Jbnf_common_notes =
+ [(map_uniqueN, [fold_maps map_unique_thm])] @
+ map2 (fn i => fn thm => (mk_set_inductN i, [thm])) ls' set_induct_thms
+ |> map (fn (thmN, thms) =>
+ ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]));
+
+ val Jbnf_notes =
+ [(map_simpsN, map single folded_map_simp_thms),
+ (rel_simpN, map single Jrel_simp_thms),
+ (set_inclN, set_incl_thmss),
+ (set_set_inclN, map flat set_set_incl_thmsss),
+ (srel_simpN, map single Jsrel_simp_thms)] @
+ map2 (fn i => fn thms => (mk_set_simpsN i, map single thms)) ls' folded_set_simp_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; lthy |> Local_Theory.notes (Jbnf_common_notes @ Jbnf_notes) |> snd
+ end;
+
+ val common_notes =
+ [(dtor_coinductN, [dtor_coinduct_thm]),
+ (dtor_strong_coinductN, [dtor_strong_coinduct_thm]),
+ (rel_coinductN, [rel_coinduct_thm]),
+ (rel_strong_coinductN, [rel_strong_coinduct_thm]),
+ (srel_coinductN, [srel_coinduct_thm]),
+ (srel_strong_coinductN, [srel_strong_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_dtor_unfoldsN, ctor_dtor_unfold_thms),
+ (ctor_dtor_corecsN, ctor_dtor_corec_thms),
+ (ctor_exhaustN, ctor_exhaust_thms),
+ (ctor_injectN, ctor_inject_thms),
+ (dtor_corecsN, dtor_corec_thms),
+ (dtor_ctorN, dtor_ctor_thms),
+ (dtor_exhaustN, dtor_exhaust_thms),
+ (dtor_injectN, dtor_inject_thms),
+ (dtor_unfold_uniqueN, dtor_unfold_unique_thms),
+ (dtor_unfoldsN, dtor_unfold_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)
+ in
+ ((dtors, ctors, unfolds, corecs, dtor_coinduct_thm, dtor_ctor_thms, ctor_dtor_thms,
+ ctor_inject_thms, ctor_dtor_unfold_thms, ctor_dtor_corec_thms),
+ lthy |> Local_Theory.notes (common_notes @ notes) |> snd)
+ end;
+
+val _ =
+ Outer_Syntax.local_theory @{command_spec "codata_raw"}
+ "define BNF-based coinductive datatypes (low-level)"
+ (Parse.and_list1
+ ((Parse.binding --| @{keyword ":"}) -- (Parse.typ --| @{keyword "="} -- Parse.typ)) >>
+ (snd oo fp_bnf_cmd bnf_gfp o apsnd split_list o split_list));
+
+val _ =
+ Outer_Syntax.local_theory @{command_spec "codata"} "define BNF-based coinductive datatypes"
+ (parse_datatype_cmd false bnf_gfp);
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Tools/bnf_gfp_tactics.ML Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,1554 @@
+(* Title: HOL/BNF/Tools/bnf_gfp_tactics.ML
+ Author: Dmitriy Traytel, TU Muenchen
+ Author: Andrei Popescu, TU Muenchen
+ Author: Jasmin Blanchette, TU Muenchen
+ Copyright 2012
+
+Tactics for the codatatype construction.
+*)
+
+signature BNF_GFP_TACTICS =
+sig
+ val mk_Lev_sbd_tac: cterm option list -> thm list -> thm list -> thm list list -> tactic
+ val mk_bd_card_order_tac: thm -> tactic
+ val mk_bd_cinfinite_tac: thm -> tactic
+ val mk_bis_Gr_tac: thm -> thm list -> thm list -> thm list -> thm list ->
+ {prems: 'a, context: Proof.context} -> tactic
+ val mk_bis_O_tac: int -> thm -> thm list -> thm list -> tactic
+ val mk_bis_Union_tac: thm -> thm list -> {prems: 'a, context: Proof.context} -> tactic
+ val mk_bis_converse_tac: int -> thm -> thm list -> thm list -> tactic
+ val mk_bis_srel_tac: int -> thm -> thm list -> thm list -> thm list -> thm list list -> tactic
+ val mk_carT_set_tac: int -> int -> thm -> thm -> thm -> thm ->
+ {prems: 'a, context: Proof.context} -> tactic
+ val mk_card_of_carT_tac: int -> thm list -> thm -> thm -> thm -> thm -> thm -> thm list -> tactic
+ val mk_coalgT_tac: int -> thm list -> thm list -> thm list list ->
+ {prems: 'a, context: Proof.context} -> tactic
+ val mk_coalg_final_tac: int -> thm -> thm list -> thm list -> thm list list -> thm list list ->
+ tactic
+ val mk_coalg_set_tac: thm -> tactic
+ val mk_coalg_thePull_tac: int -> thm -> thm list -> thm list list -> (int -> tactic) list ->
+ {prems: 'a, context: Proof.context} -> tactic
+ val mk_coind_wit_tac: thm -> thm list -> thm list -> thm list ->
+ {prems: 'a, context: Proof.context} -> tactic
+ val mk_col_bd_tac: int -> int -> cterm option list -> thm list -> thm list -> thm -> thm ->
+ thm list list -> tactic
+ val mk_col_natural_tac: cterm option list -> thm list -> thm list -> thm list -> thm list list ->
+ {prems: 'a, context: Proof.context} -> tactic
+ val mk_congruent_str_final_tac: int -> thm -> thm -> thm -> thm list -> tactic
+ val mk_corec_tac: int -> thm list -> thm -> thm -> thm list ->
+ {prems: 'a, context: Proof.context} -> tactic
+ val mk_dtor_coinduct_tac: int -> int list -> thm -> thm -> tactic
+ val mk_dtor_strong_coinduct_tac: int list -> ctyp option list -> cterm option list -> thm ->
+ thm -> thm -> tactic
+ val mk_dtor_o_ctor_tac: thm -> thm -> thm -> thm -> thm list ->
+ {prems: 'a, context: Proof.context} -> tactic
+ val mk_equiv_lsbis_tac: thm -> thm -> thm -> thm -> thm -> thm -> tactic
+ val mk_hset_minimal_tac: int -> thm list -> thm -> {prems: 'a, context: Proof.context} -> tactic
+ val mk_hset_rec_minimal_tac: int -> cterm option list -> thm list -> thm list ->
+ {prems: 'a, context: Proof.context} -> tactic
+ val mk_in_bd_tac: thm -> thm list -> thm -> thm -> thm -> thm -> thm list -> thm -> thm -> thm ->
+ thm -> thm -> thm -> tactic
+ val mk_incl_lsbis_tac: int -> int -> thm -> tactic
+ val mk_isNode_hset_tac: int -> thm -> thm list -> {prems: 'a, context: Proof.context} -> tactic
+ val mk_length_Lev'_tac: thm -> tactic
+ val mk_length_Lev_tac: cterm option list -> thm list -> thm list -> tactic
+ val mk_map_comp_tac: int -> int -> thm list -> thm list -> thm list -> thm -> tactic
+ val mk_mcong_tac: int -> (int -> tactic) -> thm list -> thm list -> thm list -> thm list list ->
+ thm list list -> thm list list list -> tactic
+ val mk_map_id_tac: thm list -> thm -> thm -> tactic
+ val mk_map_tac: int -> int -> ctyp option -> thm -> thm -> thm -> tactic
+ val mk_map_unique_tac: thm -> thm list -> {prems: 'a, context: Proof.context} -> tactic
+ val mk_mor_Abs_tac: thm list -> thm list -> {prems: 'a, context: Proof.context} -> tactic
+ val mk_mor_Rep_tac: int -> thm list -> thm list -> thm list -> thm list list -> thm list ->
+ thm list -> {prems: 'a, context: Proof.context} -> tactic
+ val mk_mor_T_final_tac: thm -> thm list -> thm list -> tactic
+ val mk_mor_UNIV_tac: thm list -> thm -> tactic
+ val mk_mor_beh_tac: int -> thm -> thm -> thm list -> thm list -> thm list -> thm list ->
+ thm list list -> thm list list -> thm list -> thm list -> thm list -> thm list -> thm list ->
+ thm list -> thm list -> thm list -> thm list list -> thm list list list -> thm list list list ->
+ thm list list list -> thm list list -> thm list list -> thm list -> thm list -> thm list ->
+ {prems: 'a, context: Proof.context} -> tactic
+ val mk_mor_comp_tac: thm -> thm list -> thm list -> thm list -> tactic
+ val mk_mor_elim_tac: thm -> tactic
+ val mk_mor_hset_rec_tac: int -> int -> cterm option list -> int -> thm list -> thm list ->
+ thm list -> thm list list -> thm list list -> tactic
+ val mk_mor_hset_tac: thm -> thm -> tactic
+ val mk_mor_incl_tac: thm -> thm list -> tactic
+ val mk_mor_str_tac: 'a list -> thm -> tactic
+ val mk_mor_sum_case_tac: 'a list -> thm -> tactic
+ val mk_mor_thePull_fst_tac: int -> thm -> thm list -> thm list -> (int -> tactic) list ->
+ {prems: thm list, context: Proof.context} -> tactic
+ val mk_mor_thePull_snd_tac: int -> thm -> thm list -> thm list -> (int -> tactic) list ->
+ {prems: thm list, context: Proof.context} -> tactic
+ val mk_mor_thePull_pick_tac: thm -> thm list -> thm list ->
+ {prems: 'a, context: Proof.context} -> tactic
+ val mk_mor_unfold_tac: int -> thm -> thm list -> thm list -> thm list -> thm list -> thm list ->
+ thm list -> tactic
+ val mk_prefCl_Lev_tac: cterm option list -> thm list -> thm list -> tactic
+ val mk_pickWP_assms_tac: thm list -> thm list -> thm -> (int -> tactic)
+ val mk_pick_col_tac: int -> int -> cterm option list -> thm list -> thm list -> thm list ->
+ thm list list -> thm list -> (int -> tactic) list -> {prems: 'a, context: Proof.context} ->
+ tactic
+ val mk_raw_coind_tac: thm -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> thm list ->
+ thm list -> thm list -> thm -> thm list -> tactic
+ val mk_rv_last_tac: ctyp option list -> cterm option list -> thm list -> thm list -> tactic
+ val mk_sbis_lsbis_tac: thm list -> thm -> thm -> tactic
+ val mk_set_Lev_tac: cterm option list -> thm list -> thm list -> thm list -> thm list ->
+ thm list list -> tactic
+ val mk_set_bd_tac: thm -> thm -> thm -> tactic
+ val mk_set_hset_incl_hset_tac: int -> thm list -> thm -> int -> tactic
+ val mk_set_image_Lev_tac: cterm option list -> thm list -> thm list -> thm list -> thm list ->
+ thm list list -> thm list list -> tactic
+ val mk_set_incl_hin_tac: thm list -> tactic
+ val mk_set_incl_hset_tac: thm -> thm -> tactic
+ val mk_set_le_tac: int -> thm -> thm list -> thm list list -> tactic
+ val mk_set_natural_tac: thm -> thm -> tactic
+ val mk_set_rv_Lev_tac: int -> cterm option list -> thm list -> thm list -> thm list -> thm list ->
+ thm list list -> thm list list -> tactic
+ val mk_set_simp_tac: int -> thm -> thm -> thm list -> tactic
+ val mk_srel_coinduct_tac: 'a list -> thm -> thm -> tactic
+ val mk_srel_strong_coinduct_tac: int -> ctyp option list -> cterm option list -> thm ->
+ thm list -> thm list -> tactic
+ val mk_srel_simp_tac: thm list -> int -> thm -> thm -> thm -> thm -> thm list -> thm -> thm ->
+ thm list -> thm list -> thm list list -> tactic
+ val mk_strT_hset_tac: int -> int -> int -> ctyp option list -> ctyp option list ->
+ cterm option list -> thm list -> thm list -> thm list -> thm list -> thm list list ->
+ thm list list -> thm list list -> thm -> thm list list -> tactic
+ val mk_unfold_unique_mor_tac: thm list -> thm -> thm -> thm list -> tactic
+ val mk_unique_mor_tac: thm list -> thm -> tactic
+ val mk_wit_tac: int -> thm list -> thm list -> thm list -> thm list ->
+ {prems: 'a, context: Proof.context} -> tactic
+ val mk_wpull_tac: int -> thm -> thm -> thm -> thm -> thm -> thm list -> thm list -> tactic
+end;
+
+structure BNF_GFP_Tactics : BNF_GFP_TACTICS =
+struct
+
+open BNF_Tactics
+open BNF_Util
+open BNF_FP
+open BNF_GFP_Util
+
+val fst_convol_fun_cong_sym = @{thm fst_convol} RS fun_cong RS sym;
+val list_inject_iffD1 = @{thm list.inject[THEN iffD1]};
+val nat_induct = @{thm nat_induct};
+val o_apply_trans_sym = o_apply RS trans RS sym;
+val ord_eq_le_trans = @{thm ord_eq_le_trans};
+val ord_eq_le_trans_trans_fun_cong_image_id_id_apply =
+ @{thm ord_eq_le_trans[OF trans[OF fun_cong[OF image_id] id_apply]]};
+val ordIso_ordLeq_trans = @{thm ordIso_ordLeq_trans};
+val snd_convol_fun_cong_sym = @{thm snd_convol} RS fun_cong RS sym;
+val sum_case_weak_cong = @{thm sum_case_weak_cong};
+val trans_fun_cong_image_id_id_apply = @{thm trans[OF fun_cong[OF image_id] id_apply]};
+
+fun mk_coalg_set_tac coalg_def =
+ dtac (coalg_def RS iffD1) 1 THEN
+ REPEAT_DETERM (etac conjE 1) THEN
+ EVERY' [dtac @{thm rev_bspec}, atac] 1 THEN
+ REPEAT_DETERM (eresolve_tac [CollectE, conjE] 1) THEN atac 1;
+
+fun mk_mor_elim_tac mor_def =
+ (dtac (subst OF [mor_def]) THEN'
+ REPEAT o etac conjE THEN'
+ TRY o rtac @{thm image_subsetI} THEN'
+ etac bspec THEN'
+ atac) 1;
+
+fun mk_mor_incl_tac mor_def map_id's =
+ (stac mor_def THEN'
+ rtac conjI THEN'
+ CONJ_WRAP' (K (EVERY' [rtac ballI, etac set_mp, stac @{thm id_apply}, atac]))
+ map_id's THEN'
+ CONJ_WRAP' (fn thm =>
+ (EVERY' [rtac ballI, rtac (thm RS trans), rtac sym, rtac (@{thm id_apply} RS arg_cong)]))
+ map_id's) 1;
+
+fun mk_mor_comp_tac mor_def mor_images morEs map_comp_ids =
+ let
+ fun fbetw_tac image = EVERY' [rtac ballI, stac o_apply, etac image, etac image, atac];
+ fun mor_tac ((mor_image, morE), map_comp_id) =
+ EVERY' [rtac ballI, stac o_apply, rtac trans, rtac (map_comp_id RS sym), rtac trans,
+ etac (morE RS arg_cong), atac, etac morE, etac mor_image, atac];
+ in
+ (stac mor_def THEN' rtac conjI THEN'
+ CONJ_WRAP' fbetw_tac mor_images THEN'
+ CONJ_WRAP' mor_tac ((mor_images ~~ morEs) ~~ map_comp_ids)) 1
+ end;
+
+fun mk_mor_UNIV_tac morEs mor_def =
+ let
+ val n = length morEs;
+ fun mor_tac morE = EVERY' [rtac ext, rtac trans, rtac o_apply, rtac trans, etac morE,
+ rtac UNIV_I, rtac sym, rtac o_apply];
+ in
+ EVERY' [rtac iffI, CONJ_WRAP' mor_tac morEs,
+ stac mor_def, rtac conjI, CONJ_WRAP' (K (rtac ballI THEN' rtac UNIV_I)) morEs,
+ CONJ_WRAP' (fn i =>
+ EVERY' [dtac (mk_conjunctN n i), rtac ballI, etac @{thm pointfreeE}]) (1 upto n)] 1
+ end;
+
+fun mk_mor_str_tac ks mor_UNIV =
+ (stac mor_UNIV THEN' CONJ_WRAP' (K (rtac refl)) ks) 1;
+
+fun mk_mor_sum_case_tac ks mor_UNIV =
+ (stac mor_UNIV THEN' CONJ_WRAP' (K (rtac @{thm sum_case_comp_Inl[symmetric]})) ks) 1;
+
+fun mk_set_incl_hset_tac def rec_Suc =
+ EVERY' (stac def ::
+ map rtac [@{thm incl_UNION_I}, UNIV_I, @{thm ord_le_eq_trans}, @{thm Un_upper1},
+ sym, rec_Suc]) 1;
+
+fun mk_set_hset_incl_hset_tac n defs rec_Suc i =
+ EVERY' (map (TRY oo stac) defs @
+ map rtac [@{thm UN_least}, subsetI, @{thm UN_I}, UNIV_I, set_mp, equalityD2, rec_Suc, UnI2,
+ mk_UnIN n i] @
+ [etac @{thm UN_I}, atac]) 1;
+
+fun mk_set_incl_hin_tac incls =
+ if null incls then rtac subset_UNIV 1
+ else EVERY' [rtac subsetI, rtac CollectI,
+ CONJ_WRAP' (fn incl => EVERY' [rtac subset_trans, etac incl, atac]) incls] 1;
+
+fun mk_hset_rec_minimal_tac m cts rec_0s rec_Sucs {context = ctxt, prems = _} =
+ EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn thm => EVERY'
+ [rtac ord_eq_le_trans, rtac thm, rtac @{thm empty_subsetI}]) rec_0s,
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn rec_Suc => EVERY'
+ [rtac ord_eq_le_trans, rtac rec_Suc,
+ if m = 0 then K all_tac
+ else (rtac @{thm Un_least} THEN' Goal.assume_rule_tac ctxt),
+ CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_least}))
+ (K (EVERY' [rtac @{thm UN_least}, REPEAT_DETERM o eresolve_tac [allE, conjE],
+ rtac subset_trans, atac, Goal.assume_rule_tac ctxt])) rec_0s])
+ rec_Sucs] 1;
+
+fun mk_hset_minimal_tac n hset_defs hset_rec_minimal {context = ctxt, prems = _} =
+ (CONJ_WRAP' (fn def => (EVERY' [rtac ord_eq_le_trans, rtac def,
+ rtac @{thm UN_least}, rtac rev_mp, rtac hset_rec_minimal,
+ EVERY' (replicate ((n + 1) * n) (Goal.assume_rule_tac ctxt)), rtac impI,
+ REPEAT_DETERM o eresolve_tac [allE, conjE], atac])) hset_defs) 1
+
+fun mk_mor_hset_rec_tac m n cts j rec_0s rec_Sucs morEs set_naturalss coalg_setss =
+ EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn thm => EVERY' (map rtac [impI, thm RS trans, thm RS sym])) rec_0s,
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP'
+ (fn (rec_Suc, (morE, ((passive_set_naturals, active_set_naturals), coalg_sets))) =>
+ EVERY' [rtac impI, rtac (rec_Suc RS trans), rtac (rec_Suc RS trans RS sym),
+ if m = 0 then K all_tac
+ else EVERY' [rtac @{thm Un_cong}, rtac box_equals,
+ rtac (nth passive_set_naturals (j - 1) RS sym),
+ rtac trans_fun_cong_image_id_id_apply, etac (morE RS arg_cong), atac],
+ CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_cong}))
+ (fn (i, (set_natural, coalg_set)) =>
+ EVERY' [rtac sym, rtac trans, rtac (refl RSN (2, @{thm UN_cong})),
+ etac (morE RS sym RS arg_cong RS trans), atac, rtac set_natural,
+ rtac (@{thm UN_simps(10)} RS trans), rtac (refl RS @{thm UN_cong}),
+ ftac coalg_set, atac, dtac set_mp, atac, rtac mp, rtac (mk_conjunctN n i),
+ REPEAT_DETERM o etac allE, atac, atac])
+ (rev ((1 upto n) ~~ (active_set_naturals ~~ coalg_sets)))])
+ (rec_Sucs ~~ (morEs ~~ (map (chop m) set_naturalss ~~ map (drop m) coalg_setss)))] 1;
+
+fun mk_mor_hset_tac hset_def mor_hset_rec =
+ EVERY' [rtac (hset_def RS trans), rtac (refl RS @{thm UN_cong} RS trans), etac mor_hset_rec,
+ atac, atac, rtac (hset_def RS sym)] 1
+
+fun mk_bis_srel_tac m bis_def srel_O_Grs map_comps map_congs set_naturalss =
+ let
+ val n = length srel_O_Grs;
+ val thms = ((1 upto n) ~~ map_comps ~~ map_congs ~~ set_naturalss ~~ srel_O_Grs);
+
+ fun mk_if_tac ((((i, map_comp), map_cong), set_naturals), srel_O_Gr) =
+ EVERY' [rtac allI, rtac allI, rtac impI, dtac (mk_conjunctN n i),
+ etac allE, etac allE, etac impE, atac, etac bexE, etac conjE,
+ rtac (srel_O_Gr RS equalityD2 RS set_mp),
+ rtac @{thm relcompI}, rtac @{thm converseI},
+ EVERY' (map (fn thm =>
+ EVERY' [rtac @{thm GrI}, REPEAT_DETERM o eresolve_tac [CollectE, conjE],
+ rtac CollectI,
+ CONJ_WRAP' (fn (i, thm) =>
+ if i <= m
+ then EVERY' [rtac ord_eq_le_trans, rtac thm, rtac subset_trans,
+ etac @{thm image_mono}, rtac @{thm image_subsetI}, etac @{thm diagI}]
+ else EVERY' [rtac ord_eq_le_trans, rtac trans, rtac thm,
+ rtac trans_fun_cong_image_id_id_apply, atac])
+ (1 upto (m + n) ~~ set_naturals),
+ rtac trans, rtac trans, rtac map_comp, rtac map_cong, REPEAT_DETERM_N m o rtac thm,
+ REPEAT_DETERM_N n o rtac (@{thm o_id} RS fun_cong), atac])
+ @{thms fst_diag_id snd_diag_id})];
+
+ fun mk_only_if_tac ((((i, map_comp), map_cong), set_naturals), srel_O_Gr) =
+ EVERY' [dtac (mk_conjunctN n i), rtac allI, rtac allI, rtac impI,
+ etac allE, etac allE, etac impE, atac,
+ dtac (srel_O_Gr RS equalityD1 RS set_mp),
+ REPEAT_DETERM o eresolve_tac [CollectE, @{thm relcompE}, @{thm converseE}],
+ REPEAT_DETERM o eresolve_tac [@{thm GrE}, exE, conjE],
+ REPEAT_DETERM o dtac Pair_eqD,
+ REPEAT_DETERM o etac conjE,
+ hyp_subst_tac,
+ REPEAT_DETERM o eresolve_tac [CollectE, conjE],
+ rtac bexI, rtac conjI, rtac trans, rtac map_comp,
+ REPEAT_DETERM_N m o stac @{thm id_o},
+ REPEAT_DETERM_N n o stac @{thm o_id},
+ etac sym, rtac trans, rtac map_comp,
+ REPEAT_DETERM_N m o stac @{thm id_o},
+ REPEAT_DETERM_N n o stac @{thm o_id},
+ rtac trans, rtac map_cong,
+ REPEAT_DETERM_N m o EVERY' [rtac @{thm diagE'}, etac set_mp, atac],
+ REPEAT_DETERM_N n o rtac refl,
+ etac sym, rtac CollectI,
+ CONJ_WRAP' (fn (i, thm) =>
+ if i <= m
+ then EVERY' [rtac ord_eq_le_trans, rtac thm, rtac @{thm image_subsetI},
+ rtac @{thm diag_fst}, etac set_mp, atac]
+ else EVERY' [rtac ord_eq_le_trans, rtac trans, rtac thm,
+ rtac trans_fun_cong_image_id_id_apply, atac])
+ (1 upto (m + n) ~~ set_naturals)];
+ in
+ EVERY' [rtac (bis_def RS trans),
+ rtac iffI, etac conjE, etac conjI, CONJ_WRAP' mk_if_tac thms,
+ etac conjE, etac conjI, CONJ_WRAP' mk_only_if_tac thms] 1
+ end;
+
+fun mk_bis_converse_tac m bis_srel srel_congs srel_converses =
+ EVERY' [stac bis_srel, dtac (bis_srel RS iffD1),
+ REPEAT_DETERM o etac conjE, rtac conjI,
+ CONJ_WRAP' (K (EVERY' [rtac @{thm converse_shift}, etac subset_trans,
+ rtac equalityD2, rtac @{thm converse_Times}])) srel_congs,
+ CONJ_WRAP' (fn (srel_cong, srel_converse) =>
+ EVERY' [rtac allI, rtac allI, rtac impI, rtac @{thm set_mp[OF equalityD2]},
+ rtac (srel_cong RS trans),
+ REPEAT_DETERM_N m o rtac @{thm diag_converse},
+ REPEAT_DETERM_N (length srel_congs) o rtac refl,
+ rtac srel_converse,
+ REPEAT_DETERM o etac allE,
+ rtac @{thm converseI}, etac mp, etac @{thm converseD}]) (srel_congs ~~ srel_converses)] 1;
+
+fun mk_bis_O_tac m bis_srel srel_congs srel_Os =
+ EVERY' [stac bis_srel, REPEAT_DETERM o dtac (bis_srel RS iffD1),
+ REPEAT_DETERM o etac conjE, rtac conjI,
+ CONJ_WRAP' (K (EVERY' [etac @{thm relcomp_subset_Sigma}, atac])) srel_congs,
+ CONJ_WRAP' (fn (srel_cong, srel_O) =>
+ EVERY' [rtac allI, rtac allI, rtac impI, rtac @{thm set_mp[OF equalityD2]},
+ rtac (srel_cong RS trans),
+ REPEAT_DETERM_N m o rtac @{thm diag_Comp},
+ REPEAT_DETERM_N (length srel_congs) o rtac refl,
+ rtac srel_O,
+ etac @{thm relcompE},
+ REPEAT_DETERM o dtac Pair_eqD,
+ etac conjE, hyp_subst_tac,
+ REPEAT_DETERM o etac allE, rtac @{thm relcompI},
+ etac mp, atac, etac mp, atac]) (srel_congs ~~ srel_Os)] 1;
+
+fun mk_bis_Gr_tac bis_srel srel_Grs mor_images morEs coalg_ins
+ {context = ctxt, prems = _} =
+ unfold_thms_tac ctxt (bis_srel :: @{thm diag_Gr} :: srel_Grs) THEN
+ EVERY' [rtac conjI,
+ CONJ_WRAP' (fn thm => rtac (@{thm Gr_incl} RS ssubst) THEN' etac thm) mor_images,
+ CONJ_WRAP' (fn (coalg_in, morE) =>
+ EVERY' [rtac allI, rtac allI, rtac impI, rtac @{thm GrI}, etac coalg_in,
+ etac @{thm GrD1}, etac (morE RS trans), etac @{thm GrD1},
+ etac (@{thm GrD2} RS arg_cong)]) (coalg_ins ~~ morEs)] 1;
+
+fun mk_bis_Union_tac bis_def in_monos {context = ctxt, prems = _} =
+ let
+ val n = length in_monos;
+ val ks = 1 upto n;
+ in
+ unfold_thms_tac ctxt [bis_def] THEN
+ EVERY' [rtac conjI,
+ CONJ_WRAP' (fn i =>
+ EVERY' [rtac @{thm UN_least}, dtac bspec, atac,
+ dtac conjunct1, etac (mk_conjunctN n i)]) ks,
+ CONJ_WRAP' (fn (i, in_mono) =>
+ EVERY' [rtac allI, rtac allI, rtac impI, etac @{thm UN_E}, dtac bspec, atac,
+ dtac conjunct2, dtac (mk_conjunctN n i), etac allE, etac allE, dtac mp,
+ atac, etac bexE, rtac bexI, atac, rtac in_mono,
+ REPEAT_DETERM_N n o etac @{thm incl_UNION_I[OF _ subset_refl]},
+ atac]) (ks ~~ in_monos)] 1
+ end;
+
+fun mk_sbis_lsbis_tac lsbis_defs bis_Union bis_cong =
+ let
+ val n = length lsbis_defs;
+ in
+ EVERY' [rtac (Thm.permute_prems 0 1 bis_cong), EVERY' (map rtac lsbis_defs),
+ rtac bis_Union, rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, conjE, exE],
+ hyp_subst_tac, etac bis_cong, EVERY' (map (rtac o mk_nth_conv n) (1 upto n))] 1
+ end;
+
+fun mk_incl_lsbis_tac n i lsbis_def =
+ EVERY' [rtac @{thm xt1(3)}, rtac lsbis_def, rtac @{thm incl_UNION_I}, rtac CollectI,
+ REPEAT_DETERM_N n o rtac exI, rtac conjI, rtac refl, atac, rtac equalityD2,
+ rtac (mk_nth_conv n i)] 1;
+
+fun mk_equiv_lsbis_tac sbis_lsbis lsbis_incl incl_lsbis bis_diag bis_converse bis_O =
+ EVERY' [rtac (@{thm equiv_def} RS iffD2),
+
+ rtac conjI, rtac (@{thm refl_on_def} RS iffD2),
+ rtac conjI, rtac lsbis_incl, rtac ballI, rtac set_mp,
+ rtac incl_lsbis, rtac bis_diag, atac, etac @{thm diagI},
+
+ rtac conjI, rtac (@{thm sym_def} RS iffD2),
+ rtac allI, rtac allI, rtac impI, rtac set_mp,
+ rtac incl_lsbis, rtac bis_converse, rtac sbis_lsbis, etac @{thm converseI},
+
+ rtac (@{thm trans_def} RS iffD2),
+ rtac allI, rtac allI, rtac allI, rtac impI, rtac impI, rtac set_mp,
+ rtac incl_lsbis, rtac bis_O, rtac sbis_lsbis, rtac sbis_lsbis,
+ etac @{thm relcompI}, atac] 1;
+
+fun mk_coalgT_tac m defs strT_defs set_naturalss {context = ctxt, prems = _} =
+ let
+ val n = length strT_defs;
+ val ks = 1 upto n;
+ fun coalg_tac (i, ((passive_sets, active_sets), def)) =
+ EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE],
+ hyp_subst_tac, rtac (def RS trans RS @{thm ssubst_mem}), etac (arg_cong RS trans),
+ rtac (mk_sum_casesN n i), rtac CollectI,
+ EVERY' (map (fn thm => EVERY' [rtac conjI, rtac (thm RS ord_eq_le_trans),
+ etac ((trans OF [@{thm image_id} RS fun_cong, @{thm id_apply}]) RS ord_eq_le_trans)])
+ passive_sets),
+ CONJ_WRAP' (fn (i, thm) => EVERY' [rtac (thm RS ord_eq_le_trans),
+ rtac @{thm image_subsetI}, rtac CollectI, rtac exI, rtac exI, rtac conjI, rtac refl,
+ rtac conjI,
+ rtac conjI, etac @{thm empty_Shift}, dtac set_rev_mp,
+ etac equalityD1, etac CollectD,
+ rtac conjI, etac @{thm Shift_clists},
+ rtac conjI, etac @{thm Shift_prefCl},
+ rtac conjI, rtac ballI,
+ rtac conjI, dtac @{thm iffD1[OF ball_conj_distrib]}, dtac conjunct1,
+ SELECT_GOAL (unfold_thms_tac ctxt @{thms Succ_Shift shift_def}),
+ etac bspec, etac @{thm ShiftD},
+ CONJ_WRAP' (fn i => EVERY' [rtac ballI, etac CollectE, dtac @{thm ShiftD},
+ dtac bspec, etac thin_rl, atac, dtac conjunct2, dtac (mk_conjunctN n i),
+ dtac bspec, rtac CollectI, etac @{thm set_mp[OF equalityD1[OF Succ_Shift]]},
+ REPEAT_DETERM o eresolve_tac [exE, conjE], rtac exI,
+ rtac conjI, rtac (@{thm shift_def} RS fun_cong RS trans),
+ rtac (@{thm append_Cons} RS sym RS arg_cong RS trans), atac,
+ REPEAT_DETERM_N m o (rtac conjI THEN' atac),
+ CONJ_WRAP' (K (EVERY' [etac trans, rtac @{thm Collect_cong},
+ rtac @{thm eqset_imp_iff}, rtac sym, rtac trans, rtac @{thm Succ_Shift},
+ rtac (@{thm append_Cons} RS sym RS arg_cong)])) ks]) ks,
+ rtac allI, rtac impI, REPEAT_DETERM o eresolve_tac [allE, impE],
+ etac @{thm not_in_Shift}, rtac trans, rtac (@{thm shift_def} RS fun_cong), atac,
+ dtac bspec, atac, dtac conjunct2, dtac (mk_conjunctN n i), dtac bspec,
+ etac @{thm set_mp[OF equalityD1]}, atac,
+ REPEAT_DETERM o eresolve_tac [exE, conjE], rtac exI,
+ rtac conjI, rtac (@{thm shift_def} RS fun_cong RS trans),
+ etac (@{thm append_Nil} RS sym RS arg_cong RS trans),
+ REPEAT_DETERM_N m o (rtac conjI THEN' atac),
+ CONJ_WRAP' (K (EVERY' [etac trans, rtac @{thm Collect_cong},
+ rtac @{thm eqset_imp_iff}, rtac sym, rtac trans, rtac @{thm Succ_Shift},
+ rtac (@{thm append_Nil} RS sym RS arg_cong)])) ks]) (ks ~~ active_sets)];
+ in
+ unfold_thms_tac ctxt defs THEN
+ CONJ_WRAP' coalg_tac (ks ~~ (map (chop m) set_naturalss ~~ strT_defs)) 1
+ end;
+
+fun mk_card_of_carT_tac m isNode_defs sbd_sbd
+ sbd_card_order sbd_Card_order sbd_Cinfinite sbd_Cnotzero in_sbds =
+ let
+ val n = length isNode_defs;
+ in
+ EVERY' [rtac (Thm.permute_prems 0 1 ctrans),
+ rtac @{thm card_of_Sigma_ordLeq_Cinfinite}, rtac @{thm Cinfinite_cexp},
+ if m = 0 then rtac @{thm ordLeq_refl} else rtac @{thm ordLeq_csum2},
+ rtac @{thm Card_order_ctwo}, rtac @{thm Cinfinite_cexp},
+ rtac @{thm ctwo_ordLeq_Cinfinite}, rtac sbd_Cinfinite, rtac sbd_Cinfinite,
+ rtac ctrans, rtac @{thm card_of_diff},
+ rtac ordIso_ordLeq_trans, rtac @{thm card_of_Field_ordIso},
+ rtac @{thm Card_order_cpow}, rtac ordIso_ordLeq_trans,
+ rtac @{thm cpow_cexp_ctwo}, rtac ctrans, rtac @{thm cexp_mono1_Cnotzero},
+ if m = 0 then rtac @{thm ordLeq_refl} else rtac @{thm ordLeq_csum2},
+ rtac @{thm Card_order_ctwo}, rtac @{thm ctwo_Cnotzero}, rtac @{thm Card_order_clists},
+ rtac @{thm cexp_mono2_Cnotzero}, rtac ordIso_ordLeq_trans,
+ rtac @{thm clists_Cinfinite},
+ if n = 1 then rtac sbd_Cinfinite else rtac (sbd_Cinfinite RS @{thm Cinfinite_csum1}),
+ rtac ordIso_ordLeq_trans, rtac sbd_sbd, rtac @{thm infinite_ordLeq_cexp},
+ rtac sbd_Cinfinite,
+ if m = 0 then rtac @{thm ctwo_Cnotzero} else rtac @{thm csum_Cnotzero2[OF ctwo_Cnotzero]},
+ rtac @{thm Cnotzero_clists},
+ rtac ballI, rtac ordIso_ordLeq_trans, rtac @{thm card_of_Func_Ffunc},
+ rtac ordIso_ordLeq_trans, rtac @{thm Func_cexp},
+ rtac ctrans, rtac @{thm cexp_mono},
+ rtac @{thm ordLeq_ordIso_trans},
+ CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1
+ (sbd_Cinfinite RS @{thm Cinfinite_cexp[OF ordLeq_csum2[OF Card_order_ctwo]]}
+ RSN (3, @{thm Un_Cinfinite_bound}))))
+ (fn thm => EVERY' [rtac ctrans, rtac @{thm card_of_image}, rtac thm]) (rev in_sbds),
+ rtac @{thm cexp_cong1_Cnotzero}, rtac @{thm csum_cong1},
+ REPEAT_DETERM_N m o rtac @{thm csum_cong2},
+ CONJ_WRAP_GEN' (rtac @{thm csum_cong})
+ (K (rtac (sbd_Card_order RS @{thm card_of_Field_ordIso}))) in_sbds,
+ rtac sbd_Card_order,
+ rtac @{thm csum_Cnotzero2}, rtac @{thm ctwo_Cnotzero},
+ rtac @{thm csum_Cnotzero2}, rtac @{thm ctwo_Cnotzero},
+ rtac @{thm ordLeq_ordIso_trans}, etac @{thm clists_bound},
+ rtac @{thm clists_Cinfinite}, TRY o rtac @{thm Cinfinite_csum1}, rtac sbd_Cinfinite,
+ rtac disjI2, rtac @{thm cone_ordLeq_cexp}, rtac @{thm cone_ordLeq_cexp},
+ rtac ctrans, rtac @{thm cone_ordLeq_ctwo}, rtac @{thm ordLeq_csum2},
+ rtac @{thm Card_order_ctwo}, rtac FalseE, etac @{thm cpow_clists_czero}, atac,
+ rtac @{thm card_of_Card_order},
+ rtac ordIso_ordLeq_trans, rtac @{thm cexp_cprod},
+ rtac @{thm csum_Cnotzero2}, rtac @{thm ctwo_Cnotzero},
+ rtac ordIso_ordLeq_trans, rtac @{thm cexp_cong2_Cnotzero},
+ rtac @{thm ordIso_transitive}, rtac @{thm cprod_cong2}, rtac sbd_sbd,
+ rtac @{thm cprod_infinite}, rtac sbd_Cinfinite,
+ rtac @{thm csum_Cnotzero2}, rtac @{thm ctwo_Cnotzero}, rtac @{thm Card_order_cprod},
+ rtac ctrans, rtac @{thm cexp_mono1_Cnotzero},
+ rtac ordIso_ordLeq_trans, rtac @{thm ordIso_transitive}, rtac @{thm csum_cong1},
+ rtac @{thm ordIso_transitive},
+ REPEAT_DETERM_N m o rtac @{thm csum_cong2},
+ rtac sbd_sbd,
+ BNF_Tactics.mk_rotate_eq_tac (rtac @{thm ordIso_refl} THEN'
+ FIRST' [rtac @{thm card_of_Card_order},
+ rtac @{thm Card_order_csum}, rtac sbd_Card_order])
+ @{thm ordIso_transitive} @{thm csum_assoc} @{thm csum_com} @{thm csum_cong}
+ (1 upto m + 1) (m + 1 :: (1 upto m)),
+ if m = 0 then K all_tac else EVERY' [rtac @{thm ordIso_transitive}, rtac @{thm csum_assoc}],
+ rtac @{thm csum_com}, rtac @{thm csum_cexp'}, rtac sbd_Cinfinite,
+ if m = 0 then rtac @{thm Card_order_ctwo} else rtac @{thm Card_order_csum},
+ if m = 0 then rtac @{thm ordLeq_refl} else rtac @{thm ordLeq_csum2},
+ rtac @{thm Card_order_ctwo},
+ rtac @{thm csum_Cnotzero2}, rtac @{thm ctwo_Cnotzero}, rtac sbd_Card_order,
+ rtac ordIso_ordLeq_trans, rtac @{thm cexp_cprod_ordLeq},
+ if m = 0 then rtac @{thm ctwo_Cnotzero} else rtac @{thm csum_Cnotzero2[OF ctwo_Cnotzero]},
+ rtac sbd_Cinfinite, rtac sbd_Cnotzero, rtac @{thm ordLeq_refl}, rtac sbd_Card_order,
+ rtac @{thm cexp_mono2_Cnotzero}, rtac @{thm infinite_ordLeq_cexp},
+ rtac sbd_Cinfinite,
+ if m = 0 then rtac @{thm ctwo_Cnotzero} else rtac @{thm csum_Cnotzero2[OF ctwo_Cnotzero]},
+ rtac sbd_Cnotzero,
+ rtac @{thm card_of_mono1}, rtac subsetI,
+ REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE, @{thm prod_caseE}], hyp_subst_tac,
+ rtac @{thm SigmaI}, rtac @{thm DiffI}, rtac set_mp, rtac equalityD2,
+ rtac (@{thm cpow_def} RS arg_cong RS trans), rtac (@{thm Pow_def} RS arg_cong RS trans),
+ rtac @{thm Field_card_of}, rtac CollectI, atac, rtac notI, etac @{thm singletonE},
+ hyp_subst_tac, etac @{thm emptyE}, rtac (@{thm Ffunc_def} RS equalityD2 RS set_mp),
+ rtac CollectI, rtac conjI, rtac ballI, dtac bspec, etac thin_rl, atac, dtac conjunct1,
+ CONJ_WRAP_GEN' (etac disjE) (fn (i, def) => EVERY'
+ [rtac (mk_UnIN n i), dtac (def RS iffD1),
+ REPEAT_DETERM o eresolve_tac [exE, conjE], rtac @{thm image_eqI}, atac, rtac CollectI,
+ REPEAT_DETERM_N m o (rtac conjI THEN' atac),
+ CONJ_WRAP' (K (EVERY' [etac ord_eq_le_trans, rtac subset_trans,
+ rtac subset_UNIV, rtac equalityD2, rtac @{thm Field_card_order},
+ rtac sbd_card_order])) isNode_defs]) (1 upto n ~~ isNode_defs),
+ atac] 1
+ end;
+
+fun mk_carT_set_tac n i carT_def strT_def isNode_def set_natural {context = ctxt, prems = _}=
+ EVERY' [dtac (carT_def RS equalityD1 RS set_mp),
+ REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE],
+ dtac Pair_eqD,
+ etac conjE, hyp_subst_tac,
+ dtac (isNode_def RS iffD1),
+ REPEAT_DETERM o eresolve_tac [exE, conjE],
+ rtac (equalityD2 RS set_mp),
+ rtac (strT_def RS arg_cong RS trans),
+ etac (arg_cong RS trans),
+ fo_rtac (mk_sum_casesN n i RS arg_cong RS trans) ctxt,
+ rtac set_natural, rtac imageI, etac (equalityD2 RS set_mp), rtac CollectI,
+ etac @{thm prefCl_Succ}, atac] 1;
+
+fun mk_strT_hset_tac n m j arg_cong_cTs cTs cts carT_defs strT_defs isNode_defs
+ set_incl_hsets set_hset_incl_hsetss coalg_setss carT_setss coalgT set_naturalss =
+ let
+ val set_naturals = map (fn xs => nth xs (j - 1)) set_naturalss;
+ val ks = 1 upto n;
+ fun base_tac (i, (cT, (strT_def, (set_incl_hset, set_natural)))) =
+ CONJ_WRAP' (fn (i', (carT_def, isNode_def)) => rtac impI THEN' etac conjE THEN'
+ (if i = i'
+ then EVERY' [rtac @{thm xt1(4)}, rtac set_incl_hset,
+ rtac (strT_def RS arg_cong RS trans), etac (arg_cong RS trans),
+ rtac (Thm.permute_prems 0 1 (set_natural RS box_equals)),
+ rtac (trans OF [@{thm image_id} RS fun_cong, @{thm id_apply}]),
+ rtac (mk_sum_casesN n i RS (Drule.instantiate' [cT] [] arg_cong) RS sym)]
+ else EVERY' [dtac (carT_def RS equalityD1 RS set_mp),
+ REPEAT_DETERM o eresolve_tac [CollectE, exE], etac conjE,
+ dtac conjunct2, dtac Pair_eqD, etac conjE,
+ hyp_subst_tac, dtac (isNode_def RS iffD1),
+ REPEAT_DETERM o eresolve_tac [exE, conjE],
+ rtac (mk_InN_not_InM i i' RS notE), etac (sym RS trans), atac]))
+ (ks ~~ (carT_defs ~~ isNode_defs));
+ fun step_tac (i, (coalg_sets, (carT_sets, set_hset_incl_hsets))) =
+ dtac (mk_conjunctN n i) THEN'
+ CONJ_WRAP' (fn (coalg_set, (carT_set, set_hset_incl_hset)) =>
+ EVERY' [rtac impI, etac conjE, etac impE, rtac conjI,
+ rtac (coalgT RS coalg_set RS set_mp), atac, etac carT_set, atac, atac,
+ etac (@{thm shift_def} RS fun_cong RS trans), etac subset_trans,
+ rtac set_hset_incl_hset, etac carT_set, atac, atac])
+ (coalg_sets ~~ (carT_sets ~~ set_hset_incl_hsets));
+ in
+ EVERY' [rtac (Drule.instantiate' cTs cts @{thm list.induct}),
+ REPEAT_DETERM o rtac allI, rtac impI,
+ CONJ_WRAP' base_tac
+ (ks ~~ (arg_cong_cTs ~~ (strT_defs ~~ (set_incl_hsets ~~ set_naturals)))),
+ REPEAT_DETERM o rtac allI, rtac impI,
+ REPEAT_DETERM o eresolve_tac [allE, impE], etac @{thm ShiftI},
+ CONJ_WRAP' (fn i => dtac (mk_conjunctN n i) THEN' rtac (mk_sumEN n) THEN'
+ CONJ_WRAP_GEN' (K all_tac) step_tac
+ (ks ~~ (drop m coalg_setss ~~ (carT_setss ~~ set_hset_incl_hsetss)))) ks] 1
+ end;
+
+fun mk_isNode_hset_tac n isNode_def strT_hsets {context = ctxt, prems = _} =
+ let
+ val m = length strT_hsets;
+ in
+ if m = 0 then atac 1
+ else (unfold_thms_tac ctxt [isNode_def] THEN
+ EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE],
+ rtac exI, rtac conjI, atac,
+ CONJ_WRAP' (fn (thm, i) => if i > m then atac
+ else EVERY' [rtac (thm RS subset_trans), atac, rtac conjI, atac, atac, atac])
+ (strT_hsets @ (replicate n mp) ~~ (1 upto (m + n)))] 1)
+ end;
+
+fun mk_Lev_sbd_tac cts Lev_0s Lev_Sucs to_sbdss =
+ let
+ val n = length Lev_0s;
+ val ks = 1 upto n;
+ in
+ EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn Lev_0 =>
+ EVERY' (map rtac [ord_eq_le_trans, Lev_0, @{thm Nil_clists}])) Lev_0s,
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn (Lev_Suc, to_sbds) =>
+ EVERY' [rtac ord_eq_le_trans, rtac Lev_Suc,
+ CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_least}))
+ (fn (i, to_sbd) => EVERY' [rtac subsetI,
+ REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac,
+ rtac @{thm Cons_clists}, rtac (mk_InN_Field n i), etac to_sbd,
+ etac set_rev_mp, REPEAT_DETERM o etac allE,
+ etac (mk_conjunctN n i)])
+ (rev (ks ~~ to_sbds))])
+ (Lev_Sucs ~~ to_sbdss)] 1
+ end;
+
+fun mk_length_Lev_tac cts Lev_0s Lev_Sucs =
+ let
+ val n = length Lev_0s;
+ val ks = n downto 1;
+ in
+ EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn Lev_0 =>
+ EVERY' [rtac impI, dtac (Lev_0 RS equalityD1 RS set_mp),
+ etac @{thm singletonE}, etac ssubst, rtac @{thm list.size(3)}]) Lev_0s,
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn Lev_Suc =>
+ EVERY' [rtac impI, dtac (Lev_Suc RS equalityD1 RS set_mp),
+ CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE))
+ (fn i => EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac,
+ rtac trans, rtac @{thm length_Cons}, rtac @{thm arg_cong[of _ _ Suc]},
+ REPEAT_DETERM o etac allE, dtac (mk_conjunctN n i), etac mp, atac]) ks])
+ Lev_Sucs] 1
+ end;
+
+fun mk_length_Lev'_tac length_Lev =
+ EVERY' [ftac length_Lev, etac ssubst, atac] 1;
+
+fun mk_prefCl_Lev_tac cts Lev_0s Lev_Sucs =
+ let
+ val n = length Lev_0s;
+ val ks = n downto 1;
+ in
+ EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn Lev_0 =>
+ EVERY' [rtac impI, etac conjE, dtac (Lev_0 RS equalityD1 RS set_mp),
+ etac @{thm singletonE}, hyp_subst_tac, dtac @{thm prefix_Nil[THEN subst, of "%x. x"]},
+ hyp_subst_tac, rtac @{thm set_mp[OF equalityD2[OF trans[OF arg_cong[OF list.size(3)]]]]},
+ rtac Lev_0, rtac @{thm singletonI}]) Lev_0s,
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn (Lev_0, Lev_Suc) =>
+ EVERY' [rtac impI, etac conjE, dtac (Lev_Suc RS equalityD1 RS set_mp),
+ CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE))
+ (fn i => EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac,
+ dtac @{thm prefix_Cons[THEN subst, of "%x. x"]}, etac disjE, hyp_subst_tac,
+ rtac @{thm set_mp[OF equalityD2[OF trans[OF arg_cong[OF list.size(3)]]]]},
+ rtac Lev_0, rtac @{thm singletonI},
+ REPEAT_DETERM o eresolve_tac [exE, conjE], hyp_subst_tac,
+ rtac @{thm set_mp[OF equalityD2[OF trans[OF arg_cong[OF length_Cons]]]]},
+ rtac Lev_Suc, rtac (mk_UnIN n i), rtac CollectI, REPEAT_DETERM o rtac exI, rtac conjI,
+ rtac refl, etac conjI, REPEAT_DETERM o etac allE, dtac (mk_conjunctN n i),
+ etac mp, etac conjI, atac]) ks])
+ (Lev_0s ~~ Lev_Sucs)] 1
+ end;
+
+fun mk_rv_last_tac cTs cts rv_Nils rv_Conss =
+ let
+ val n = length rv_Nils;
+ val ks = 1 upto n;
+ in
+ EVERY' [rtac (Drule.instantiate' cTs cts @{thm list.induct}),
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn rv_Cons =>
+ CONJ_WRAP' (fn (i, rv_Nil) => (EVERY' [rtac exI,
+ rtac (@{thm append_Nil} RS arg_cong RS trans),
+ rtac (rv_Cons RS trans), rtac (mk_sum_casesN n i RS arg_cong RS trans), rtac rv_Nil]))
+ (ks ~~ rv_Nils))
+ rv_Conss,
+ REPEAT_DETERM o rtac allI, rtac (mk_sumEN n),
+ EVERY' (map (fn i =>
+ CONJ_WRAP' (fn rv_Cons => EVERY' [REPEAT_DETERM o etac allE, dtac (mk_conjunctN n i),
+ CONJ_WRAP' (fn i' => EVERY' [dtac (mk_conjunctN n i'), etac exE, rtac exI,
+ rtac (@{thm append_Cons} RS arg_cong RS trans),
+ rtac (rv_Cons RS trans), etac (sum_case_weak_cong RS arg_cong RS trans),
+ rtac (mk_sum_casesN n i RS arg_cong RS trans), atac])
+ ks])
+ rv_Conss)
+ ks)] 1
+ end;
+
+fun mk_set_rv_Lev_tac m cts Lev_0s Lev_Sucs rv_Nils rv_Conss coalg_setss from_to_sbdss =
+ let
+ val n = length Lev_0s;
+ val ks = 1 upto n;
+ in
+ EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn (i, ((Lev_0, rv_Nil), coalg_sets)) =>
+ EVERY' [rtac impI, REPEAT_DETERM o etac conjE,
+ dtac (Lev_0 RS equalityD1 RS set_mp), etac @{thm singletonE}, etac ssubst,
+ rtac (rv_Nil RS arg_cong RS iffD2),
+ rtac (mk_sum_casesN n i RS iffD2),
+ CONJ_WRAP' (fn thm => etac thm THEN' atac) (take m coalg_sets)])
+ (ks ~~ ((Lev_0s ~~ rv_Nils) ~~ coalg_setss)),
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn ((Lev_Suc, rv_Cons), (from_to_sbds, coalg_sets)) =>
+ EVERY' [rtac impI, etac conjE, dtac (Lev_Suc RS equalityD1 RS set_mp),
+ CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE))
+ (fn (i, (from_to_sbd, coalg_set)) =>
+ EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac,
+ rtac (rv_Cons RS arg_cong RS iffD2),
+ rtac (mk_sum_casesN n i RS arg_cong RS trans RS iffD2),
+ etac (from_to_sbd RS arg_cong), REPEAT_DETERM o etac allE,
+ dtac (mk_conjunctN n i), etac mp, etac conjI, etac set_rev_mp,
+ etac coalg_set, atac])
+ (rev (ks ~~ (from_to_sbds ~~ drop m coalg_sets)))])
+ ((Lev_Sucs ~~ rv_Conss) ~~ (from_to_sbdss ~~ coalg_setss))] 1
+ end;
+
+fun mk_set_Lev_tac cts Lev_0s Lev_Sucs rv_Nils rv_Conss from_to_sbdss =
+ let
+ val n = length Lev_0s;
+ val ks = 1 upto n;
+ in
+ EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn ((i, (Lev_0, Lev_Suc)), rv_Nil) =>
+ EVERY' [rtac impI, dtac (Lev_0 RS equalityD1 RS set_mp),
+ etac @{thm singletonE}, hyp_subst_tac,
+ CONJ_WRAP' (fn i' => rtac impI THEN' dtac (sym RS trans) THEN' rtac rv_Nil THEN'
+ (if i = i'
+ then EVERY' [dtac (mk_InN_inject n i), hyp_subst_tac,
+ CONJ_WRAP' (fn (i'', Lev_0'') =>
+ EVERY' [rtac impI, rtac @{thm ssubst_mem[OF append_Nil]},
+ rtac (Lev_Suc RS equalityD2 RS set_mp), rtac (mk_UnIN n i''),
+ rtac CollectI, REPEAT_DETERM o rtac exI, rtac conjI, rtac refl,
+ etac conjI, rtac (Lev_0'' RS equalityD2 RS set_mp),
+ rtac @{thm singletonI}])
+ (ks ~~ Lev_0s)]
+ else etac (mk_InN_not_InM i' i RS notE)))
+ ks])
+ ((ks ~~ (Lev_0s ~~ Lev_Sucs)) ~~ rv_Nils),
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn ((Lev_Suc, rv_Cons), from_to_sbds) =>
+ EVERY' [rtac impI, dtac (Lev_Suc RS equalityD1 RS set_mp),
+ CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE))
+ (fn (i, from_to_sbd) =>
+ EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac,
+ CONJ_WRAP' (fn i' => rtac impI THEN'
+ CONJ_WRAP' (fn i'' =>
+ EVERY' [rtac impI, rtac (Lev_Suc RS equalityD2 RS set_mp),
+ rtac @{thm ssubst_mem[OF append_Cons]}, rtac (mk_UnIN n i),
+ rtac CollectI, REPEAT_DETERM o rtac exI, rtac conjI, rtac refl,
+ rtac conjI, atac, dtac (sym RS trans RS sym),
+ rtac (rv_Cons RS trans), rtac (mk_sum_casesN n i RS trans),
+ etac (from_to_sbd RS arg_cong), REPEAT_DETERM o etac allE,
+ dtac (mk_conjunctN n i), dtac mp, atac,
+ dtac (mk_conjunctN n i'), dtac mp, atac,
+ dtac (mk_conjunctN n i''), etac mp, atac])
+ ks)
+ ks])
+ (rev (ks ~~ from_to_sbds))])
+ ((Lev_Sucs ~~ rv_Conss) ~~ from_to_sbdss)] 1
+ end;
+
+fun mk_set_image_Lev_tac cts Lev_0s Lev_Sucs rv_Nils rv_Conss from_to_sbdss to_sbd_injss =
+ let
+ val n = length Lev_0s;
+ val ks = 1 upto n;
+ in
+ EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn ((i, (Lev_0, Lev_Suc)), rv_Nil) =>
+ EVERY' [rtac impI, dtac (Lev_0 RS equalityD1 RS set_mp),
+ etac @{thm singletonE}, hyp_subst_tac,
+ CONJ_WRAP' (fn i' => rtac impI THEN'
+ CONJ_WRAP' (fn i'' => rtac impI THEN' dtac (sym RS trans) THEN' rtac rv_Nil THEN'
+ (if i = i''
+ then EVERY' [dtac @{thm ssubst_mem[OF sym[OF append_Nil]]},
+ dtac (Lev_Suc RS equalityD1 RS set_mp), dtac (mk_InN_inject n i),
+ hyp_subst_tac,
+ CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE))
+ (fn k => REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE] THEN'
+ dtac list_inject_iffD1 THEN' etac conjE THEN'
+ (if k = i'
+ then EVERY' [dtac (mk_InN_inject n k), hyp_subst_tac, etac imageI]
+ else etac (mk_InN_not_InM i' k RS notE)))
+ (rev ks)]
+ else etac (mk_InN_not_InM i'' i RS notE)))
+ ks)
+ ks])
+ ((ks ~~ (Lev_0s ~~ Lev_Sucs)) ~~ rv_Nils),
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn ((Lev_Suc, rv_Cons), (from_to_sbds, to_sbd_injs)) =>
+ EVERY' [rtac impI, dtac (Lev_Suc RS equalityD1 RS set_mp),
+ CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE))
+ (fn (i, (from_to_sbd, to_sbd_inj)) =>
+ REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE] THEN' hyp_subst_tac THEN'
+ CONJ_WRAP' (fn i' => rtac impI THEN'
+ dtac @{thm ssubst_mem[OF sym[OF append_Cons]]} THEN'
+ dtac (Lev_Suc RS equalityD1 RS set_mp) THEN'
+ CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE)) (fn k =>
+ REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE] THEN'
+ dtac list_inject_iffD1 THEN' etac conjE THEN'
+ (if k = i
+ then EVERY' [dtac (mk_InN_inject n i),
+ dtac (Thm.permute_prems 0 2 (to_sbd_inj RS iffD1)),
+ atac, atac, hyp_subst_tac] THEN'
+ CONJ_WRAP' (fn i'' =>
+ EVERY' [rtac impI, dtac (sym RS trans),
+ rtac (rv_Cons RS trans), rtac (mk_sum_casesN n i RS arg_cong RS trans),
+ etac (from_to_sbd RS arg_cong),
+ REPEAT_DETERM o etac allE,
+ dtac (mk_conjunctN n i), dtac mp, atac,
+ dtac (mk_conjunctN n i'), dtac mp, atac,
+ dtac (mk_conjunctN n i''), etac mp, etac sym])
+ ks
+ else etac (mk_InN_not_InM i k RS notE)))
+ (rev ks))
+ ks)
+ (rev (ks ~~ (from_to_sbds ~~ to_sbd_injs)))])
+ ((Lev_Sucs ~~ rv_Conss) ~~ (from_to_sbdss ~~ to_sbd_injss))] 1
+ end;
+
+fun mk_mor_beh_tac m mor_def mor_cong beh_defs carT_defs strT_defs isNode_defs
+ to_sbd_injss from_to_sbdss Lev_0s Lev_Sucs rv_Nils rv_Conss Lev_sbds length_Levs length_Lev's
+ prefCl_Levs rv_lastss set_rv_Levsss set_Levsss set_image_Levsss set_naturalss coalg_setss
+ map_comp_ids map_congs map_arg_congs {context = ctxt, prems = _} =
+ let
+ val n = length beh_defs;
+ val ks = 1 upto n;
+
+ fun fbetw_tac (i, (carT_def, (isNode_def, (Lev_0, (rv_Nil, (Lev_sbd,
+ ((length_Lev, length_Lev'), (prefCl_Lev, (rv_lasts, (set_naturals,
+ (coalg_sets, (set_rv_Levss, (set_Levss, set_image_Levss))))))))))))) =
+ EVERY' [rtac ballI, rtac (carT_def RS equalityD2 RS set_mp),
+ rtac CollectI, REPEAT_DETERM o rtac exI, rtac conjI, rtac refl, rtac conjI,
+ rtac conjI,
+ rtac @{thm UN_I}, rtac UNIV_I, rtac (Lev_0 RS equalityD2 RS set_mp),
+ rtac @{thm singletonI},
+ rtac conjI,
+ rtac @{thm UN_least}, rtac Lev_sbd,
+ rtac conjI,
+ rtac @{thm prefCl_UN}, rtac ssubst, rtac @{thm PrefCl_def}, REPEAT_DETERM o rtac allI,
+ rtac impI, etac conjE, rtac exI, rtac conjI, rtac @{thm ord_le_eq_trans},
+ etac @{thm prefix_length_le}, etac length_Lev, rtac prefCl_Lev, etac conjI, atac,
+ rtac conjI,
+ rtac ballI, etac @{thm UN_E}, rtac conjI,
+ if n = 1 then K all_tac else rtac (mk_sumEN n),
+ EVERY' (map6 (fn i => fn isNode_def => fn set_naturals =>
+ fn set_rv_Levs => fn set_Levs => fn set_image_Levs =>
+ EVERY' [rtac (mk_disjIN n i), rtac (isNode_def RS ssubst),
+ rtac exI, rtac conjI,
+ (if n = 1 then rtac @{thm if_P} THEN' etac length_Lev'
+ else rtac (@{thm if_P} RS arg_cong RS trans) THEN' etac length_Lev' THEN'
+ etac (sum_case_weak_cong RS trans) THEN' rtac (mk_sum_casesN n i)),
+ EVERY' (map2 (fn set_natural => fn set_rv_Lev =>
+ EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac (set_natural RS trans),
+ rtac trans_fun_cong_image_id_id_apply,
+ etac set_rv_Lev, TRY o atac, etac conjI, atac])
+ (take m set_naturals) set_rv_Levs),
+ CONJ_WRAP' (fn (set_natural, (set_Lev, set_image_Lev)) =>
+ EVERY' [rtac (set_natural RS trans), rtac equalityI, rtac @{thm image_subsetI},
+ rtac CollectI, rtac @{thm SuccI}, rtac @{thm UN_I}, rtac UNIV_I, etac set_Lev,
+ if n = 1 then rtac refl else atac, atac, rtac subsetI,
+ REPEAT_DETERM o eresolve_tac [CollectE, @{thm SuccE}, @{thm UN_E}],
+ rtac set_image_Lev, atac, dtac length_Lev, hyp_subst_tac, dtac length_Lev',
+ etac @{thm set_mp[OF equalityD1[OF arg_cong[OF length_append_singleton]]]},
+ if n = 1 then rtac refl else atac])
+ (drop m set_naturals ~~ (set_Levs ~~ set_image_Levs))])
+ ks isNode_defs set_naturalss set_rv_Levss set_Levss set_image_Levss),
+ CONJ_WRAP' (fn (i, (rv_last, (isNode_def, (set_naturals,
+ (set_rv_Levs, (set_Levs, set_image_Levs)))))) =>
+ EVERY' [rtac ballI,
+ REPEAT_DETERM o eresolve_tac [CollectE, @{thm SuccE}, @{thm UN_E}],
+ rtac (rev_mp OF [rv_last, impI]), etac exE, rtac (isNode_def RS ssubst),
+ rtac exI, rtac conjI,
+ (if n = 1 then rtac @{thm if_P} THEN' etac length_Lev'
+ else rtac (@{thm if_P} RS trans) THEN' etac length_Lev' THEN'
+ etac (sum_case_weak_cong RS trans) THEN' rtac (mk_sum_casesN n i)),
+ EVERY' (map2 (fn set_natural => fn set_rv_Lev =>
+ EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac (set_natural RS trans),
+ rtac trans_fun_cong_image_id_id_apply,
+ etac set_rv_Lev, TRY o atac, etac conjI, atac])
+ (take m set_naturals) set_rv_Levs),
+ CONJ_WRAP' (fn (set_natural, (set_Lev, set_image_Lev)) =>
+ EVERY' [rtac (set_natural RS trans), rtac equalityI, rtac @{thm image_subsetI},
+ rtac CollectI, rtac @{thm SuccI}, rtac @{thm UN_I}, rtac UNIV_I, etac set_Lev,
+ if n = 1 then rtac refl else atac, atac, rtac subsetI,
+ REPEAT_DETERM o eresolve_tac [CollectE, @{thm SuccE}, @{thm UN_E}],
+ REPEAT_DETERM_N 4 o etac thin_rl,
+ rtac set_image_Lev,
+ atac, dtac length_Lev, hyp_subst_tac, dtac length_Lev',
+ etac @{thm set_mp[OF equalityD1[OF arg_cong[OF length_append_singleton]]]},
+ if n = 1 then rtac refl else atac])
+ (drop m set_naturals ~~ (set_Levs ~~ set_image_Levs))])
+ (ks ~~ (rv_lasts ~~ (isNode_defs ~~ (set_naturalss ~~
+ (set_rv_Levss ~~ (set_Levss ~~ set_image_Levss)))))),
+ (**)
+ rtac allI, rtac impI, rtac @{thm if_not_P}, rtac notI,
+ etac notE, etac @{thm UN_I[OF UNIV_I]},
+ (*root isNode*)
+ rtac (isNode_def RS ssubst), rtac exI, rtac conjI, rtac (@{thm if_P} RS trans),
+ rtac length_Lev', rtac (Lev_0 RS equalityD2 RS set_mp), rtac @{thm singletonI},
+ CONVERSION (Conv.top_conv
+ (K (Conv.try_conv (Conv.rewr_conv (rv_Nil RS eq_reflection)))) ctxt),
+ if n = 1 then rtac refl else rtac (mk_sum_casesN n i),
+ EVERY' (map2 (fn set_natural => fn coalg_set =>
+ EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac (set_natural RS trans),
+ rtac trans_fun_cong_image_id_id_apply, etac coalg_set, atac])
+ (take m set_naturals) (take m coalg_sets)),
+ CONJ_WRAP' (fn (set_natural, (set_Lev, set_image_Lev)) =>
+ EVERY' [rtac (set_natural RS trans), rtac equalityI, rtac @{thm image_subsetI},
+ rtac CollectI, rtac @{thm SuccI}, rtac @{thm UN_I}, rtac UNIV_I, rtac set_Lev,
+ rtac (Lev_0 RS equalityD2 RS set_mp), rtac @{thm singletonI}, rtac rv_Nil,
+ atac, rtac subsetI,
+ REPEAT_DETERM o eresolve_tac [CollectE, @{thm SuccE}, @{thm UN_E}],
+ rtac set_image_Lev, rtac (Lev_0 RS equalityD2 RS set_mp),
+ rtac @{thm singletonI}, dtac length_Lev',
+ etac @{thm set_mp[OF equalityD1[OF arg_cong[OF
+ trans[OF length_append_singleton arg_cong[of _ _ Suc, OF list.size(3)]]]]]},
+ rtac rv_Nil])
+ (drop m set_naturals ~~ (nth set_Levss (i - 1) ~~ nth set_image_Levss (i - 1)))];
+
+ fun mor_tac (i, (strT_def, (((Lev_0, Lev_Suc), (rv_Nil, rv_Cons)),
+ ((map_comp_id, (map_cong, map_arg_cong)), (length_Lev', (from_to_sbds, to_sbd_injs)))))) =
+ EVERY' [rtac ballI, rtac sym, rtac trans, rtac strT_def,
+ rtac (@{thm if_P} RS
+ (if n = 1 then map_arg_cong else sum_case_weak_cong) RS trans),
+ rtac (@{thm list.size(3)} RS arg_cong RS trans RS equalityD2 RS set_mp),
+ rtac Lev_0, rtac @{thm singletonI},
+ CONVERSION (Conv.top_conv
+ (K (Conv.try_conv (Conv.rewr_conv (rv_Nil RS eq_reflection)))) ctxt),
+ if n = 1 then K all_tac
+ else (rtac (sum_case_weak_cong RS trans) THEN'
+ rtac (mk_sum_casesN n i) THEN' rtac (mk_sum_casesN n i RS trans)),
+ rtac (map_comp_id RS trans), rtac (map_cong OF replicate m refl),
+ EVERY' (map3 (fn i' => fn to_sbd_inj => fn from_to_sbd =>
+ DETERM o EVERY' [rtac trans, rtac o_apply, rtac Pair_eqI, rtac conjI,
+ rtac trans, rtac @{thm Shift_def},
+ rtac equalityI, rtac subsetI, etac thin_rl, etac thin_rl,
+ REPEAT_DETERM o eresolve_tac [CollectE, @{thm UN_E}], dtac length_Lev', dtac asm_rl,
+ etac thin_rl, dtac @{thm set_rev_mp[OF _ equalityD1]},
+ rtac (@{thm length_Cons} RS arg_cong RS trans), rtac Lev_Suc,
+ CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE)) (fn i'' =>
+ EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE],
+ dtac list_inject_iffD1, etac conjE,
+ if i' = i'' then EVERY' [dtac (mk_InN_inject n i'),
+ dtac (Thm.permute_prems 0 2 (to_sbd_inj RS iffD1)),
+ atac, atac, hyp_subst_tac, etac @{thm UN_I[OF UNIV_I]}]
+ else etac (mk_InN_not_InM i' i'' RS notE)])
+ (rev ks),
+ rtac @{thm UN_least}, rtac subsetI, rtac CollectI, rtac @{thm UN_I[OF UNIV_I]},
+ rtac (Lev_Suc RS equalityD2 RS set_mp), rtac (mk_UnIN n i'), rtac CollectI,
+ REPEAT_DETERM o rtac exI, rtac conjI, rtac refl, etac conjI, atac,
+ rtac trans, rtac @{thm shift_def}, rtac ssubst, rtac @{thm fun_eq_iff}, rtac allI,
+ rtac @{thm if_cong}, rtac (@{thm length_Cons} RS arg_cong RS trans), rtac iffI,
+ dtac asm_rl, dtac asm_rl, dtac asm_rl,
+ dtac (Lev_Suc RS equalityD1 RS set_mp),
+ CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE)) (fn i'' =>
+ EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE],
+ dtac list_inject_iffD1, etac conjE,
+ if i' = i'' then EVERY' [dtac (mk_InN_inject n i'),
+ dtac (Thm.permute_prems 0 2 (to_sbd_inj RS iffD1)),
+ atac, atac, hyp_subst_tac, atac]
+ else etac (mk_InN_not_InM i' i'' RS notE)])
+ (rev ks),
+ rtac (Lev_Suc RS equalityD2 RS set_mp), rtac (mk_UnIN n i'), rtac CollectI,
+ REPEAT_DETERM o rtac exI, rtac conjI, rtac refl, etac conjI, atac,
+ CONVERSION (Conv.top_conv
+ (K (Conv.try_conv (Conv.rewr_conv (rv_Cons RS eq_reflection)))) ctxt),
+ if n = 1 then K all_tac
+ else rtac sum_case_weak_cong THEN' rtac (mk_sum_casesN n i' RS trans),
+ SELECT_GOAL (unfold_thms_tac ctxt [from_to_sbd]), rtac refl,
+ rtac refl])
+ ks to_sbd_injs from_to_sbds)];
+ in
+ (rtac mor_cong THEN'
+ EVERY' (map (fn thm => rtac (thm RS ext)) beh_defs) THEN'
+ stac mor_def THEN' rtac conjI THEN'
+ CONJ_WRAP' fbetw_tac
+ (ks ~~ (carT_defs ~~ (isNode_defs ~~ (Lev_0s ~~ (rv_Nils ~~ (Lev_sbds ~~
+ ((length_Levs ~~ length_Lev's) ~~ (prefCl_Levs ~~ (rv_lastss ~~
+ (set_naturalss ~~ (coalg_setss ~~
+ (set_rv_Levsss ~~ (set_Levsss ~~ set_image_Levsss))))))))))))) THEN'
+ CONJ_WRAP' mor_tac
+ (ks ~~ (strT_defs ~~ (((Lev_0s ~~ Lev_Sucs) ~~ (rv_Nils ~~ rv_Conss)) ~~
+ ((map_comp_ids ~~ (map_congs ~~ map_arg_congs)) ~~
+ (length_Lev's ~~ (from_to_sbdss ~~ to_sbd_injss))))))) 1
+ end;
+
+fun mk_congruent_str_final_tac m lsbisE map_comp_id map_cong equiv_LSBISs =
+ EVERY' [rtac @{thm congruentI}, dtac lsbisE,
+ REPEAT_DETERM o eresolve_tac [CollectE, conjE, bexE], rtac (o_apply RS trans),
+ etac (sym RS arg_cong RS trans), rtac (map_comp_id RS trans),
+ rtac (map_cong RS trans), REPEAT_DETERM_N m o rtac refl,
+ EVERY' (map (fn equiv_LSBIS =>
+ EVERY' [rtac @{thm equiv_proj}, rtac equiv_LSBIS, etac set_mp, atac])
+ equiv_LSBISs), rtac sym, rtac (o_apply RS trans),
+ etac (sym RS arg_cong RS trans), rtac map_comp_id] 1;
+
+fun mk_coalg_final_tac m coalg_def congruent_str_finals equiv_LSBISs set_naturalss coalgT_setss =
+ EVERY' [stac coalg_def,
+ CONJ_WRAP' (fn ((set_naturals, coalgT_sets), (equiv_LSBIS, congruent_str_final)) =>
+ EVERY' [rtac @{thm univ_preserves}, rtac equiv_LSBIS, rtac congruent_str_final,
+ rtac ballI, rtac @{thm ssubst_mem}, rtac o_apply, rtac CollectI,
+ EVERY' (map2 (fn set_natural => fn coalgT_set =>
+ EVERY' [rtac conjI, rtac (set_natural RS ord_eq_le_trans),
+ rtac ord_eq_le_trans_trans_fun_cong_image_id_id_apply,
+ etac coalgT_set])
+ (take m set_naturals) (take m coalgT_sets)),
+ CONJ_WRAP' (fn (equiv_LSBIS, (set_natural, coalgT_set)) =>
+ EVERY' [rtac (set_natural RS ord_eq_le_trans),
+ rtac @{thm image_subsetI}, rtac ssubst, rtac @{thm proj_in_iff},
+ rtac equiv_LSBIS, etac set_rev_mp, etac coalgT_set])
+ (equiv_LSBISs ~~ drop m (set_naturals ~~ coalgT_sets))])
+ ((set_naturalss ~~ coalgT_setss) ~~ (equiv_LSBISs ~~ congruent_str_finals))] 1;
+
+fun mk_mor_T_final_tac mor_def congruent_str_finals equiv_LSBISs =
+ EVERY' [stac mor_def, rtac conjI,
+ CONJ_WRAP' (fn equiv_LSBIS =>
+ EVERY' [rtac ballI, rtac ssubst, rtac @{thm proj_in_iff}, rtac equiv_LSBIS, atac])
+ equiv_LSBISs,
+ CONJ_WRAP' (fn (equiv_LSBIS, congruent_str_final) =>
+ EVERY' [rtac ballI, rtac sym, rtac trans, rtac @{thm univ_commute}, rtac equiv_LSBIS,
+ rtac congruent_str_final, atac, rtac o_apply])
+ (equiv_LSBISs ~~ congruent_str_finals)] 1;
+
+fun mk_mor_Rep_tac m defs Reps Abs_inverses coalg_final_setss map_comp_ids map_congLs
+ {context = ctxt, prems = _} =
+ unfold_thms_tac ctxt defs THEN
+ EVERY' [rtac conjI,
+ CONJ_WRAP' (fn thm => rtac ballI THEN' rtac thm) Reps,
+ CONJ_WRAP' (fn (Rep, ((map_comp_id, map_congL), coalg_final_sets)) =>
+ EVERY' [rtac ballI, rtac (map_comp_id RS trans), rtac map_congL,
+ EVERY' (map2 (fn Abs_inverse => fn coalg_final_set =>
+ EVERY' [rtac ballI, rtac (o_apply RS trans), rtac Abs_inverse,
+ etac set_rev_mp, rtac coalg_final_set, rtac Rep])
+ Abs_inverses (drop m coalg_final_sets))])
+ (Reps ~~ ((map_comp_ids ~~ map_congLs) ~~ coalg_final_setss))] 1;
+
+fun mk_mor_Abs_tac defs Abs_inverses {context = ctxt, prems = _} =
+ unfold_thms_tac ctxt defs THEN
+ EVERY' [rtac conjI,
+ CONJ_WRAP' (K (rtac ballI THEN' rtac UNIV_I)) Abs_inverses,
+ CONJ_WRAP' (fn thm => rtac ballI THEN' etac (thm RS arg_cong RS sym)) Abs_inverses] 1;
+
+fun mk_mor_unfold_tac m mor_UNIV dtor_defs unfold_defs Abs_inverses morEs map_comp_ids map_congs =
+ EVERY' [rtac iffD2, rtac mor_UNIV,
+ CONJ_WRAP' (fn ((Abs_inverse, morE), ((dtor_def, unfold_def), (map_comp_id, map_cong))) =>
+ EVERY' [rtac ext, rtac (o_apply RS trans RS sym), rtac (dtor_def RS trans),
+ rtac (unfold_def RS arg_cong RS trans), rtac (Abs_inverse RS arg_cong RS trans),
+ rtac (morE RS arg_cong RS trans), rtac (map_comp_id RS trans),
+ rtac (o_apply RS trans RS sym), rtac map_cong,
+ REPEAT_DETERM_N m o rtac refl,
+ EVERY' (map (fn thm => rtac (thm RS trans) THEN' rtac (o_apply RS sym)) unfold_defs)])
+ ((Abs_inverses ~~ morEs) ~~ ((dtor_defs ~~ unfold_defs) ~~ (map_comp_ids ~~ map_congs)))] 1;
+
+fun mk_raw_coind_tac bis_def bis_cong bis_O bis_converse bis_Gr tcoalg coalgT mor_T_final
+ sbis_lsbis lsbis_incls incl_lsbiss equiv_LSBISs mor_Rep Rep_injects =
+ let
+ val n = length Rep_injects;
+ in
+ EVERY' [rtac rev_mp, ftac (bis_def RS iffD1),
+ REPEAT_DETERM o etac conjE, rtac bis_cong, rtac bis_O, rtac bis_converse,
+ rtac bis_Gr, rtac tcoalg, rtac mor_Rep, rtac bis_O, atac, rtac bis_Gr, rtac tcoalg,
+ rtac mor_Rep, REPEAT_DETERM_N n o etac @{thm relImage_Gr},
+ rtac impI, rtac rev_mp, rtac bis_cong, rtac bis_O, rtac bis_Gr, rtac coalgT,
+ rtac mor_T_final, rtac bis_O, rtac sbis_lsbis, rtac bis_converse, rtac bis_Gr, rtac coalgT,
+ rtac mor_T_final, EVERY' (map (fn thm => rtac (thm RS @{thm relInvImage_Gr})) lsbis_incls),
+ rtac impI,
+ CONJ_WRAP' (fn (Rep_inject, (equiv_LSBIS , (incl_lsbis, lsbis_incl))) =>
+ EVERY' [rtac subset_trans, rtac @{thm relInvImage_UNIV_relImage}, rtac subset_trans,
+ rtac @{thm relInvImage_mono}, rtac subset_trans, etac incl_lsbis,
+ rtac ord_eq_le_trans, rtac @{thm sym[OF relImage_relInvImage]},
+ rtac @{thm xt1(3)}, rtac @{thm Sigma_cong},
+ rtac @{thm proj_image}, rtac @{thm proj_image}, rtac lsbis_incl,
+ rtac subset_trans, rtac @{thm relImage_mono}, rtac incl_lsbis, atac,
+ rtac @{thm relImage_proj}, rtac equiv_LSBIS, rtac @{thm relInvImage_diag},
+ rtac Rep_inject])
+ (Rep_injects ~~ (equiv_LSBISs ~~ (incl_lsbiss ~~ lsbis_incls)))] 1
+ end;
+
+fun mk_unique_mor_tac raw_coinds bis =
+ CONJ_WRAP' (fn raw_coind =>
+ EVERY' [rtac impI, rtac (bis RS raw_coind RS set_mp RS @{thm IdD}), atac,
+ etac conjunct1, atac, etac conjunct2, rtac @{thm image2_eqI}, rtac refl, rtac refl, atac])
+ raw_coinds 1;
+
+fun mk_unfold_unique_mor_tac raw_coinds bis mor unfold_defs =
+ CONJ_WRAP' (fn (raw_coind, unfold_def) =>
+ EVERY' [rtac ext, etac (bis RS raw_coind RS set_mp RS @{thm IdD}), rtac mor,
+ rtac @{thm image2_eqI}, rtac refl, rtac (unfold_def RS arg_cong RS trans),
+ rtac (o_apply RS sym), rtac UNIV_I]) (raw_coinds ~~ unfold_defs) 1;
+
+fun mk_dtor_o_ctor_tac ctor_def unfold map_comp_id map_congL unfold_o_dtors
+ {context = ctxt, prems = _} =
+ unfold_thms_tac ctxt [ctor_def] THEN EVERY' [rtac ext, rtac trans, rtac o_apply,
+ rtac trans, rtac unfold, rtac trans, rtac map_comp_id, rtac trans, rtac map_congL,
+ EVERY' (map (fn thm =>
+ rtac ballI THEN' rtac (trans OF [thm RS fun_cong, @{thm id_apply}])) unfold_o_dtors),
+ rtac sym, rtac @{thm id_apply}] 1;
+
+fun mk_corec_tac m corec_defs unfold map_cong corec_Inls {context = ctxt, prems = _} =
+ unfold_thms_tac ctxt corec_defs THEN EVERY' [rtac trans, rtac (o_apply RS arg_cong),
+ rtac trans, rtac unfold, fo_rtac (@{thm sum.cases(2)} RS arg_cong RS trans) ctxt, rtac map_cong,
+ REPEAT_DETERM_N m o rtac refl,
+ EVERY' (map (fn thm => rtac @{thm sum_case_expand_Inr} THEN' rtac thm) corec_Inls)] 1;
+
+fun mk_srel_coinduct_tac ks raw_coind bis_srel =
+ EVERY' [rtac rev_mp, rtac raw_coind, rtac ssubst, rtac bis_srel, rtac conjI,
+ CONJ_WRAP' (K (rtac @{thm ord_le_eq_trans[OF subset_UNIV UNIV_Times_UNIV[THEN sym]]})) ks,
+ CONJ_WRAP' (K (EVERY' [rtac allI, rtac allI, rtac impI,
+ REPEAT_DETERM o etac allE, etac mp, etac CollectE, etac @{thm splitD}])) ks,
+ rtac impI, REPEAT_DETERM o etac conjE,
+ CONJ_WRAP' (K (EVERY' [rtac impI, rtac @{thm IdD}, etac set_mp,
+ rtac CollectI, etac @{thm prod_caseI}])) ks] 1;
+
+fun mk_srel_strong_coinduct_tac m cTs cts srel_coinduct srel_monos srel_Ids =
+ EVERY' [rtac rev_mp, rtac (Drule.instantiate' cTs cts srel_coinduct),
+ EVERY' (map2 (fn srel_mono => fn srel_Id =>
+ EVERY' [REPEAT_DETERM o resolve_tac [allI, impI], REPEAT_DETERM o etac allE,
+ etac disjE, etac mp, atac, hyp_subst_tac, rtac (srel_mono RS set_mp),
+ REPEAT_DETERM_N m o rtac @{thm subset_refl},
+ REPEAT_DETERM_N (length srel_monos) o rtac @{thm Id_subset},
+ rtac (srel_Id RS equalityD2 RS set_mp), rtac @{thm IdI}])
+ srel_monos srel_Ids),
+ rtac impI, REPEAT_DETERM o etac conjE,
+ CONJ_WRAP' (K (rtac impI THEN' etac mp THEN' etac disjI1)) srel_Ids] 1;
+
+fun mk_dtor_coinduct_tac m ks raw_coind bis_def =
+ let
+ val n = length ks;
+ in
+ EVERY' [rtac rev_mp, rtac raw_coind, rtac ssubst, rtac bis_def, rtac conjI,
+ CONJ_WRAP' (K (rtac @{thm ord_le_eq_trans[OF subset_UNIV UNIV_Times_UNIV[THEN sym]]})) ks,
+ CONJ_WRAP' (fn i => EVERY' [select_prem_tac n (dtac asm_rl) i, REPEAT_DETERM o rtac allI,
+ rtac impI, REPEAT_DETERM o dtac @{thm meta_spec}, etac CollectE, etac @{thm meta_impE},
+ atac, etac exE, etac conjE, etac conjE, rtac bexI, rtac conjI,
+ etac @{thm fst_conv[THEN subst]}, etac @{thm snd_conv[THEN subst]},
+ rtac CollectI, REPEAT_DETERM_N m o (rtac conjI THEN' rtac subset_UNIV),
+ CONJ_WRAP' (fn i' => EVERY' [rtac subsetI, rtac CollectI, dtac (mk_conjunctN n i'),
+ REPEAT_DETERM o etac allE, etac mp, rtac @{thm ssubst_mem[OF pair_collapse]}, atac])
+ ks])
+ ks,
+ rtac impI,
+ CONJ_WRAP' (fn i => EVERY' [rtac impI, dtac (mk_conjunctN n i),
+ rtac @{thm subst[OF pair_in_Id_conv]}, etac set_mp,
+ rtac CollectI, etac (refl RSN (2, @{thm subst_Pair}))]) ks] 1
+ end;
+
+fun mk_dtor_strong_coinduct_tac ks cTs cts dtor_coinduct bis_def bis_diag =
+ EVERY' [rtac rev_mp, rtac (Drule.instantiate' cTs cts dtor_coinduct),
+ EVERY' (map (fn i =>
+ EVERY' [etac disjE, REPEAT_DETERM o dtac @{thm meta_spec}, etac @{thm meta_mp},
+ atac, rtac rev_mp, rtac subst, rtac bis_def, rtac bis_diag,
+ rtac impI, dtac conjunct2, dtac (mk_conjunctN (length ks) i), REPEAT_DETERM o etac allE,
+ etac impE, etac @{thm diag_UNIV_I}, REPEAT_DETERM o eresolve_tac [bexE, conjE, CollectE],
+ rtac exI, rtac conjI, etac conjI, atac,
+ CONJ_WRAP' (K (EVERY' [REPEAT_DETERM o resolve_tac [allI, impI],
+ rtac disjI2, rtac @{thm diagE}, etac set_mp, atac])) ks])
+ ks),
+ rtac impI, REPEAT_DETERM o etac conjE,
+ CONJ_WRAP' (K (rtac impI THEN' etac mp THEN' etac disjI1)) ks] 1;
+
+fun mk_map_tac m n cT unfold map_comp' map_cong =
+ EVERY' [rtac ext, rtac (o_apply RS trans RS sym), rtac (o_apply RS trans RS sym),
+ rtac (unfold RS trans), rtac (Thm.permute_prems 0 1 (map_comp' RS box_equals)), rtac map_cong,
+ REPEAT_DETERM_N m o rtac (@{thm id_o} RS fun_cong),
+ REPEAT_DETERM_N n o rtac (@{thm o_id} RS fun_cong),
+ rtac (o_apply RS (Drule.instantiate' [cT] [] arg_cong) RS sym)] 1;
+
+fun mk_set_le_tac n hset_minimal set_hsets set_hset_hsetss =
+ EVERY' [rtac hset_minimal,
+ REPEAT_DETERM_N n o rtac @{thm Un_upper1},
+ REPEAT_DETERM_N n o
+ EVERY' (map3 (fn i => fn set_hset => fn set_hset_hsets =>
+ EVERY' [rtac subsetI, rtac @{thm UnI2}, rtac (mk_UnIN n i), etac @{thm UN_I},
+ etac UnE, etac set_hset, REPEAT_DETERM_N (n - 1) o etac UnE,
+ EVERY' (map (fn thm => EVERY' [etac @{thm UN_E}, etac thm, atac]) set_hset_hsets)])
+ (1 upto n) set_hsets set_hset_hsetss)] 1;
+
+fun mk_set_simp_tac n set_le set_incl_hset set_hset_incl_hsets =
+ EVERY' [rtac equalityI, rtac set_le, rtac @{thm Un_least}, rtac set_incl_hset,
+ REPEAT_DETERM_N (n - 1) o rtac @{thm Un_least},
+ EVERY' (map (fn thm => rtac @{thm UN_least} THEN' etac thm) set_hset_incl_hsets)] 1;
+
+fun mk_map_id_tac maps unfold_unique unfold_dtor =
+ EVERY' [rtac (unfold_unique RS trans), EVERY' (map (fn thm => rtac (thm RS sym)) maps),
+ rtac unfold_dtor] 1;
+
+fun mk_map_comp_tac m n maps map_comps map_congs unfold_unique =
+ EVERY' [rtac unfold_unique,
+ EVERY' (map3 (fn map_thm => fn map_comp => fn map_cong =>
+ EVERY' (map rtac
+ ([@{thm o_assoc} RS trans,
+ @{thm arg_cong2[of _ _ _ _ "op o"]} OF [map_comp RS sym, refl] RS trans,
+ @{thm o_assoc} RS trans RS sym,
+ @{thm arg_cong2[of _ _ _ _ "op o"]} OF [map_thm, refl] RS trans,
+ @{thm o_assoc} RS sym RS trans, map_thm RS arg_cong RS trans, @{thm o_assoc} RS trans,
+ @{thm arg_cong2[of _ _ _ _ "op o"]} OF [map_comp RS sym, refl] RS trans,
+ ext, o_apply RS trans, o_apply RS trans RS sym, map_cong] @
+ replicate m (@{thm id_o} RS fun_cong) @
+ replicate n (@{thm o_id} RS fun_cong))))
+ maps map_comps map_congs)] 1;
+
+fun mk_mcong_tac m coinduct_tac map_comp's map_simps map_congs set_naturalss set_hsetss
+ set_hset_hsetsss =
+ let
+ val n = length map_comp's;
+ val ks = 1 upto n;
+ in
+ EVERY' ([rtac rev_mp,
+ coinduct_tac] @
+ maps (fn (((((map_comp'_trans, map_simps_trans), map_cong), set_naturals), set_hsets),
+ set_hset_hsetss) =>
+ [REPEAT_DETERM o eresolve_tac [exE, conjE], hyp_subst_tac, rtac exI, rtac conjI, rtac conjI,
+ rtac map_comp'_trans, rtac sym, rtac map_simps_trans, rtac map_cong,
+ REPEAT_DETERM_N m o (rtac o_apply_trans_sym THEN' rtac @{thm id_apply}),
+ REPEAT_DETERM_N n o rtac fst_convol_fun_cong_sym,
+ rtac map_comp'_trans, rtac sym, rtac map_simps_trans, rtac map_cong,
+ EVERY' (maps (fn set_hset =>
+ [rtac o_apply_trans_sym, rtac (@{thm id_apply} RS trans), etac CollectE,
+ REPEAT_DETERM o etac conjE, etac bspec, etac set_hset]) set_hsets),
+ REPEAT_DETERM_N n o rtac snd_convol_fun_cong_sym,
+ CONJ_WRAP' (fn (set_natural, set_hset_hsets) =>
+ EVERY' [REPEAT_DETERM o rtac allI, rtac impI, rtac @{thm image_convolD},
+ etac set_rev_mp, rtac ord_eq_le_trans, rtac set_natural,
+ rtac @{thm image_mono}, rtac subsetI, rtac CollectI, etac CollectE,
+ REPEAT_DETERM o etac conjE,
+ CONJ_WRAP' (fn set_hset_hset =>
+ EVERY' [rtac ballI, etac bspec, etac set_hset_hset, atac]) set_hset_hsets])
+ (drop m set_naturals ~~ set_hset_hsetss)])
+ (map (fn th => th RS trans) map_comp's ~~ map (fn th => th RS trans) map_simps ~~
+ map_congs ~~ set_naturalss ~~ set_hsetss ~~ set_hset_hsetsss) @
+ [rtac impI,
+ CONJ_WRAP' (fn k =>
+ EVERY' [rtac impI, dtac (mk_conjunctN n k), etac mp, rtac exI, rtac conjI, etac CollectI,
+ rtac conjI, rtac refl, rtac refl]) ks]) 1
+ end
+
+fun mk_map_unique_tac unfold_unique map_comps {context = ctxt, prems = _} =
+ rtac unfold_unique 1 THEN
+ unfold_thms_tac ctxt (map (fn thm => thm RS sym) map_comps @ @{thms o_assoc id_o o_id}) THEN
+ ALLGOALS (etac sym);
+
+fun mk_col_natural_tac cts rec_0s rec_Sucs map_simps set_naturalss
+ {context = ctxt, prems = _} =
+ let
+ val n = length map_simps;
+ in
+ EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
+ REPEAT_DETERM o rtac allI, SELECT_GOAL (unfold_thms_tac ctxt rec_0s),
+ CONJ_WRAP' (K (rtac @{thm image_empty})) rec_0s,
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn (rec_Suc, (map_simp, set_nats)) => EVERY'
+ [SELECT_GOAL (unfold_thms_tac ctxt
+ (rec_Suc :: map_simp :: set_nats @ @{thms image_Un image_UN UN_simps(10)})),
+ rtac @{thm Un_cong}, rtac refl,
+ CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_cong}))
+ (fn i => EVERY' [rtac @{thm UN_cong[OF refl]},
+ REPEAT_DETERM o etac allE, etac (mk_conjunctN n i)]) (n downto 1)])
+ (rec_Sucs ~~ (map_simps ~~ set_naturalss))] 1
+ end;
+
+fun mk_set_natural_tac hset_def col_natural =
+ EVERY' (map rtac [ext, (o_apply RS trans), (hset_def RS trans), sym,
+ (o_apply RS trans), (@{thm image_cong} OF [hset_def, refl] RS trans),
+ (@{thm image_UN} RS trans), (refl RS @{thm UN_cong}), col_natural]) 1;
+
+fun mk_col_bd_tac m j cts rec_0s rec_Sucs sbd_Card_order sbd_Cinfinite set_sbdss =
+ let
+ val n = length rec_0s;
+ in
+ EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn rec_0 => EVERY' (map rtac [ordIso_ordLeq_trans,
+ @{thm card_of_ordIso_subst}, rec_0, @{thm Card_order_empty}, sbd_Card_order])) rec_0s,
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn (rec_Suc, set_sbds) => EVERY'
+ [rtac ordIso_ordLeq_trans, rtac @{thm card_of_ordIso_subst}, rtac rec_Suc,
+ rtac (sbd_Cinfinite RSN (3, @{thm Un_Cinfinite_bound})), rtac (nth set_sbds (j - 1)),
+ REPEAT_DETERM_N (n - 1) o rtac (sbd_Cinfinite RSN (3, @{thm Un_Cinfinite_bound})),
+ EVERY' (map2 (fn i => fn set_sbd => EVERY' [rtac @{thm UNION_Cinfinite_bound},
+ rtac set_sbd, rtac ballI, REPEAT_DETERM o etac allE,
+ etac (mk_conjunctN n i), rtac sbd_Cinfinite]) (1 upto n) (drop m set_sbds))])
+ (rec_Sucs ~~ set_sbdss)] 1
+ end;
+
+fun mk_set_bd_tac sbd_Cinfinite hset_def col_bd =
+ EVERY' (map rtac [ordIso_ordLeq_trans, @{thm card_of_ordIso_subst}, hset_def,
+ ctrans, @{thm UNION_Cinfinite_bound}, ordIso_ordLeq_trans, @{thm card_of_nat},
+ @{thm natLeq_ordLeq_cinfinite}, sbd_Cinfinite, ballI, col_bd, sbd_Cinfinite,
+ ctrans, @{thm infinite_ordLeq_cexp}, sbd_Cinfinite, @{thm cexp_ordLeq_ccexp}]) 1;
+
+fun mk_in_bd_tac isNode_hset isNode_hsets carT_def card_of_carT mor_image Rep_inverse mor_hsets
+ sbd_Cnotzero sbd_Card_order mor_Rep coalgT mor_T_final tcoalg =
+ let
+ val n = length isNode_hsets;
+ val in_hin_tac = rtac CollectI THEN'
+ CONJ_WRAP' (fn mor_hset => EVERY' (map etac
+ [mor_hset OF [coalgT, mor_T_final] RS sym RS ord_eq_le_trans,
+ arg_cong RS sym RS ord_eq_le_trans,
+ mor_hset OF [tcoalg, mor_Rep, UNIV_I] RS ord_eq_le_trans])) mor_hsets;
+ in
+ EVERY' [rtac (Thm.permute_prems 0 1 @{thm ordLeq_transitive}), rtac ctrans,
+ rtac @{thm card_of_image}, rtac ordIso_ordLeq_trans,
+ rtac @{thm card_of_ordIso_subst}, rtac @{thm sym[OF proj_image]}, rtac ctrans,
+ rtac @{thm card_of_image}, rtac ctrans, rtac card_of_carT, rtac @{thm cexp_mono2_Cnotzero},
+ rtac @{thm cexp_ordLeq_ccexp}, rtac @{thm csum_Cnotzero2}, rtac @{thm ctwo_Cnotzero},
+ rtac @{thm Cnotzero_cexp}, rtac sbd_Cnotzero, rtac sbd_Card_order,
+ rtac @{thm card_of_mono1}, rtac subsetI, rtac @{thm image_eqI}, rtac sym,
+ rtac Rep_inverse, REPEAT_DETERM o eresolve_tac [CollectE, conjE],
+ rtac set_mp, rtac equalityD2, rtac @{thm sym[OF proj_image]}, rtac imageE,
+ rtac set_rev_mp, rtac mor_image, rtac mor_Rep, rtac UNIV_I, rtac equalityD2,
+ rtac @{thm proj_image}, rtac @{thm image_eqI}, atac,
+ ftac (carT_def RS equalityD1 RS set_mp),
+ REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac,
+ rtac (carT_def RS equalityD2 RS set_mp), rtac CollectI, REPEAT_DETERM o rtac exI,
+ rtac conjI, rtac refl, rtac conjI, etac conjI, etac conjI, etac conjI, rtac conjI,
+ rtac ballI, dtac bspec, atac, REPEAT_DETERM o etac conjE, rtac conjI,
+ CONJ_WRAP_GEN' (etac disjE) (fn (i, isNode_hset) =>
+ EVERY' [rtac (mk_disjIN n i), rtac isNode_hset, atac, atac, atac, in_hin_tac])
+ (1 upto n ~~ isNode_hsets),
+ CONJ_WRAP' (fn isNode_hset =>
+ EVERY' [rtac ballI, rtac isNode_hset, atac, ftac CollectD, etac @{thm SuccD},
+ etac bspec, atac, in_hin_tac])
+ isNode_hsets,
+ atac, rtac isNode_hset, atac, atac, atac, in_hin_tac] 1
+ end;
+
+fun mk_bd_card_order_tac sbd_card_order =
+ EVERY' (map rtac [@{thm card_order_ccexp}, sbd_card_order, sbd_card_order]) 1;
+
+fun mk_bd_cinfinite_tac sbd_Cinfinite =
+ EVERY' (map rtac [@{thm cinfinite_ccexp}, @{thm ctwo_ordLeq_Cinfinite},
+ sbd_Cinfinite, sbd_Cinfinite]) 1;
+
+fun mk_pickWP_assms_tac set_incl_hsets set_incl_hins map_eq =
+ let
+ val m = length set_incl_hsets;
+ in
+ EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac CollectI,
+ EVERY' (map (fn thm => rtac conjI THEN' etac (thm RS @{thm subset_trans})) set_incl_hsets),
+ CONJ_WRAP' (fn thm => rtac thm THEN' REPEAT_DETERM_N m o atac) set_incl_hins,
+ REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac CollectI,
+ EVERY' (map (fn thm => rtac conjI THEN' etac (thm RS @{thm subset_trans})) set_incl_hsets),
+ CONJ_WRAP' (fn thm => rtac thm THEN' REPEAT_DETERM_N m o atac) set_incl_hins,
+ REPEAT_DETERM o eresolve_tac [CollectE, conjE], etac map_eq]
+ end;
+
+fun mk_coalg_thePull_tac m coalg_def map_wpulls set_naturalss pickWP_assms_tacs
+ {context = ctxt, prems = _} =
+ unfold_thms_tac ctxt [coalg_def] THEN
+ CONJ_WRAP' (fn (map_wpull, (pickWP_assms_tac, set_naturals)) =>
+ EVERY' [rtac ballI, dtac @{thm set_mp[OF equalityD1[OF thePull_def]]},
+ REPEAT_DETERM o eresolve_tac [CollectE, @{thm prod_caseE}],
+ hyp_subst_tac, rtac rev_mp, rtac (map_wpull RS @{thm pickWP(1)}),
+ EVERY' (map (etac o mk_conjunctN m) (1 upto m)),
+ pickWP_assms_tac,
+ SELECT_GOAL (unfold_thms_tac ctxt @{thms o_apply prod.cases}), rtac impI,
+ REPEAT_DETERM o eresolve_tac [CollectE, conjE],
+ rtac CollectI,
+ REPEAT_DETERM_N m o (rtac conjI THEN' rtac subset_UNIV),
+ CONJ_WRAP' (fn set_natural =>
+ EVERY' [rtac ord_eq_le_trans, rtac trans, rtac set_natural,
+ rtac trans_fun_cong_image_id_id_apply, atac])
+ (drop m set_naturals)])
+ (map_wpulls ~~ (pickWP_assms_tacs ~~ set_naturalss)) 1;
+
+fun mk_mor_thePull_nth_tac conv pick m mor_def map_wpulls map_comps pickWP_assms_tacs
+ {context = ctxt, prems = _} =
+ let
+ val n = length map_comps;
+ in
+ unfold_thms_tac ctxt [mor_def] THEN
+ EVERY' [rtac conjI,
+ CONJ_WRAP' (K (rtac ballI THEN' rtac UNIV_I)) (1 upto n),
+ CONJ_WRAP' (fn (map_wpull, (pickWP_assms_tac, map_comp)) =>
+ EVERY' [rtac ballI, dtac @{thm set_mp[OF equalityD1[OF thePull_def]]},
+ REPEAT_DETERM o eresolve_tac [CollectE, @{thm prod_caseE}, conjE],
+ hyp_subst_tac,
+ SELECT_GOAL (unfold_thms_tac ctxt @{thms o_apply prod.cases}),
+ rtac (map_comp RS trans),
+ SELECT_GOAL (unfold_thms_tac ctxt (conv :: @{thms o_id id_o})),
+ rtac (map_wpull RS pick), REPEAT_DETERM_N m o atac,
+ pickWP_assms_tac])
+ (map_wpulls ~~ (pickWP_assms_tacs ~~ map_comps))] 1
+ end;
+
+val mk_mor_thePull_fst_tac = mk_mor_thePull_nth_tac @{thm fst_conv} @{thm pickWP(2)};
+val mk_mor_thePull_snd_tac = mk_mor_thePull_nth_tac @{thm snd_conv} @{thm pickWP(3)};
+
+fun mk_mor_thePull_pick_tac mor_def unfolds map_comps {context = ctxt, prems = _} =
+ unfold_thms_tac ctxt [mor_def, @{thm thePull_def}] THEN rtac conjI 1 THEN
+ CONJ_WRAP' (K (rtac ballI THEN' rtac UNIV_I)) unfolds 1 THEN
+ CONJ_WRAP' (fn (unfold, map_comp) =>
+ EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, @{thm prod_caseE}, conjE],
+ hyp_subst_tac,
+ SELECT_GOAL (unfold_thms_tac ctxt (unfold :: map_comp :: @{thms comp_def id_def})),
+ rtac refl])
+ (unfolds ~~ map_comps) 1;
+
+fun mk_pick_col_tac m j cts rec_0s rec_Sucs unfolds set_naturalss map_wpulls pickWP_assms_tacs
+ {context = ctxt, prems = _} =
+ let
+ val n = length rec_0s;
+ val ks = n downto 1;
+ in
+ EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn thm => EVERY'
+ [rtac impI, rtac ord_eq_le_trans, rtac thm, rtac @{thm empty_subsetI}]) rec_0s,
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn (rec_Suc, ((unfold, set_naturals), (map_wpull, pickWP_assms_tac))) =>
+ EVERY' [rtac impI, dtac @{thm set_mp[OF equalityD1[OF thePull_def]]},
+ REPEAT_DETERM o eresolve_tac [CollectE, @{thm prod_caseE}],
+ hyp_subst_tac, rtac rev_mp, rtac (map_wpull RS @{thm pickWP(1)}),
+ EVERY' (map (etac o mk_conjunctN m) (1 upto m)),
+ pickWP_assms_tac,
+ rtac impI, REPEAT_DETERM o eresolve_tac [CollectE, conjE],
+ rtac ord_eq_le_trans, rtac rec_Suc,
+ rtac @{thm Un_least},
+ SELECT_GOAL (unfold_thms_tac ctxt [unfold, nth set_naturals (j - 1),
+ @{thm prod.cases}]),
+ etac ord_eq_le_trans_trans_fun_cong_image_id_id_apply,
+ CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_least})) (fn (i, set_natural) =>
+ EVERY' [rtac @{thm UN_least},
+ SELECT_GOAL (unfold_thms_tac ctxt [unfold, set_natural, @{thm prod.cases}]),
+ etac imageE, hyp_subst_tac, REPEAT_DETERM o etac allE,
+ dtac (mk_conjunctN n i), etac mp, etac set_mp, atac])
+ (ks ~~ rev (drop m set_naturals))])
+ (rec_Sucs ~~ ((unfolds ~~ set_naturalss) ~~ (map_wpulls ~~ pickWP_assms_tacs)))] 1
+ end;
+
+fun mk_wpull_tac m coalg_thePull mor_thePull_fst mor_thePull_snd mor_thePull_pick
+ mor_unique pick_cols hset_defs =
+ EVERY' [rtac (@{thm wpull_def} RS iffD2), REPEAT_DETERM o rtac allI, rtac impI,
+ REPEAT_DETERM o etac conjE, rtac bexI, rtac conjI,
+ rtac box_equals, rtac mor_unique,
+ rtac coalg_thePull, REPEAT_DETERM_N (m - 1) o etac conjI, atac,
+ rtac mor_thePull_pick, REPEAT_DETERM_N (m - 1) o etac conjI, atac,
+ rtac mor_thePull_fst, REPEAT_DETERM_N (m - 1) o etac conjI, atac,
+ rtac @{thm set_mp[OF equalityD2[OF thePull_def]]}, rtac CollectI,
+ rtac @{thm prod_caseI}, etac conjI, etac conjI, atac, rtac o_apply, rtac @{thm fst_conv},
+ rtac box_equals, rtac mor_unique,
+ rtac coalg_thePull, REPEAT_DETERM_N (m - 1) o etac conjI, atac,
+ rtac mor_thePull_pick, REPEAT_DETERM_N (m - 1) o etac conjI, atac,
+ rtac mor_thePull_snd, REPEAT_DETERM_N (m - 1) o etac conjI, atac,
+ rtac @{thm set_mp[OF equalityD2[OF thePull_def]]}, rtac CollectI,
+ rtac @{thm prod_caseI}, etac conjI, etac conjI, atac, rtac o_apply, rtac @{thm snd_conv},
+ rtac CollectI,
+ CONJ_WRAP' (fn (pick, def) =>
+ EVERY' [rtac (def RS ord_eq_le_trans), rtac @{thm UN_least},
+ rtac pick, REPEAT_DETERM_N (m - 1) o etac conjI, atac,
+ rtac @{thm set_mp[OF equalityD2[OF thePull_def]]}, rtac CollectI,
+ rtac @{thm prod_caseI}, etac conjI, etac conjI, atac])
+ (pick_cols ~~ hset_defs)] 1;
+
+fun mk_wit_tac n dtor_ctors set_simp wit coind_wits {context = ctxt, prems = _} =
+ ALLGOALS (TRY o (eresolve_tac coind_wits THEN' rtac refl)) THEN
+ REPEAT_DETERM (atac 1 ORELSE
+ EVERY' [dtac set_rev_mp, rtac equalityD1, resolve_tac set_simp,
+ K (unfold_thms_tac ctxt dtor_ctors),
+ REPEAT_DETERM_N n o etac UnE,
+ REPEAT_DETERM o
+ (TRY o REPEAT_DETERM o etac UnE THEN' TRY o etac @{thm UN_E} THEN'
+ (eresolve_tac wit ORELSE'
+ (dresolve_tac wit THEN'
+ (etac FalseE ORELSE'
+ EVERY' [hyp_subst_tac, dtac set_rev_mp, rtac equalityD1, resolve_tac set_simp,
+ K (unfold_thms_tac ctxt dtor_ctors), REPEAT_DETERM_N n o etac UnE]))))] 1);
+
+fun mk_coind_wit_tac induct unfolds set_nats wits {context = ctxt, prems = _} =
+ rtac induct 1 THEN ALLGOALS (TRY o rtac impI THEN' TRY o hyp_subst_tac) THEN
+ unfold_thms_tac ctxt (unfolds @ set_nats @ @{thms image_id id_apply}) THEN
+ ALLGOALS (REPEAT_DETERM o etac imageE THEN' TRY o hyp_subst_tac) THEN
+ ALLGOALS (TRY o
+ FIRST' [rtac TrueI, rtac refl, etac (refl RSN (2, mp)), dresolve_tac wits THEN' etac FalseE])
+
+fun mk_srel_simp_tac in_Jsrels i in_srel map_comp map_cong map_simp set_simps dtor_inject dtor_ctor
+ set_naturals set_incls set_set_inclss =
+ let
+ val m = length set_incls;
+ val n = length set_set_inclss;
+ val (passive_set_naturals, active_set_naturals) = chop m set_naturals;
+ val in_Jsrel = nth in_Jsrels (i - 1);
+ val if_tac =
+ EVERY' [dtac (in_Jsrel RS iffD1), REPEAT_DETERM o eresolve_tac [exE, conjE, CollectE],
+ rtac (in_srel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
+ EVERY' (map2 (fn set_natural => fn set_incl =>
+ EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac set_natural,
+ rtac ord_eq_le_trans, rtac trans_fun_cong_image_id_id_apply,
+ etac (set_incl RS @{thm subset_trans})])
+ passive_set_naturals set_incls),
+ CONJ_WRAP' (fn (in_Jsrel, (set_natural, set_set_incls)) =>
+ EVERY' [rtac ord_eq_le_trans, rtac set_natural, rtac @{thm image_subsetI},
+ rtac (in_Jsrel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
+ CONJ_WRAP' (fn thm => etac (thm RS @{thm subset_trans}) THEN' atac) set_set_incls,
+ rtac conjI, rtac refl, rtac refl])
+ (in_Jsrels ~~ (active_set_naturals ~~ set_set_inclss)),
+ CONJ_WRAP' (fn conv =>
+ EVERY' [rtac trans, rtac map_comp, rtac trans, rtac map_cong,
+ REPEAT_DETERM_N m o rtac @{thm fun_cong[OF o_id]},
+ REPEAT_DETERM_N n o EVERY' (map rtac [trans, o_apply, conv]),
+ rtac trans, rtac sym, rtac map_simp, rtac (dtor_inject RS iffD2), atac])
+ @{thms fst_conv snd_conv}];
+ val only_if_tac =
+ EVERY' [dtac (in_srel RS iffD1), REPEAT_DETERM o eresolve_tac [exE, conjE, CollectE],
+ rtac (in_Jsrel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
+ CONJ_WRAP' (fn (set_simp, passive_set_natural) =>
+ EVERY' [rtac ord_eq_le_trans, rtac set_simp, rtac @{thm Un_least},
+ rtac ord_eq_le_trans, rtac box_equals, rtac passive_set_natural,
+ rtac (dtor_ctor RS sym RS arg_cong), rtac trans_fun_cong_image_id_id_apply, atac,
+ CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_least}))
+ (fn (active_set_natural, in_Jsrel) => EVERY' [rtac ord_eq_le_trans,
+ rtac @{thm UN_cong[OF _ refl]}, rtac @{thm box_equals[OF _ _ refl]},
+ rtac active_set_natural, rtac (dtor_ctor RS sym RS arg_cong), rtac @{thm UN_least},
+ dtac set_rev_mp, etac @{thm image_mono}, etac imageE,
+ dtac @{thm ssubst_mem[OF pair_collapse]}, dtac (in_Jsrel RS iffD1),
+ dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE,
+ dtac (Thm.permute_prems 0 1 @{thm ssubst_mem}), atac,
+ hyp_subst_tac, REPEAT_DETERM o eresolve_tac [CollectE, conjE], atac])
+ (rev (active_set_naturals ~~ in_Jsrels))])
+ (set_simps ~~ passive_set_naturals),
+ rtac conjI,
+ REPEAT_DETERM_N 2 o EVERY'[rtac (dtor_inject RS iffD1), rtac trans, rtac map_simp,
+ rtac box_equals, rtac map_comp, rtac (dtor_ctor RS sym RS arg_cong), rtac trans,
+ rtac map_cong, REPEAT_DETERM_N m o rtac @{thm fun_cong[OF o_id]},
+ EVERY' (map (fn in_Jsrel => EVERY' [rtac trans, rtac o_apply, dtac set_rev_mp, atac,
+ dtac @{thm ssubst_mem[OF pair_collapse]}, dtac (in_Jsrel RS iffD1),
+ dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE, atac]) in_Jsrels),
+ atac]]
+ in
+ EVERY' [rtac iffI, if_tac, only_if_tac] 1
+ end;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Tools/bnf_gfp_util.ML Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,197 @@
+(* Title: HOL/BNF/Tools/bnf_gfp_util.ML
+ Author: Dmitriy Traytel, TU Muenchen
+ Copyright 2012
+
+Library for the codatatype construction.
+*)
+
+signature BNF_GFP_UTIL =
+sig
+ val mk_rec_simps: int -> thm -> thm list -> thm list list
+
+ val dest_listT: typ -> typ
+
+ val mk_Cons: term -> term -> term
+ val mk_Shift: term -> term -> term
+ val mk_Succ: term -> term -> term
+ val mk_Times: term * term -> term
+ val mk_append: term * term -> term
+ val mk_congruent: term -> term -> term
+ val mk_clists: term -> term
+ val mk_diag: term -> term
+ val mk_equiv: term -> term -> term
+ val mk_fromCard: term -> term -> term
+ val mk_list_rec: term -> term -> term
+ val mk_nat_rec: term -> term -> term
+ val mk_pickWP: term -> term -> term -> term -> term -> term
+ val mk_prefCl: term -> term
+ val mk_proj: term -> term
+ val mk_quotient: term -> term -> term
+ val mk_shift: term -> term -> term
+ val mk_size: term -> term
+ val mk_thePull: term -> term -> term -> term -> term
+ val mk_toCard: term -> term -> term
+ val mk_undefined: typ -> term
+ val mk_univ: term -> term
+
+ val mk_specN: int -> thm -> thm
+
+ val mk_InN_Field: int -> int -> thm
+ val mk_InN_inject: int -> int -> thm
+ val mk_InN_not_InM: int -> int -> thm
+end;
+
+structure BNF_GFP_Util : BNF_GFP_UTIL =
+struct
+
+open BNF_Util
+
+val mk_append = HOLogic.mk_binop @{const_name append};
+
+fun mk_equiv B R =
+ Const (@{const_name equiv}, fastype_of B --> fastype_of R --> HOLogic.boolT) $ B $ R;
+
+fun mk_Sigma (A, B) =
+ let
+ val AT = fastype_of A;
+ val BT = fastype_of B;
+ val ABT = mk_relT (HOLogic.dest_setT AT, HOLogic.dest_setT (range_type BT));
+ in Const (@{const_name Sigma}, AT --> BT --> ABT) $ A $ B end;
+
+fun mk_diag A =
+ let
+ val AT = fastype_of A;
+ val AAT = mk_relT (HOLogic.dest_setT AT, HOLogic.dest_setT AT);
+ in Const (@{const_name diag}, AT --> AAT) $ A end;
+
+fun mk_Times (A, B) =
+ let val AT = HOLogic.dest_setT (fastype_of A);
+ in mk_Sigma (A, Term.absdummy AT B) end;
+
+fun dest_listT (Type (@{type_name list}, [T])) = T
+ | dest_listT T = raise TYPE ("dest_setT: set type expected", [T], []);
+
+fun mk_Succ Kl kl =
+ let val T = fastype_of kl;
+ in
+ Const (@{const_name Succ},
+ HOLogic.mk_setT T --> T --> HOLogic.mk_setT (dest_listT T)) $ Kl $ kl
+ end;
+
+fun mk_Shift Kl k =
+ let val T = fastype_of Kl;
+ in
+ Const (@{const_name Shift}, T --> dest_listT (HOLogic.dest_setT T) --> T) $ Kl $ k
+ end;
+
+fun mk_shift lab k =
+ let val T = fastype_of lab;
+ in
+ Const (@{const_name shift}, T --> dest_listT (Term.domain_type T) --> T) $ lab $ k
+ end;
+
+fun mk_prefCl A =
+ Const (@{const_name prefCl}, fastype_of A --> HOLogic.boolT) $ A;
+
+fun mk_clists r =
+ let val T = fastype_of r;
+ in Const (@{const_name clists}, T --> mk_relT (`I (HOLogic.listT (fst (dest_relT T))))) $ r end;
+
+fun mk_toCard A r =
+ let
+ val AT = fastype_of A;
+ val rT = fastype_of r;
+ in
+ Const (@{const_name toCard},
+ AT --> rT --> HOLogic.dest_setT AT --> fst (dest_relT rT)) $ A $ r
+ end;
+
+fun mk_fromCard A r =
+ let
+ val AT = fastype_of A;
+ val rT = fastype_of r;
+ in
+ Const (@{const_name fromCard},
+ AT --> rT --> fst (dest_relT rT) --> HOLogic.dest_setT AT) $ A $ r
+ end;
+
+fun mk_Cons x xs =
+ let val T = fastype_of xs;
+ in Const (@{const_name Cons}, dest_listT T --> T --> T) $ x $ xs end;
+
+fun mk_size t = HOLogic.size_const (fastype_of t) $ t;
+
+fun mk_quotient A R =
+ let val T = fastype_of A;
+ in Const (@{const_name quotient}, T --> fastype_of R --> HOLogic.mk_setT T) $ A $ R end;
+
+fun mk_proj R =
+ let val ((AT, BT), T) = `dest_relT (fastype_of R);
+ in Const (@{const_name proj}, T --> AT --> HOLogic.mk_setT BT) $ R end;
+
+fun mk_univ f =
+ let val ((AT, BT), T) = `dest_funT (fastype_of f);
+ in Const (@{const_name univ}, T --> HOLogic.mk_setT AT --> BT) $ f end;
+
+fun mk_congruent R f =
+ Const (@{const_name congruent}, fastype_of R --> fastype_of f --> HOLogic.boolT) $ R $ f;
+
+fun mk_undefined T = Const (@{const_name undefined}, T);
+
+fun mk_nat_rec Zero Suc =
+ let val T = fastype_of Zero;
+ in Const (@{const_name nat_rec}, T --> fastype_of Suc --> HOLogic.natT --> T) $ Zero $ Suc end;
+
+fun mk_list_rec Nil Cons =
+ let
+ val T = fastype_of Nil;
+ val (U, consT) = `(Term.domain_type) (fastype_of Cons);
+ in
+ Const (@{const_name list_rec}, T --> consT --> HOLogic.listT U --> T) $ Nil $ Cons
+ end;
+
+fun mk_thePull B1 B2 f1 f2 =
+ let
+ val fT1 = fastype_of f1;
+ val fT2 = fastype_of f2;
+ val BT1 = domain_type fT1;
+ val BT2 = domain_type fT2;
+ in
+ Const (@{const_name thePull}, HOLogic.mk_setT BT1 --> HOLogic.mk_setT BT2 --> fT1 --> fT2 -->
+ mk_relT (BT1, BT2)) $ B1 $ B2 $ f1 $ f2
+ end;
+
+fun mk_pickWP A f1 f2 b1 b2 =
+ let
+ val fT1 = fastype_of f1;
+ val fT2 = fastype_of f2;
+ val AT = domain_type fT1;
+ val BT1 = range_type fT1;
+ val BT2 = range_type fT2;
+ in
+ Const (@{const_name pickWP}, HOLogic.mk_setT AT --> fT1 --> fT2 --> BT1 --> BT2 --> AT) $
+ A $ f1 $ f2 $ b1 $ b2
+ end;
+
+fun mk_InN_not_InM 1 _ = @{thm Inl_not_Inr}
+ | mk_InN_not_InM n m =
+ if n > m then mk_InN_not_InM m n RS @{thm not_sym}
+ else mk_InN_not_InM (n - 1) (m - 1) RS @{thm not_arg_cong_Inr};
+
+fun mk_InN_Field 1 1 = @{thm TrueE[OF TrueI]}
+ | mk_InN_Field _ 1 = @{thm Inl_Field_csum}
+ | mk_InN_Field 2 2 = @{thm Inr_Field_csum}
+ | mk_InN_Field n m = mk_InN_Field (n - 1) (m - 1) RS @{thm Inr_Field_csum};
+
+fun mk_InN_inject 1 _ = @{thm TrueE[OF TrueI]}
+ | mk_InN_inject _ 1 = @{thm Inl_inject}
+ | mk_InN_inject 2 2 = @{thm Inr_inject}
+ | mk_InN_inject n m = @{thm Inr_inject} RS mk_InN_inject (n - 1) (m - 1);
+
+fun mk_specN 0 thm = thm
+ | mk_specN n thm = mk_specN (n - 1) (thm RS spec);
+
+fun mk_rec_simps n rec_thm defs = map (fn i =>
+ map (fn def => def RS rec_thm RS mk_nthI n i RS fun_cong) defs) (1 upto n);
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Tools/bnf_lfp.ML Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,1838 @@
+(* Title: HOL/BNF/Tools/bnf_lfp.ML
+ Author: Dmitriy Traytel, TU Muenchen
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Datatype construction.
+*)
+
+signature BNF_LFP =
+sig
+ val bnf_lfp: mixfix list -> (string * sort) list option -> binding list ->
+ typ list * typ list list -> BNF_Def.BNF list -> local_theory ->
+ (term list * term list * term list * term list * thm * thm list * thm list * thm list *
+ thm list * thm list) * local_theory
+end;
+
+structure BNF_LFP : BNF_LFP =
+struct
+
+open BNF_Def
+open BNF_Util
+open BNF_Tactics
+open BNF_FP
+open BNF_FP_Sugar
+open BNF_LFP_Util
+open BNF_LFP_Tactics
+
+(*all BNFs have the same lives*)
+fun bnf_lfp mixfixes resBs bs (resDs, Dss) bnfs lthy =
+ let
+ 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 b = Binding.name (mk_common_name (map Binding.name_of bs));
+
+ (* TODO: check if m, n, etc., are sane *)
+
+ val deads = fold (union (op =)) Dss resDs;
+ val names_lthy = fold Variable.declare_typ deads lthy;
+
+ (* tvars *)
+ val (((((((passiveAs, activeAs), allAs)), (passiveBs, activeBs)),
+ activeCs), passiveXs), passiveYs) = names_lthy
+ |> mk_TFrees live
+ |> apfst (`(chop m))
+ ||> mk_TFrees live
+ ||>> apfst (chop m)
+ ||>> mk_TFrees n
+ ||>> mk_TFrees m
+ ||> fst o mk_TFrees m;
+
+ 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';
+
+ (* typs *)
+ val dead_poss =
+ (case resBs of
+ NONE => map SOME deads @ replicate m NONE
+ | SOME Ts => map (fn T => if member (op =) deads (TFree T) then SOME (TFree T) else NONE) Ts);
+ 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 (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';
+ val prodBsAs = map2 (curry HOLogic.mk_prodT) activeBs activeAs;
+ val prodFTs = mk_FTs (passiveAs @ prodBsAs);
+ val prod_sTs = map2 (curry (op -->)) prodFTs activeAs;
+
+ (* terms *)
+ val mapsAsAs = map4 mk_map_of_bnf Dss Ass Ass bnfs;
+ val mapsAsBs = map4 mk_map_of_bnf Dss Ass Bss bnfs;
+ val mapsBsAs = map4 mk_map_of_bnf Dss Bss Ass bnfs;
+ val mapsBsCs' = map4 mk_map_of_bnf Dss Bss Css' bnfs;
+ val mapsAsCs' = map4 mk_map_of_bnf Dss Ass Css' bnfs;
+ val map_fsts = map4 mk_map_of_bnf Dss (replicate n (passiveAs @ prodBsAs)) Bss bnfs;
+ val map_fsts_rev = map4 mk_map_of_bnf Dss Bss (replicate n (passiveAs @ prodBsAs)) bnfs;
+ fun mk_setss Ts = map3 mk_sets_of_bnf (map (replicate live) Dss)
+ (map (replicate live) (replicate n Ts)) bnfs;
+ val setssAs = mk_setss allAs;
+ val bds = map3 mk_bd_of_bnf Dss Ass bnfs;
+ val witss = map wits_of_bnf bnfs;
+
+ val (((((((((((((((((((zs, zs'), As), Bs), Bs_copy), B's), B''s), ss), prod_ss), s's), s''s),
+ fs), fs_copy), inv_fs), self_fs), gs), all_gs), (xFs, xFs')), (yFs, yFs')),
+ names_lthy) = lthy
+ |> mk_Frees' "z" activeAs
+ ||>> mk_Frees "A" ATs
+ ||>> mk_Frees "B" BTs
+ ||>> mk_Frees "B" BTs
+ ||>> mk_Frees "B'" B'Ts
+ ||>> mk_Frees "B''" B''Ts
+ ||>> mk_Frees "s" sTs
+ ||>> mk_Frees "prods" prod_sTs
+ ||>> mk_Frees "s'" s'Ts
+ ||>> mk_Frees "s''" s''Ts
+ ||>> mk_Frees "f" fTs
+ ||>> mk_Frees "f" fTs
+ ||>> mk_Frees "f" inv_fTs
+ ||>> mk_Frees "f" self_fTs
+ ||>> mk_Frees "g" gTs
+ ||>> mk_Frees "g" all_gTs
+ ||>> mk_Frees' "x" FTsAs
+ ||>> mk_Frees' "y" FTsBs;
+
+ val passive_UNIVs = map HOLogic.mk_UNIV passiveAs;
+ val active_UNIVs = map HOLogic.mk_UNIV activeAs;
+ val prod_UNIVs = map HOLogic.mk_UNIV prodBsAs;
+ val passive_ids = map HOLogic.id_const passiveAs;
+ val active_ids = map HOLogic.id_const activeAs;
+ val fsts = map fst_const prodBsAs;
+
+ (* thms *)
+ val bd_card_orders = map bd_card_order_of_bnf bnfs;
+ val bd_Card_orders = map bd_Card_order_of_bnf bnfs;
+ val bd_Card_order = hd bd_Card_orders;
+ val bd_Cinfinite = bd_Cinfinite_of_bnf (hd bnfs);
+ val bd_Cnotzeros = map bd_Cnotzero_of_bnf bnfs;
+ val bd_Cnotzero = hd bd_Cnotzeros;
+ val in_bds = map in_bd_of_bnf bnfs;
+ val map_comp's = map map_comp'_of_bnf bnfs;
+ val map_congs = map map_cong_of_bnf bnfs;
+ val map_ids = map map_id_of_bnf bnfs;
+ val map_id's = map map_id'_of_bnf bnfs;
+ val map_wpulls = map map_wpull_of_bnf bnfs;
+ val set_bdss = map set_bd_of_bnf bnfs;
+ val set_natural'ss = map set_natural'_of_bnf bnfs;
+
+ val timer = time (timer "Extracted terms & thms");
+
+ (* nonemptiness check *)
+ fun new_wit X wit = 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 :: _ => error ("Cannot define empty datatype " ^ quote (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_comp =
+ 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;
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (x :: fs @ all_gs) (mk_Trueprop_eq (lhs, rhs)))
+ (K (mk_map_comp_id_tac map_comp))
+ |> Thm.close_derivation
+ end;
+
+ val map_comp_id_thms = map5 mk_map_comp_id xFs mapsAsBs mapsBsCs' mapsAsCs' map_comp's;
+
+ (*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_congL x mapAsAs sets map_cong 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 = map4 mk_prem (drop m sets) self_fs zs zs';
+ val goal = mk_Trueprop_eq (Term.list_comb (mapAsAs, passive_ids @ self_fs) $ x, x);
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (x :: self_fs) (Logic.list_implies (prems, goal)))
+ (K (mk_map_congL_tac m map_cong map_id'))
+ |> Thm.close_derivation
+ end;
+
+ val map_congL_thms = map5 mk_map_congL xFs mapsAsAs setssAs map_congs map_id's;
+ 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 = Binding.suffix_name ("_" ^ algN) b;
+ val alg_name = Binding.name_of alg_bind;
+ val alg_def_bind = (Thm.def_binding alg_bind, []);
+
+ (*forall i = 1 ... n: (\<forall>x \<in> Fi_in A1 .. Am B1 ... Bn. si x \<in> Bi)*)
+ val alg_spec =
+ let
+ val algT = Library.foldr (op -->) (ATs @ BTs @ sTs, HOLogic.boolT);
+
+ val ins = map3 mk_in (replicate n (As @ 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 lhs = Term.list_comb (Free (alg_name, algT), As @ Bs @ ss);
+ val rhs = Library.foldr1 HOLogic.mk_conj (map5 mk_alg_conjunct Bs ss ins xFs xFs')
+ in
+ mk_Trueprop_eq (lhs, rhs)
+ end;
+
+ val ((alg_free, (_, alg_def_free)), (lthy, lthy_old)) =
+ lthy
+ |> Specification.definition (SOME (alg_bind, NONE, NoSyn), (alg_def_bind, alg_spec))
+ ||> `Local_Theory.restore;
+
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+ val alg = fst (Term.dest_Const (Morphism.term phi alg_free));
+ val alg_def = Morphism.thm phi alg_def_free;
+
+ fun mk_alg As Bs ss =
+ let
+ val args = As @ 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 alg_set_thms =
+ let
+ val alg_prem = HOLogic.mk_Trueprop (mk_alg As Bs ss);
+ fun mk_prem x set B = HOLogic.mk_Trueprop (mk_subset (set $ x) B);
+ fun mk_concl s x B = HOLogic.mk_Trueprop (HOLogic.mk_mem (s $ x, B));
+ val premss = map2 ((fn x => fn sets => map2 (mk_prem x) sets (As @ Bs))) xFs setssAs;
+ val concls = map3 mk_concl ss xFs Bs;
+ val goals = map3 (fn x => fn prems => fn concl =>
+ fold_rev Logic.all (x :: As @ Bs @ ss)
+ (Logic.list_implies (alg_prem :: prems, concl))) xFs premss concls;
+ in
+ map (fn goal =>
+ Skip_Proof.prove lthy [] [] goal (K (mk_alg_set_tac alg_def)) |> Thm.close_derivation)
+ goals
+ end;
+
+ fun mk_talg ATs BTs = mk_alg (map HOLogic.mk_UNIV ATs) (map HOLogic.mk_UNIV BTs);
+
+ val talg_thm =
+ let
+ val goal = fold_rev Logic.all ss
+ (HOLogic.mk_Trueprop (mk_talg passiveAs activeAs ss))
+ in
+ Skip_Proof.prove lthy [] [] goal
+ (K (stac alg_def 1 THEN CONJ_WRAP (K (EVERY' [rtac ballI, rtac UNIV_I] 1)) ss))
+ |> Thm.close_derivation
+ end;
+
+ val timer = time (timer "Algebra definition & thms");
+
+ val alg_not_empty_thms =
+ let
+ val alg_prem =
+ HOLogic.mk_Trueprop (mk_alg passive_UNIVs Bs ss);
+ val concls = map (HOLogic.mk_Trueprop o mk_not_empty) Bs;
+ val goals =
+ map (fn concl =>
+ fold_rev Logic.all (Bs @ ss) (Logic.mk_implies (alg_prem, concl))) concls;
+ in
+ map2 (fn goal => fn alg_set =>
+ Skip_Proof.prove lthy [] []
+ goal (K (mk_alg_not_empty_tac alg_set alg_set_thms wit_thms))
+ |> Thm.close_derivation)
+ goals alg_set_thms
+ end;
+
+ val timer = time (timer "Proved nonemptiness");
+
+ (* morphism *)
+
+ val mor_bind = Binding.suffix_name ("_" ^ morN) b;
+ val mor_name = Binding.name_of mor_bind;
+ val mor_def_bind = (Thm.def_binding mor_bind, []);
+
+ (*fbetw) forall i = 1 ... n: (\<forall>x \<in> Bi. f x \<in> B'i)*)
+ (*mor) forall i = 1 ... n: (\<forall>x \<in> Fi_in UNIV ... UNIV B1 ... Bn.
+ f (s1 x) = s1' (Fi_map id ... id f1 ... fn x))*)
+ val mor_spec =
+ let
+ val morT = Library.foldr (op -->) (BTs @ sTs @ B'Ts @ s'Ts @ fTs, HOLogic.boolT);
+
+ 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 lhs = Term.list_comb (Free (mor_name, morT), Bs @ ss @ B's @ s's @ fs);
+ val rhs = HOLogic.mk_conj
+ (Library.foldr1 HOLogic.mk_conj (map5 mk_fbetw fs Bs B's zs zs'),
+ Library.foldr1 HOLogic.mk_conj
+ (map8 mk_mor setssAs mapsAsBs fs ss s's FTsAs xFs xFs'))
+ in
+ mk_Trueprop_eq (lhs, rhs)
+ end;
+
+ val ((mor_free, (_, mor_def_free)), (lthy, lthy_old)) =
+ lthy
+ |> Specification.definition (SOME (mor_bind, NONE, NoSyn), (mor_def_bind, mor_spec))
+ ||> `Local_Theory.restore;
+
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+ val mor = fst (Term.dest_Const (Morphism.term phi mor_free));
+ val mor_def = 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 (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 = fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs)
+ (Logic.mk_implies (prem, HOLogic.mk_Trueprop (mk_subset (mk_image f $ B1) B2)));
+ val image_goals = map3 mk_image_goal fs Bs B's;
+ 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 =
+ fold_rev Logic.all (x :: Bs @ ss @ B's @ s's @ fs)
+ (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 = map7 mk_elim_goal setssAs mapsAsBs fs ss s's xFs FTsAs;
+ fun prove goal =
+ Skip_Proof.prove lthy [] [] goal (K (mk_mor_elim_tac mor_def)) |> Thm.close_derivation;
+ in
+ (map prove image_goals, map prove elim_goals)
+ end;
+
+ val mor_incl_thm =
+ let
+ val prems = map2 (HOLogic.mk_Trueprop oo mk_subset) Bs Bs_copy;
+ val concl = HOLogic.mk_Trueprop (mk_mor Bs ss Bs_copy ss active_ids);
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (Bs @ ss @ Bs_copy) (Logic.list_implies (prems, concl)))
+ (K (mk_mor_incl_tac mor_def map_id's))
+ |> Thm.close_derivation
+ 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));
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (Bs @ ss @ B's @ s's @ B''s @ s''s @ fs @ gs)
+ (Logic.list_implies (prems, concl)))
+ (K (mk_mor_comp_tac mor_def set_natural'ss map_comp_id_thms))
+ |> Thm.close_derivation
+ end;
+
+ val mor_inv_thm =
+ let
+ fun mk_inv_prem f inv_f B B' = HOLogic.mk_conj (mk_subset (mk_image inv_f $ B') B,
+ HOLogic.mk_conj (mk_inver inv_f f B, mk_inver f inv_f B'));
+ val prems = map HOLogic.mk_Trueprop
+ ([mk_mor Bs ss B's s's fs,
+ mk_alg passive_UNIVs Bs ss,
+ mk_alg passive_UNIVs B's s's] @
+ map4 mk_inv_prem fs inv_fs Bs B's);
+ val concl = HOLogic.mk_Trueprop (mk_mor B's s's Bs ss inv_fs);
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs @ inv_fs)
+ (Logic.list_implies (prems, concl)))
+ (K (mk_mor_inv_tac alg_def mor_def
+ set_natural'ss morE_thms map_comp_id_thms map_congL_thms))
+ |> Thm.close_derivation
+ 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);
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs @ fs_copy)
+ (Logic.list_implies (prems, concl)))
+ (K ((hyp_subst_tac THEN' atac) 1))
+ |> Thm.close_derivation
+ 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;
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all ss (HOLogic.mk_Trueprop
+ (mk_mor (map HOLogic.mk_UNIV FTsAs) maps active_UNIVs ss ss)))
+ (K (mk_mor_str_tac ks mor_def))
+ |> Thm.close_derivation
+ end;
+
+ val mor_convol_thm =
+ let
+ val maps = map3 (fn s => fn prod_s => fn mapx =>
+ mk_convol (HOLogic.mk_comp (s, Term.list_comb (mapx, passive_ids @ fsts)), prod_s))
+ s's prod_ss map_fsts;
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (s's @ prod_ss) (HOLogic.mk_Trueprop
+ (mk_mor prod_UNIVs maps (map HOLogic.mk_UNIV activeBs) s's fsts)))
+ (K (mk_mor_convol_tac ks mor_def))
+ |> Thm.close_derivation
+ 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 (map4 mk_conjunct mapsAsBs fs ss s's);
+ in
+ Skip_Proof.prove lthy [] [] (fold_rev Logic.all (ss @ s's @ fs) (mk_Trueprop_eq (lhs, rhs)))
+ (K (mk_mor_UNIV_tac m morE_thms mor_def))
+ |> Thm.close_derivation
+ end;
+
+ val timer = time (timer "Morphism definition & thms");
+
+ (* isomorphism *)
+
+ (*mor Bs1 ss1 Bs2 ss2 fs \<and> (\<exists>gs. mor Bs2 ss2 Bs1 ss1 fs \<and>
+ forall i = 1 ... n. (inver gs[i] fs[i] Bs1[i] \<and> inver fs[i] gs[i] Bs2[i]))*)
+ fun mk_iso Bs1 ss1 Bs2 ss2 fs gs =
+ let
+ val ex_inv_mor = list_exists_free gs
+ (HOLogic.mk_conj (mk_mor Bs2 ss2 Bs1 ss1 gs,
+ Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_conj)
+ (map3 mk_inver gs fs Bs1) (map3 mk_inver fs gs Bs2))));
+ in
+ HOLogic.mk_conj (mk_mor Bs1 ss1 Bs2 ss2 fs, ex_inv_mor)
+ end;
+
+ val iso_alt_thm =
+ let
+ val prems = map HOLogic.mk_Trueprop
+ [mk_alg passive_UNIVs Bs ss,
+ mk_alg passive_UNIVs B's s's]
+ val concl = mk_Trueprop_eq (mk_iso Bs ss B's s's fs inv_fs,
+ HOLogic.mk_conj (mk_mor Bs ss B's s's fs,
+ Library.foldr1 HOLogic.mk_conj (map3 mk_bij_betw fs Bs B's)));
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs) (Logic.list_implies (prems, concl)))
+ (K (mk_iso_alt_tac mor_image_thms mor_inv_thm))
+ |> Thm.close_derivation
+ end;
+
+ val timer = time (timer "Isomorphism definition & thms");
+
+ (* algebra copies *)
+
+ val (copy_alg_thm, ex_copy_alg_thm) =
+ let
+ val prems = map HOLogic.mk_Trueprop
+ (mk_alg passive_UNIVs Bs ss :: map3 mk_bij_betw inv_fs B's Bs);
+ val inver_prems = map HOLogic.mk_Trueprop
+ (map3 mk_inver inv_fs fs Bs @ map3 mk_inver fs inv_fs B's);
+ val all_prems = prems @ inver_prems;
+ fun mk_s f s mapT y y' = Term.absfree y' (f $ (s $
+ (Term.list_comb (mapT, passive_ids @ inv_fs) $ y)));
+
+ val alg = HOLogic.mk_Trueprop
+ (mk_alg passive_UNIVs B's (map5 mk_s fs ss mapsBsAs yFs yFs'));
+ val copy_str_thm = Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (Bs @ ss @ B's @ inv_fs @ fs)
+ (Logic.list_implies (all_prems, alg)))
+ (K (mk_copy_str_tac set_natural'ss alg_def alg_set_thms))
+ |> Thm.close_derivation;
+
+ val iso = HOLogic.mk_Trueprop
+ (mk_iso B's (map5 mk_s fs ss mapsBsAs yFs yFs') Bs ss inv_fs fs_copy);
+ val copy_alg_thm = Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (Bs @ ss @ B's @ inv_fs @ fs)
+ (Logic.list_implies (all_prems, iso)))
+ (K (mk_copy_alg_tac set_natural'ss alg_set_thms mor_def iso_alt_thm copy_str_thm))
+ |> Thm.close_derivation;
+
+ val ex = HOLogic.mk_Trueprop
+ (list_exists_free s's
+ (HOLogic.mk_conj (mk_alg passive_UNIVs B's s's,
+ mk_iso B's s's Bs ss inv_fs fs_copy)));
+ val ex_copy_alg_thm = Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (Bs @ ss @ B's @ inv_fs @ fs)
+ (Logic.list_implies (prems, ex)))
+ (K (mk_ex_copy_alg_tac n copy_str_thm copy_alg_thm))
+ |> Thm.close_derivation;
+ in
+ (copy_alg_thm, ex_copy_alg_thm)
+ end;
+
+ val timer = time (timer "Copy thms");
+
+
+ (* bounds *)
+
+ val sum_Card_order = if n = 1 then bd_Card_order else @{thm Card_order_csum};
+ val sum_Cnotzero = if n = 1 then bd_Cnotzero else bd_Cnotzero RS @{thm csum_Cnotzero1};
+ val sum_Cinfinite = if n = 1 then bd_Cinfinite else bd_Cinfinite RS @{thm Cinfinite_csum1};
+ fun mk_set_bd_sums i bd_Card_order bds =
+ if n = 1 then bds
+ else map (fn thm => bd_Card_order RS mk_ordLeq_csum n i thm) bds;
+ val set_bd_sumss = map3 mk_set_bd_sums ks bd_Card_orders set_bdss;
+
+ fun mk_in_bd_sum i Co Cnz bd =
+ if n = 1 then bd
+ else Cnz RS ((Co RS mk_ordLeq_csum n i (Co RS @{thm ordLeq_refl})) RS
+ (bd RS @{thm ordLeq_transitive[OF _
+ cexp_mono2_Cnotzero[OF _ csum_Cnotzero2[OF ctwo_Cnotzero]]]}));
+ val in_bd_sums = map4 mk_in_bd_sum ks bd_Card_orders bd_Cnotzeros in_bds;
+
+ val sum_bd = Library.foldr1 (uncurry mk_csum) bds;
+ val suc_bd = mk_cardSuc sum_bd;
+ 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 = if n = 1 then bd_Card_order RS @{thm cardSuc_Card_order}
+ else @{thm cardSuc_Card_order[OF Card_order_csum]};
+ val suc_bd_Cinfinite = if n = 1 then bd_Cinfinite RS @{thm Cinfinite_cardSuc}
+ else bd_Cinfinite RS @{thm Cinfinite_cardSuc[OF Cinfinite_csum1]};
+ 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 Asuc_bd_Cnotzero = Asuc_bd_Cinfinite RS @{thm Cinfinite_Cnotzero};
+
+ val suc_bd_Asuc_bd = @{thm ordLess_ordLeq_trans[OF
+ ordLess_ctwo_cexp
+ cexp_mono1_Cnotzero[OF _ ctwo_Cnotzero]]} OF
+ [suc_bd_Card_order, basis_Asuc, suc_bd_Card_order];
+
+ val Asuc_bdT = fst (dest_relT (fastype_of (mk_Asuc_bd As)));
+ 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 (((((((idxs, Asi_name), (idx, idx')), (jdx, jdx')), II_Bs), II_ss), Asuc_fs),
+ names_lthy) = names_lthy
+ |> 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
+ ||>> mk_Frees "IIB" II_BTs
+ ||>> mk_Frees "IIs" II_sTs
+ ||>> mk_Frees "f" (map (fn T => Asuc_bdT --> T) activeAs);
+
+ 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))));
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all idxs (Logic.list_implies ([prem], concl)))
+ (K (mk_bd_limit_tac n suc_bd_Cinfinite))
+ |> Thm.close_derivation
+ 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 As Asi i sets Ts s k =
+ HOLogic.mk_binop @{const_name "sup"}
+ (mk_minG Asi i k, mk_image s $ mk_in (As @ map (mk_minG Asi i) ks) sets Ts);
+
+ fun mk_min_algs As ss =
+ let
+ val BTs = map (range_type o fastype_of) ss;
+ val Ts = map (HOLogic.dest_setT o fastype_of) As @ 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
+ (map4 (mk_minH_component As 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 As 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
+ (map4 (mk_minH_component As min_algs idx) setssAs FTsAs ss ks)));
+ val goal = fold_rev Logic.all (idx :: As @ ss)
+ (Logic.mk_implies (HOLogic.mk_Trueprop i_field, concl));
+
+ val min_algs_thm = Skip_Proof.prove lthy [] [] goal
+ (K (mk_min_algs_tac suc_bd_worel in_cong'_thms))
+ |> Thm.close_derivation;
+
+ val min_algs_thms = map (fn k => min_algs_thm RS mk_nthI n k) ks;
+
+ fun mk_mono_goal min_alg =
+ fold_rev Logic.all (As @ ss) (HOLogic.mk_Trueprop (mk_relChain suc_bd
+ (Term.absfree idx' min_alg)));
+
+ val monos =
+ map2 (fn goal => fn min_algs =>
+ Skip_Proof.prove lthy [] [] goal (K (mk_min_algs_mono_tac min_algs))
+ |> Thm.close_derivation)
+ (map mk_mono_goal min_algss) min_algs_thms;
+
+ val Asuc_bd = mk_Asuc_bd As;
+
+ 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 = certifyT lthy suc_bdT;
+ val card_ct = certify lthy (Term.absfree idx' card_conjunction);
+
+ val card_of = singleton (Proof_Context.export names_lthy lthy)
+ (Skip_Proof.prove lthy [] []
+ (HOLogic.mk_Trueprop (HOLogic.mk_imp (i_field, card_conjunction)))
+ (K (mk_min_algs_card_of_tac card_cT card_ct
+ m suc_bd_worel min_algs_thms in_bd_sums
+ sum_Card_order sum_Cnotzero suc_bd_Card_order suc_bd_Cinfinite suc_bd_Cnotzero
+ suc_bd_Asuc_bd Asuc_bd_Cinfinite Asuc_bd_Cnotzero)))
+ |> Thm.close_derivation;
+
+ val least_prem = HOLogic.mk_Trueprop (mk_alg As Bs ss);
+ val least_conjunction = Library.foldr1 HOLogic.mk_conj (map2 mk_subset min_algss Bs);
+ val least_cT = certifyT lthy suc_bdT;
+ val least_ct = certify lthy (Term.absfree idx' least_conjunction);
+
+ val least = singleton (Proof_Context.export names_lthy lthy)
+ (Skip_Proof.prove lthy [] []
+ (Logic.mk_implies (least_prem,
+ HOLogic.mk_Trueprop (HOLogic.mk_imp (i_field, least_conjunction))))
+ (K (mk_min_algs_least_tac least_cT least_ct
+ suc_bd_worel min_algs_thms alg_set_thms)))
+ |> Thm.close_derivation;
+ in
+ (min_algs_thms, monos, card_of, least)
+ end;
+
+ val timer = time (timer "min_algs definition & thms");
+
+ fun min_alg_bind i = Binding.suffix_name
+ ("_" ^ min_algN ^ (if n = 1 then "" else string_of_int i)) b;
+ val min_alg_name = Binding.name_of o min_alg_bind;
+ val min_alg_def_bind = rpair [] o Thm.def_binding o min_alg_bind;
+
+ fun min_alg_spec i =
+ let
+ val min_algT =
+ Library.foldr (op -->) (ATs @ sTs, HOLogic.mk_setT (nth activeAs (i - 1)));
+
+ val lhs = Term.list_comb (Free (min_alg_name i, min_algT), As @ ss);
+ val rhs = mk_UNION (field_suc_bd)
+ (Term.absfree idx' (mk_nthN n (mk_min_algs As ss $ idx) i));
+ in
+ mk_Trueprop_eq (lhs, rhs)
+ end;
+
+ val ((min_alg_frees, (_, min_alg_def_frees)), (lthy, lthy_old)) =
+ lthy
+ |> fold_map (fn i => Specification.definition
+ (SOME (min_alg_bind i, NONE, NoSyn), (min_alg_def_bind i, min_alg_spec i))) ks
+ |>> apsnd split_list o split_list
+ ||> `Local_Theory.restore;
+
+ 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 (Morphism.thm phi) min_alg_def_frees;
+
+ fun mk_min_alg As ss i =
+ let
+ val T = HOLogic.mk_setT (range_type (fastype_of (nth ss (i - 1))))
+ val args = As @ ss;
+ val Ts = map fastype_of args;
+ val min_algT = Library.foldr (op -->) (Ts, T);
+ in
+ Term.list_comb (Const (nth min_algs (i - 1), min_algT), args)
+ end;
+
+ val (alg_min_alg_thm, card_of_min_alg_thms, least_min_alg_thms, mor_incl_min_alg_thm) =
+ let
+ val min_algs = map (mk_min_alg As ss) ks;
+
+ val goal = fold_rev Logic.all (As @ ss) (HOLogic.mk_Trueprop (mk_alg As min_algs ss));
+ val alg_min_alg = Skip_Proof.prove lthy [] [] goal
+ (K (mk_alg_min_alg_tac m alg_def min_alg_defs suc_bd_limit_thm sum_Cinfinite
+ set_bd_sumss min_algs_thms min_algs_mono_thms))
+ |> Thm.close_derivation;
+
+ val Asuc_bd = mk_Asuc_bd As;
+ fun mk_card_of_thm min_alg def = Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (As @ ss)
+ (HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of min_alg) Asuc_bd)))
+ (K (mk_card_of_min_alg_tac def card_of_min_algs_thm
+ suc_bd_Card_order suc_bd_Asuc_bd Asuc_bd_Cinfinite))
+ |> Thm.close_derivation;
+
+ val least_prem = HOLogic.mk_Trueprop (mk_alg As Bs ss);
+ fun mk_least_thm min_alg B def = Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (As @ Bs @ ss)
+ (Logic.mk_implies (least_prem, HOLogic.mk_Trueprop (mk_subset min_alg B))))
+ (K (mk_least_min_alg_tac def least_min_algs_thm))
+ |> Thm.close_derivation;
+
+ val leasts = map3 mk_least_thm min_algs Bs min_alg_defs;
+
+ val incl_prem = HOLogic.mk_Trueprop (mk_alg passive_UNIVs Bs ss);
+ val incl_min_algs = map (mk_min_alg passive_UNIVs ss) ks;
+ val incl = Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (Bs @ ss)
+ (Logic.mk_implies (incl_prem,
+ HOLogic.mk_Trueprop (mk_mor incl_min_algs ss Bs ss active_ids))))
+ (K (EVERY' (rtac mor_incl_thm :: map etac leasts) 1))
+ |> Thm.close_derivation;
+ 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 = Binding.suffix_name ("_" ^ IITN) b;
+
+ val ((IIT_name, (IIT_glob_info, IIT_loc_info)), lthy) =
+ typedef false NONE (IIT_bind, params, NoSyn)
+ (HOLogic.mk_UNIV II_repT) NONE (EVERY' [rtac exI, rtac 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 (((((((iidx, iidx'), init_xs), (init_xFs, init_xFs')),
+ init_fs), init_fs_copy), init_phis), names_lthy) = names_lthy
+ |> 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 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 passive_UNIVs 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;
+
+ fun str_init_bind i = Binding.suffix_name ("_" ^ str_initN ^ (if n = 1 then "" else
+ string_of_int i)) b;
+ val str_init_name = Binding.name_of o str_init_bind;
+ val str_init_def_bind = rpair [] o Thm.def_binding o str_init_bind;
+
+ fun str_init_spec i =
+ let
+ val T = nth init_FTs (i - 1);
+ 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 str_initT = T --> IIT --> Asuc_bdT;
+
+ val lhs = Term.list_comb (Free (str_init_name i, str_initT), [init_xF, iidx]);
+ val rhs = select_s $ (Term.list_comb (map, map_args) $ init_xF);
+ in
+ mk_Trueprop_eq (lhs, rhs)
+ end;
+
+ val ((str_init_frees, (_, str_init_def_frees)), (lthy, lthy_old)) =
+ lthy
+ |> fold_map (fn i => Specification.definition
+ (SOME (str_init_bind i, NONE, NoSyn), (str_init_def_bind i, str_init_spec i))) ks
+ |>> apsnd split_list o split_list
+ ||> `Local_Theory.restore;
+
+ 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 (Morphism.thm phi) str_init_def_frees;
+
+ val car_inits = map (mk_min_alg passive_UNIVs str_inits) ks;
+
+ (*TODO: replace with instantiate? (problem: figure out right type instantiation)*)
+ val alg_init_thm = Skip_Proof.prove lthy [] []
+ (HOLogic.mk_Trueprop (mk_alg passive_UNIVs car_inits str_inits))
+ (K (rtac alg_min_alg_thm 1))
+ |> Thm.close_derivation;
+
+ val alg_select_thm = Skip_Proof.prove lthy [] []
+ (HOLogic.mk_Trueprop (mk_Ball II
+ (Term.absfree iidx' (mk_alg passive_UNIVs select_Bs select_ss))))
+ (mk_alg_select_tac Abs_IIT_inverse_thm)
+ |> Thm.close_derivation;
+
+ val mor_select_thm =
+ let
+ val alg_prem = HOLogic.mk_Trueprop (mk_alg passive_UNIVs Bs ss);
+ val i_prem = HOLogic.mk_Trueprop (HOLogic.mk_mem (iidx, II));
+ val mor_prem = HOLogic.mk_Trueprop (mk_mor select_Bs select_ss Bs ss Asuc_fs);
+ val prems = [alg_prem, i_prem, mor_prem];
+ val concl = HOLogic.mk_Trueprop
+ (mk_mor car_inits str_inits Bs ss
+ (map (fn f => HOLogic.mk_comp (f, mk_rapp iidx Asuc_bdT)) Asuc_fs));
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (iidx :: Bs @ ss @ Asuc_fs) (Logic.list_implies (prems, concl)))
+ (K (mk_mor_select_tac mor_def mor_cong_thm mor_comp_thm mor_incl_min_alg_thm alg_def
+ alg_select_thm alg_set_thms set_natural'ss str_init_defs))
+ |> Thm.close_derivation
+ end;
+
+ val (init_ex_mor_thm, init_unique_mor_thms) =
+ let
+ val prem = HOLogic.mk_Trueprop (mk_alg passive_UNIVs Bs ss);
+ val concl = HOLogic.mk_Trueprop
+ (list_exists_free init_fs (mk_mor car_inits str_inits Bs ss init_fs));
+ val ex_mor = Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (Bs @ ss) (Logic.mk_implies (prem, concl)))
+ (mk_init_ex_mor_tac Abs_IIT_inverse_thm ex_copy_alg_thm alg_min_alg_thm
+ card_of_min_alg_thms mor_comp_thm mor_select_thm mor_incl_min_alg_thm)
+ |> Thm.close_derivation;
+
+ 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 (map3 mk_fun_eq init_fs init_fs_copy init_xs));
+ val unique_mor = Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (init_xs @ Bs @ ss @ init_fs @ init_fs_copy)
+ (Logic.list_implies (prems @ mor_prems, unique)))
+ (K (mk_init_unique_mor_tac m alg_def alg_init_thm least_min_alg_thms
+ in_mono'_thms alg_set_thms morE_thms map_congs))
+ |> Thm.close_derivation;
+ in
+ (ex_mor, 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
+ (map6 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));
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all init_phis (Logic.mk_implies (prem, concl)))
+ (K (mk_init_induct_tac m alg_def alg_init_thm least_min_alg_thms alg_set_thms))
+ |> Thm.close_derivation
+ end;
+
+ val timer = time (timer "Initiality definition & thms");
+
+ val ((T_names, (T_glob_infos, T_loc_infos)), lthy) =
+ lthy
+ |> fold_map3 (fn b => fn mx => fn car_init => typedef false NONE (b, params, mx) car_init NONE
+ (EVERY' [rtac ssubst, rtac @{thm ex_in_conv}, resolve_tac alg_not_empty_thms,
+ rtac 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_casess = map #Rep_cases T_loc_infos;
+ val Rep_injects = map #Rep_inject T_loc_infos;
+ val Rep_inverses = map #Rep_inverse T_loc_infos;
+ val Abs_inverses = map #Abs_inverse T_loc_infos;
+
+ fun mk_inver_thm mk_tac rep abs X thm =
+ Skip_Proof.prove lthy [] []
+ (HOLogic.mk_Trueprop (mk_inver rep abs X))
+ (K (EVERY' [rtac ssubst, rtac @{thm inver_def}, rtac ballI, mk_tac thm] 1))
+ |> Thm.close_derivation;
+
+ val inver_Reps = map4 (mk_inver_thm rtac) Abs_Ts Rep_Ts (map HOLogic.mk_UNIV Ts) Rep_inverses;
+ val inver_Abss = map4 (mk_inver_thm etac) Rep_Ts Abs_Ts car_inits Abs_inverses;
+
+ 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 rec_sTs = map (Term.typ_subst_atomic (activeBs ~~ Ts)) prod_sTs;
+ val rec_maps = map (Term.subst_atomic_types (activeBs ~~ Ts)) map_fsts;
+ val rec_maps_rev = map (Term.subst_atomic_types (activeBs ~~ Ts)) map_fsts_rev;
+ val rec_fsts = map (Term.subst_atomic_types (activeBs ~~ Ts)) fsts;
+
+ val (((((((((Izs1, Izs1'), (Izs2, Izs2')), (xFs, xFs')), yFs), (AFss, AFss')),
+ (fold_f, fold_f')), fs), rec_ss), names_lthy) = names_lthy
+ |> mk_Frees' "z1" Ts
+ ||>> mk_Frees' "z2" Ts'
+ ||>> mk_Frees' "x" FTs
+ ||>> mk_Frees "y" FTs'
+ ||>> mk_Freess' "z" setFTss
+ ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "f") foldT
+ ||>> mk_Frees "f" fTs
+ ||>> mk_Frees "s" rec_sTs;
+
+ val Izs = map2 retype_free Ts zs;
+ val phis = map2 retype_free (map mk_pred1T Ts) init_phis;
+ val phi2s = map2 retype_free (map2 mk_pred2T Ts Ts') init_phis;
+
+ fun ctor_bind i = Binding.suffix_name ("_" ^ ctorN) (nth bs (i - 1));
+ val ctor_name = Binding.name_of o ctor_bind;
+ val ctor_def_bind = rpair [] o Thm.def_binding o ctor_bind;
+
+ fun ctor_spec i abs str map_FT_init x x' =
+ let
+ val ctorT = nth FTs (i - 1) --> nth Ts (i - 1);
+
+ val lhs = Free (ctor_name i, ctorT);
+ val rhs = Term.absfree x' (abs $ (str $
+ (Term.list_comb (map_FT_init, map HOLogic.id_const passiveAs @ Rep_Ts) $ x)));
+ in
+ mk_Trueprop_eq (lhs, rhs)
+ end;
+
+ val ((ctor_frees, (_, ctor_def_frees)), (lthy, lthy_old)) =
+ lthy
+ |> fold_map6 (fn i => fn abs => fn str => fn mapx => fn x => fn x' =>
+ Specification.definition
+ (SOME (ctor_bind i, NONE, NoSyn), (ctor_def_bind i, ctor_spec i abs str mapx x x')))
+ ks Abs_Ts str_inits map_FT_inits xFs xFs'
+ |>> apsnd split_list o split_list
+ ||> `Local_Theory.restore;
+
+ 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 (Morphism.thm phi) ctor_def_frees;
+
+ val (mor_Rep_thm, mor_Abs_thm) =
+ let
+ val copy = alg_init_thm RS copy_alg_thm;
+ fun mk_bij inj Rep cases = @{thm bij_betwI'} OF [inj, Rep, cases];
+ val bijs = map3 mk_bij Rep_injects Reps Rep_casess;
+ val mor_Rep =
+ Skip_Proof.prove lthy [] []
+ (HOLogic.mk_Trueprop (mk_mor UNIVs ctors car_inits str_inits Rep_Ts))
+ (mk_mor_Rep_tac ctor_defs copy bijs inver_Abss inver_Reps)
+ |> Thm.close_derivation;
+
+ val inv = mor_inv_thm OF [mor_Rep, talg_thm, alg_init_thm];
+ val mor_Abs =
+ Skip_Proof.prove lthy [] []
+ (HOLogic.mk_Trueprop (mk_mor car_inits str_inits UNIVs ctors Abs_Ts))
+ (K (mk_mor_Abs_tac inv inver_Abss inver_Reps))
+ |> Thm.close_derivation;
+ 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 = Binding.suffix_name ("_" ^ ctor_foldN) (nth bs (i - 1));
+ val fold_name = Binding.name_of o fold_bind;
+ val fold_def_bind = rpair [] o Thm.def_binding o fold_bind;
+
+ fun fold_spec i T AT =
+ let
+ val foldT = Library.foldr (op -->) (sTs, T --> AT);
+
+ val lhs = Term.list_comb (Free (fold_name i, foldT), ss);
+ val rhs = mk_nthN n foldx i;
+ in
+ mk_Trueprop_eq (lhs, rhs)
+ end;
+
+ val ((fold_frees, (_, fold_def_frees)), (lthy, lthy_old)) =
+ lthy
+ |> fold_map3 (fn i => fn T => fn AT =>
+ Specification.definition
+ (SOME (fold_bind i, NONE, NoSyn), (fold_def_bind i, fold_spec i T AT)))
+ ks Ts activeAs
+ |>> apsnd split_list o split_list
+ ||> `Local_Theory.restore;
+
+ 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_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 (Morphism.thm phi) fold_def_frees;
+
+ val mor_fold_thm =
+ let
+ val ex_mor = talg_thm RS init_ex_mor_thm;
+ val mor_cong = mor_cong_thm OF (map (mk_nth_conv n) ks);
+ val mor_comp = mor_Rep_thm RS mor_comp_thm;
+ val cT = certifyT lthy foldT;
+ val ct = certify lthy fold_fun
+ in
+ singleton (Proof_Context.export names_lthy lthy)
+ (Skip_Proof.prove lthy [] []
+ (HOLogic.mk_Trueprop (mk_mor UNIVs ctors active_UNIVs ss (map (mk_fold Ts ss) ks)))
+ (K (mk_mor_fold_tac cT ct fold_defs ex_mor (mor_comp RS mor_cong))))
+ |> Thm.close_derivation
+ end;
+
+ val ctor_fold_thms = map (fn morE => rule_by_tactic lthy
+ ((rtac CollectI THEN' CONJ_WRAP' (K (rtac @{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 unique_mor = Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (ss @ fs) (Logic.mk_implies (prem, unique)))
+ (K (mk_fold_unique_mor_tac type_defs init_unique_mor_thms Reps
+ mor_comp_thm mor_Abs_thm mor_fold_thm))
+ |> Thm.close_derivation;
+ in
+ `split_conj_thm unique_mor
+ end;
+
+ val ctor_fold_unique_thms =
+ split_conj_thm (mk_conjIN n RS
+ (mor_UNIV_thm RS @{thm ssubst[of _ _ "%x. x"]} 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 = Binding.suffix_name ("_" ^ dtorN) (nth bs (i - 1));
+ val dtor_name = Binding.name_of o dtor_bind;
+ val dtor_def_bind = rpair [] o Thm.def_binding o dtor_bind;
+
+ fun dtor_spec i FT T =
+ let
+ val dtorT = T --> FT;
+
+ val lhs = Free (dtor_name i, dtorT);
+ val rhs = mk_fold Ts map_ctors i;
+ in
+ mk_Trueprop_eq (lhs, rhs)
+ end;
+
+ val ((dtor_frees, (_, dtor_def_frees)), (lthy, lthy_old)) =
+ lthy
+ |> fold_map3 (fn i => fn FT => fn T =>
+ Specification.definition
+ (SOME (dtor_bind i, NONE, NoSyn), (dtor_def_bind i, dtor_spec i FT T))) ks FTs Ts
+ |>> apsnd split_list o split_list
+ ||> `Local_Theory.restore;
+
+ 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 (Morphism.thm phi) dtor_def_frees;
+
+ val ctor_o_dtor_thms = map2 (fold_thms 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 = map3 mk_goal dtors ctors FTs;
+ in
+ map5 (fn goal => fn dtor_def => fn foldx => fn map_comp_id => fn map_congL =>
+ Skip_Proof.prove lthy [] [] goal
+ (K (mk_dtor_o_ctor_tac dtor_def foldx map_comp_id map_congL ctor_o_fold_thms))
+ |> Thm.close_derivation)
+ goals dtor_defs ctor_fold_thms map_comp_id_thms map_congL_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 fst_rec_pair_thms =
+ let
+ val mor = mor_comp_thm OF [mor_fold_thm, mor_convol_thm];
+ in
+ map2 (fn unique => fn fold_ctor =>
+ trans OF [mor RS unique, fold_ctor]) fold_unique_mor_thms fold_ctor_thms
+ end;
+
+ fun rec_bind i = Binding.suffix_name ("_" ^ ctor_recN) (nth bs (i - 1));
+ val rec_name = Binding.name_of o rec_bind;
+ val rec_def_bind = rpair [] o Thm.def_binding o rec_bind;
+
+ fun rec_spec i T AT =
+ let
+ val recT = Library.foldr (op -->) (rec_sTs, T --> AT);
+ val maps = map3 (fn ctor => fn prod_s => fn mapx =>
+ mk_convol (HOLogic.mk_comp (ctor, Term.list_comb (mapx, passive_ids @ rec_fsts)), prod_s))
+ ctors rec_ss rec_maps;
+
+ val lhs = Term.list_comb (Free (rec_name i, recT), rec_ss);
+ val rhs = HOLogic.mk_comp (snd_const (HOLogic.mk_prodT (T, AT)), mk_fold Ts maps i);
+ in
+ mk_Trueprop_eq (lhs, rhs)
+ end;
+
+ val ((rec_frees, (_, rec_def_frees)), (lthy, lthy_old)) =
+ lthy
+ |> fold_map3 (fn i => fn T => fn AT =>
+ Specification.definition
+ (SOME (rec_bind i, NONE, NoSyn), (rec_def_bind i, rec_spec i T AT)))
+ ks Ts activeAs
+ |>> apsnd split_list o split_list
+ ||> `Local_Theory.restore;
+
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+ val recs = map (Morphism.term phi) rec_frees;
+ val rec_names = map (fst o dest_Const) recs;
+ fun mk_rec ss i = Term.list_comb (Const (nth rec_names (i - 1), Library.foldr (op -->)
+ (map fastype_of ss, nth Ts (i - 1) --> range_type (fastype_of (nth ss (i - 1))))), ss);
+ val rec_defs = map (Morphism.thm phi) rec_def_frees;
+
+ val convols = map2 (fn T => fn i => mk_convol (HOLogic.id_const T, mk_rec rec_ss i)) Ts ks;
+ val ctor_rec_thms =
+ let
+ fun mk_goal i rec_s rec_map ctor x =
+ let
+ val lhs = mk_rec rec_ss i $ (ctor $ x);
+ val rhs = rec_s $ (Term.list_comb (rec_map, passive_ids @ convols) $ x);
+ in
+ fold_rev Logic.all (x :: rec_ss) (mk_Trueprop_eq (lhs, rhs))
+ end;
+ val goals = map5 mk_goal ks rec_ss rec_maps_rev ctors xFs;
+ in
+ map2 (fn goal => fn foldx =>
+ Skip_Proof.prove lthy [] [] goal (mk_rec_tac rec_defs foldx fst_rec_pair_thms)
+ |> Thm.close_derivation)
+ goals ctor_fold_thms
+ end;
+
+ val timer = time (timer "rec definitions & thms");
+
+ val (ctor_induct_thm, induct_params) =
+ let
+ fun mk_prem phi ctor sets x =
+ let
+ fun mk_IH phi set z =
+ let
+ val prem = HOLogic.mk_Trueprop (HOLogic.mk_mem (z, set $ x));
+ val concl = HOLogic.mk_Trueprop (phi $ z);
+ in
+ Logic.all z (Logic.mk_implies (prem, concl))
+ end;
+
+ val IHs = map3 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 = map4 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);
+ in
+ (Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (phis @ Izs) goal)
+ (K (mk_ctor_induct_tac m set_natural'ss init_induct_thm morE_thms mor_Abs_thm
+ Rep_inverses Abs_inverses Reps))
+ |> Thm.close_derivation,
+ rev (Term.add_tfrees goal []))
+ end;
+
+ val cTs = map (SOME o certifyT lthy o TFree) induct_params;
+
+ val weak_ctor_induct_thms =
+ let fun insts i = (replicate (i - 1) TrueI) @ (@{thm 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 = HOLogic.mk_Trueprop (HOLogic.mk_mem (z1, (set $ x)));
+ val prem2 = HOLogic.mk_Trueprop (HOLogic.mk_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 = map5 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 = map7 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
+ (map3 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 = map3 (SOME o certify lthy ooo mk_t) phi2s (Izs1 ~~ Izs1') (Izs2 ~~ Izs2');
+ val goal = Logic.list_implies (prems, concl);
+ in
+ (singleton (Proof_Context.export names_lthy lthy)
+ (Skip_Proof.prove lthy [] [] goal
+ (mk_ctor_induct2_tac cTs cts ctor_induct_thm weak_ctor_induct_thms))
+ |> Thm.close_derivation,
+ rev (Term.add_tfrees goal []))
+ end;
+
+ val timer = time (timer "induction");
+
+ (*register new datatypes as BNFs*)
+ val lthy = if m = 0 then lthy else
+ let
+ val fTs = map2 (curry op -->) passiveAs passiveBs;
+ val f1Ts = map2 (curry op -->) passiveAs passiveYs;
+ val f2Ts = map2 (curry op -->) passiveBs passiveYs;
+ val p1Ts = map2 (curry op -->) passiveXs passiveAs;
+ val p2Ts = map2 (curry op -->) passiveXs passiveBs;
+ val uTs = map2 (curry op -->) Ts Ts';
+ val B1Ts = map HOLogic.mk_setT passiveAs;
+ val B2Ts = map HOLogic.mk_setT passiveBs;
+ val AXTs = map HOLogic.mk_setT passiveXs;
+ val XTs = mk_Ts passiveXs;
+ val YTs = mk_Ts passiveYs;
+ val IRTs = map2 (curry mk_relT) passiveAs passiveBs;
+ val IphiTs = map2 mk_pred2T passiveAs passiveBs;
+
+ val (((((((((((((((fs, fs'), fs_copy), us),
+ B1s), B2s), AXs), (xs, xs')), f1s), f2s), p1s), p2s), (ys, ys')), IRs), Iphis),
+ names_lthy) = names_lthy
+ |> mk_Frees' "f" fTs
+ ||>> mk_Frees "f" fTs
+ ||>> mk_Frees "u" uTs
+ ||>> mk_Frees "B1" B1Ts
+ ||>> mk_Frees "B2" B2Ts
+ ||>> mk_Frees "A" AXTs
+ ||>> mk_Frees' "x" XTs
+ ||>> mk_Frees "f1" f1Ts
+ ||>> mk_Frees "f2" f2Ts
+ ||>> mk_Frees "p1" p1Ts
+ ||>> mk_Frees "p2" p2Ts
+ ||>> mk_Frees' "y" passiveAs
+ ||>> mk_Frees "R" IRTs
+ ||>> mk_Frees "P" IphiTs;
+
+ 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 fs_copy_maps = map (mk_map Ts fs_copy Ts' ctor's pmapsABT') ks;
+ val Yctors = mk_ctors passiveYs;
+ val f1s_maps = map (mk_map Ts f1s YTs Yctors (mk_passive_maps passiveAs passiveYs)) ks;
+ val f2s_maps = map (mk_map Ts' f2s YTs Yctors (mk_passive_maps passiveBs passiveYs)) ks;
+ val p1s_maps = map (mk_map XTs p1s Ts ctors (mk_passive_maps passiveXs passiveAs)) ks;
+ val p2s_maps = map (mk_map XTs p2s Ts' ctor's (mk_passive_maps passiveXs passiveBs)) ks;
+
+ val map_simp_thms =
+ let
+ fun mk_goal fs_map map ctor ctor' = fold_rev Logic.all fs
+ (mk_Trueprop_eq (HOLogic.mk_comp (fs_map, ctor),
+ HOLogic.mk_comp (ctor', Term.list_comb (map, fs @ fs_maps))));
+ val goals = map4 mk_goal fs_maps map_FTFT's ctors ctor's;
+ val maps =
+ map4 (fn goal => fn foldx => fn map_comp_id => fn map_cong =>
+ Skip_Proof.prove lthy [] [] goal (K (mk_map_tac m n foldx map_comp_id map_cong))
+ |> Thm.close_derivation)
+ goals ctor_fold_thms map_comp_id_thms map_congs;
+ in
+ map (fn thm => thm RS @{thm pointfreeE}) maps
+ end;
+
+ val (map_unique_thms, map_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 = map4 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_maps));
+ val unique = Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (us @ fs) (Logic.list_implies (prems, goal)))
+ (K (mk_map_unique_tac m mor_def fold_unique_mor_thm map_comp_id_thms map_congs))
+ |> Thm.close_derivation;
+ in
+ `split_conj_thm unique
+ end;
+
+ val timer = time (timer "map functions for the new datatypes");
+
+ val bd = mk_cpow sum_bd;
+ val bd_Cinfinite = sum_Cinfinite RS @{thm Cinfinite_cpow};
+ fun mk_cpow_bd thm = @{thm ordLeq_transitive} OF
+ [thm, sum_Card_order RS @{thm cpow_greater_eq}];
+ val set_bd_cpowss = map (map mk_cpow_bd) set_bd_sumss;
+
+ val timer = time (timer "bounds for the new datatypes");
+
+ val ls = 1 upto m;
+ val setsss = map (mk_setss o mk_set_Ts) passiveAs;
+ val map_setss = map (fn T => map2 (fn Ds =>
+ mk_map_of_bnf Ds (passiveAs @ Ts) (mk_set_Ts T)) Dss bnfs) 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 = map5 (fn l => fn T => map3 (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_simp_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 =
+ map3 (fn sets => map4 (mk_goal sets) ctors sets) setss_by_range colss map_setss;
+ val setss = map (map2 (fn foldx => fn goal =>
+ Skip_Proof.prove lthy [] [] goal (K (mk_set_tac foldx)) |> Thm.close_derivation)
+ ctor_fold_thms) goalss;
+
+ fun mk_simp_goal pas_set act_sets sets ctor z set =
+ Logic.all z (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 =>
+ map4 (fn Fsets => mk_simp_goal (nth Fsets (i - 1)) (drop m Fsets) sets)
+ FTs_setss ctors xFs sets)
+ ls setss_by_range;
+
+ val set_simpss = map3 (fn i => map3 (fn set_nats => fn goal => fn set =>
+ Skip_Proof.prove lthy [] [] goal
+ (K (mk_set_simp_tac set (nth set_nats (i - 1)) (drop m set_nats)))
+ |> Thm.close_derivation)
+ set_natural'ss) ls simp_goalss setss;
+ in
+ set_simpss
+ end;
+
+ fun mk_set_thms set_simp = (@{thm xt1(3)} OF [set_simp, @{thm Un_upper1}]) ::
+ map (fn i => (@{thm xt1(3)} OF [set_simp, @{thm Un_upper2}]) RS
+ (mk_Un_upper n i RS subset_trans) RSN
+ (2, @{thm UN_upper} RS subset_trans))
+ (1 upto n);
+ val Fset_set_thmsss = transpose (map (map mk_set_thms) set_simp_thmss);
+
+ val timer = time (timer "set functions for the new datatypes");
+
+ val cxs = map (SOME o certify lthy) Izs;
+ val setss_by_bnf' =
+ map (map (Term.subst_atomic_types (passiveAs ~~ passiveBs))) setss_by_bnf;
+ val setss_by_range' = transpose setss_by_bnf';
+
+ val set_natural_thmss =
+ let
+ fun mk_set_natural f map z set set' =
+ HOLogic.mk_eq (mk_image f $ (set $ z), set' $ (map $ z));
+
+ fun mk_cphi f map z set set' = certify lthy
+ (Term.absfree (dest_Free z) (mk_set_natural f map z set set'));
+
+ val csetss = map (map (certify lthy)) setss_by_range';
+
+ val cphiss = map3 (fn f => fn sets => fn sets' =>
+ (map4 (mk_cphi f) fs_maps Izs sets sets')) fs setss_by_range setss_by_range';
+
+ val inducts = map (fn cphis =>
+ Drule.instantiate' cTs (map SOME cphis @ cxs) ctor_induct_thm) cphiss;
+
+ val goals =
+ map3 (fn f => fn sets => fn sets' =>
+ HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+ (map4 (mk_set_natural f) fs_maps Izs sets sets')))
+ fs setss_by_range setss_by_range';
+
+ fun mk_tac induct = mk_set_nat_tac m (rtac induct) set_natural'ss map_simp_thms;
+ val thms =
+ map5 (fn goal => fn csets => fn set_simps => fn induct => fn i =>
+ singleton (Proof_Context.export names_lthy lthy)
+ (Skip_Proof.prove lthy [] [] goal (mk_tac induct csets set_simps i))
+ |> Thm.close_derivation)
+ goals csetss set_simp_thmss inducts ls;
+ in
+ map split_conj_thm thms
+ end;
+
+ val set_bd_thmss =
+ let
+ fun mk_set_bd z set = mk_ordLeq (mk_card_of (set $ z)) bd;
+
+ fun mk_cphi z set = certify lthy (Term.absfree (dest_Free z) (mk_set_bd z set));
+
+ val cphiss = map (map2 mk_cphi Izs) setss_by_range;
+
+ val inducts = map (fn cphis =>
+ Drule.instantiate' cTs (map SOME cphis @ cxs) ctor_induct_thm) cphiss;
+
+ val goals =
+ map (fn sets =>
+ HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+ (map2 mk_set_bd Izs sets))) setss_by_range;
+
+ fun mk_tac induct = mk_set_bd_tac m (rtac induct) bd_Cinfinite set_bd_cpowss;
+ val thms =
+ map4 (fn goal => fn set_simps => fn induct => fn i =>
+ singleton (Proof_Context.export names_lthy lthy)
+ (Skip_Proof.prove lthy [] [] goal (mk_tac induct set_simps i))
+ |> Thm.close_derivation)
+ goals set_simp_thmss inducts ls;
+ in
+ map split_conj_thm thms
+ end;
+
+ val map_cong_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_cong sets z fmap gmap =
+ HOLogic.mk_imp
+ (Library.foldr1 HOLogic.mk_conj (map5 (mk_prem z) sets fs fs_copy ys ys'),
+ HOLogic.mk_eq (fmap $ z, gmap $ z));
+
+ fun mk_cphi sets z fmap gmap =
+ certify lthy (Term.absfree (dest_Free z) (mk_map_cong sets z fmap gmap));
+
+ val cphis = map4 mk_cphi setss_by_bnf Izs fs_maps fs_copy_maps;
+
+ val induct = Drule.instantiate' cTs (map SOME cphis @ cxs) ctor_induct_thm;
+
+ val goal =
+ HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+ (map4 mk_map_cong setss_by_bnf Izs fs_maps fs_copy_maps));
+
+ val thm = singleton (Proof_Context.export names_lthy lthy)
+ (Skip_Proof.prove lthy [] [] goal
+ (mk_mcong_tac (rtac induct) Fset_set_thmsss map_congs map_simp_thms))
+ |> Thm.close_derivation;
+ in
+ split_conj_thm thm
+ end;
+
+ val in_incl_min_alg_thms =
+ let
+ fun mk_prem z sets =
+ HOLogic.mk_mem (z, mk_in As sets (fastype_of z));
+
+ fun mk_incl z sets i =
+ HOLogic.mk_imp (mk_prem z sets, HOLogic.mk_mem (z, mk_min_alg As ctors i));
+
+ fun mk_cphi z sets i =
+ certify lthy (Term.absfree (dest_Free z) (mk_incl z sets i));
+
+ val cphis = map3 mk_cphi Izs setss_by_bnf ks;
+
+ val induct = Drule.instantiate' cTs (map SOME cphis @ cxs) ctor_induct_thm;
+
+ val goal =
+ HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+ (map3 mk_incl Izs setss_by_bnf ks));
+
+ val thm = singleton (Proof_Context.export names_lthy lthy)
+ (Skip_Proof.prove lthy [] [] goal
+ (mk_incl_min_alg_tac (rtac induct) Fset_set_thmsss alg_set_thms alg_min_alg_thm))
+ |> Thm.close_derivation;
+ in
+ split_conj_thm thm
+ end;
+
+ val Xsetss = map (map (Term.subst_atomic_types (passiveAs ~~ passiveXs))) setss_by_bnf;
+
+ val map_wpull_thms =
+ let
+ val cTs = map (SOME o certifyT lthy o TFree) induct2_params;
+ val cxs = map (SOME o certify lthy) (interleave Izs1 Izs2);
+
+ fun mk_prem z1 z2 sets1 sets2 map1 map2 =
+ HOLogic.mk_conj
+ (HOLogic.mk_mem (z1, mk_in B1s sets1 (fastype_of z1)),
+ HOLogic.mk_conj
+ (HOLogic.mk_mem (z2, mk_in B2s sets2 (fastype_of z2)),
+ HOLogic.mk_eq (map1 $ z1, map2 $ z2)));
+
+ val prems = map6 mk_prem Izs1 Izs2 setss_by_bnf setss_by_bnf' f1s_maps f2s_maps;
+
+ fun mk_concl z1 z2 sets map1 map2 T x x' =
+ mk_Bex (mk_in AXs sets T) (Term.absfree x'
+ (HOLogic.mk_conj (HOLogic.mk_eq (map1 $ x, z1), HOLogic.mk_eq (map2 $ x, z2))));
+
+ val concls = map8 mk_concl Izs1 Izs2 Xsetss p1s_maps p2s_maps XTs xs xs';
+
+ val goals = map2 (curry HOLogic.mk_imp) prems concls;
+
+ fun mk_cphi z1 z2 goal = certify lthy (Term.absfree z1 (Term.absfree z2 goal));
+
+ val cphis = map3 mk_cphi Izs1' Izs2' goals;
+
+ val induct = Drule.instantiate' cTs (map SOME cphis @ cxs) ctor_induct2_thm;
+
+ val goal = Logic.list_implies (map HOLogic.mk_Trueprop
+ (map8 mk_wpull AXs B1s B2s f1s f2s (replicate m NONE) p1s p2s),
+ HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj goals));
+
+ val thm = singleton (Proof_Context.export names_lthy lthy)
+ (Skip_Proof.prove lthy [] [] goal
+ (K (mk_lfp_map_wpull_tac m (rtac induct) map_wpulls map_simp_thms
+ (transpose set_simp_thmss) Fset_set_thmsss ctor_inject_thms)))
+ |> Thm.close_derivation;
+ in
+ split_conj_thm thm
+ end;
+
+ val timer = time (timer "helpers for BNF properties");
+
+ val map_id_tacs = map (K o mk_map_id_tac map_ids) map_unique_thms;
+ val map_comp_tacs =
+ map2 (K oo mk_map_comp_tac map_comp's map_simp_thms) map_unique_thms ks;
+ val map_cong_tacs = map (mk_map_cong_tac m) map_cong_thms;
+ val set_nat_tacss = map (map (K o mk_set_natural_tac)) (transpose set_natural_thmss);
+ val bd_co_tacs = replicate n (K (mk_bd_card_order_tac bd_card_orders));
+ val bd_cinf_tacs = replicate n (K (rtac (bd_Cinfinite RS conjunct1) 1));
+ val set_bd_tacss = map (map (fn thm => K (rtac thm 1))) (transpose set_bd_thmss);
+ val in_bd_tacs = map2 (K oo mk_in_bd_tac sum_Card_order suc_bd_Cnotzero)
+ in_incl_min_alg_thms card_of_min_alg_thms;
+ val map_wpull_tacs = map (K o mk_wpull_tac) map_wpull_thms;
+
+ val srel_O_Gr_tacs = replicate n (simple_srel_O_Gr_tac o #context);
+
+ val tacss = map10 zip_axioms map_id_tacs map_comp_tacs map_cong_tacs set_nat_tacss
+ bd_co_tacs bd_cinf_tacs set_bd_tacss in_bd_tacs map_wpull_tacs srel_O_Gr_tacs;
+
+ 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
+ map3 (fn ctor => fn i => map close_wit o minimize_wits o maps (mk_wit witss ctor i))
+ ctors (0 upto n - 1) witss
+ end;
+
+ fun wit_tac _ = mk_wit_tac n (flat set_simp_thmss) (maps wit_thms_of_bnf bnfs);
+
+ val policy = user_policy Derive_All_Facts_Note_Most;
+
+ val (Ibnfs, lthy) =
+ fold_map6 (fn tacs => fn b => fn mapx => fn sets => fn T => fn wits => fn lthy =>
+ bnf_def Dont_Inline policy I tacs wit_tac (SOME deads)
+ (((((b, fold_rev Term.absfree fs' mapx), sets), absdummy T bd), wits), NONE) lthy
+ |> register_bnf (Local_Theory.full_name lthy b))
+ tacss bs fs_maps setss_by_bnf Ts ctor_witss lthy;
+
+ val fold_maps = fold_thms lthy (map (fn bnf =>
+ mk_unabs_def m (map_def_of_bnf bnf RS @{thm meta_eq_to_obj_eq})) Ibnfs);
+
+ val fold_sets = fold_thms lthy (maps (fn bnf =>
+ map (fn thm => thm RS @{thm meta_eq_to_obj_eq}) (set_defs_of_bnf bnf)) Ibnfs);
+
+ val timer = time (timer "registered new datatypes as BNFs");
+
+ val srels = map2 (fn Ds => mk_srel_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs;
+ val Isrels = map (mk_srel_of_bnf deads passiveAs passiveBs) Ibnfs;
+ val rels = map2 (fn Ds => mk_rel_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs;
+ val Irels = map (mk_rel_of_bnf deads passiveAs passiveBs) Ibnfs;
+
+ val IrelRs = map (fn Isrel => Term.list_comb (Isrel, IRs)) Isrels;
+ val relRs = map (fn srel => Term.list_comb (srel, IRs @ IrelRs)) srels;
+ val Ipredphis = map (fn Isrel => Term.list_comb (Isrel, Iphis)) Irels;
+ val predphis = map (fn srel => Term.list_comb (srel, Iphis @ Ipredphis)) rels;
+
+ val in_srels = map in_srel_of_bnf bnfs;
+ val in_Isrels = map in_srel_of_bnf Ibnfs;
+ val srel_defs = map srel_def_of_bnf bnfs;
+ val Isrel_defs = map srel_def_of_bnf Ibnfs;
+ val Irel_defs = map rel_def_of_bnf Ibnfs;
+
+ val set_incl_thmss = map (map (fold_sets o hd)) Fset_set_thmsss;
+ val set_set_incl_thmsss = map (transpose o map (map fold_sets o tl)) Fset_set_thmsss;
+ val folded_map_simp_thms = map fold_maps map_simp_thms;
+ val folded_set_simp_thmss = map (map fold_sets) set_simp_thmss;
+ val folded_set_simp_thmss' = transpose folded_set_simp_thmss;
+
+ val Isrel_simp_thms =
+ let
+ fun mk_goal xF yF ctor ctor' IrelR relR = fold_rev Logic.all (xF :: yF :: IRs)
+ (mk_Trueprop_eq (HOLogic.mk_mem (HOLogic.mk_prod (ctor $ xF, ctor' $ yF), IrelR),
+ HOLogic.mk_mem (HOLogic.mk_prod (xF, yF), relR)));
+ val goals = map6 mk_goal xFs yFs ctors ctor's IrelRs relRs;
+ in
+ map12 (fn i => fn goal => fn in_srel => fn map_comp => fn map_cong =>
+ fn map_simp => fn set_simps => fn ctor_inject => fn ctor_dtor =>
+ fn set_naturals => fn set_incls => fn set_set_inclss =>
+ Skip_Proof.prove lthy [] [] goal
+ (K (mk_srel_simp_tac in_Isrels i in_srel map_comp map_cong map_simp set_simps
+ ctor_inject ctor_dtor set_naturals set_incls set_set_inclss))
+ |> Thm.close_derivation)
+ ks goals in_srels map_comp's map_congs folded_map_simp_thms folded_set_simp_thmss'
+ ctor_inject_thms ctor_dtor_thms set_natural'ss set_incl_thmss set_set_incl_thmsss
+ end;
+
+ val Irel_simp_thms =
+ let
+ fun mk_goal xF yF ctor ctor' Ipredphi predphi = fold_rev Logic.all (xF :: yF :: Iphis)
+ (mk_Trueprop_eq (Ipredphi $ (ctor $ xF) $ (ctor' $ yF), predphi $ xF $ yF));
+ val goals = map6 mk_goal xFs yFs ctors ctor's Ipredphis predphis;
+ in
+ map3 (fn goal => fn srel_def => fn Isrel_simp =>
+ Skip_Proof.prove lthy [] [] goal
+ (mk_rel_simp_tac srel_def Irel_defs Isrel_defs Isrel_simp)
+ |> Thm.close_derivation)
+ goals srel_defs Isrel_simp_thms
+ end;
+
+ val timer = time (timer "additional properties");
+
+ val ls' = if m = 1 then [0] else ls
+
+ val Ibnf_common_notes =
+ [(map_uniqueN, [fold_maps map_unique_thm])]
+ |> map (fn (thmN, thms) =>
+ ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]));
+
+ val Ibnf_notes =
+ [(map_simpsN, map single folded_map_simp_thms),
+ (set_inclN, set_incl_thmss),
+ (set_set_inclN, map flat set_set_incl_thmsss),
+ (srel_simpN, map single Isrel_simp_thms),
+ (rel_simpN, map single Irel_simp_thms)] @
+ map2 (fn i => fn thms => (mk_set_simpsN i, map single thms)) ls' folded_set_simp_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; lthy |> Local_Theory.notes (Ibnf_common_notes @ Ibnf_notes) |> snd
+ end;
+
+ val common_notes =
+ [(ctor_inductN, [ctor_induct_thm]),
+ (ctor_induct2N, [ctor_induct2_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_fold_uniqueN, ctor_fold_unique_thms),
+ (ctor_foldsN, ctor_fold_thms),
+ (ctor_injectN, ctor_inject_thms),
+ (ctor_recsN, ctor_rec_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)
+ in
+ ((dtors, ctors, folds, recs, ctor_induct_thm, dtor_ctor_thms, ctor_dtor_thms, ctor_inject_thms,
+ ctor_fold_thms, ctor_rec_thms),
+ lthy |> Local_Theory.notes (common_notes @ notes) |> snd)
+ end;
+
+val _ =
+ Outer_Syntax.local_theory @{command_spec "data_raw"}
+ "define BNF-based inductive datatypes (low-level)"
+ (Parse.and_list1
+ ((Parse.binding --| @{keyword ":"}) -- (Parse.typ --| @{keyword "="} -- Parse.typ)) >>
+ (snd oo fp_bnf_cmd bnf_lfp o apsnd split_list o split_list));
+
+val _ =
+ Outer_Syntax.local_theory @{command_spec "data"} "define BNF-based inductive datatypes"
+ (parse_datatype_cmd true bnf_lfp);
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Tools/bnf_lfp_tactics.ML Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,835 @@
+(* Title: HOL/BNF/Tools/bnf_lfp_tactics.ML
+ Author: Dmitriy Traytel, TU Muenchen
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Tactics for the datatype construction.
+*)
+
+signature BNF_LFP_TACTICS =
+sig
+ val mk_alg_min_alg_tac: int -> thm -> thm list -> thm -> thm -> thm list list -> thm list ->
+ thm list -> tactic
+ val mk_alg_not_empty_tac: thm -> thm list -> thm list -> tactic
+ val mk_alg_select_tac: thm -> {prems: 'a, context: Proof.context} -> tactic
+ val mk_alg_set_tac: thm -> tactic
+ val mk_bd_card_order_tac: thm list -> tactic
+ val mk_bd_limit_tac: int -> thm -> tactic
+ val mk_card_of_min_alg_tac: thm -> thm -> thm -> thm -> thm -> tactic
+ val mk_copy_alg_tac: thm list list -> thm list -> thm -> thm -> thm -> tactic
+ val mk_copy_str_tac: thm list list -> thm -> thm list -> tactic
+ val mk_ctor_induct_tac: int -> thm list list -> thm -> thm list -> thm -> thm list -> thm list ->
+ thm list -> tactic
+ val mk_ctor_induct2_tac: ctyp option list -> cterm option list -> thm -> thm list ->
+ {prems: 'a, context: Proof.context} -> tactic
+ val mk_dtor_o_ctor_tac: thm -> thm -> thm -> thm -> thm list -> tactic
+ val mk_ex_copy_alg_tac: int -> thm -> thm -> tactic
+ val mk_in_bd_tac: thm -> thm -> thm -> thm -> tactic
+ val mk_incl_min_alg_tac: (int -> tactic) -> thm list list list -> thm list -> thm ->
+ {prems: 'a, context: Proof.context} -> tactic
+ val mk_init_ex_mor_tac: thm -> thm -> thm -> thm list -> thm -> thm -> thm ->
+ {prems: 'a, context: Proof.context} -> tactic
+ val mk_init_induct_tac: int -> thm -> thm -> thm list -> thm list -> tactic
+ val mk_init_unique_mor_tac: int -> thm -> thm -> thm list -> thm list -> thm list -> thm list ->
+ thm list -> tactic
+ val mk_iso_alt_tac: thm list -> thm -> tactic
+ val mk_fold_unique_mor_tac: thm list -> thm list -> thm list -> thm -> thm -> thm -> tactic
+ val mk_least_min_alg_tac: thm -> thm -> tactic
+ val mk_lfp_map_wpull_tac: int -> (int -> tactic) -> thm list -> thm list -> thm list list ->
+ thm list list list -> thm list -> tactic
+ val mk_map_comp_tac: thm list -> thm list -> thm -> int -> tactic
+ val mk_map_id_tac: thm list -> thm -> tactic
+ val mk_map_tac: int -> int -> thm -> thm -> thm -> tactic
+ val mk_map_unique_tac: int -> thm -> thm -> thm list -> thm list -> tactic
+ val mk_mcong_tac: (int -> tactic) -> thm list list list -> thm list -> thm list ->
+ {prems: 'a, context: Proof.context} -> tactic
+ val mk_min_algs_card_of_tac: ctyp -> cterm -> int -> thm -> thm list -> thm list -> thm -> thm ->
+ thm -> thm -> thm -> thm -> thm -> thm -> tactic
+ val mk_min_algs_least_tac: ctyp -> cterm -> thm -> thm list -> thm list -> tactic
+ val mk_min_algs_mono_tac: thm -> tactic
+ val mk_min_algs_tac: thm -> thm list -> tactic
+ val mk_mor_Abs_tac: thm -> thm list -> thm list -> tactic
+ val mk_mor_Rep_tac: thm list -> thm -> thm list -> thm list -> thm list ->
+ {prems: 'a, context: Proof.context} -> tactic
+ val mk_mor_UNIV_tac: int -> thm list -> thm -> tactic
+ val mk_mor_comp_tac: thm -> thm list list -> thm list -> tactic
+ val mk_mor_convol_tac: 'a list -> thm -> tactic
+ val mk_mor_elim_tac: thm -> tactic
+ val mk_mor_incl_tac: thm -> thm list -> tactic
+ val mk_mor_inv_tac: thm -> thm -> thm list list -> thm list -> thm list -> thm list -> tactic
+ val mk_mor_fold_tac: ctyp -> cterm -> thm list -> thm -> thm -> tactic
+ val mk_mor_select_tac: thm -> thm -> thm -> thm -> thm -> thm -> thm list -> thm list list ->
+ thm list -> tactic
+ val mk_mor_str_tac: 'a list -> thm -> tactic
+ val mk_rec_tac: thm list -> thm -> thm list -> {prems: 'a, context: Proof.context} -> tactic
+ val mk_set_bd_tac: int -> (int -> tactic) -> thm -> thm list list -> thm list -> int ->
+ {prems: 'a, context: Proof.context} -> tactic
+ val mk_set_nat_tac: int -> (int -> tactic) -> thm list list -> thm list -> cterm list ->
+ thm list -> int -> {prems: 'a, context: Proof.context} -> tactic
+ val mk_set_natural_tac: thm -> tactic
+ val mk_set_simp_tac: thm -> thm -> thm list -> tactic
+ val mk_set_tac: thm -> tactic
+ val mk_srel_simp_tac: thm list -> int -> thm -> thm -> thm -> thm -> thm list -> thm ->
+ thm -> thm list -> thm list -> thm list list -> tactic
+ val mk_wit_tac: int -> thm list -> thm list -> tactic
+ val mk_wpull_tac: thm -> tactic
+end;
+
+structure BNF_LFP_Tactics : BNF_LFP_TACTICS =
+struct
+
+open BNF_Tactics
+open BNF_LFP_Util
+open BNF_Util
+
+val fst_snd_convs = @{thms fst_conv snd_conv};
+val id_apply = @{thm id_apply};
+val meta_mp = @{thm meta_mp};
+val ord_eq_le_trans = @{thm ord_eq_le_trans};
+val subset_trans = @{thm subset_trans};
+val trans_fun_cong_image_id_id_apply = @{thm trans[OF fun_cong[OF image_id] id_apply]};
+
+fun mk_alg_set_tac alg_def =
+ dtac (alg_def RS iffD1) 1 THEN
+ REPEAT_DETERM (etac conjE 1) THEN
+ EVERY' [etac bspec, rtac CollectI] 1 THEN
+ REPEAT_DETERM (etac conjI 1) THEN atac 1;
+
+fun mk_alg_not_empty_tac alg_set alg_sets wits =
+ (EVERY' [rtac notI, hyp_subst_tac, ftac alg_set] THEN'
+ REPEAT_DETERM o FIRST'
+ [rtac subset_UNIV,
+ EVERY' [rtac @{thm subset_emptyI}, eresolve_tac wits],
+ EVERY' [rtac subsetI, rtac FalseE, eresolve_tac wits],
+ EVERY' [rtac subsetI, dresolve_tac wits, hyp_subst_tac,
+ FIRST' (map (fn thm => rtac thm THEN' atac) alg_sets)]] THEN'
+ etac @{thm emptyE}) 1;
+
+fun mk_mor_elim_tac mor_def =
+ (dtac (subst OF [mor_def]) THEN'
+ REPEAT o etac conjE THEN'
+ TRY o rtac @{thm image_subsetI} THEN'
+ etac bspec THEN'
+ atac) 1;
+
+fun mk_mor_incl_tac mor_def map_id's =
+ (stac mor_def THEN'
+ rtac conjI THEN'
+ CONJ_WRAP' (K (EVERY' [rtac ballI, etac set_mp, stac id_apply, atac])) map_id's THEN'
+ CONJ_WRAP' (fn thm =>
+ (EVERY' [rtac ballI, rtac trans, rtac id_apply, stac thm, rtac refl])) map_id's) 1;
+
+fun mk_mor_comp_tac mor_def set_natural's map_comp_ids =
+ let
+ val fbetw_tac = EVERY' [rtac ballI, stac o_apply, etac bspec, etac bspec, atac];
+ fun mor_tac (set_natural', map_comp_id) =
+ EVERY' [rtac ballI, stac o_apply, rtac trans,
+ rtac trans, dtac @{thm rev_bspec}, atac, etac arg_cong,
+ REPEAT o eresolve_tac [CollectE, conjE], etac bspec, rtac CollectI] THEN'
+ CONJ_WRAP' (fn thm =>
+ FIRST' [rtac subset_UNIV,
+ (EVERY' [rtac ord_eq_le_trans, rtac thm, rtac @{thm image_subsetI},
+ etac bspec, etac set_mp, atac])]) set_natural' THEN'
+ rtac (map_comp_id RS arg_cong);
+ in
+ (dtac (mor_def RS subst) THEN' dtac (mor_def RS subst) THEN' stac mor_def THEN'
+ REPEAT o etac conjE THEN'
+ rtac conjI THEN'
+ CONJ_WRAP' (K fbetw_tac) set_natural's THEN'
+ CONJ_WRAP' mor_tac (set_natural's ~~ map_comp_ids)) 1
+ end;
+
+fun mk_mor_inv_tac alg_def mor_def set_natural's morEs map_comp_ids map_congLs =
+ let
+ val fbetw_tac = EVERY' [rtac ballI, etac set_mp, etac imageI];
+ fun Collect_tac set_natural' =
+ CONJ_WRAP' (fn thm =>
+ FIRST' [rtac subset_UNIV,
+ (EVERY' [rtac ord_eq_le_trans, rtac thm, rtac subset_trans,
+ etac @{thm image_mono}, atac])]) set_natural';
+ fun mor_tac (set_natural', ((morE, map_comp_id), map_congL)) =
+ EVERY' [rtac ballI, ftac @{thm rev_bspec}, atac,
+ REPEAT o eresolve_tac [CollectE, conjE], rtac sym, rtac trans, rtac sym,
+ etac @{thm inverE}, etac bspec, rtac CollectI, Collect_tac set_natural',
+ rtac trans, etac (morE RS arg_cong), rtac CollectI, Collect_tac set_natural',
+ rtac trans, rtac (map_comp_id RS arg_cong), rtac (map_congL RS arg_cong),
+ REPEAT_DETERM_N (length morEs) o
+ (EVERY' [rtac subst, rtac @{thm inver_pointfree}, etac @{thm inver_mono}, atac])];
+ in
+ (stac mor_def THEN'
+ dtac (alg_def RS iffD1) THEN'
+ dtac (alg_def RS iffD1) THEN'
+ REPEAT o etac conjE THEN'
+ rtac conjI THEN'
+ CONJ_WRAP' (K fbetw_tac) set_natural's THEN'
+ CONJ_WRAP' mor_tac (set_natural's ~~ (morEs ~~ map_comp_ids ~~ map_congLs))) 1
+ end;
+
+fun mk_mor_str_tac ks mor_def =
+ (stac mor_def THEN' rtac conjI THEN'
+ CONJ_WRAP' (K (EVERY' [rtac ballI, rtac UNIV_I])) ks THEN'
+ CONJ_WRAP' (K (EVERY' [rtac ballI, rtac refl])) ks) 1;
+
+fun mk_mor_convol_tac ks mor_def =
+ (stac mor_def THEN' rtac conjI THEN'
+ CONJ_WRAP' (K (EVERY' [rtac ballI, rtac UNIV_I])) ks THEN'
+ CONJ_WRAP' (K (EVERY' [rtac ballI, rtac trans, rtac @{thm fst_convol'}, rtac o_apply])) ks) 1;
+
+fun mk_mor_UNIV_tac m morEs mor_def =
+ let
+ val n = length morEs;
+ fun mor_tac morE = EVERY' [rtac ext, rtac trans, rtac o_apply, rtac trans, etac morE,
+ rtac CollectI, CONJ_WRAP' (K (rtac subset_UNIV)) (1 upto m + n),
+ rtac sym, rtac o_apply];
+ in
+ EVERY' [rtac iffI, CONJ_WRAP' mor_tac morEs,
+ stac mor_def, rtac conjI, CONJ_WRAP' (K (rtac ballI THEN' rtac UNIV_I)) morEs,
+ REPEAT_DETERM o etac conjE, REPEAT_DETERM_N n o dtac (@{thm fun_eq_iff} RS subst),
+ CONJ_WRAP' (K (EVERY' [rtac ballI, REPEAT_DETERM o etac allE, rtac trans,
+ etac (o_apply RS subst), rtac o_apply])) morEs] 1
+ end;
+
+fun mk_iso_alt_tac mor_images mor_inv =
+ let
+ val n = length mor_images;
+ fun if_wrap_tac thm =
+ EVERY' [rtac ssubst, rtac @{thm bij_betw_iff_ex}, rtac exI, rtac conjI,
+ rtac @{thm inver_surj}, etac thm, etac thm, atac, etac conjI, atac]
+ val if_tac =
+ EVERY' [etac thin_rl, etac thin_rl, REPEAT o eresolve_tac [conjE, exE],
+ rtac conjI, atac, CONJ_WRAP' if_wrap_tac mor_images];
+ val only_if_tac =
+ EVERY' [rtac conjI, etac conjunct1, EVERY' (map (fn thm =>
+ EVERY' [rtac exE, rtac @{thm bij_betw_ex_weakE}, etac (conjunct2 RS thm)])
+ (map (mk_conjunctN n) (1 upto n))), REPEAT o rtac exI, rtac conjI, rtac mor_inv,
+ etac conjunct1, atac, atac, REPEAT_DETERM_N n o atac,
+ CONJ_WRAP' (K (etac conjunct2)) mor_images];
+ in
+ (rtac iffI THEN' if_tac THEN' only_if_tac) 1
+ end;
+
+fun mk_copy_str_tac set_natural's alg_def alg_sets =
+ let
+ val n = length alg_sets;
+ val bij_betw_inv_tac =
+ EVERY' [etac thin_rl, REPEAT_DETERM_N n o EVERY' [dtac @{thm bij_betwI}, atac, atac],
+ REPEAT_DETERM_N (2 * n) o etac thin_rl, REPEAT_DETERM_N (n - 1) o etac conjI, atac];
+ fun set_tac thms =
+ EVERY' [rtac ord_eq_le_trans, resolve_tac thms, rtac subset_trans,
+ etac @{thm image_mono}, rtac equalityD1, etac @{thm bij_betw_imageE}];
+ val copy_str_tac =
+ CONJ_WRAP' (fn (thms, thm) =>
+ EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac set_mp,
+ rtac equalityD1, etac @{thm bij_betw_imageE}, rtac imageI, etac thm,
+ REPEAT_DETERM o rtac subset_UNIV, REPEAT_DETERM_N n o (set_tac thms)])
+ (set_natural's ~~ alg_sets);
+ in
+ (rtac rev_mp THEN' DETERM o bij_betw_inv_tac THEN' rtac impI THEN'
+ stac alg_def THEN' copy_str_tac) 1
+ end;
+
+fun mk_copy_alg_tac set_natural's alg_sets mor_def iso_alt copy_str =
+ let
+ val n = length alg_sets;
+ val fbetw_tac = CONJ_WRAP' (K (etac @{thm bij_betwE})) alg_sets;
+ fun set_tac thms =
+ EVERY' [rtac ord_eq_le_trans, resolve_tac thms, rtac subset_trans,
+ REPEAT_DETERM o etac conjE, etac @{thm image_mono},
+ rtac equalityD1, etac @{thm bij_betw_imageE}];
+ val mor_tac =
+ CONJ_WRAP' (fn (thms, thm) =>
+ EVERY' [rtac ballI, etac CollectE, etac @{thm inverE}, etac thm,
+ REPEAT_DETERM o rtac subset_UNIV, REPEAT_DETERM_N n o (set_tac thms)])
+ (set_natural's ~~ alg_sets);
+ in
+ (rtac (iso_alt RS @{thm ssubst[of _ _ "%x. x"]}) THEN'
+ etac copy_str THEN' REPEAT_DETERM o atac THEN'
+ rtac conjI THEN' stac mor_def THEN' rtac conjI THEN' fbetw_tac THEN' mor_tac THEN'
+ CONJ_WRAP' (K atac) alg_sets) 1
+ end;
+
+fun mk_ex_copy_alg_tac n copy_str copy_alg =
+ EVERY' [REPEAT_DETERM_N n o rtac exI, rtac conjI, etac copy_str,
+ REPEAT_DETERM_N n o atac,
+ REPEAT_DETERM_N n o etac @{thm bij_betw_inver2},
+ REPEAT_DETERM_N n o etac @{thm bij_betw_inver1}, etac copy_alg,
+ REPEAT_DETERM_N n o atac,
+ REPEAT_DETERM_N n o etac @{thm bij_betw_inver2},
+ REPEAT_DETERM_N n o etac @{thm bij_betw_inver1}] 1;
+
+fun mk_bd_limit_tac n bd_Cinfinite =
+ EVERY' [REPEAT_DETERM o etac conjE, rtac rev_mp, rtac @{thm Cinfinite_limit_finite},
+ REPEAT_DETERM_N n o rtac @{thm finite.insertI}, rtac @{thm finite.emptyI},
+ REPEAT_DETERM_N n o etac @{thm insert_subsetI}, rtac @{thm empty_subsetI},
+ rtac bd_Cinfinite, rtac impI, etac bexE, rtac bexI,
+ CONJ_WRAP' (fn i =>
+ EVERY' [etac bspec, REPEAT_DETERM_N i o rtac @{thm insertI2}, rtac @{thm insertI1}])
+ (0 upto n - 1),
+ atac] 1;
+
+fun mk_min_algs_tac worel in_congs =
+ let
+ val minG_tac = EVERY' [rtac @{thm UN_cong}, rtac refl, dtac bspec, atac, etac arg_cong];
+ fun minH_tac thm =
+ EVERY' [rtac @{thm Un_cong}, minG_tac, rtac @{thm image_cong}, rtac thm,
+ REPEAT_DETERM_N (length in_congs) o minG_tac, rtac refl];
+ in
+ (rtac (worel RS (@{thm wo_rel.worec_fixpoint} RS fun_cong)) THEN' rtac ssubst THEN'
+ rtac meta_eq_to_obj_eq THEN' rtac (worel RS @{thm wo_rel.adm_wo_def}) THEN'
+ REPEAT_DETERM_N 3 o rtac allI THEN' rtac impI THEN'
+ CONJ_WRAP_GEN' (EVERY' [rtac Pair_eqI, rtac conjI]) minH_tac in_congs) 1
+ end;
+
+fun mk_min_algs_mono_tac min_algs = EVERY' [stac @{thm relChain_def}, rtac allI, rtac allI,
+ rtac impI, rtac @{thm case_split}, rtac @{thm xt1(3)}, rtac min_algs, etac @{thm FieldI2},
+ rtac subsetI, rtac UnI1, rtac @{thm UN_I}, etac @{thm underS_I}, atac, atac,
+ rtac equalityD1, dtac @{thm notnotD}, hyp_subst_tac, rtac refl] 1;
+
+fun mk_min_algs_card_of_tac cT ct m worel min_algs in_bds bd_Card_order bd_Cnotzero
+ suc_Card_order suc_Cinfinite suc_Cnotzero suc_Asuc Asuc_Cinfinite Asuc_Cnotzero =
+ let
+ val induct = worel RS
+ Drule.instantiate' [SOME cT] [NONE, SOME ct] @{thm well_order_induct_imp};
+ val src = 1 upto m + 1;
+ val dest = (m + 1) :: (1 upto m);
+ val absorbAs_tac = if m = 0 then K (all_tac)
+ else EVERY' [rtac @{thm ordIso_transitive}, rtac @{thm csum_cong1},
+ rtac @{thm ordIso_transitive},
+ BNF_Tactics.mk_rotate_eq_tac (rtac @{thm ordIso_refl} THEN'
+ FIRST' [rtac @{thm card_of_Card_order}, rtac @{thm Card_order_csum},
+ rtac @{thm Card_order_cexp}])
+ @{thm ordIso_transitive} @{thm csum_assoc} @{thm csum_com} @{thm csum_cong}
+ src dest,
+ rtac @{thm csum_absorb1}, rtac Asuc_Cinfinite, rtac ctrans, rtac @{thm ordLeq_csum1},
+ FIRST' [rtac @{thm Card_order_csum}, rtac @{thm card_of_Card_order}],
+ rtac @{thm ordLeq_cexp1}, rtac suc_Cnotzero, rtac @{thm Card_order_csum}];
+
+ val minG_tac = EVERY' [rtac @{thm UNION_Cinfinite_bound}, rtac @{thm ordLess_imp_ordLeq},
+ rtac @{thm ordLess_transitive}, rtac @{thm card_of_underS}, rtac suc_Card_order,
+ atac, rtac suc_Asuc, rtac ballI, etac allE, dtac mp, etac @{thm underS_E},
+ dtac mp, etac @{thm underS_Field}, REPEAT o etac conjE, atac, rtac Asuc_Cinfinite]
+
+ fun mk_minH_tac (min_alg, in_bd) = EVERY' [rtac @{thm ordIso_ordLeq_trans},
+ rtac @{thm card_of_ordIso_subst}, etac min_alg, rtac @{thm Un_Cinfinite_bound},
+ minG_tac, rtac ctrans, rtac @{thm card_of_image}, rtac ctrans, rtac in_bd, rtac ctrans,
+ rtac @{thm cexp_mono1_Cnotzero}, rtac @{thm csum_mono1},
+ REPEAT_DETERM_N m o rtac @{thm csum_mono2},
+ CONJ_WRAP_GEN' (rtac @{thm csum_cinfinite_bound}) (K minG_tac) min_algs,
+ REPEAT_DETERM o FIRST'
+ [rtac @{thm card_of_Card_order}, rtac @{thm Card_order_csum}, rtac Asuc_Cinfinite],
+ rtac @{thm csum_Cnotzero2}, rtac @{thm ctwo_Cnotzero}, rtac bd_Card_order,
+ rtac @{thm ordIso_ordLeq_trans}, rtac @{thm cexp_cong1_Cnotzero}, absorbAs_tac,
+ rtac @{thm csum_absorb1}, rtac Asuc_Cinfinite, rtac @{thm ctwo_ordLeq_Cinfinite},
+ rtac Asuc_Cinfinite, rtac bd_Card_order,
+ rtac @{thm csum_Cnotzero2}, rtac @{thm ctwo_Cnotzero}, rtac Asuc_Cnotzero,
+ rtac @{thm ordIso_imp_ordLeq}, rtac @{thm cexp_cprod_ordLeq},
+ TRY o rtac @{thm csum_Cnotzero2}, rtac @{thm ctwo_Cnotzero}, rtac suc_Cinfinite,
+ rtac bd_Cnotzero, rtac @{thm cardSuc_ordLeq}, rtac bd_Card_order, rtac Asuc_Cinfinite];
+ in
+ (rtac induct THEN'
+ rtac impI THEN'
+ CONJ_WRAP' mk_minH_tac (min_algs ~~ in_bds)) 1
+ end;
+
+fun mk_min_algs_least_tac cT ct worel min_algs alg_sets =
+ let
+ val induct = worel RS
+ Drule.instantiate' [SOME cT] [NONE, SOME ct] @{thm well_order_induct_imp};
+
+ val minG_tac = EVERY' [rtac @{thm UN_least}, etac allE, dtac mp, etac @{thm underS_E},
+ dtac mp, etac @{thm underS_Field}, REPEAT_DETERM o etac conjE, atac];
+
+ fun mk_minH_tac (min_alg, alg_set) = EVERY' [rtac ord_eq_le_trans, etac min_alg,
+ rtac @{thm Un_least}, minG_tac, rtac @{thm image_subsetI},
+ REPEAT_DETERM o eresolve_tac [CollectE, conjE], etac alg_set,
+ REPEAT_DETERM o FIRST' [atac, etac subset_trans THEN' minG_tac]];
+ in
+ (rtac induct THEN'
+ rtac impI THEN'
+ CONJ_WRAP' mk_minH_tac (min_algs ~~ alg_sets)) 1
+ end;
+
+fun mk_alg_min_alg_tac m alg_def min_alg_defs bd_limit bd_Cinfinite
+ set_bdss min_algs min_alg_monos =
+ let
+ val n = length min_algs;
+ fun mk_cardSuc_UNION_tac set_bds (mono, def) = EVERY'
+ [rtac bexE, rtac @{thm cardSuc_UNION_Cinfinite}, rtac bd_Cinfinite, rtac mono,
+ etac (def RSN (2, @{thm subset_trans[OF _ equalityD1]})), resolve_tac set_bds];
+ fun mk_conjunct_tac (set_bds, (min_alg, min_alg_def)) =
+ EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, conjE],
+ EVERY' (map (mk_cardSuc_UNION_tac set_bds) (min_alg_monos ~~ min_alg_defs)), rtac bexE,
+ rtac bd_limit, REPEAT_DETERM_N (n - 1) o etac conjI, atac,
+ rtac (min_alg_def RS @{thm set_mp[OF equalityD2]}),
+ rtac @{thm UN_I}, REPEAT_DETERM_N (m + 3 * n) o etac thin_rl, atac, rtac set_mp,
+ rtac equalityD2, rtac min_alg, atac, rtac UnI2, rtac @{thm image_eqI}, rtac refl,
+ rtac CollectI, REPEAT_DETERM_N m o dtac asm_rl, REPEAT_DETERM_N n o etac thin_rl,
+ REPEAT_DETERM o etac conjE,
+ CONJ_WRAP' (K (FIRST' [atac,
+ EVERY' [etac subset_trans, rtac subsetI, rtac @{thm UN_I},
+ etac @{thm underS_I}, atac, atac]]))
+ set_bds];
+ in
+ (rtac (alg_def RS @{thm ssubst[of _ _ "%x. x"]}) THEN'
+ CONJ_WRAP' mk_conjunct_tac (set_bdss ~~ (min_algs ~~ min_alg_defs))) 1
+ end;
+
+fun mk_card_of_min_alg_tac min_alg_def card_of suc_Card_order suc_Asuc Asuc_Cinfinite =
+ EVERY' [stac min_alg_def, rtac @{thm UNION_Cinfinite_bound},
+ rtac @{thm ordIso_ordLeq_trans}, rtac @{thm card_of_Field_ordIso}, rtac suc_Card_order,
+ rtac @{thm ordLess_imp_ordLeq}, rtac suc_Asuc, rtac ballI, dtac rev_mp, rtac card_of,
+ REPEAT_DETERM o etac conjE, atac, rtac Asuc_Cinfinite] 1;
+
+fun mk_least_min_alg_tac min_alg_def least =
+ EVERY' [stac min_alg_def, rtac @{thm UN_least}, dtac least, dtac mp, atac,
+ REPEAT_DETERM o etac conjE, atac] 1;
+
+fun mk_alg_select_tac Abs_inverse {context = ctxt, prems = _} =
+ EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac] 1 THEN
+ unfold_thms_tac ctxt (Abs_inverse :: fst_snd_convs) THEN atac 1;
+
+fun mk_mor_select_tac mor_def mor_cong mor_comp mor_incl_min_alg alg_def alg_select
+ alg_sets set_natural's str_init_defs =
+ let
+ val n = length alg_sets;
+ val fbetw_tac =
+ CONJ_WRAP' (K (EVERY' [rtac ballI, etac @{thm rev_bspec}, etac CollectE, atac])) alg_sets;
+ val mor_tac =
+ CONJ_WRAP' (fn thm => EVERY' [rtac ballI, rtac thm]) str_init_defs;
+ fun alg_epi_tac ((alg_set, str_init_def), set_natural') =
+ EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac CollectI,
+ rtac ballI, ftac (alg_select RS bspec), stac str_init_def, etac alg_set,
+ REPEAT_DETERM o FIRST' [rtac subset_UNIV,
+ EVERY' [rtac ord_eq_le_trans, resolve_tac set_natural', rtac subset_trans,
+ etac @{thm image_mono}, rtac @{thm image_Collect_subsetI}, etac bspec, atac]]];
+ in
+ (rtac mor_cong THEN' REPEAT_DETERM_N n o (rtac sym THEN' rtac @{thm o_id}) THEN'
+ rtac (Thm.permute_prems 0 1 mor_comp) THEN' etac (Thm.permute_prems 0 1 mor_comp) THEN'
+ stac mor_def THEN' rtac conjI THEN' fbetw_tac THEN' mor_tac THEN' rtac mor_incl_min_alg THEN'
+ stac alg_def THEN' CONJ_WRAP' alg_epi_tac ((alg_sets ~~ str_init_defs) ~~ set_natural's)) 1
+ end;
+
+fun mk_init_ex_mor_tac Abs_inverse copy_alg_ex alg_min_alg card_of_min_algs
+ mor_comp mor_select mor_incl_min_alg {context = ctxt, prems = _} =
+ let
+ val n = length card_of_min_algs;
+ val card_of_ordIso_tac = EVERY' [rtac ssubst, rtac @{thm card_of_ordIso},
+ rtac @{thm ordIso_symmetric}, rtac conjunct1, rtac conjunct2, atac];
+ fun internalize_tac card_of = EVERY' [rtac subst, rtac @{thm internalize_card_of_ordLeq2},
+ rtac @{thm ordLeq_ordIso_trans}, rtac card_of, rtac subst,
+ rtac @{thm Card_order_iff_ordIso_card_of}, rtac @{thm Card_order_cexp}];
+ in
+ (rtac rev_mp THEN'
+ REPEAT_DETERM_N (2 * n) o (rtac mp THEN' rtac @{thm ex_mono} THEN' rtac impI) THEN'
+ REPEAT_DETERM_N (n + 1) o etac thin_rl THEN' rtac (alg_min_alg RS copy_alg_ex) THEN'
+ REPEAT_DETERM_N n o atac THEN'
+ REPEAT_DETERM_N n o card_of_ordIso_tac THEN'
+ EVERY' (map internalize_tac card_of_min_algs) THEN'
+ rtac impI THEN'
+ REPEAT_DETERM o eresolve_tac [exE, conjE] THEN'
+ REPEAT_DETERM o rtac exI THEN'
+ rtac mor_select THEN' atac THEN' rtac CollectI THEN'
+ REPEAT_DETERM o rtac exI THEN'
+ rtac conjI THEN' rtac refl THEN' atac THEN'
+ K (unfold_thms_tac ctxt (Abs_inverse :: fst_snd_convs)) THEN'
+ etac mor_comp THEN' etac mor_incl_min_alg) 1
+ end;
+
+fun mk_init_unique_mor_tac m
+ alg_def alg_min_alg least_min_algs in_monos alg_sets morEs map_congs =
+ let
+ val n = length least_min_algs;
+ val ks = (1 upto n);
+
+ fun mor_tac morE in_mono = EVERY' [etac morE, rtac set_mp, rtac in_mono,
+ REPEAT_DETERM_N n o rtac @{thm Collect_restrict}, rtac CollectI,
+ REPEAT_DETERM_N (m + n) o (TRY o rtac conjI THEN' atac)];
+ fun cong_tac map_cong = EVERY' [rtac (map_cong RS arg_cong),
+ REPEAT_DETERM_N m o rtac refl,
+ REPEAT_DETERM_N n o (etac @{thm prop_restrict} THEN' atac)];
+
+ fun mk_alg_tac (alg_set, (in_mono, (morE, map_cong))) = EVERY' [rtac ballI, rtac CollectI,
+ REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac conjI, rtac (alg_min_alg RS alg_set),
+ REPEAT_DETERM_N m o rtac subset_UNIV,
+ REPEAT_DETERM_N n o (etac subset_trans THEN' rtac @{thm Collect_restrict}),
+ rtac trans, mor_tac morE in_mono,
+ rtac trans, cong_tac map_cong,
+ rtac sym, mor_tac morE in_mono];
+
+ fun mk_unique_tac (k, least_min_alg) =
+ select_prem_tac n (etac @{thm prop_restrict}) k THEN' rtac least_min_alg THEN'
+ stac alg_def THEN'
+ CONJ_WRAP' mk_alg_tac (alg_sets ~~ (in_monos ~~ (morEs ~~ map_congs)));
+ in
+ CONJ_WRAP' mk_unique_tac (ks ~~ least_min_algs) 1
+ end;
+
+fun mk_init_induct_tac m alg_def alg_min_alg least_min_algs alg_sets =
+ let
+ val n = length least_min_algs;
+
+ fun mk_alg_tac alg_set = EVERY' [rtac ballI, rtac CollectI,
+ REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac conjI, rtac (alg_min_alg RS alg_set),
+ REPEAT_DETERM_N m o rtac subset_UNIV,
+ REPEAT_DETERM_N n o (etac subset_trans THEN' rtac @{thm Collect_restrict}),
+ rtac mp, etac bspec, rtac CollectI,
+ REPEAT_DETERM_N m o (rtac conjI THEN' atac),
+ CONJ_WRAP' (K (etac subset_trans THEN' rtac @{thm Collect_restrict})) alg_sets,
+ CONJ_WRAP' (K (rtac ballI THEN' etac @{thm prop_restrict} THEN' atac)) alg_sets];
+
+ fun mk_induct_tac least_min_alg =
+ rtac ballI THEN' etac @{thm prop_restrict} THEN' rtac least_min_alg THEN'
+ stac alg_def THEN'
+ CONJ_WRAP' mk_alg_tac alg_sets;
+ in
+ CONJ_WRAP' mk_induct_tac least_min_algs 1
+ end;
+
+fun mk_mor_Rep_tac ctor_defs copy bijs inver_Abss inver_Reps {context = ctxt, prems = _} =
+ (K (unfold_thms_tac ctxt ctor_defs) THEN' rtac conjunct1 THEN' rtac copy THEN'
+ EVERY' (map (fn bij => EVERY' [rtac bij, atac, etac bexI, rtac UNIV_I]) bijs) THEN'
+ EVERY' (map rtac inver_Abss) THEN'
+ EVERY' (map rtac inver_Reps)) 1;
+
+fun mk_mor_Abs_tac inv inver_Abss inver_Reps =
+ (rtac inv THEN'
+ EVERY' (map2 (fn inver_Abs => fn inver_Rep =>
+ EVERY' [rtac conjI, rtac subset_UNIV, rtac conjI, rtac inver_Rep, rtac inver_Abs])
+ inver_Abss inver_Reps)) 1;
+
+fun mk_mor_fold_tac cT ct fold_defs ex_mor mor =
+ (EVERY' (map stac fold_defs) THEN' EVERY' [rtac rev_mp, rtac ex_mor, rtac impI] THEN'
+ REPEAT_DETERM_N (length fold_defs) o etac exE THEN'
+ rtac (Drule.instantiate' [SOME cT] [SOME ct] @{thm someI}) THEN' etac mor) 1;
+
+fun mk_fold_unique_mor_tac type_defs init_unique_mors Reps mor_comp mor_Abs mor_fold =
+ let
+ fun mk_unique type_def =
+ EVERY' [rtac @{thm surj_fun_eq}, rtac (type_def RS @{thm type_definition.Abs_image}),
+ rtac ballI, resolve_tac init_unique_mors,
+ EVERY' (map (fn thm => atac ORELSE' rtac thm) Reps),
+ rtac mor_comp, rtac mor_Abs, atac,
+ rtac mor_comp, rtac mor_Abs, rtac mor_fold];
+ in
+ CONJ_WRAP' mk_unique type_defs 1
+ end;
+
+fun mk_dtor_o_ctor_tac dtor_def foldx map_comp_id map_congL ctor_o_folds =
+ EVERY' [stac dtor_def, rtac ext, rtac trans, rtac o_apply, rtac trans, rtac foldx,
+ rtac trans, rtac map_comp_id, rtac trans, rtac map_congL,
+ EVERY' (map (fn thm => rtac ballI THEN' rtac (trans OF [thm RS fun_cong, id_apply]))
+ ctor_o_folds),
+ rtac sym, rtac id_apply] 1;
+
+fun mk_rec_tac rec_defs foldx fst_recs {context = ctxt, prems = _}=
+ unfold_thms_tac ctxt
+ (rec_defs @ map (fn thm => thm RS @{thm convol_expand_snd}) fst_recs) THEN
+ EVERY' [rtac trans, rtac o_apply, rtac trans, rtac (foldx RS @{thm arg_cong[of _ _ snd]}),
+ rtac @{thm snd_convol'}] 1;
+
+fun mk_ctor_induct_tac m set_natural'ss init_induct morEs mor_Abs Rep_invs Abs_invs Reps =
+ let
+ val n = length set_natural'ss;
+ val ks = 1 upto n;
+
+ fun mk_IH_tac Rep_inv Abs_inv set_natural' =
+ DETERM o EVERY' [dtac meta_mp, rtac (Rep_inv RS arg_cong RS subst), etac bspec,
+ dtac set_rev_mp, rtac equalityD1, rtac set_natural', etac imageE,
+ hyp_subst_tac, rtac (Abs_inv RS ssubst), etac set_mp, atac, atac];
+
+ fun mk_closed_tac (k, (morE, set_natural's)) =
+ EVERY' [select_prem_tac n (dtac asm_rl) k, rtac ballI, rtac impI,
+ rtac (mor_Abs RS morE RS arg_cong RS ssubst), atac,
+ REPEAT_DETERM o eresolve_tac [CollectE, conjE], dtac @{thm meta_spec},
+ EVERY' (map3 mk_IH_tac Rep_invs Abs_invs (drop m set_natural's)), atac];
+
+ fun mk_induct_tac (Rep, Rep_inv) =
+ EVERY' [rtac (Rep_inv RS arg_cong RS subst), etac (Rep RSN (2, bspec))];
+ in
+ (rtac mp THEN' rtac impI THEN'
+ DETERM o CONJ_WRAP_GEN' (etac conjE THEN' rtac conjI) mk_induct_tac (Reps ~~ Rep_invs) THEN'
+ rtac init_induct THEN'
+ DETERM o CONJ_WRAP' mk_closed_tac (ks ~~ (morEs ~~ set_natural'ss))) 1
+ end;
+
+fun mk_ctor_induct2_tac cTs cts ctor_induct weak_ctor_inducts {context = ctxt, prems = _} =
+ let
+ val n = length weak_ctor_inducts;
+ val ks = 1 upto n;
+ fun mk_inner_induct_tac induct i =
+ EVERY' [rtac allI, fo_rtac induct ctxt,
+ select_prem_tac n (dtac @{thm meta_spec2}) i,
+ REPEAT_DETERM_N n o
+ EVERY' [dtac meta_mp THEN_ALL_NEW Goal.norm_hhf_tac,
+ REPEAT_DETERM o dtac @{thm meta_spec}, etac (spec RS meta_mp), atac],
+ atac];
+ in
+ EVERY' [rtac rev_mp, rtac (Drule.instantiate' cTs cts ctor_induct),
+ EVERY' (map2 mk_inner_induct_tac weak_ctor_inducts ks), rtac impI,
+ REPEAT_DETERM o eresolve_tac [conjE, allE],
+ CONJ_WRAP' (K atac) ks] 1
+ end;
+
+fun mk_map_tac m n foldx map_comp_id map_cong =
+ EVERY' [rtac ext, rtac trans, rtac o_apply, rtac trans, rtac foldx, rtac trans, rtac o_apply,
+ rtac trans, rtac (map_comp_id RS arg_cong), rtac trans, rtac (map_cong RS arg_cong),
+ REPEAT_DETERM_N m o rtac refl,
+ REPEAT_DETERM_N n o (EVERY' (map rtac [trans, o_apply, id_apply])),
+ rtac sym, rtac o_apply] 1;
+
+fun mk_map_unique_tac m mor_def fold_unique_mor map_comp_ids map_congs =
+ let
+ val n = length map_congs;
+ fun mk_mor (comp_id, cong) = EVERY' [rtac ballI, rtac trans, etac @{thm pointfreeE},
+ rtac sym, rtac trans, rtac o_apply, rtac trans, rtac (comp_id RS arg_cong),
+ rtac (cong RS arg_cong),
+ REPEAT_DETERM_N m o rtac refl,
+ REPEAT_DETERM_N n o (EVERY' (map rtac [trans, o_apply, id_apply]))];
+ in
+ EVERY' [rtac fold_unique_mor, rtac ssubst, rtac mor_def, rtac conjI,
+ CONJ_WRAP' (K (EVERY' [rtac ballI, rtac UNIV_I])) map_congs,
+ CONJ_WRAP' mk_mor (map_comp_ids ~~ map_congs)] 1
+ end;
+
+fun mk_set_tac foldx = EVERY' [rtac ext, rtac trans, rtac o_apply,
+ rtac trans, rtac foldx, rtac sym, rtac o_apply] 1;
+
+fun mk_set_simp_tac set set_natural' set_natural's =
+ let
+ val n = length set_natural's;
+ fun mk_UN thm = rtac (thm RS @{thm arg_cong[of _ _ Union]} RS trans) THEN'
+ rtac @{thm Union_image_eq};
+ in
+ EVERY' [rtac (set RS @{thm pointfreeE} RS trans), rtac @{thm Un_cong},
+ rtac (trans OF [set_natural', trans_fun_cong_image_id_id_apply]),
+ REPEAT_DETERM_N (n - 1) o rtac @{thm Un_cong},
+ EVERY' (map mk_UN set_natural's)] 1
+ end;
+
+fun mk_set_nat_tac m induct_tac set_natural'ss
+ map_simps csets set_simps i {context = ctxt, prems = _} =
+ let
+ val n = length map_simps;
+
+ fun useIH set_nat = EVERY' [rtac trans, rtac @{thm image_UN}, rtac trans, rtac @{thm UN_cong},
+ rtac refl, Goal.assume_rule_tac ctxt, rtac sym, rtac trans, rtac @{thm UN_cong},
+ rtac set_nat, rtac refl, rtac @{thm UN_simps(10)}];
+
+ fun mk_set_nat cset map_simp set_simp set_nats =
+ EVERY' [rtac trans, rtac @{thm image_cong}, rtac set_simp, rtac refl,
+ rtac sym, rtac (trans OF [map_simp RS HOL_arg_cong cset, set_simp RS trans]),
+ rtac sym, EVERY' (map rtac (trans :: @{thms image_Un Un_cong})),
+ rtac sym, rtac (nth set_nats (i - 1)),
+ REPEAT_DETERM_N (n - 1) o EVERY' (map rtac (trans :: @{thms image_Un Un_cong})),
+ EVERY' (map useIH (drop m set_nats))];
+ in
+ (induct_tac THEN' EVERY' (map4 mk_set_nat csets map_simps set_simps set_natural'ss)) 1
+ end;
+
+fun mk_set_bd_tac m induct_tac bd_Cinfinite set_bdss set_simps i {context = ctxt, prems = _} =
+ let
+ val n = length set_simps;
+
+ fun useIH set_bd = EVERY' [rtac @{thm UNION_Cinfinite_bound}, rtac set_bd, rtac ballI,
+ Goal.assume_rule_tac ctxt, rtac bd_Cinfinite];
+
+ fun mk_set_nat set_simp set_bds =
+ EVERY' [rtac @{thm ordIso_ordLeq_trans}, rtac @{thm card_of_ordIso_subst}, rtac set_simp,
+ rtac (bd_Cinfinite RSN (3, @{thm Un_Cinfinite_bound})), rtac (nth set_bds (i - 1)),
+ REPEAT_DETERM_N (n - 1) o rtac (bd_Cinfinite RSN (3, @{thm Un_Cinfinite_bound})),
+ EVERY' (map useIH (drop m set_bds))];
+ in
+ (induct_tac THEN' EVERY' (map2 mk_set_nat set_simps set_bdss)) 1
+ end;
+
+fun mk_mcong_tac induct_tac set_setsss map_congs map_simps {context = ctxt, prems = _} =
+ let
+ fun use_asm thm = EVERY' [etac bspec, etac set_rev_mp, rtac thm];
+
+ fun useIH set_sets = EVERY' [rtac mp, Goal.assume_rule_tac ctxt,
+ CONJ_WRAP' (fn thm =>
+ EVERY' [rtac ballI, etac bspec, etac set_rev_mp, etac thm]) set_sets];
+
+ fun mk_map_cong map_simp map_cong set_setss =
+ EVERY' [rtac impI, REPEAT_DETERM o etac conjE,
+ rtac trans, rtac map_simp, rtac trans, rtac (map_cong RS arg_cong),
+ EVERY' (map use_asm (map hd set_setss)),
+ EVERY' (map useIH (transpose (map tl set_setss))),
+ rtac sym, rtac map_simp];
+ in
+ (induct_tac THEN' EVERY' (map3 mk_map_cong map_simps map_congs set_setsss)) 1
+ end;
+
+fun mk_incl_min_alg_tac induct_tac set_setsss alg_sets alg_min_alg {context = ctxt, prems = _} =
+ let
+ fun use_asm thm = etac (thm RS subset_trans);
+
+ fun useIH set_sets = EVERY' [rtac subsetI, rtac mp, Goal.assume_rule_tac ctxt,
+ rtac CollectI, CONJ_WRAP' (fn thm => EVERY' [etac (thm RS subset_trans), atac]) set_sets];
+
+ fun mk_incl alg_set set_setss =
+ EVERY' [rtac impI, REPEAT_DETERM o eresolve_tac [CollectE, conjE],
+ rtac (alg_min_alg RS alg_set),
+ EVERY' (map use_asm (map hd set_setss)),
+ EVERY' (map useIH (transpose (map tl set_setss)))];
+ in
+ (induct_tac THEN' EVERY' (map2 mk_incl alg_sets set_setsss)) 1
+ end;
+
+fun mk_lfp_map_wpull_tac m induct_tac wpulls map_simps set_simpss set_setsss ctor_injects =
+ let
+ val n = length wpulls;
+ val ks = 1 upto n;
+ val ls = 1 upto m;
+
+ fun use_pass_asm thm = rtac conjI THEN' etac (thm RS subset_trans);
+ fun use_act_asm thm = etac (thm RS subset_trans) THEN' atac;
+
+ fun useIH set_sets i = EVERY' [rtac ssubst, rtac @{thm wpull_def},
+ REPEAT_DETERM_N m o etac thin_rl, select_prem_tac n (dtac asm_rl) i,
+ rtac allI, rtac allI, rtac impI, REPEAT_DETERM o etac conjE,
+ REPEAT_DETERM o dtac @{thm meta_spec},
+ dtac meta_mp, atac,
+ dtac meta_mp, atac, etac mp,
+ rtac conjI, rtac CollectI, CONJ_WRAP' use_act_asm set_sets,
+ rtac conjI, rtac CollectI, CONJ_WRAP' use_act_asm set_sets,
+ atac];
+
+ fun mk_subset thm = EVERY' [rtac ord_eq_le_trans, rtac thm, rtac @{thm Un_least}, atac,
+ REPEAT_DETERM_N (n - 1) o rtac @{thm Un_least},
+ REPEAT_DETERM_N n o
+ EVERY' [rtac @{thm UN_least}, rtac CollectE, etac set_rev_mp, atac,
+ REPEAT_DETERM o etac conjE, atac]];
+
+ fun mk_wpull wpull map_simp set_simps set_setss ctor_inject =
+ EVERY' [rtac impI, REPEAT_DETERM o eresolve_tac [CollectE, conjE],
+ rtac rev_mp, rtac wpull,
+ EVERY' (map (fn i => REPEAT_DETERM_N (i - 1) o etac thin_rl THEN' atac) ls),
+ EVERY' (map2 useIH (transpose (map tl set_setss)) ks),
+ rtac impI, REPEAT_DETERM_N (m + n) o etac thin_rl,
+ dtac @{thm subst[OF wpull_def, of "%x. x"]}, etac allE, etac allE, etac impE,
+ rtac conjI, rtac CollectI, EVERY' (map (use_pass_asm o hd) set_setss),
+ CONJ_WRAP' (K (rtac subset_refl)) ks,
+ rtac conjI, rtac CollectI, EVERY' (map (use_pass_asm o hd) set_setss),
+ CONJ_WRAP' (K (rtac subset_refl)) ks,
+ rtac subst, rtac ctor_inject, rtac trans, rtac sym, rtac map_simp,
+ rtac trans, atac, rtac map_simp, REPEAT_DETERM o eresolve_tac [CollectE, conjE, bexE],
+ hyp_subst_tac, rtac bexI, rtac conjI, rtac map_simp, rtac map_simp, rtac CollectI,
+ CONJ_WRAP' mk_subset set_simps];
+ in
+ (induct_tac THEN' EVERY' (map5 mk_wpull wpulls map_simps set_simpss set_setsss ctor_injects)) 1
+ end;
+
+(* BNF tactics *)
+
+fun mk_map_id_tac map_ids unique =
+ (rtac sym THEN' rtac unique THEN'
+ EVERY' (map (fn thm =>
+ EVERY' [rtac trans, rtac @{thm id_o}, rtac trans, rtac sym, rtac @{thm o_id},
+ rtac (thm RS sym RS arg_cong)]) map_ids)) 1;
+
+fun mk_map_comp_tac map_comps map_simps unique iplus1 =
+ let
+ val i = iplus1 - 1;
+ val unique' = Thm.permute_prems 0 i unique;
+ val map_comps' = drop i map_comps @ take i map_comps;
+ val map_simps' = drop i map_simps @ take i map_simps;
+ fun mk_comp comp simp =
+ EVERY' [rtac ext, rtac trans, rtac o_apply, rtac trans, rtac o_apply,
+ rtac trans, rtac (simp RS arg_cong), rtac trans, rtac simp,
+ rtac trans, rtac (comp RS arg_cong), rtac sym, rtac o_apply];
+ in
+ (rtac sym THEN' rtac unique' THEN' EVERY' (map2 mk_comp map_comps' map_simps')) 1
+ end;
+
+fun mk_set_natural_tac set_nat =
+ EVERY' (map rtac [ext, trans, o_apply, sym, trans, o_apply, set_nat]) 1;
+
+fun mk_in_bd_tac sum_Card_order sucbd_Cnotzero incl card_of_min_alg =
+ EVERY' [rtac ctrans, rtac @{thm card_of_mono1}, rtac subsetI, etac rev_mp,
+ rtac incl, rtac ctrans, rtac card_of_min_alg, rtac @{thm cexp_mono2_Cnotzero},
+ rtac @{thm cardSuc_ordLeq_cpow}, rtac sum_Card_order, rtac @{thm csum_Cnotzero2},
+ rtac @{thm ctwo_Cnotzero}, rtac sucbd_Cnotzero] 1;
+
+fun mk_bd_card_order_tac bd_card_orders =
+ (rtac @{thm card_order_cpow} THEN'
+ CONJ_WRAP_GEN' (rtac @{thm card_order_csum}) rtac bd_card_orders) 1;
+
+fun mk_wpull_tac wpull =
+ EVERY' [rtac ssubst, rtac @{thm wpull_def}, rtac allI, rtac allI,
+ rtac wpull, REPEAT_DETERM o atac] 1;
+
+fun mk_wit_tac n set_simp wit =
+ REPEAT_DETERM (atac 1 ORELSE
+ EVERY' [dtac set_rev_mp, rtac equalityD1, resolve_tac set_simp,
+ REPEAT_DETERM o
+ (TRY o REPEAT_DETERM o etac UnE THEN' TRY o etac @{thm UN_E} THEN'
+ (eresolve_tac wit ORELSE'
+ (dresolve_tac wit THEN'
+ (etac FalseE ORELSE'
+ EVERY' [hyp_subst_tac, dtac set_rev_mp, rtac equalityD1, resolve_tac set_simp,
+ REPEAT_DETERM_N n o etac UnE]))))] 1);
+
+fun mk_srel_simp_tac in_Isrels i in_srel map_comp map_cong map_simp set_simps ctor_inject
+ ctor_dtor set_naturals set_incls set_set_inclss =
+ let
+ val m = length set_incls;
+ val n = length set_set_inclss;
+
+ val (passive_set_naturals, active_set_naturals) = chop m set_naturals;
+ val in_Isrel = nth in_Isrels (i - 1);
+ val le_arg_cong_ctor_dtor = ctor_dtor RS arg_cong RS ord_eq_le_trans;
+ val eq_arg_cong_ctor_dtor = ctor_dtor RS arg_cong RS trans;
+ val if_tac =
+ EVERY' [dtac (in_Isrel RS iffD1), REPEAT_DETERM o eresolve_tac [exE, conjE, CollectE],
+ rtac (in_srel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
+ EVERY' (map2 (fn set_natural => fn set_incl =>
+ EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac set_natural,
+ rtac ord_eq_le_trans, rtac trans_fun_cong_image_id_id_apply,
+ rtac (set_incl RS subset_trans), etac le_arg_cong_ctor_dtor])
+ passive_set_naturals set_incls),
+ CONJ_WRAP' (fn (in_Isrel, (set_natural, set_set_incls)) =>
+ EVERY' [rtac ord_eq_le_trans, rtac set_natural, rtac @{thm image_subsetI},
+ rtac (in_Isrel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
+ CONJ_WRAP' (fn thm =>
+ EVERY' (map etac [thm RS subset_trans, le_arg_cong_ctor_dtor]))
+ set_set_incls,
+ rtac conjI, rtac refl, rtac refl])
+ (in_Isrels ~~ (active_set_naturals ~~ set_set_inclss)),
+ CONJ_WRAP' (fn conv =>
+ EVERY' [rtac trans, rtac map_comp, rtac trans, rtac map_cong,
+ REPEAT_DETERM_N m o rtac @{thm fun_cong[OF o_id]},
+ REPEAT_DETERM_N n o EVERY' (map rtac [trans, o_apply, conv]),
+ rtac (ctor_inject RS iffD1), rtac trans, rtac sym, rtac map_simp,
+ etac eq_arg_cong_ctor_dtor])
+ fst_snd_convs];
+ val only_if_tac =
+ EVERY' [dtac (in_srel RS iffD1), REPEAT_DETERM o eresolve_tac [exE, conjE, CollectE],
+ rtac (in_Isrel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
+ CONJ_WRAP' (fn (set_simp, passive_set_natural) =>
+ EVERY' [rtac ord_eq_le_trans, rtac set_simp, rtac @{thm Un_least},
+ rtac ord_eq_le_trans, rtac @{thm box_equals[OF _ refl]},
+ rtac passive_set_natural, rtac trans_fun_cong_image_id_id_apply, atac,
+ CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_least}))
+ (fn (active_set_natural, in_Isrel) => EVERY' [rtac ord_eq_le_trans,
+ rtac @{thm UN_cong[OF _ refl]}, rtac active_set_natural, rtac @{thm UN_least},
+ dtac set_rev_mp, etac @{thm image_mono}, etac imageE,
+ dtac @{thm ssubst_mem[OF pair_collapse]}, dtac (in_Isrel RS iffD1),
+ dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE,
+ dtac (Thm.permute_prems 0 1 @{thm ssubst_mem}), atac,
+ hyp_subst_tac, REPEAT_DETERM o eresolve_tac [CollectE, conjE], atac])
+ (rev (active_set_naturals ~~ in_Isrels))])
+ (set_simps ~~ passive_set_naturals),
+ rtac conjI,
+ REPEAT_DETERM_N 2 o EVERY' [rtac trans, rtac map_simp, rtac (ctor_inject RS iffD2),
+ rtac trans, rtac map_comp, rtac trans, rtac map_cong,
+ REPEAT_DETERM_N m o rtac @{thm fun_cong[OF o_id]},
+ EVERY' (map (fn in_Isrel => EVERY' [rtac trans, rtac o_apply, dtac set_rev_mp, atac,
+ dtac @{thm ssubst_mem[OF pair_collapse]}, dtac (in_Isrel RS iffD1),
+ dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE, atac]) in_Isrels),
+ atac]]
+ in
+ EVERY' [rtac iffI, if_tac, only_if_tac] 1
+ end;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Tools/bnf_lfp_util.ML Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,76 @@
+(* Title: HOL/BNF/Tools/bnf_lfp_util.ML
+ Author: Dmitriy Traytel, TU Muenchen
+ Author: Jasmin Blanchette, TU Muenchen
+ Copyright 2012
+
+Library for the datatype construction.
+*)
+
+signature BNF_LFP_UTIL =
+sig
+ val HOL_arg_cong: cterm -> thm
+
+ val mk_bij_betw: term -> term -> term -> term
+ val mk_cardSuc: term -> term
+ val mk_convol: term * term -> term
+ val mk_cpow: term -> term
+ val mk_inver: term -> term -> term -> term
+ val mk_not_empty: term -> term
+ val mk_not_eq: term -> term -> term
+ val mk_rapp: term -> typ -> term
+ val mk_relChain: term -> term -> term
+ val mk_underS: term -> term
+ val mk_worec: term -> term -> term
+end;
+
+structure BNF_LFP_Util : BNF_LFP_UTIL =
+struct
+
+open BNF_Util
+
+fun HOL_arg_cong ct = Drule.instantiate'
+ (map SOME (Thm.dest_ctyp (Thm.ctyp_of_term ct))) [NONE, NONE, SOME ct] arg_cong;
+
+(*reverse application*)
+fun mk_rapp arg T = Term.absdummy (fastype_of arg --> T) (Bound 0 $ arg);
+
+fun mk_underS r =
+ let val T = fst (dest_relT (fastype_of r));
+ in Const (@{const_name rel.underS}, mk_relT (T, T) --> T --> HOLogic.mk_setT T) $ r end;
+
+fun mk_worec r f =
+ let val (A, AB) = apfst domain_type (dest_funT (fastype_of f));
+ in Const (@{const_name wo_rel.worec}, mk_relT (A, A) --> (AB --> AB) --> AB) $ r $ f end;
+
+fun mk_relChain r f =
+ let val (A, AB) = `domain_type (fastype_of f);
+ in Const (@{const_name relChain}, mk_relT (A, A) --> AB --> HOLogic.boolT) $ r $ f end;
+
+fun mk_cardSuc r =
+ let val T = fst (dest_relT (fastype_of r));
+ in Const (@{const_name cardSuc}, mk_relT (T, T) --> mk_relT (`I (HOLogic.mk_setT T))) $ r end;
+
+fun mk_cpow r =
+ let val T = fst (dest_relT (fastype_of r));
+ in Const (@{const_name cpow}, mk_relT (T, T) --> mk_relT (`I (HOLogic.mk_setT T))) $ r end;
+
+fun mk_bij_betw f A B =
+ Const (@{const_name bij_betw},
+ fastype_of f --> fastype_of A --> fastype_of B --> HOLogic.boolT) $ f $ A $ B;
+
+fun mk_inver f g A =
+ Const (@{const_name inver}, fastype_of f --> fastype_of g --> fastype_of A --> HOLogic.boolT) $
+ f $ g $ A;
+
+fun mk_not_eq x y = HOLogic.mk_not (HOLogic.mk_eq (x, y));
+
+fun mk_not_empty B = mk_not_eq B (HOLogic.mk_set (HOLogic.dest_setT (fastype_of B)) []);
+
+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;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Tools/bnf_tactics.ML Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,125 @@
+(* Title: HOL/BNF/Tools/bnf_tactics.ML
+ Author: Dmitriy Traytel, TU Muenchen
+ Author: Jasmin Blanchette, TU Muenchen
+ Copyright 2012
+
+General tactics for bounded natural functors.
+*)
+
+signature BNF_TACTICS =
+sig
+ val ss_only: thm list -> simpset
+
+ val prefer_tac: int -> tactic
+ val select_prem_tac: int -> (int -> tactic) -> int -> int -> tactic
+ val fo_rtac: thm -> Proof.context -> int -> tactic
+ val subst_asm_tac: Proof.context -> thm list -> int -> tactic
+ val subst_tac: Proof.context -> thm list -> int -> tactic
+ val substs_tac: Proof.context -> thm list -> int -> tactic
+ val unfold_thms_tac: Proof.context -> thm list -> tactic
+ val mk_unfold_thms_then_tac: Proof.context -> thm list -> ('a -> tactic) -> 'a -> tactic
+
+ val mk_flatten_assoc_tac: (int -> tactic) -> thm -> thm -> thm -> tactic
+ val mk_rotate_eq_tac: (int -> tactic) -> thm -> thm -> thm -> thm -> ''a list -> ''a list ->
+ int -> tactic
+
+ val mk_Abs_bij_thm: Proof.context -> thm -> thm -> thm
+ val mk_Abs_inj_thm: thm -> thm
+
+ val simple_srel_O_Gr_tac: Proof.context -> tactic
+ val mk_rel_simp_tac: thm -> thm list -> thm list -> thm -> {prems: 'a, context: Proof.context} ->
+ tactic
+
+ val mk_map_comp_id_tac: thm -> tactic
+ val mk_map_cong_tac: int -> thm -> {prems: 'a, context: Proof.context} -> tactic
+ val mk_map_congL_tac: int -> thm -> thm -> tactic
+end;
+
+structure BNF_Tactics : BNF_TACTICS =
+struct
+
+open BNF_Util
+
+fun ss_only thms = Simplifier.clear_ss HOL_basic_ss addsimps thms;
+
+(* FIXME: why not in "Pure"? *)
+fun prefer_tac i = defer_tac i THEN PRIMITIVE (Thm.permute_prems 0 ~1);
+
+fun select_prem_tac n tac k = DETERM o (EVERY' [REPEAT_DETERM_N (k - 1) o etac thin_rl,
+ tac, REPEAT_DETERM_N (n - k) o etac thin_rl]);
+
+(*stolen from Christian Urban's Cookbook*)
+fun fo_rtac thm = Subgoal.FOCUS (fn {concl, ...} =>
+ let
+ val concl_pat = Drule.strip_imp_concl (cprop_of thm)
+ val insts = Thm.first_order_match (concl_pat, concl)
+ in
+ rtac (Drule.instantiate_normalize insts thm) 1
+ end);
+
+fun unfold_thms_tac ctxt thms = Local_Defs.unfold_tac ctxt (distinct Thm.eq_thm_prop thms);
+
+fun mk_unfold_thms_then_tac lthy defs tac x = unfold_thms_tac lthy defs THEN tac x;
+
+(*unlike "unfold_thms_tac", succeeds when the RHS contains schematic variables not in the LHS*)
+fun subst_asm_tac ctxt = EqSubst.eqsubst_asm_tac ctxt [0];
+fun subst_tac ctxt = EqSubst.eqsubst_tac ctxt [0];
+fun substs_tac ctxt = REPEAT_DETERM oo subst_tac ctxt;
+
+
+(* Theorems for open typedefs with UNIV as representing set *)
+
+fun mk_Abs_inj_thm inj = inj OF (replicate 2 UNIV_I);
+fun mk_Abs_bij_thm ctxt Abs_inj_thm surj = rule_by_tactic ctxt ((rtac surj THEN' etac exI) 1)
+ (Abs_inj_thm RS @{thm bijI});
+
+
+
+(* General tactic generators *)
+
+(*applies assoc rule to the lhs of an equation as long as possible*)
+fun mk_flatten_assoc_tac refl_tac trans assoc cong = rtac trans 1 THEN
+ REPEAT_DETERM (CHANGED ((FIRST' [rtac trans THEN' rtac assoc, rtac cong THEN' refl_tac]) 1)) THEN
+ refl_tac 1;
+
+(*proves two sides of an equation to be equal assuming both are flattened and rhs can be obtained
+from lhs by the given permutation of monoms*)
+fun mk_rotate_eq_tac refl_tac trans assoc com cong =
+ let
+ fun gen_tac [] [] = K all_tac
+ | gen_tac [x] [y] = if x = y then refl_tac else error "mk_rotate_eq_tac: different lists"
+ | gen_tac (x :: xs) (y :: ys) = if x = y
+ then rtac cong THEN' refl_tac THEN' gen_tac xs ys
+ else rtac trans THEN' rtac com THEN'
+ K (mk_flatten_assoc_tac refl_tac trans assoc cong) THEN'
+ gen_tac (xs @ [x]) (y :: ys)
+ | gen_tac _ _ = error "mk_rotate_eq_tac: different lists";
+ in
+ gen_tac
+ end;
+
+fun simple_srel_O_Gr_tac ctxt =
+ unfold_thms_tac ctxt @{thms Collect_fst_snd_mem_eq Collect_pair_mem_eq} THEN rtac refl 1;
+
+fun mk_rel_simp_tac srel_def IJrel_defs IJsrel_defs srel_simp {context = ctxt, prems = _} =
+ unfold_thms_tac ctxt IJrel_defs THEN
+ subst_tac ctxt [unfold_thms ctxt (IJrel_defs @ IJsrel_defs @
+ @{thms Collect_pair_mem_eq mem_Collect_eq fst_conv snd_conv}) srel_simp] 1 THEN
+ unfold_thms_tac ctxt (srel_def ::
+ @{thms Collect_fst_snd_mem_eq mem_Collect_eq pair_mem_Collect_split fst_conv snd_conv
+ split_conv}) THEN
+ rtac refl 1;
+
+fun mk_map_comp_id_tac map_comp =
+ (rtac trans THEN' rtac map_comp THEN' REPEAT_DETERM o stac @{thm o_id} THEN' rtac refl) 1;
+
+fun mk_map_cong_tac m map_cong {context = ctxt, prems = _} =
+ EVERY' [rtac mp, rtac map_cong,
+ CONJ_WRAP' (K (rtac ballI THEN' Goal.assume_rule_tac ctxt)) (1 upto m)] 1;
+
+fun mk_map_congL_tac passive map_cong map_id' =
+ (rtac trans THEN' rtac map_cong THEN' EVERY' (replicate passive (rtac refl))) 1 THEN
+ REPEAT_DETERM (EVERY' [rtac trans, etac bspec, atac, rtac sym, rtac @{thm id_apply}] 1) THEN
+ rtac map_id' 1;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Tools/bnf_util.ML Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,619 @@
+(* Title: HOL/BNF/Tools/bnf_util.ML
+ Author: Dmitriy Traytel, TU Muenchen
+ Copyright 2012
+
+Library for bounded natural functors.
+*)
+
+signature BNF_UTIL =
+sig
+ val map3: ('a -> 'b -> 'c -> 'd) -> 'a list -> 'b list -> 'c list -> 'd list
+ val map4: ('a -> 'b -> 'c -> 'd -> 'e) -> 'a list -> 'b list -> 'c list -> 'd list -> 'e list
+ val map5: ('a -> 'b -> 'c -> 'd -> 'e -> 'f) ->
+ 'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list
+ val map6: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g) ->
+ 'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list
+ val map7: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h) ->
+ 'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list
+ val map8: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i) ->
+ 'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list -> 'i list
+ val map9: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i -> 'j) ->
+ 'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list ->
+ 'i list -> 'j list
+ val map10: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i -> 'j -> 'k) ->
+ 'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list ->
+ 'i list -> 'j list -> 'k list
+ val map11: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i -> 'j -> 'k -> 'l) ->
+ 'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list ->
+ 'i list -> 'j list -> 'k list -> 'l list
+ val map12: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i -> 'j -> 'k -> 'l -> 'm) ->
+ 'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list ->
+ 'i list -> 'j list -> 'k list -> 'l list -> 'm list
+ val fold_map2: ('a -> 'b -> 'c -> 'd * 'c) -> 'a list -> 'b list -> 'c -> 'd list * 'c
+ val fold_map3: ('a -> 'b -> 'c -> 'd -> 'e * 'd) ->
+ 'a list -> 'b list -> 'c list -> 'd -> 'e list * 'd
+ val fold_map4: ('a -> 'b -> 'c -> 'd -> 'e -> 'f * 'e) ->
+ 'a list -> 'b list -> 'c list -> 'd list -> 'e -> 'f list * 'e
+ val fold_map5: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g * 'f) ->
+ 'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f -> 'g list * 'f
+ val fold_map6: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h * 'g) ->
+ 'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g -> 'h list * 'g
+ val fold_map7: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i * 'h) ->
+ 'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h -> 'i list * 'h
+ val interleave: 'a list -> 'a list -> 'a list
+ val transpose: 'a list list -> 'a list list
+ val seq_conds: (bool -> 'a -> 'b) -> int -> int -> 'a list -> 'b list
+
+ val mk_fresh_names: Proof.context -> int -> string -> string list * Proof.context
+ val mk_TFrees: int -> Proof.context -> typ list * Proof.context
+ val mk_TFreess: int list -> Proof.context -> typ list list * Proof.context
+ val mk_TFrees': sort list -> Proof.context -> typ list * Proof.context
+ val mk_Frees: string -> typ list -> Proof.context -> term list * Proof.context
+ val mk_Freess: string -> typ list list -> Proof.context -> term list list * Proof.context
+ val mk_Freesss: string -> typ list list list -> Proof.context ->
+ term list list list * Proof.context
+ val mk_Freessss: string -> typ list list list list -> Proof.context ->
+ term list list list list * Proof.context
+ val mk_Frees': string -> typ list -> Proof.context ->
+ (term list * (string * typ) list) * Proof.context
+ val mk_Freess': string -> typ list list -> Proof.context ->
+ (term list list * (string * typ) list list) * Proof.context
+ val nonzero_string_of_int: int -> string
+
+ val strip_typeN: int -> typ -> typ list * typ
+
+ val mk_predT: typ list -> typ
+ val mk_pred1T: typ -> typ
+ val mk_pred2T: typ -> typ -> typ
+ val mk_optionT: typ -> typ
+ val mk_relT: typ * typ -> typ
+ val dest_relT: typ -> typ * typ
+ val mk_sumT: typ * typ -> typ
+
+ val ctwo: term
+ val fst_const: typ -> term
+ val snd_const: typ -> term
+ val Id_const: typ -> term
+
+ val mk_Ball: term -> term -> term
+ val mk_Bex: term -> term -> term
+ val mk_Card_order: term -> term
+ val mk_Field: term -> term
+ val mk_Gr: term -> term -> term
+ val mk_IfN: typ -> term list -> term list -> term
+ val mk_Trueprop_eq: term * term -> term
+ val mk_UNION: term -> term -> term
+ val mk_Union: typ -> term
+ val mk_card_binop: string -> (typ * typ -> typ) -> term -> term -> term
+ val mk_card_of: term -> term
+ val mk_card_order: term -> term
+ val mk_ccexp: term -> term -> term
+ val mk_cexp: term -> term -> term
+ val mk_cinfinite: term -> term
+ val mk_collect: term list -> typ -> term
+ val mk_converse: term -> term
+ val mk_cprod: term -> term -> term
+ val mk_csum: term -> term -> term
+ val mk_dir_image: term -> term -> term
+ val mk_image: term -> term
+ val mk_in: term list -> term list -> typ -> term
+ val mk_ordLeq: term -> term -> term
+ val mk_rel_comp: term * term -> term
+ val mk_subset: term -> term -> term
+ val mk_wpull: term -> term -> term -> term -> term -> (term * term) option -> term -> term -> term
+
+ val list_all_free: term list -> term -> term
+ val list_exists_free: term list -> term -> term
+
+ (*parameterized terms*)
+ val mk_nthN: int -> term -> int -> term
+
+ (*parameterized thms*)
+ val mk_Un_upper: int -> int -> thm
+ val mk_conjIN: int -> thm
+ val mk_conjunctN: int -> int -> thm
+ val conj_dests: int -> thm -> thm list
+ val mk_disjIN: int -> int -> thm
+ val mk_nthI: int -> int -> thm
+ val mk_nth_conv: int -> int -> thm
+ val mk_ordLeq_csum: int -> int -> thm -> thm
+ val mk_UnIN: int -> int -> thm
+
+ val ctrans: thm
+ val o_apply: thm
+ val set_mp: thm
+ val set_rev_mp: thm
+ val subset_UNIV: thm
+ val Pair_eqD: thm
+ val Pair_eqI: thm
+ val mk_sym: thm -> thm
+ val mk_trans: thm -> thm -> thm
+ val mk_unabs_def: int -> thm -> thm
+
+ val is_refl: thm -> bool
+ val no_refl: thm list -> thm list
+ val no_reflexive: thm list -> thm list
+
+ val fold_thms: Proof.context -> thm list -> thm -> thm
+ val unfold_thms: Proof.context -> thm list -> thm -> thm
+
+ val mk_permute: ''a list -> ''a list -> 'b list -> 'b list
+ val find_indices: ''a list -> ''a list -> int list
+
+ val certifyT: Proof.context -> typ -> ctyp
+ val certify: Proof.context -> term -> cterm
+
+ val parse_binding_colon: Token.T list -> binding * Token.T list
+ val parse_opt_binding_colon: Token.T list -> binding * Token.T list
+
+ val typedef: bool -> binding option -> binding * (string * sort) list * mixfix -> term ->
+ (binding * binding) option -> tactic -> local_theory -> (string * Typedef.info) * local_theory
+
+ val WRAP: ('a -> tactic) -> ('a -> tactic) -> 'a list -> tactic -> tactic
+ val WRAP': ('a -> int -> tactic) -> ('a -> int -> tactic) -> 'a list -> (int -> tactic) -> int ->
+ tactic
+ val CONJ_WRAP_GEN: tactic -> ('a -> tactic) -> 'a list -> tactic
+ val CONJ_WRAP_GEN': (int -> tactic) -> ('a -> int -> tactic) -> 'a list -> int -> tactic
+ val CONJ_WRAP: ('a -> tactic) -> 'a list -> tactic
+ val CONJ_WRAP': ('a -> int -> tactic) -> 'a list -> int -> tactic
+end;
+
+structure BNF_Util : BNF_UTIL =
+struct
+
+(* Library proper *)
+
+fun map3 _ [] [] [] = []
+ | map3 f (x1::x1s) (x2::x2s) (x3::x3s) = f x1 x2 x3 :: map3 f x1s x2s x3s
+ | map3 _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun map4 _ [] [] [] [] = []
+ | map4 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) = f x1 x2 x3 x4 :: map4 f x1s x2s x3s x4s
+ | map4 _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun map5 _ [] [] [] [] [] = []
+ | map5 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) =
+ f x1 x2 x3 x4 x5 :: map5 f x1s x2s x3s x4s x5s
+ | map5 _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun map6 _ [] [] [] [] [] [] = []
+ | map6 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) =
+ f x1 x2 x3 x4 x5 x6 :: map6 f x1s x2s x3s x4s x5s x6s
+ | map6 _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun map7 _ [] [] [] [] [] [] [] = []
+ | map7 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) =
+ f x1 x2 x3 x4 x5 x6 x7 :: map7 f x1s x2s x3s x4s x5s x6s x7s
+ | map7 _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun map8 _ [] [] [] [] [] [] [] [] = []
+ | map8 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) (x8::x8s) =
+ f x1 x2 x3 x4 x5 x6 x7 x8 :: map8 f x1s x2s x3s x4s x5s x6s x7s x8s
+ | map8 _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun map9 _ [] [] [] [] [] [] [] [] [] = []
+ | map9 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s)
+ (x6::x6s) (x7::x7s) (x8::x8s) (x9::x9s) =
+ f x1 x2 x3 x4 x5 x6 x7 x8 x9 :: map9 f x1s x2s x3s x4s x5s x6s x7s x8s x9s
+ | map9 _ _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun map10 _ [] [] [] [] [] [] [] [] [] [] = []
+ | map10 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s)
+ (x6::x6s) (x7::x7s) (x8::x8s) (x9::x9s) (x10::x10s) =
+ f x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 :: map10 f x1s x2s x3s x4s x5s x6s x7s x8s x9s x10s
+ | map10 _ _ _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun map11 _ [] [] [] [] [] [] [] [] [] [] [] = []
+ | map11 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s)
+ (x6::x6s) (x7::x7s) (x8::x8s) (x9::x9s) (x10::x10s) (x11::x11s) =
+ f x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 :: map11 f x1s x2s x3s x4s x5s x6s x7s x8s x9s x10s x11s
+ | map11 _ _ _ _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun map12 _ [] [] [] [] [] [] [] [] [] [] [] [] = []
+ | map12 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s)
+ (x6::x6s) (x7::x7s) (x8::x8s) (x9::x9s) (x10::x10s) (x11::x11s) (x12::x12s) =
+ f x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 ::
+ map12 f x1s x2s x3s x4s x5s x6s x7s x8s x9s x10s x11s x12s
+ | map12 _ _ _ _ _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun fold_map2 _ [] [] acc = ([], acc)
+ | fold_map2 f (x1::x1s) (x2::x2s) acc =
+ let
+ val (x, acc') = f x1 x2 acc;
+ val (xs, acc'') = fold_map2 f x1s x2s acc';
+ in (x :: xs, acc'') end
+ | fold_map2 _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun fold_map3 _ [] [] [] acc = ([], acc)
+ | fold_map3 f (x1::x1s) (x2::x2s) (x3::x3s) acc =
+ let
+ val (x, acc') = f x1 x2 x3 acc;
+ val (xs, acc'') = fold_map3 f x1s x2s x3s acc';
+ in (x :: xs, acc'') end
+ | fold_map3 _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun fold_map4 _ [] [] [] [] acc = ([], acc)
+ | fold_map4 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) acc =
+ let
+ val (x, acc') = f x1 x2 x3 x4 acc;
+ val (xs, acc'') = fold_map4 f x1s x2s x3s x4s acc';
+ in (x :: xs, acc'') end
+ | fold_map4 _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun fold_map5 _ [] [] [] [] [] acc = ([], acc)
+ | fold_map5 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) acc =
+ let
+ val (x, acc') = f x1 x2 x3 x4 x5 acc;
+ val (xs, acc'') = fold_map5 f x1s x2s x3s x4s x5s acc';
+ in (x :: xs, acc'') end
+ | fold_map5 _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun fold_map6 _ [] [] [] [] [] [] acc = ([], acc)
+ | fold_map6 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) acc =
+ let
+ val (x, acc') = f x1 x2 x3 x4 x5 x6 acc;
+ val (xs, acc'') = fold_map6 f x1s x2s x3s x4s x5s x6s acc';
+ in (x :: xs, acc'') end
+ | fold_map6 _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+fun fold_map7 _ [] [] [] [] [] [] [] acc = ([], acc)
+ | fold_map7 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) acc =
+ let
+ val (x, acc') = f x1 x2 x3 x4 x5 x6 x7 acc;
+ val (xs, acc'') = fold_map7 f x1s x2s x3s x4s x5s x6s x7s acc';
+ in (x :: xs, acc'') end
+ | fold_map7 _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
+
+(*stolen from ~~/src/HOL/Tools/SMT/smt_utils.ML*)
+fun certify ctxt = Thm.cterm_of (Proof_Context.theory_of ctxt);
+fun certifyT ctxt = Thm.ctyp_of (Proof_Context.theory_of ctxt);
+
+val parse_binding_colon = Parse.binding --| @{keyword ":"};
+val parse_opt_binding_colon = Scan.optional parse_binding_colon Binding.empty;
+
+(*TODO: is this really different from Typedef.add_typedef_global?*)
+fun typedef def opt_name typ set opt_morphs tac lthy =
+ let
+ val ((name, info), (lthy, lthy_old)) =
+ lthy
+ |> Typedef.add_typedef def opt_name typ set opt_morphs tac
+ ||> `Local_Theory.restore;
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+ in
+ ((name, Typedef.transform_info phi info), lthy)
+ end;
+
+(*Tactical WRAP surrounds a static given tactic (core) with two deterministic chains of tactics*)
+fun WRAP gen_before gen_after xs core_tac =
+ fold_rev (fn x => fn tac => gen_before x THEN tac THEN gen_after x) xs core_tac;
+
+fun WRAP' gen_before gen_after xs core_tac =
+ fold_rev (fn x => fn tac => gen_before x THEN' tac THEN' gen_after x) xs core_tac;
+
+fun CONJ_WRAP_GEN conj_tac gen_tac xs =
+ let val (butlast, last) = split_last xs;
+ in WRAP (fn thm => conj_tac THEN gen_tac thm) (K all_tac) butlast (gen_tac last) end;
+
+fun CONJ_WRAP_GEN' conj_tac gen_tac xs =
+ let val (butlast, last) = split_last xs;
+ in WRAP' (fn thm => conj_tac THEN' gen_tac thm) (K (K all_tac)) butlast (gen_tac last) end;
+
+(*not eta-converted because of monotype restriction*)
+fun CONJ_WRAP gen_tac = CONJ_WRAP_GEN (rtac conjI 1) gen_tac;
+fun CONJ_WRAP' gen_tac = CONJ_WRAP_GEN' (rtac conjI) gen_tac;
+
+
+
+(* Term construction *)
+
+(** Fresh variables **)
+
+fun nonzero_string_of_int 0 = ""
+ | nonzero_string_of_int n = string_of_int n;
+
+val mk_TFrees' = apfst (map TFree) oo Variable.invent_types;
+
+fun mk_TFrees n = mk_TFrees' (replicate n HOLogic.typeS);
+val mk_TFreess = fold_map mk_TFrees;
+
+fun mk_names n x = if n = 1 then [x] else map (fn i => x ^ string_of_int i) (1 upto n);
+
+fun mk_fresh_names ctxt = (fn xs => Variable.variant_fixes xs ctxt) oo mk_names;
+fun mk_Frees x Ts ctxt = mk_fresh_names ctxt (length Ts) x |>> (fn xs => map2 (curry Free) xs Ts);
+fun mk_Freess x Tss = fold_map2 mk_Frees (mk_names (length Tss) x) Tss;
+fun mk_Freesss x Tsss = fold_map2 mk_Freess (mk_names (length Tsss) x) Tsss;
+fun mk_Freessss x Tssss = fold_map2 mk_Freesss (mk_names (length Tssss) x) Tssss;
+fun mk_Frees' x Ts ctxt = mk_fresh_names ctxt (length Ts) x |>> (fn xs => `(map Free) (xs ~~ Ts));
+fun mk_Freess' x Tss = fold_map2 mk_Frees' (mk_names (length Tss) x) Tss #>> split_list;
+
+
+(** Types **)
+
+fun strip_typeN 0 T = ([], T)
+ | strip_typeN n (Type (@{type_name fun}, [T, T'])) = strip_typeN (n - 1) T' |>> cons T
+ | strip_typeN _ T = raise TYPE ("strip_typeN", [T], []);
+
+fun mk_predT Ts = Ts ---> HOLogic.boolT;
+fun mk_pred1T T = mk_predT [T];
+fun mk_pred2T T U = mk_predT [T, U];
+fun mk_optionT T = Type (@{type_name option}, [T]);
+val mk_relT = HOLogic.mk_setT o HOLogic.mk_prodT;
+val dest_relT = HOLogic.dest_prodT o HOLogic.dest_setT;
+fun mk_sumT (LT, RT) = Type (@{type_name Sum_Type.sum}, [LT, RT]);
+fun mk_partial_funT (ranT, domT) = domT --> mk_optionT ranT;
+
+
+(** Constants **)
+
+fun fst_const T = Const (@{const_name fst}, T --> fst (HOLogic.dest_prodT T));
+fun snd_const T = Const (@{const_name snd}, T --> snd (HOLogic.dest_prodT T));
+fun Id_const T = Const (@{const_name Id}, mk_relT (T, T));
+
+
+(** Operators **)
+
+val mk_Trueprop_eq = HOLogic.mk_Trueprop o HOLogic.mk_eq;
+
+fun mk_IfN _ _ [t] = t
+ | mk_IfN T (c :: cs) (t :: ts) =
+ Const (@{const_name If}, HOLogic.boolT --> T --> T --> T) $ c $ t $ mk_IfN T cs ts;
+
+fun mk_converse R =
+ let
+ val RT = dest_relT (fastype_of R);
+ val RST = mk_relT (snd RT, fst RT);
+ in Const (@{const_name converse}, fastype_of R --> RST) $ R end;
+
+fun mk_rel_comp (R, S) =
+ let
+ val RT = fastype_of R;
+ val ST = fastype_of S;
+ val RST = mk_relT (fst (dest_relT RT), snd (dest_relT ST));
+ in Const (@{const_name relcomp}, RT --> ST --> RST) $ R $ S end;
+
+fun mk_Gr A f =
+ let val ((AT, BT), FT) = `dest_funT (fastype_of f);
+ in Const (@{const_name Gr}, HOLogic.mk_setT AT --> FT --> mk_relT (AT, BT)) $ A $ f end;
+
+fun mk_image f =
+ let val (T, U) = dest_funT (fastype_of f);
+ in Const (@{const_name image},
+ (T --> U) --> (HOLogic.mk_setT T) --> (HOLogic.mk_setT U)) $ f end;
+
+fun mk_Ball X f =
+ Const (@{const_name Ball}, fastype_of X --> fastype_of f --> HOLogic.boolT) $ X $ f;
+
+fun mk_Bex X f =
+ Const (@{const_name Bex}, fastype_of X --> fastype_of f --> HOLogic.boolT) $ X $ f;
+
+fun mk_UNION X f =
+ let val (T, U) = dest_funT (fastype_of f);
+ in Const (@{const_name SUPR}, fastype_of X --> (T --> U) --> U) $ X $ f end;
+
+fun mk_Union T =
+ Const (@{const_name Sup}, HOLogic.mk_setT (HOLogic.mk_setT T) --> HOLogic.mk_setT T);
+
+fun mk_Field r =
+ let val T = fst (dest_relT (fastype_of r));
+ in Const (@{const_name Field}, mk_relT (T, T) --> HOLogic.mk_setT T) $ r end;
+
+fun mk_card_order bd =
+ let
+ val T = fastype_of bd;
+ val AT = fst (dest_relT T);
+ in
+ Const (@{const_name card_order_on}, HOLogic.mk_setT AT --> T --> HOLogic.boolT) $
+ (HOLogic.mk_UNIV AT) $ bd
+ end;
+
+fun mk_Card_order bd =
+ let
+ val T = fastype_of bd;
+ val AT = fst (dest_relT T);
+ in
+ Const (@{const_name card_order_on}, HOLogic.mk_setT AT --> T --> HOLogic.boolT) $
+ mk_Field bd $ bd
+ end;
+
+fun mk_cinfinite bd =
+ Const (@{const_name cinfinite}, fastype_of bd --> HOLogic.boolT) $ bd;
+
+fun mk_ordLeq t1 t2 =
+ HOLogic.mk_mem (HOLogic.mk_prod (t1, t2),
+ Const (@{const_name ordLeq}, mk_relT (fastype_of t1, fastype_of t2)));
+
+fun mk_card_of A =
+ let
+ val AT = fastype_of A;
+ val T = HOLogic.dest_setT AT;
+ in
+ Const (@{const_name card_of}, AT --> mk_relT (T, T)) $ A
+ end;
+
+fun mk_dir_image r f =
+ let val (T, U) = dest_funT (fastype_of f);
+ in Const (@{const_name dir_image}, mk_relT (T, T) --> (T --> U) --> mk_relT (U, U)) $ r $ f end;
+
+(*FIXME: "x"?*)
+(*(nth sets i) must be of type "T --> 'ai set"*)
+fun mk_in As sets T =
+ let
+ fun in_single set A =
+ let val AT = fastype_of A;
+ in Const (@{const_name less_eq},
+ AT --> AT --> HOLogic.boolT) $ (set $ Free ("x", T)) $ A end;
+ in
+ if length sets > 0
+ then HOLogic.mk_Collect ("x", T, foldr1 (HOLogic.mk_conj) (map2 in_single sets As))
+ else HOLogic.mk_UNIV T
+ end;
+
+fun mk_wpull A B1 B2 f1 f2 pseudo p1 p2 =
+ let
+ val AT = fastype_of A;
+ val BT1 = fastype_of B1;
+ val BT2 = fastype_of B2;
+ val FT1 = fastype_of f1;
+ val FT2 = fastype_of f2;
+ val PT1 = fastype_of p1;
+ val PT2 = fastype_of p2;
+ val T1 = HOLogic.dest_setT BT1;
+ val T2 = HOLogic.dest_setT BT2;
+ val domP = domain_type PT1;
+ val ranF = range_type FT1;
+ val _ = if is_some pseudo orelse
+ (HOLogic.dest_setT AT = domP andalso
+ domain_type FT1 = T1 andalso
+ domain_type FT2 = T2 andalso
+ domain_type PT2 = domP andalso
+ range_type PT1 = T1 andalso
+ range_type PT2 = T2 andalso
+ range_type FT2 = ranF)
+ then () else raise TYPE ("mk_wpull", [BT1, BT2, FT1, FT2, PT1, PT2], []);
+ in
+ (case pseudo of
+ NONE => Const (@{const_name wpull},
+ AT --> BT1 --> BT2 --> FT1 --> FT2 --> PT1 --> PT2 --> HOLogic.boolT) $
+ A $ B1 $ B2 $ f1 $ f2 $ p1 $ p2
+ | SOME (e1, e2) => Const (@{const_name wppull},
+ AT --> BT1 --> BT2 --> FT1 --> FT2 --> fastype_of e1 --> fastype_of e2 -->
+ PT1 --> PT2 --> HOLogic.boolT) $
+ A $ B1 $ B2 $ f1 $ f2 $ e1 $ e2 $ p1 $ p2)
+ end;
+
+fun mk_subset t1 t2 =
+ Const (@{const_name less_eq}, (fastype_of t1) --> (fastype_of t2) --> HOLogic.boolT) $ t1 $ t2;
+
+fun mk_card_binop binop typop t1 t2 =
+ let
+ val (T1, relT1) = `(fst o dest_relT) (fastype_of t1);
+ val (T2, relT2) = `(fst o dest_relT) (fastype_of t2);
+ in
+ Const (binop, relT1 --> relT2 --> mk_relT (typop (T1, T2), typop (T1, T2))) $ t1 $ t2
+ end;
+
+val mk_csum = mk_card_binop @{const_name csum} mk_sumT;
+val mk_cprod = mk_card_binop @{const_name cprod} HOLogic.mk_prodT;
+val mk_cexp = mk_card_binop @{const_name cexp} mk_partial_funT;
+val mk_ccexp = mk_card_binop @{const_name ccexp} mk_partial_funT;
+val ctwo = @{term ctwo};
+
+fun mk_collect xs defT =
+ let val T = (case xs of [] => defT | (x::_) => fastype_of x);
+ in Const (@{const_name collect}, HOLogic.mk_setT T --> T) $ (HOLogic.mk_set T xs) end;
+
+fun mk_permute src dest xs = map (nth xs o (fn x => find_index ((curry op =) x) src)) dest;
+
+val list_all_free =
+ fold_rev (fn free => fn P =>
+ let val (x, T) = Term.dest_Free free;
+ in HOLogic.all_const T $ Term.absfree (x, T) P end);
+
+val list_exists_free =
+ fold_rev (fn free => fn P =>
+ let val (x, T) = Term.dest_Free free;
+ in HOLogic.exists_const T $ Term.absfree (x, T) P end);
+
+fun find_indices xs ys = map_filter I
+ (map_index (fn (i, y) => if member (op =) xs y then SOME i else NONE) ys);
+
+fun mk_trans thm1 thm2 = trans OF [thm1, thm2];
+fun mk_sym thm = sym OF [thm];
+
+(*TODO: antiquote heavily used theorems once*)
+val ctrans = @{thm ordLeq_transitive};
+val o_apply = @{thm o_apply};
+val set_mp = @{thm set_mp};
+val set_rev_mp = @{thm set_rev_mp};
+val subset_UNIV = @{thm subset_UNIV};
+val Pair_eqD = @{thm iffD1[OF Pair_eq]};
+val Pair_eqI = @{thm iffD2[OF Pair_eq]};
+
+fun mk_nthN 1 t 1 = t
+ | mk_nthN _ t 1 = HOLogic.mk_fst t
+ | mk_nthN 2 t 2 = HOLogic.mk_snd t
+ | mk_nthN n t m = mk_nthN (n - 1) (HOLogic.mk_snd t) (m - 1);
+
+fun mk_nth_conv n m =
+ let
+ fun thm b = if b then @{thm fst_snd} else @{thm snd_snd}
+ fun mk_nth_conv _ 1 1 = refl
+ | mk_nth_conv _ _ 1 = @{thm fst_conv}
+ | mk_nth_conv _ 2 2 = @{thm snd_conv}
+ | mk_nth_conv b _ 2 = @{thm snd_conv} RS thm b
+ | mk_nth_conv b n m = mk_nth_conv false (n - 1) (m - 1) RS thm b;
+ in mk_nth_conv (not (m = n)) n m end;
+
+fun mk_nthI 1 1 = @{thm TrueE[OF TrueI]}
+ | mk_nthI n m = fold (curry op RS) (replicate (m - 1) @{thm sndI})
+ (if m = n then @{thm TrueE[OF TrueI]} else @{thm fstI});
+
+fun mk_conjunctN 1 1 = @{thm TrueE[OF TrueI]}
+ | mk_conjunctN _ 1 = conjunct1
+ | mk_conjunctN 2 2 = conjunct2
+ | mk_conjunctN n m = conjunct2 RS (mk_conjunctN (n - 1) (m - 1));
+
+fun conj_dests n thm = map (fn k => thm RS mk_conjunctN n k) (1 upto n);
+
+fun mk_conjIN 1 = @{thm TrueE[OF TrueI]}
+ | mk_conjIN n = mk_conjIN (n - 1) RSN (2, conjI);
+
+fun mk_disjIN 1 1 = @{thm TrueE[OF TrueI]}
+ | mk_disjIN _ 1 = disjI1
+ | mk_disjIN 2 2 = disjI2
+ | mk_disjIN n m = (mk_disjIN (n - 1) (m - 1)) RS disjI2;
+
+fun mk_ordLeq_csum 1 1 thm = thm
+ | mk_ordLeq_csum _ 1 thm = @{thm ordLeq_transitive} OF [thm, @{thm ordLeq_csum1}]
+ | mk_ordLeq_csum 2 2 thm = @{thm ordLeq_transitive} OF [thm, @{thm ordLeq_csum2}]
+ | mk_ordLeq_csum n m thm = @{thm ordLeq_transitive} OF
+ [mk_ordLeq_csum (n - 1) (m - 1) thm, @{thm ordLeq_csum2[OF Card_order_csum]}];
+
+local
+ fun mk_Un_upper' 0 = subset_refl
+ | mk_Un_upper' 1 = @{thm Un_upper1}
+ | mk_Un_upper' k = Library.foldr (op RS o swap)
+ (replicate (k - 1) @{thm subset_trans[OF Un_upper1]}, @{thm Un_upper1});
+in
+ fun mk_Un_upper 1 1 = subset_refl
+ | mk_Un_upper n 1 = mk_Un_upper' (n - 2) RS @{thm subset_trans[OF Un_upper1]}
+ | mk_Un_upper n m = mk_Un_upper' (n - m) RS @{thm subset_trans[OF Un_upper2]};
+end;
+
+local
+ fun mk_UnIN' 0 = @{thm UnI2}
+ | mk_UnIN' m = mk_UnIN' (m - 1) RS @{thm UnI1};
+in
+ fun mk_UnIN 1 1 = @{thm TrueE[OF TrueI]}
+ | mk_UnIN n 1 = Library.foldr1 (op RS o swap) (replicate (n - 1) @{thm UnI1})
+ | mk_UnIN n m = mk_UnIN' (n - m)
+end;
+
+fun interleave xs ys = flat (map2 (fn x => fn y => [x, y]) xs ys);
+
+fun transpose [] = []
+ | transpose ([] :: xss) = transpose xss
+ | transpose xss = map hd xss :: transpose (map tl xss);
+
+fun seq_conds f n k xs =
+ if k = n then
+ map (f false) (take (k - 1) xs)
+ else
+ let val (negs, pos) = split_last (take k xs) in
+ map (f false) negs @ [f true pos]
+ end;
+
+fun mk_unabs_def 0 thm = thm
+ | mk_unabs_def n thm = mk_unabs_def (n - 1) thm RS @{thm spec[OF iffD1[OF fun_eq_iff]]};
+
+fun is_refl thm =
+ op aconv (HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of thm)))
+ handle TERM _ => false;
+
+val no_refl = filter_out is_refl;
+val no_reflexive = filter_out Thm.is_reflexive;
+
+fun fold_thms ctxt thms = Local_Defs.fold ctxt (distinct Thm.eq_thm_prop thms);
+fun unfold_thms ctxt thms = Local_Defs.unfold ctxt (distinct Thm.eq_thm_prop thms);
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Tools/bnf_wrap.ML Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,665 @@
+(* Title: HOL/BNF/Tools/bnf_wrap.ML
+ Author: Jasmin Blanchette, TU Muenchen
+ Copyright 2012
+
+Wrapping existing datatypes.
+*)
+
+signature BNF_WRAP =
+sig
+ val mk_half_pairss: 'a list -> ('a * 'a) list list
+ val mk_ctr: typ list -> term -> term
+ val mk_disc_or_sel: typ list -> term -> term
+ val base_name_of_ctr: term -> string
+ val wrap_datatype: ({prems: thm list, context: Proof.context} -> tactic) list list ->
+ ((bool * term list) * term) *
+ (binding list * (binding list list * (binding * term) list list)) -> local_theory ->
+ (term list * term list list * thm list * thm list * thm list * thm list list * thm list *
+ thm list list) * local_theory
+ val parse_wrap_options: bool parser
+ val parse_bound_term: (binding * string) parser
+end;
+
+structure BNF_Wrap : BNF_WRAP =
+struct
+
+open BNF_Util
+open BNF_Wrap_Tactics
+
+val isN = "is_";
+val unN = "un_";
+fun mk_unN 1 1 suf = unN ^ suf
+ | mk_unN _ l suf = unN ^ suf ^ string_of_int l;
+
+val case_congN = "case_cong";
+val case_eqN = "case_eq";
+val casesN = "cases";
+val collapseN = "collapse";
+val disc_excludeN = "disc_exclude";
+val disc_exhaustN = "disc_exhaust";
+val discsN = "discs";
+val distinctN = "distinct";
+val exhaustN = "exhaust";
+val expandN = "expand";
+val injectN = "inject";
+val nchotomyN = "nchotomy";
+val selsN = "sels";
+val splitN = "split";
+val split_asmN = "split_asm";
+val weak_case_cong_thmsN = "weak_case_cong";
+
+val std_binding = @{binding _};
+
+val induct_simp_attrs = @{attributes [induct_simp]};
+val cong_attrs = @{attributes [cong]};
+val iff_attrs = @{attributes [iff]};
+val safe_elim_attrs = @{attributes [elim!]};
+val simp_attrs = @{attributes [simp]};
+
+fun pad_list x n xs = xs @ replicate (n - length xs) x;
+
+fun unflat_lookup eq ys zs = map (map (fn x => nth zs (find_index (curry eq x) ys)));
+
+fun mk_half_pairss' _ [] = []
+ | mk_half_pairss' indent (x :: xs) =
+ indent @ fold_rev (cons o single o pair x) xs (mk_half_pairss' ([] :: indent) xs);
+
+fun mk_half_pairss xs = mk_half_pairss' [[]] xs;
+
+fun join_halves n half_xss other_half_xss =
+ let
+ val xsss =
+ map2 (map2 append) (Library.chop_groups n half_xss)
+ (transpose (Library.chop_groups n other_half_xss))
+ val xs = interleave (flat half_xss) (flat other_half_xss);
+ in (xs, xsss |> `transpose) end;
+
+fun mk_undefined T = Const (@{const_name undefined}, T);
+
+fun mk_ctr Ts t =
+ let val Type (_, Ts0) = body_type (fastype_of t) in
+ Term.subst_atomic_types (Ts0 ~~ Ts) t
+ end;
+
+fun mk_disc_or_sel Ts t =
+ Term.subst_atomic_types (snd (Term.dest_Type (domain_type (fastype_of t))) ~~ Ts) t;
+
+fun base_name_of_ctr c =
+ Long_Name.base_name (case head_of c of
+ Const (s, _) => s
+ | Free (s, _) => s
+ | _ => error "Cannot extract name of constructor");
+
+fun rapp u t = betapply (t, u);
+
+fun eta_expand_arg xs f_xs = fold_rev Term.lambda xs f_xs;
+
+fun prepare_wrap_datatype prep_term (((no_dests, raw_ctrs), raw_case),
+ (raw_disc_bindings, (raw_sel_bindingss, raw_sel_defaultss))) no_defs_lthy =
+ let
+ (* TODO: sanity checks on arguments *)
+ (* TODO: case syntax *)
+
+ val n = length raw_ctrs;
+ val ks = 1 upto n;
+
+ val _ = if n > 0 then () else error "No constructors specified";
+
+ val ctrs0 = map (prep_term no_defs_lthy) raw_ctrs;
+ val case0 = prep_term no_defs_lthy raw_case;
+ val sel_defaultss =
+ pad_list [] n (map (map (apsnd (prep_term no_defs_lthy))) raw_sel_defaultss);
+
+ val Type (dataT_name, As0) = body_type (fastype_of (hd ctrs0));
+ val data_b = Binding.qualified_name dataT_name;
+ val data_b_name = Binding.name_of data_b;
+
+ val (As, B) =
+ no_defs_lthy
+ |> mk_TFrees' (map Type.sort_of_atyp As0)
+ ||> the_single o fst o mk_TFrees 1;
+
+ val dataT = Type (dataT_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;
+
+ val raw_disc_bindings' = pad_list Binding.empty n raw_disc_bindings;
+
+ fun can_really_rely_on_disc k =
+ not (Binding.eq_name (nth raw_disc_bindings' (k - 1), Binding.empty)) orelse
+ nth ms (k - 1) = 0;
+ fun can_rely_on_disc k =
+ can_really_rely_on_disc k orelse (k = 1 andalso not (can_really_rely_on_disc 2));
+ fun can_omit_disc_binding k m =
+ n = 1 orelse m = 0 orelse (n = 2 andalso can_rely_on_disc (3 - k));
+
+ val std_disc_binding =
+ Binding.qualify false data_b_name o Binding.name o prefix isN o base_name_of_ctr;
+
+ val disc_bindings =
+ raw_disc_bindings'
+ |> map4 (fn k => fn m => fn ctr => fn disc =>
+ Option.map (Binding.qualify false data_b_name)
+ (if Binding.eq_name (disc, Binding.empty) then
+ if can_omit_disc_binding k m then NONE else SOME (std_disc_binding ctr)
+ else if Binding.eq_name (disc, std_binding) then
+ SOME (std_disc_binding ctr)
+ else
+ SOME disc)) ks ms ctrs0;
+
+ val no_discs = map is_none disc_bindings;
+ val no_discs_at_all = forall I no_discs;
+
+ fun std_sel_binding m l = Binding.name o mk_unN m l o base_name_of_ctr;
+
+ val sel_bindingss =
+ pad_list [] n raw_sel_bindingss
+ |> map3 (fn ctr => fn m => map2 (fn l => fn sel =>
+ Binding.qualify false data_b_name
+ (if Binding.eq_name (sel, Binding.empty) orelse Binding.eq_name (sel, std_binding) then
+ std_sel_binding m l ctr
+ else
+ sel)) (1 upto m) o pad_list Binding.empty m) ctrs0 ms;
+
+ fun mk_case Ts T =
+ let
+ val (bindings, body) = strip_type (fastype_of case0)
+ val Type (_, Ts0) = List.last bindings
+ in Term.subst_atomic_types ((body, T) :: (Ts0 ~~ Ts)) case0 end;
+
+ val casex = mk_case As B;
+ val case_Ts = map (fn Ts => Ts ---> B) ctr_Tss;
+
+ val (((((((xss, xss'), yss), fs), gs), [u', v']), (p, p')), names_lthy) = no_defs_lthy |>
+ mk_Freess' "x" ctr_Tss
+ ||>> mk_Freess "y" ctr_Tss
+ ||>> mk_Frees "f" case_Ts
+ ||>> mk_Frees "g" case_Ts
+ ||>> (apfst (map (rpair dataT)) oo Variable.variant_fixes) [data_b_name, data_b_name ^ "'"]
+ ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "P") HOLogic.boolT;
+
+ val u = Free u';
+ val v = Free v';
+ 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;
+
+ val eta_fs = map2 eta_expand_arg xss xfs;
+ val eta_gs = map2 eta_expand_arg xss xgs;
+
+ val fcase = Term.list_comb (casex, eta_fs);
+ val gcase = Term.list_comb (casex, eta_gs);
+
+ val ufcase = fcase $ u;
+ val vfcase = fcase $ v;
+ val vgcase = gcase $ v;
+
+ fun mk_u_eq_u () = HOLogic.mk_eq (u, u);
+
+ val u_eq_v = mk_Trueprop_eq (u, v);
+
+ 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
+ (case nth disc_bindings (k - 1) of
+ NONE => nth exist_xs_u_eq_ctrs (k - 1)
+ | SOME b => get_udisc b (k - 1));
+
+ val (all_sels_distinct, discs, selss, udiscs, uselss, vdiscs, vselss, disc_defs, sel_defs,
+ sel_defss, lthy') =
+ if no_dests then
+ (true, [], [], [], [], [], [], [], [], [], no_defs_lthy)
+ else
+ let
+ fun disc_free b = Free (Binding.name_of b, mk_pred1T dataT);
+
+ 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_default T t =
+ let
+ val Ts0 = map TFree (Term.add_tfreesT (fastype_of t) []);
+ val Ts = map TFree (Term.add_tfreesT T []);
+ in Term.subst_atomic_types (Ts0 ~~ Ts) t end;
+
+ fun mk_sel_case_args b proto_sels T =
+ map2 (fn Ts => fn k =>
+ (case AList.lookup (op =) proto_sels k of
+ NONE =>
+ (case AList.lookup Binding.eq_name (rev (nth sel_defaultss (k - 1))) b of
+ NONE => fold_rev (Term.lambda o curry Free Name.uu) Ts (mk_undefined T)
+ | SOME t => mk_default (Ts ---> T) t)
+ | SOME (xs, x) => fold_rev Term.lambda xs x)) 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 no_defs_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 no_defs_lthy T) ^
+ " vs. " ^ quote (Syntax.string_of_typ no_defs_lthy T')));
+ in
+ mk_Trueprop_eq (Free (Binding.name_of b, dataT --> T) $ u,
+ Term.list_comb (mk_case As T, mk_sel_case_args b proto_sels T) $ u)
+ end;
+
+ val sel_bindings = flat sel_bindingss;
+ val uniq_sel_bindings = distinct Binding.eq_name sel_bindings;
+ val all_sels_distinct = (length uniq_sel_bindings = length sel_bindings);
+
+ val sel_binding_index =
+ if all_sels_distinct then 1 upto length sel_bindings
+ else map (fn b => find_index (curry Binding.eq_name b) uniq_sel_bindings) sel_bindings;
+
+ val proto_sels = flat (map3 (fn k => fn xs => map (fn x => (k, (xs, x)))) ks xss xss);
+ val sel_infos =
+ AList.group (op =) (sel_binding_index ~~ proto_sels)
+ |> sort (int_ord o pairself fst)
+ |> map snd |> curry (op ~~) uniq_sel_bindings;
+ val sel_bindings = map fst sel_infos;
+
+ 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)) =
+ no_defs_lthy
+ |> apfst split_list o fold_map4 (fn k => fn m => fn exist_xs_u_eq_ctr =>
+ fn NONE =>
+ if n = 1 then pair (Term.lambda u (mk_u_eq_u ()), unique_disc_no_def)
+ else if m = 0 then pair (Term.lambda u exist_xs_u_eq_ctr, refl)
+ else pair (alternate_disc k, alternate_disc_no_def)
+ | SOME b => Specification.definition (SOME (b, NONE, NoSyn),
+ ((Thm.def_binding b, []), disc_spec b exist_xs_u_eq_ctr)) #>> apsnd snd)
+ ks ms 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.restore;
+
+ 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;
+
+ val udiscs = map (rapp u) discs;
+ val uselss = map (map (rapp u)) selss;
+
+ val vdiscs = map (rapp v) discs;
+ val vselss = map (map (rapp v)) selss;
+ in
+ (all_sels_distinct, discs, selss, udiscs, uselss, vdiscs, vselss, 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 (u, xctr)]) in
+ fold_rev Logic.all [p, u] (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
+ map4 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 (xss ~~ xctrs))
+ end;
+
+ val cases_goal =
+ map3 (fn xs => fn xctr => fn xf =>
+ fold_rev Logic.all (fs @ xs) (mk_Trueprop_eq (fcase $ xctr, xf))) xss xctrs xfs;
+
+ val goalss = [exhaust_goal] :: inject_goalss @ half_distinct_goalss @ [cases_goal];
+
+ fun after_qed thmss lthy =
+ let
+ val ([exhaust_thm], (inject_thmss, (half_distinct_thmss, [case_thms]))) =
+ (hd thmss, apsnd (chop (n * n)) (chop n (tl thmss)));
+
+ val inject_thms = flat inject_thmss;
+
+ val Tinst = map (pairself (certifyT lthy)) (map Logic.varifyT_global As ~~ As);
+
+ fun inst_thm t thm =
+ Drule.instantiate' [] [SOME (certify lthy t)]
+ (Thm.instantiate (Tinst, []) (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;
+
+ 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
+ Skip_Proof.prove lthy [] [] goal (fn _ => mk_nchotomy_tac n exhaust_thm)
+ end;
+
+ val (all_sel_thms, sel_thmss, disc_thmss, disc_thms, discI_thms, disc_exclude_thms,
+ disc_exhaust_thms, collapse_thms, expand_thms, case_eq_thms) =
+ if no_dests then
+ ([], [], [], [], [], [], [], [], [], [])
+ else
+ let
+ fun make_sel_thm xs' case_thm sel_def =
+ zero_var_indexes (Drule.gen_all (Drule.rename_bvars' (map (SOME o fst) xs')
+ (Drule.forall_intr_vars (case_thm RS (sel_def RS trans)))));
+
+ fun has_undefined_rhs thm =
+ (case snd (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of thm))) of
+ Const (@{const_name undefined}, _) => true
+ | _ => false);
+
+ val sel_thmss = map3 (map oo make_sel_thm) xss' case_thms sel_defss;
+
+ val all_sel_thms =
+ (if all_sels_distinct andalso forall null sel_defaultss 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_u_eq_u (), the_single exist_xs_u_eq_ctrs);
+ in
+ Skip_Proof.prove lthy [] [] goal (fn _ => mk_unique_disc_def_tac m uexhaust_thm)
+ |> singleton (Proof_Context.export names_lthy lthy)
+ |> Thm.close_derivation
+ 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));
+ in
+ Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
+ mk_alternate_disc_def_tac ctxt k (nth disc_defs (2 - k))
+ (nth distinct_thms (2 - k)) uexhaust_thm)
+ |> singleton (Proof_Context.export names_lthy lthy)
+ |> Thm.close_derivation
+ end;
+
+ val has_alternate_disc_def =
+ exists (fn def => Thm.eq_thm_prop (def, alternate_disc_no_def)) disc_defs;
+
+ 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
+ map3 mk_thms discI_thms not_discI_thms distinct_thmsss' |> `transpose
+ end;
+
+ val disc_thms = flat (map2 (fn true => K [] | false => I) no_discs disc_thmss);
+
+ val (disc_exclude_thms, (disc_exclude_thmsss', disc_exclude_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 = Skip_Proof.prove lthy [] [] goal (K tac);
+
+ val infos = ms ~~ discD_thms ~~ udiscs;
+ val half_pairss = mk_half_pairss infos;
+
+ val half_goalss = map mk_goal half_pairss;
+ val half_thmss =
+ map3 (fn [] => K (K []) | [goal] => fn [(((m, discD), _), _)] =>
+ fn disc_thm => [prove (mk_half_disc_exclude_tac 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 (prove o mk_other_half_disc_exclude_tac)) half_thmss
+ other_half_goalss;
+ in
+ join_halves n half_thmss other_half_thmss
+ |>> has_alternate_disc_def ? K []
+ end;
+
+ val disc_exhaust_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
+ Skip_Proof.prove lthy [] [] goal (fn _ =>
+ mk_disc_exhaust_tac n exhaust_thm discI_thms)
+ end;
+
+ val disc_exhaust_thms =
+ if has_alternate_disc_def orelse no_discs_at_all then [] else [disc_exhaust_thm];
+
+ val (collapse_thms, collapse_thm_opts) =
+ let
+ fun mk_goal ctr udisc usels =
+ let
+ val prem = HOLogic.mk_Trueprop udisc;
+ val concl =
+ mk_Trueprop_eq ((null usels ? swap) (Term.list_comb (ctr, usels), u));
+ in
+ if prem aconv concl then NONE
+ else SOME (Logic.all u (Logic.mk_implies (prem, concl)))
+ end;
+ val goals = map3 mk_goal ctrs udiscs uselss;
+ in
+ map4 (fn m => fn discD => fn sel_thms => Option.map (fn goal =>
+ Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
+ mk_collapse_tac ctxt m discD sel_thms)
+ |> perhaps (try (fn thm => refl RS thm)))) ms discD_thms sel_thmss goals
+ |> `(map_filter I)
+ end;
+
+ val expand_thms =
+ 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 uncollapse_thms =
+ map (fn NONE => Drule.dummy_thm | SOME thm => thm RS sym) collapse_thm_opts;
+
+ val goal =
+ Library.foldr Logic.list_implies
+ (map5 mk_prems ks udiscs uselss vdiscs vselss, u_eq_v);
+ in
+ [Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
+ mk_expand_tac ctxt n ms (inst_thm u disc_exhaust_thm)
+ (inst_thm v disc_exhaust_thm) uncollapse_thms disc_exclude_thmsss
+ disc_exclude_thmsss')]
+ |> Proof_Context.export names_lthy lthy
+ end;
+
+ val case_eq_thms =
+ let
+ fun mk_body f usels = Term.list_comb (f, usels);
+ val goal = mk_Trueprop_eq (ufcase, mk_IfN B udiscs (map2 mk_body fs uselss));
+ in
+ [Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
+ mk_case_eq_tac ctxt n uexhaust_thm case_thms disc_thmss' sel_thmss)]
+ |> Proof_Context.export names_lthy lthy
+ end;
+ in
+ (all_sel_thms, sel_thmss, disc_thmss, disc_thms, discI_thms, disc_exclude_thms,
+ disc_exhaust_thms, collapse_thms, expand_thms, case_eq_thms)
+ end;
+
+ val (case_cong_thm, weak_case_cong_thm) =
+ let
+ fun mk_prem xctr xs f g =
+ fold_rev Logic.all xs (Logic.mk_implies (mk_Trueprop_eq (v, xctr),
+ mk_Trueprop_eq (f, g)));
+
+ val goal =
+ Logic.list_implies (u_eq_v :: map4 mk_prem xctrs xss fs gs,
+ mk_Trueprop_eq (ufcase, vgcase));
+ val weak_goal = Logic.mk_implies (u_eq_v, mk_Trueprop_eq (ufcase, vfcase));
+ in
+ (Skip_Proof.prove lthy [] [] goal (fn _ => mk_case_cong_tac uexhaust_thm case_thms),
+ Skip_Proof.prove lthy [] [] weak_goal (K (etac arg_cong 1)))
+ |> pairself (singleton (Proof_Context.export names_lthy lthy))
+ end;
+
+ val (split_thm, split_asm_thm) =
+ let
+ fun mk_conjunct xctr xs f_xs =
+ list_all_free xs (HOLogic.mk_imp (HOLogic.mk_eq (u, xctr), q $ f_xs));
+ fun mk_disjunct xctr xs f_xs =
+ list_exists_free xs (HOLogic.mk_conj (HOLogic.mk_eq (u, xctr),
+ HOLogic.mk_not (q $ f_xs)));
+
+ val lhs = q $ ufcase;
+
+ val goal =
+ mk_Trueprop_eq (lhs, Library.foldr1 HOLogic.mk_conj (map3 mk_conjunct xctrs xss xfs));
+ val asm_goal =
+ mk_Trueprop_eq (lhs, HOLogic.mk_not (Library.foldr1 HOLogic.mk_disj
+ (map3 mk_disjunct xctrs xss xfs)));
+
+ val split_thm =
+ Skip_Proof.prove lthy [] [] goal
+ (fn _ => mk_split_tac uexhaust_thm case_thms inject_thmss distinct_thmsss)
+ |> singleton (Proof_Context.export names_lthy lthy)
+ val split_asm_thm =
+ Skip_Proof.prove lthy [] [] asm_goal (fn {context = ctxt, ...} =>
+ mk_split_asm_tac ctxt split_thm)
+ |> singleton (Proof_Context.export names_lthy lthy)
+ in
+ (split_thm, split_asm_thm)
+ 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 dataT_name));
+
+ val notes =
+ [(case_congN, [case_cong_thm], []),
+ (case_eqN, case_eq_thms, []),
+ (casesN, case_thms, simp_attrs),
+ (collapseN, collapse_thms, simp_attrs),
+ (discsN, disc_thms, simp_attrs),
+ (disc_excludeN, disc_exclude_thms, []),
+ (disc_exhaustN, disc_exhaust_thms, [exhaust_case_names_attr]),
+ (distinctN, distinct_thms, simp_attrs @ induct_simp_attrs),
+ (exhaustN, [exhaust_thm], [exhaust_case_names_attr, cases_type_attr]),
+ (expandN, expand_thms, []),
+ (injectN, inject_thms, iff_attrs @ induct_simp_attrs),
+ (nchotomyN, [nchotomy_thm], []),
+ (selsN, all_sel_thms, simp_attrs),
+ (splitN, [split_thm], []),
+ (split_asmN, [split_asm_thm], []),
+ (weak_case_cong_thmsN, [weak_case_cong_thm], cong_attrs)]
+ |> filter_out (null o #2)
+ |> map (fn (thmN, thms, attrs) =>
+ ((Binding.qualify true data_b_name (Binding.name thmN), attrs), [(thms, [])]));
+
+ val notes' =
+ [(map (fn th => th RS notE) distinct_thms, safe_elim_attrs)]
+ |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])]));
+ in
+ ((discs, selss, inject_thms, distinct_thms, case_thms, disc_thmss, discI_thms, sel_thmss),
+ lthy |> Local_Theory.notes (notes' @ notes) |> snd)
+ end;
+ in
+ (goalss, after_qed, lthy')
+ end;
+
+fun wrap_datatype tacss = (fn (goalss, after_qed, lthy) =>
+ map2 (map2 (Skip_Proof.prove lthy [] [])) goalss tacss
+ |> (fn thms => after_qed thms lthy)) oo prepare_wrap_datatype (K I);
+
+val wrap_datatype_cmd = (fn (goalss, after_qed, lthy) =>
+ Proof.theorem NONE (snd oo after_qed) (map (map (rpair [])) goalss) lthy) oo
+ prepare_wrap_datatype Syntax.read_term;
+
+fun parse_bracket_list parser = @{keyword "["} |-- Parse.list parser --| @{keyword "]"};
+
+val parse_bindings = parse_bracket_list Parse.binding;
+val parse_bindingss = parse_bracket_list parse_bindings;
+
+val parse_bound_term = (Parse.binding --| @{keyword ":"}) -- Parse.term;
+val parse_bound_terms = parse_bracket_list parse_bound_term;
+val parse_bound_termss = parse_bracket_list parse_bound_terms;
+
+val parse_wrap_options =
+ Scan.optional (@{keyword "("} |-- (@{keyword "no_dests"} >> K true) --| @{keyword ")"}) false;
+
+val _ =
+ Outer_Syntax.local_theory_to_proof @{command_spec "wrap_data"} "wraps an existing datatype"
+ ((parse_wrap_options -- (@{keyword "["} |-- Parse.list Parse.term --| @{keyword "]"}) --
+ Parse.term -- Scan.optional (parse_bindings -- Scan.optional (parse_bindingss --
+ Scan.optional parse_bound_termss []) ([], [])) ([], ([], [])))
+ >> wrap_datatype_cmd);
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF/Tools/bnf_wrap_tactics.ML Fri Sep 21 16:45:06 2012 +0200
@@ -0,0 +1,122 @@
+(* Title: HOL/BNF/Tools/bnf_wrap_tactics.ML
+ Author: Jasmin Blanchette, TU Muenchen
+ Copyright 2012
+
+Tactics for wrapping datatypes.
+*)
+
+signature BNF_WRAP_TACTICS =
+sig
+ val mk_alternate_disc_def_tac: Proof.context -> int -> thm -> thm -> thm -> tactic
+ val mk_case_cong_tac: thm -> thm list -> tactic
+ val mk_case_eq_tac: Proof.context -> int -> thm -> thm list -> thm list list -> thm list list ->
+ tactic
+ val mk_collapse_tac: Proof.context -> int -> thm -> thm list -> tactic
+ val mk_disc_exhaust_tac: int -> thm -> thm list -> tactic
+ val mk_expand_tac: Proof.context -> int -> int list -> thm -> thm -> thm list ->
+ thm list list list -> thm list list list -> tactic
+ val mk_half_disc_exclude_tac: int -> thm -> thm -> tactic
+ val mk_nchotomy_tac: int -> thm -> tactic
+ val mk_other_half_disc_exclude_tac: thm -> tactic
+ val mk_split_tac: thm -> thm list -> thm list list -> thm list list list -> tactic
+ val mk_split_asm_tac: Proof.context -> thm -> tactic
+ val mk_unique_disc_def_tac: int -> thm -> tactic
+end;
+
+structure BNF_Wrap_Tactics : BNF_WRAP_TACTICS =
+struct
+
+open BNF_Util
+open BNF_Tactics
+
+val meta_mp = @{thm meta_mp};
+
+fun if_P_or_not_P_OF pos thm = thm RS (if pos then @{thm if_P} else @{thm if_not_P});
+
+fun mk_nchotomy_tac n exhaust =
+ (rtac allI THEN' rtac exhaust THEN'
+ EVERY' (maps (fn k => [rtac (mk_disjIN n k), REPEAT_DETERM o rtac exI, atac]) (1 upto n))) 1;
+
+fun mk_unique_disc_def_tac m uexhaust =
+ EVERY' [rtac iffI, rtac uexhaust, REPEAT_DETERM_N m o rtac exI, atac, rtac refl] 1;
+
+fun mk_alternate_disc_def_tac ctxt k other_disc_def distinct uexhaust =
+ EVERY' ([subst_tac ctxt [other_disc_def], rtac @{thm iffI_np}, REPEAT_DETERM o etac exE,
+ hyp_subst_tac, SELECT_GOAL (unfold_thms_tac ctxt [not_ex]), REPEAT_DETERM o rtac allI,
+ rtac distinct, rtac uexhaust] @
+ (([etac notE, REPEAT_DETERM o rtac exI, atac], [REPEAT_DETERM o rtac exI, atac])
+ |> k = 1 ? swap |> op @)) 1;
+
+fun mk_half_disc_exclude_tac m discD disc' =
+ (dtac discD THEN' REPEAT_DETERM_N m o etac exE THEN' hyp_subst_tac THEN' rtac disc') 1;
+
+fun mk_other_half_disc_exclude_tac half = (etac @{thm contrapos_pn} THEN' etac half) 1;
+
+fun mk_disc_exhaust_tac n exhaust discIs =
+ (rtac exhaust THEN'
+ EVERY' (map2 (fn k => fn discI =>
+ dtac discI THEN' select_prem_tac n (etac meta_mp) k THEN' atac) (1 upto n) discIs)) 1;
+
+fun mk_collapse_tac ctxt m discD sels =
+ (dtac discD THEN'
+ (if m = 0 then
+ atac
+ else
+ REPEAT_DETERM_N m o etac exE THEN' hyp_subst_tac THEN'
+ SELECT_GOAL (unfold_thms_tac ctxt sels) THEN' rtac refl)) 1;
+
+fun mk_expand_tac ctxt n ms udisc_exhaust vdisc_exhaust uncollapses disc_excludesss
+ disc_excludesss' =
+ if ms = [0] then
+ rtac (@{thm trans_sym} OF (replicate 2 (the_single uncollapses RS sym))) 1
+ else
+ let
+ val ks = 1 upto n;
+ val maybe_atac = if n = 1 then K all_tac else atac;
+ in
+ (rtac udisc_exhaust THEN'
+ EVERY' (map5 (fn k => fn m => fn disc_excludess => fn disc_excludess' => fn uuncollapse =>
+ EVERY' [if m = 0 then K all_tac else subst_tac ctxt [uuncollapse] THEN' maybe_atac,
+ rtac sym, rtac vdisc_exhaust,
+ EVERY' (map4 (fn k' => fn disc_excludes => fn disc_excludes' => fn vuncollapse =>
+ EVERY'
+ (if k' = k then
+ if m = 0 then
+ [hyp_subst_tac, rtac refl]
+ else
+ [subst_tac ctxt [vuncollapse], maybe_atac,
+ if n = 1 then K all_tac else EVERY' [dtac meta_mp, atac, dtac meta_mp, atac],
+ REPEAT_DETERM_N (Int.max (0, m - 1)) o etac conjE, asm_simp_tac (ss_only [])]
+ else
+ [dtac (the_single (if k = n then disc_excludes else disc_excludes')),
+ etac (if k = n then @{thm iff_contradict(1)} else @{thm iff_contradict(2)}),
+ atac, atac]))
+ ks disc_excludess disc_excludess' uncollapses)])
+ ks ms disc_excludesss disc_excludesss' uncollapses)) 1
+ end;
+
+fun mk_case_eq_tac ctxt n uexhaust cases discss' selss =
+ (rtac uexhaust THEN'
+ EVERY' (map3 (fn casex => fn if_discs => fn sels =>
+ EVERY' [hyp_subst_tac, SELECT_GOAL (unfold_thms_tac ctxt (if_discs @ sels)), rtac casex])
+ cases (map2 (seq_conds if_P_or_not_P_OF n) (1 upto n) discss') selss)) 1;
+
+fun mk_case_cong_tac uexhaust cases =
+ (rtac uexhaust THEN'
+ EVERY' (maps (fn casex => [dtac sym, asm_simp_tac (ss_only [casex])]) cases)) 1;
+
+val naked_ctxt = Proof_Context.init_global @{theory HOL};
+
+fun mk_split_tac uexhaust cases injectss distinctsss =
+ rtac uexhaust 1 THEN
+ ALLGOALS (fn k => (hyp_subst_tac THEN'
+ simp_tac (ss_only (@{thms simp_thms} @ cases @ nth injectss (k - 1) @
+ flat (nth distinctsss (k - 1))))) k) THEN
+ ALLGOALS (blast_tac naked_ctxt);
+
+val split_asm_thms = @{thms imp_conv_disj de_Morgan_conj de_Morgan_disj not_not not_ex};
+
+fun mk_split_asm_tac ctxt split =
+ rtac (split RS trans) 1 THEN unfold_thms_tac ctxt split_asm_thms THEN rtac refl 1;
+
+end;
--- a/src/HOL/Codatatype/BNF.thy Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,16 +0,0 @@
-(* Title: HOL/BNF/BNF.thy
- Author: Dmitriy Traytel, TU Muenchen
- Author: Andrei Popescu, TU Muenchen
- Author: Jasmin Blanchette, TU Muenchen
- Copyright 2012
-
-Bounded natural functors for (co)datatypes.
-*)
-
-header {* Bounded Natural Functors for (Co)datatypes *}
-
-theory BNF
-imports More_BNFs
-begin
-
-end
--- a/src/HOL/Codatatype/BNF_Comp.thy Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,91 +0,0 @@
-(* Title: HOL/BNF/BNF_Comp.thy
- Author: Dmitriy Traytel, TU Muenchen
- Copyright 2012
-
-Composition of bounded natural functors.
-*)
-
-header {* Composition of Bounded Natural Functors *}
-
-theory BNF_Comp
-imports Basic_BNFs
-begin
-
-lemma empty_natural: "(\<lambda>_. {}) o f = image g o (\<lambda>_. {})"
-by (rule ext) simp
-
-lemma Union_natural: "Union o image (image f) = image f o Union"
-by (rule ext) (auto simp only: o_apply)
-
-lemma in_Union_o_assoc: "x \<in> (Union o gset o gmap) A \<Longrightarrow> x \<in> (Union o (gset o gmap)) A"
-by (unfold o_assoc)
-
-lemma comp_single_set_bd:
- assumes fbd_Card_order: "Card_order fbd" and
- fset_bd: "\<And>x. |fset x| \<le>o fbd" and
- gset_bd: "\<And>x. |gset x| \<le>o gbd"
- shows "|\<Union>fset ` gset x| \<le>o gbd *c fbd"
-apply (subst sym[OF SUP_def])
-apply (rule ordLeq_transitive)
-apply (rule card_of_UNION_Sigma)
-apply (subst SIGMA_CSUM)
-apply (rule ordLeq_transitive)
-apply (rule card_of_Csum_Times')
-apply (rule fbd_Card_order)
-apply (rule ballI)
-apply (rule fset_bd)
-apply (rule ordLeq_transitive)
-apply (rule cprod_mono1)
-apply (rule gset_bd)
-apply (rule ordIso_imp_ordLeq)
-apply (rule ordIso_refl)
-apply (rule Card_order_cprod)
-done
-
-lemma Union_image_insert: "\<Union>f ` insert a B = f a \<union> \<Union>f ` B"
-by simp
-
-lemma Union_image_empty: "A \<union> \<Union>f ` {} = A"
-by simp
-
-lemma image_o_collect: "collect ((\<lambda>f. image g o f) ` F) = image g o collect F"
-by (rule ext) (auto simp add: collect_def)
-
-lemma conj_subset_def: "A \<subseteq> {x. P x \<and> Q x} = (A \<subseteq> {x. P x} \<and> A \<subseteq> {x. Q x})"
-by blast
-
-lemma UN_image_subset: "\<Union>f ` g x \<subseteq> X = (g x \<subseteq> {x. f x \<subseteq> X})"
-by blast
-
-lemma comp_set_bd_Union_o_collect: "|\<Union>\<Union>(\<lambda>f. f x) ` X| \<le>o hbd \<Longrightarrow> |(Union \<circ> collect X) x| \<le>o hbd"
-by (unfold o_apply collect_def SUP_def)
-
-lemma wpull_cong:
-"\<lbrakk>A' = A; B1' = B1; B2' = B2; wpull A B1 B2 f1 f2 p1 p2\<rbrakk> \<Longrightarrow> wpull A' B1' B2' f1 f2 p1 p2"
-by simp
-
-lemma Id_def': "Id = {(a,b). a = b}"
-by auto
-
-lemma Gr_fst_snd: "(Gr R fst)^-1 O Gr R snd = R"
-unfolding Gr_def by auto
-
-lemma subst_rel_def: "A = B \<Longrightarrow> (Gr A f)^-1 O Gr A g = (Gr B f)^-1 O Gr B g"
-by simp
-
-lemma abs_pred_def: "\<lbrakk>\<And>x y. (x, y) \<in> rel = pred x y\<rbrakk> \<Longrightarrow> rel = Collect (split pred)"
-by auto
-
-lemma Collect_split_cong: "Collect (split pred) = Collect (split pred') \<Longrightarrow> pred = pred'"
-by blast
-
-lemma pred_def_abs: "rel = Collect (split pred) \<Longrightarrow> pred = (\<lambda>x y. (x, y) \<in> rel)"
-by auto
-
-lemma mem_Id_eq_eq: "(\<lambda>x y. (x, y) \<in> Id) = (op =)"
-by simp
-
-ML_file "Tools/bnf_comp_tactics.ML"
-ML_file "Tools/bnf_comp.ML"
-
-end
--- a/src/HOL/Codatatype/BNF_Def.thy Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,151 +0,0 @@
-(* Title: HOL/BNF/BNF_Def.thy
- Author: Dmitriy Traytel, TU Muenchen
- Copyright 2012
-
-Definition of bounded natural functors.
-*)
-
-header {* Definition of Bounded Natural Functors *}
-
-theory BNF_Def
-imports BNF_Util
-keywords
- "print_bnfs" :: diag and
- "bnf_def" :: thy_goal
-begin
-
-lemma collect_o: "collect F o g = collect ((\<lambda>f. f o g) ` F)"
-by (rule ext) (auto simp only: o_apply collect_def)
-
-lemma converse_mono:
-"R1 ^-1 \<subseteq> R2 ^-1 \<longleftrightarrow> R1 \<subseteq> R2"
-unfolding converse_def by auto
-
-lemma converse_shift:
-"R1 \<subseteq> R2 ^-1 \<Longrightarrow> R1 ^-1 \<subseteq> R2"
-unfolding converse_def by auto
-
-definition convol ("<_ , _>") where
-"<f , g> \<equiv> %a. (f a, g a)"
-
-lemma fst_convol:
-"fst o <f , g> = f"
-apply(rule ext)
-unfolding convol_def by simp
-
-lemma snd_convol:
-"snd o <f , g> = g"
-apply(rule ext)
-unfolding convol_def by simp
-
-lemma convol_memI:
-"\<lbrakk>f x = f' x; g x = g' x; P x\<rbrakk> \<Longrightarrow> <f , g> x \<in> {(f' a, g' a) |a. P a}"
-unfolding convol_def by auto
-
-definition csquare where
-"csquare A f1 f2 p1 p2 \<longleftrightarrow> (\<forall> a \<in> A. f1 (p1 a) = f2 (p2 a))"
-
-(* The pullback of sets *)
-definition thePull where
-"thePull B1 B2 f1 f2 = {(b1,b2). b1 \<in> B1 \<and> b2 \<in> B2 \<and> f1 b1 = f2 b2}"
-
-lemma wpull_thePull:
-"wpull (thePull B1 B2 f1 f2) B1 B2 f1 f2 fst snd"
-unfolding wpull_def thePull_def by auto
-
-lemma wppull_thePull:
-assumes "wppull A B1 B2 f1 f2 e1 e2 p1 p2"
-shows
-"\<exists> j. \<forall> a' \<in> thePull B1 B2 f1 f2.
- j a' \<in> A \<and>
- e1 (p1 (j a')) = e1 (fst a') \<and> e2 (p2 (j a')) = e2 (snd a')"
-(is "\<exists> j. \<forall> a' \<in> ?A'. ?phi a' (j a')")
-proof(rule bchoice[of ?A' ?phi], default)
- fix a' assume a': "a' \<in> ?A'"
- hence "fst a' \<in> B1" unfolding thePull_def by auto
- moreover
- from a' have "snd a' \<in> B2" unfolding thePull_def by auto
- moreover have "f1 (fst a') = f2 (snd a')"
- using a' unfolding csquare_def thePull_def by auto
- ultimately show "\<exists> ja'. ?phi a' ja'"
- using assms unfolding wppull_def by blast
-qed
-
-lemma wpull_wppull:
-assumes wp: "wpull A' B1 B2 f1 f2 p1' p2'" and
-1: "\<forall> a' \<in> A'. j a' \<in> A \<and> e1 (p1 (j a')) = e1 (p1' a') \<and> e2 (p2 (j a')) = e2 (p2' a')"
-shows "wppull A B1 B2 f1 f2 e1 e2 p1 p2"
-unfolding wppull_def proof safe
- fix b1 b2
- assume b1: "b1 \<in> B1" and b2: "b2 \<in> B2" and f: "f1 b1 = f2 b2"
- then obtain a' where a': "a' \<in> A'" and b1: "b1 = p1' a'" and b2: "b2 = p2' a'"
- using wp unfolding wpull_def by blast
- show "\<exists>a\<in>A. e1 (p1 a) = e1 b1 \<and> e2 (p2 a) = e2 b2"
- apply (rule bexI[of _ "j a'"]) unfolding b1 b2 using a' 1 by auto
-qed
-
-lemma wppull_id: "\<lbrakk>wpull UNIV UNIV UNIV f1 f2 p1 p2; e1 = id; e2 = id\<rbrakk> \<Longrightarrow>
- wppull UNIV UNIV UNIV f1 f2 e1 e2 p1 p2"
-by (erule wpull_wppull) auto
-
-lemma Id_alt: "Id = Gr UNIV id"
-unfolding Gr_def by auto
-
-lemma Gr_UNIV_id: "f = id \<Longrightarrow> (Gr UNIV f)^-1 O Gr UNIV f = Gr UNIV f"
-unfolding Gr_def by auto
-
-lemma Gr_mono: "A \<subseteq> B \<Longrightarrow> Gr A f \<subseteq> Gr B f"
-unfolding Gr_def by auto
-
-lemma wpull_Gr:
-"wpull (Gr A f) A (f ` A) f id fst snd"
-unfolding wpull_def Gr_def by auto
-
-definition "pick_middle P Q a c = (SOME b. (a,b) \<in> P \<and> (b,c) \<in> Q)"
-
-lemma pick_middle:
-"(a,c) \<in> P O Q \<Longrightarrow> (a, pick_middle P Q a c) \<in> P \<and> (pick_middle P Q a c, c) \<in> Q"
-unfolding pick_middle_def apply(rule someI_ex)
-using assms unfolding relcomp_def by auto
-
-definition fstO where "fstO P Q ac = (fst ac, pick_middle P Q (fst ac) (snd ac))"
-definition sndO where "sndO P Q ac = (pick_middle P Q (fst ac) (snd ac), snd ac)"
-
-lemma fstO_in: "ac \<in> P O Q \<Longrightarrow> fstO P Q ac \<in> P"
-unfolding fstO_def
-by (subst (asm) surjective_pairing) (rule pick_middle[THEN conjunct1])
-
-lemma fst_fstO: "fst bc = (fst \<circ> fstO P Q) bc"
-unfolding comp_def fstO_def by simp
-
-lemma snd_sndO: "snd bc = (snd \<circ> sndO P Q) bc"
-unfolding comp_def sndO_def by simp
-
-lemma sndO_in: "ac \<in> P O Q \<Longrightarrow> sndO P Q ac \<in> Q"
-unfolding sndO_def
-by (subst (asm) surjective_pairing) (rule pick_middle[THEN conjunct2])
-
-lemma csquare_fstO_sndO:
-"csquare (P O Q) snd fst (fstO P Q) (sndO P Q)"
-unfolding csquare_def fstO_def sndO_def using pick_middle by simp
-
-lemma wppull_fstO_sndO:
-shows "wppull (P O Q) P Q snd fst fst snd (fstO P Q) (sndO P Q)"
-using pick_middle unfolding wppull_def fstO_def sndO_def relcomp_def by auto
-
-lemma snd_fst_flip: "snd xy = (fst o (%(x, y). (y, x))) xy"
-by (simp split: prod.split)
-
-lemma fst_snd_flip: "fst xy = (snd o (%(x, y). (y, x))) xy"
-by (simp split: prod.split)
-
-lemma flip_rel: "A \<subseteq> (R ^-1) \<Longrightarrow> (%(x, y). (y, x)) ` A \<subseteq> R"
-by auto
-
-lemma pointfreeE: "f o g = f' o g' \<Longrightarrow> f (g x) = f' (g' x)"
-unfolding o_def fun_eq_iff by simp
-
-ML_file "Tools/bnf_def_tactics.ML"
-ML_file"Tools/bnf_def.ML"
-
-end
--- a/src/HOL/Codatatype/BNF_FP.thy Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,113 +0,0 @@
-(* Title: HOL/BNF/BNF_FP.thy
- Author: Dmitriy Traytel, TU Muenchen
- Author: Jasmin Blanchette, TU Muenchen
- Copyright 2012
-
-Composition of bounded natural functors.
-*)
-
-header {* Composition of Bounded Natural Functors *}
-
-theory BNF_FP
-imports BNF_Comp BNF_Wrap
-keywords
- "defaults"
-begin
-
-lemma case_unit: "(case u of () => f) = f"
-by (cases u) (hypsubst, rule unit.cases)
-
-lemma unit_all_impI: "(P () \<Longrightarrow> Q ()) \<Longrightarrow> \<forall>x. P x \<longrightarrow> Q x"
-by simp
-
-lemma prod_all_impI: "(\<And>x y. P (x, y) \<Longrightarrow> Q (x, y)) \<Longrightarrow> \<forall>x. P x \<longrightarrow> Q x"
-by clarify
-
-lemma prod_all_impI_step: "(\<And>x. \<forall>y. P (x, y) \<longrightarrow> Q (x, y)) \<Longrightarrow> \<forall>x. P x \<longrightarrow> Q x"
-by auto
-
-lemma all_unit_eq: "(\<And>x. PROP P x) \<equiv> PROP P ()"
-by simp
-
-lemma all_prod_eq: "(\<And>x. PROP P x) \<equiv> (\<And>a b. PROP P (a, b))"
-by clarsimp
-
-lemma rev_bspec: "a \<in> A \<Longrightarrow> \<forall>z \<in> A. P z \<Longrightarrow> P a"
-by simp
-
-lemma Un_cong: "\<lbrakk>A = B; C = D\<rbrakk> \<Longrightarrow> A \<union> C = B \<union> D"
-by simp
-
-lemma pointfree_idE: "f o g = id \<Longrightarrow> f (g x) = x"
-unfolding o_def fun_eq_iff by simp
-
-lemma o_bij:
- assumes gf: "g o f = id" and fg: "f o g = id"
- shows "bij f"
-unfolding bij_def inj_on_def surj_def proof safe
- fix a1 a2 assume "f a1 = f a2"
- hence "g ( f a1) = g (f a2)" by simp
- thus "a1 = a2" using gf unfolding fun_eq_iff by simp
-next
- fix b
- have "b = f (g b)"
- using fg unfolding fun_eq_iff by simp
- thus "EX a. b = f a" by blast
-qed
-
-lemma ssubst_mem: "\<lbrakk>t = s; s \<in> X\<rbrakk> \<Longrightarrow> t \<in> X" by simp
-
-lemma sum_case_step:
- "sum_case (sum_case f' g') g (Inl p) = sum_case f' g' p"
- "sum_case f (sum_case f' g') (Inr p) = sum_case f' g' p"
-by auto
-
-lemma one_pointE: "\<lbrakk>\<And>x. s = x \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
-by simp
-
-lemma obj_one_pointE: "\<forall>x. s = x \<longrightarrow> P \<Longrightarrow> P"
-by blast
-
-lemma obj_sumE_f':
-"\<lbrakk>\<forall>x. s = f (Inl x) \<longrightarrow> P; \<forall>x. s = f (Inr x) \<longrightarrow> P\<rbrakk> \<Longrightarrow> s = f x \<longrightarrow> P"
-by (cases x) blast+
-
-lemma obj_sumE_f:
-"\<lbrakk>\<forall>x. s = f (Inl x) \<longrightarrow> P; \<forall>x. s = f (Inr x) \<longrightarrow> P\<rbrakk> \<Longrightarrow> \<forall>x. s = f x \<longrightarrow> P"
-by (rule allI) (rule obj_sumE_f')
-
-lemma obj_sumE: "\<lbrakk>\<forall>x. s = Inl x \<longrightarrow> P; \<forall>x. s = Inr x \<longrightarrow> P\<rbrakk> \<Longrightarrow> P"
-by (cases s) auto
-
-lemma obj_sum_step':
-"\<lbrakk>\<forall>x. s = f (Inr (Inl x)) \<longrightarrow> P; \<forall>x. s = f (Inr (Inr x)) \<longrightarrow> P\<rbrakk> \<Longrightarrow> s = f (Inr x) \<longrightarrow> P"
-by (cases x) blast+
-
-lemma obj_sum_step:
-"\<lbrakk>\<forall>x. s = f (Inr (Inl x)) \<longrightarrow> P; \<forall>x. s = f (Inr (Inr x)) \<longrightarrow> P\<rbrakk> \<Longrightarrow> \<forall>x. s = f (Inr x) \<longrightarrow> P"
-by (rule allI) (rule obj_sum_step')
-
-lemma sum_case_if:
-"sum_case f g (if p then Inl x else Inr y) = (if p then f x else g y)"
-by simp
-
-lemma mem_UN_compreh_eq: "(z : \<Union>{y. \<exists>x\<in>A. y = F x}) = (\<exists>x\<in>A. z : F x)"
-by blast
-
-lemma prod_set_simps:
-"fsts (x, y) = {x}"
-"snds (x, y) = {y}"
-unfolding fsts_def snds_def by simp+
-
-lemma sum_set_simps:
-"setl (Inl x) = {x}"
-"setl (Inr x) = {}"
-"setr (Inl x) = {}"
-"setr (Inr x) = {x}"
-unfolding sum_set_defs by simp+
-
-ML_file "Tools/bnf_fp.ML"
-ML_file "Tools/bnf_fp_sugar_tactics.ML"
-ML_file "Tools/bnf_fp_sugar.ML"
-
-end
--- a/src/HOL/Codatatype/BNF_GFP.thy Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,331 +0,0 @@
-(* Title: HOL/BNF/BNF_GFP.thy
- Author: Dmitriy Traytel, TU Muenchen
- Copyright 2012
-
-Greatest fixed point operation on bounded natural functors.
-*)
-
-header {* Greatest Fixed Point Operation on Bounded Natural Functors *}
-
-theory BNF_GFP
-imports BNF_FP Equiv_Relations_More "~~/src/HOL/Library/Prefix_Order"
-keywords
- "codata_raw" :: thy_decl and
- "codata" :: thy_decl
-begin
-
-lemma sum_case_comp_Inl:
-"sum_case f g \<circ> Inl = f"
-unfolding comp_def by simp
-
-lemma sum_case_expand_Inr: "f o Inl = g \<Longrightarrow> f x = sum_case g (f o Inr) x"
-by (auto split: sum.splits)
-
-lemma converse_Times: "(A \<times> B) ^-1 = B \<times> A"
-by auto
-
-lemma equiv_triv1:
-assumes "equiv A R" and "(a, b) \<in> R" and "(a, c) \<in> R"
-shows "(b, c) \<in> R"
-using assms unfolding equiv_def sym_def trans_def by blast
-
-lemma equiv_triv2:
-assumes "equiv A R" and "(a, b) \<in> R" and "(b, c) \<in> R"
-shows "(a, c) \<in> R"
-using assms unfolding equiv_def trans_def by blast
-
-lemma equiv_proj:
- assumes e: "equiv A R" and "z \<in> R"
- shows "(proj R o fst) z = (proj R o snd) z"
-proof -
- from assms(2) have z: "(fst z, snd z) \<in> R" by auto
- have P: "\<And>x. (fst z, x) \<in> R \<Longrightarrow> (snd z, x) \<in> R" by (erule equiv_triv1[OF e z])
- have "\<And>x. (snd z, x) \<in> R \<Longrightarrow> (fst z, x) \<in> R" by (erule equiv_triv2[OF e z])
- with P show ?thesis unfolding proj_def[abs_def] by auto
-qed
-
-(* Operators: *)
-definition diag where "diag A \<equiv> {(a,a) | a. a \<in> A}"
-definition image2 where "image2 A f g = {(f a, g a) | a. a \<in> A}"
-
-lemma diagI: "x \<in> A \<Longrightarrow> (x, x) \<in> diag A"
-unfolding diag_def by simp
-
-lemma diagE: "(a, b) \<in> diag A \<Longrightarrow> a = b"
-unfolding diag_def by simp
-
-lemma diagE': "x \<in> diag A \<Longrightarrow> fst x = snd x"
-unfolding diag_def by auto
-
-lemma diag_fst: "x \<in> diag A \<Longrightarrow> fst x \<in> A"
-unfolding diag_def by auto
-
-lemma diag_UNIV: "diag UNIV = Id"
-unfolding diag_def by auto
-
-lemma diag_converse: "diag A = (diag A) ^-1"
-unfolding diag_def by auto
-
-lemma diag_Comp: "diag A = diag A O diag A"
-unfolding diag_def by auto
-
-lemma diag_Gr: "diag A = Gr A id"
-unfolding diag_def Gr_def by simp
-
-lemma diag_UNIV_I: "x = y \<Longrightarrow> (x, y) \<in> diag UNIV"
-unfolding diag_def by auto
-
-lemma image2_eqI: "\<lbrakk>b = f x; c = g x; x \<in> A\<rbrakk> \<Longrightarrow> (b, c) \<in> image2 A f g"
-unfolding image2_def by auto
-
-lemma Id_subset: "Id \<subseteq> {(a, b). P a b \<or> a = b}"
-by auto
-
-lemma IdD: "(a, b) \<in> Id \<Longrightarrow> a = b"
-by auto
-
-lemma image2_Gr: "image2 A f g = (Gr A f)^-1 O (Gr A g)"
-unfolding image2_def Gr_def by auto
-
-lemma GrI: "\<lbrakk>x \<in> A; f x = fx\<rbrakk> \<Longrightarrow> (x, fx) \<in> Gr A f"
-unfolding Gr_def by simp
-
-lemma GrE: "(x, fx) \<in> Gr A f \<Longrightarrow> (x \<in> A \<Longrightarrow> f x = fx \<Longrightarrow> P) \<Longrightarrow> P"
-unfolding Gr_def by simp
-
-lemma GrD1: "(x, fx) \<in> Gr A f \<Longrightarrow> x \<in> A"
-unfolding Gr_def by simp
-
-lemma GrD2: "(x, fx) \<in> Gr A f \<Longrightarrow> f x = fx"
-unfolding Gr_def by simp
-
-lemma Gr_incl: "Gr A f \<subseteq> A <*> B \<longleftrightarrow> f ` A \<subseteq> B"
-unfolding Gr_def by auto
-
-definition relImage where
-"relImage R f \<equiv> {(f a1, f a2) | a1 a2. (a1,a2) \<in> R}"
-
-definition relInvImage where
-"relInvImage A R f \<equiv> {(a1, a2) | a1 a2. a1 \<in> A \<and> a2 \<in> A \<and> (f a1, f a2) \<in> R}"
-
-lemma relImage_Gr:
-"\<lbrakk>R \<subseteq> A \<times> A\<rbrakk> \<Longrightarrow> relImage R f = (Gr A f)^-1 O R O Gr A f"
-unfolding relImage_def Gr_def relcomp_def by auto
-
-lemma relInvImage_Gr: "\<lbrakk>R \<subseteq> B \<times> B\<rbrakk> \<Longrightarrow> relInvImage A R f = Gr A f O R O (Gr A f)^-1"
-unfolding Gr_def relcomp_def image_def relInvImage_def by auto
-
-lemma relImage_mono:
-"R1 \<subseteq> R2 \<Longrightarrow> relImage R1 f \<subseteq> relImage R2 f"
-unfolding relImage_def by auto
-
-lemma relInvImage_mono:
-"R1 \<subseteq> R2 \<Longrightarrow> relInvImage A R1 f \<subseteq> relInvImage A R2 f"
-unfolding relInvImage_def by auto
-
-lemma relInvImage_diag:
-"(\<And>a1 a2. f a1 = f a2 \<longleftrightarrow> a1 = a2) \<Longrightarrow> relInvImage A (diag B) f \<subseteq> Id"
-unfolding relInvImage_def diag_def by auto
-
-lemma relInvImage_UNIV_relImage:
-"R \<subseteq> relInvImage UNIV (relImage R f) f"
-unfolding relInvImage_def relImage_def by auto
-
-lemma equiv_Image: "equiv A R \<Longrightarrow> (\<And>a b. (a, b) \<in> R \<Longrightarrow> a \<in> A \<and> b \<in> A \<and> R `` {a} = R `` {b})"
-unfolding equiv_def refl_on_def Image_def by (auto intro: transD symD)
-
-lemma relImage_proj:
-assumes "equiv A R"
-shows "relImage R (proj R) \<subseteq> diag (A//R)"
-unfolding relImage_def diag_def apply safe
-using proj_iff[OF assms]
-by (metis assms equiv_Image proj_def proj_preserves)
-
-lemma relImage_relInvImage:
-assumes "R \<subseteq> f ` A <*> f ` A"
-shows "relImage (relInvImage A R f) f = R"
-using assms unfolding relImage_def relInvImage_def by fastforce
-
-lemma subst_Pair: "P x y \<Longrightarrow> a = (x, y) \<Longrightarrow> P (fst a) (snd a)"
-by simp
-
-lemma fst_diag_id: "(fst \<circ> (%x. (x, x))) z = id z"
-by simp
-
-lemma snd_diag_id: "(snd \<circ> (%x. (x, x))) z = id z"
-by simp
-
-lemma Collect_restrict': "{(x, y) | x y. phi x y \<and> P x y} \<subseteq> {(x, y) | x y. phi x y}"
-by auto
-
-lemma image_convolD: "\<lbrakk>(a, b) \<in> <f, g> ` X\<rbrakk> \<Longrightarrow> \<exists>x. x \<in> X \<and> a = f x \<and> b = g x"
-unfolding convol_def by auto
-
-(*Extended Sublist*)
-
-definition prefCl where
- "prefCl Kl = (\<forall> kl1 kl2. kl1 \<le> kl2 \<and> kl2 \<in> Kl \<longrightarrow> kl1 \<in> Kl)"
-definition PrefCl where
- "PrefCl A n = (\<forall>kl kl'. kl \<in> A n \<and> kl' \<le> kl \<longrightarrow> (\<exists>m\<le>n. kl' \<in> A m))"
-
-lemma prefCl_UN:
- "\<lbrakk>\<And>n. PrefCl A n\<rbrakk> \<Longrightarrow> prefCl (\<Union>n. A n)"
-unfolding prefCl_def PrefCl_def by fastforce
-
-definition Succ where "Succ Kl kl = {k . kl @ [k] \<in> Kl}"
-definition Shift where "Shift Kl k = {kl. k # kl \<in> Kl}"
-definition shift where "shift lab k = (\<lambda>kl. lab (k # kl))"
-
-lemma empty_Shift: "\<lbrakk>[] \<in> Kl; k \<in> Succ Kl []\<rbrakk> \<Longrightarrow> [] \<in> Shift Kl k"
-unfolding Shift_def Succ_def by simp
-
-lemma Shift_clists: "Kl \<subseteq> Field (clists r) \<Longrightarrow> Shift Kl k \<subseteq> Field (clists r)"
-unfolding Shift_def clists_def Field_card_of by auto
-
-lemma Shift_prefCl: "prefCl Kl \<Longrightarrow> prefCl (Shift Kl k)"
-unfolding prefCl_def Shift_def
-proof safe
- fix kl1 kl2
- assume "\<forall>kl1 kl2. kl1 \<le> kl2 \<and> kl2 \<in> Kl \<longrightarrow> kl1 \<in> Kl"
- "kl1 \<le> kl2" "k # kl2 \<in> Kl"
- thus "k # kl1 \<in> Kl" using Cons_prefix_Cons[of k kl1 k kl2] by blast
-qed
-
-lemma not_in_Shift: "kl \<notin> Shift Kl x \<Longrightarrow> x # kl \<notin> Kl"
-unfolding Shift_def by simp
-
-lemma prefCl_Succ: "\<lbrakk>prefCl Kl; k # kl \<in> Kl\<rbrakk> \<Longrightarrow> k \<in> Succ Kl []"
-unfolding Succ_def proof
- assume "prefCl Kl" "k # kl \<in> Kl"
- moreover have "k # [] \<le> k # kl" by auto
- ultimately have "k # [] \<in> Kl" unfolding prefCl_def by blast
- thus "[] @ [k] \<in> Kl" by simp
-qed
-
-lemma SuccD: "k \<in> Succ Kl kl \<Longrightarrow> kl @ [k] \<in> Kl"
-unfolding Succ_def by simp
-
-lemmas SuccE = SuccD[elim_format]
-
-lemma SuccI: "kl @ [k] \<in> Kl \<Longrightarrow> k \<in> Succ Kl kl"
-unfolding Succ_def by simp
-
-lemma ShiftD: "kl \<in> Shift Kl k \<Longrightarrow> k # kl \<in> Kl"
-unfolding Shift_def by simp
-
-lemma Succ_Shift: "Succ (Shift Kl k) kl = Succ Kl (k # kl)"
-unfolding Succ_def Shift_def by auto
-
-lemma ShiftI: "k # kl \<in> Kl \<Longrightarrow> kl \<in> Shift Kl k"
-unfolding Shift_def by simp
-
-lemma Func_cexp: "|Func A B| =o |B| ^c |A|"
-unfolding cexp_def Field_card_of by (simp only: card_of_refl)
-
-lemma clists_bound: "A \<in> Field (cpow (clists r)) - {{}} \<Longrightarrow> |A| \<le>o clists r"
-unfolding cpow_def clists_def Field_card_of by (auto simp: card_of_mono1)
-
-lemma cpow_clists_czero: "\<lbrakk>A \<in> Field (cpow (clists r)) - {{}}; |A| =o czero\<rbrakk> \<Longrightarrow> False"
-unfolding cpow_def clists_def
-by (auto simp add: card_of_ordIso_czero_iff_empty[symmetric])
- (erule notE, erule ordIso_transitive, rule czero_ordIso)
-
-lemma incl_UNION_I:
-assumes "i \<in> I" and "A \<subseteq> F i"
-shows "A \<subseteq> UNION I F"
-using assms by auto
-
-lemma Nil_clists: "{[]} \<subseteq> Field (clists r)"
-unfolding clists_def Field_card_of by auto
-
-lemma Cons_clists:
- "\<lbrakk>x \<in> Field r; xs \<in> Field (clists r)\<rbrakk> \<Longrightarrow> x # xs \<in> Field (clists r)"
-unfolding clists_def Field_card_of by auto
-
-lemma length_Cons: "length (x # xs) = Suc (length xs)"
-by simp
-
-lemma length_append_singleton: "length (xs @ [x]) = Suc (length xs)"
-by simp
-
-(*injection into the field of a cardinal*)
-definition "toCard_pred A r f \<equiv> inj_on f A \<and> f ` A \<subseteq> Field r \<and> Card_order r"
-definition "toCard A r \<equiv> SOME f. toCard_pred A r f"
-
-lemma ex_toCard_pred:
-"\<lbrakk>|A| \<le>o r; Card_order r\<rbrakk> \<Longrightarrow> \<exists> f. toCard_pred A r f"
-unfolding toCard_pred_def
-using card_of_ordLeq[of A "Field r"]
- ordLeq_ordIso_trans[OF _ card_of_unique[of "Field r" r], of "|A|"]
-by blast
-
-lemma toCard_pred_toCard:
- "\<lbrakk>|A| \<le>o r; Card_order r\<rbrakk> \<Longrightarrow> toCard_pred A r (toCard A r)"
-unfolding toCard_def using someI_ex[OF ex_toCard_pred] .
-
-lemma toCard_inj: "\<lbrakk>|A| \<le>o r; Card_order r; x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow>
- toCard A r x = toCard A r y \<longleftrightarrow> x = y"
-using toCard_pred_toCard unfolding inj_on_def toCard_pred_def by blast
-
-lemma toCard: "\<lbrakk>|A| \<le>o r; Card_order r; b \<in> A\<rbrakk> \<Longrightarrow> toCard A r b \<in> Field r"
-using toCard_pred_toCard unfolding toCard_pred_def by blast
-
-definition "fromCard A r k \<equiv> SOME b. b \<in> A \<and> toCard A r b = k"
-
-lemma fromCard_toCard:
-"\<lbrakk>|A| \<le>o r; Card_order r; b \<in> A\<rbrakk> \<Longrightarrow> fromCard A r (toCard A r b) = b"
-unfolding fromCard_def by (rule some_equality) (auto simp add: toCard_inj)
-
-(* pick according to the weak pullback *)
-definition pickWP_pred where
-"pickWP_pred A p1 p2 b1 b2 a \<equiv> a \<in> A \<and> p1 a = b1 \<and> p2 a = b2"
-
-definition pickWP where
-"pickWP A p1 p2 b1 b2 \<equiv> SOME a. pickWP_pred A p1 p2 b1 b2 a"
-
-lemma pickWP_pred:
-assumes "wpull A B1 B2 f1 f2 p1 p2" and
-"b1 \<in> B1" and "b2 \<in> B2" and "f1 b1 = f2 b2"
-shows "\<exists> a. pickWP_pred A p1 p2 b1 b2 a"
-using assms unfolding wpull_def pickWP_pred_def by blast
-
-lemma pickWP_pred_pickWP:
-assumes "wpull A B1 B2 f1 f2 p1 p2" and
-"b1 \<in> B1" and "b2 \<in> B2" and "f1 b1 = f2 b2"
-shows "pickWP_pred A p1 p2 b1 b2 (pickWP A p1 p2 b1 b2)"
-unfolding pickWP_def using assms by(rule someI_ex[OF pickWP_pred])
-
-lemma pickWP:
-assumes "wpull A B1 B2 f1 f2 p1 p2" and
-"b1 \<in> B1" and "b2 \<in> B2" and "f1 b1 = f2 b2"
-shows "pickWP A p1 p2 b1 b2 \<in> A"
- "p1 (pickWP A p1 p2 b1 b2) = b1"
- "p2 (pickWP A p1 p2 b1 b2) = b2"
-using assms pickWP_pred_pickWP unfolding pickWP_pred_def by fastforce+
-
-lemma Inl_Field_csum: "a \<in> Field r \<Longrightarrow> Inl a \<in> Field (r +c s)"
-unfolding Field_card_of csum_def by auto
-
-lemma Inr_Field_csum: "a \<in> Field s \<Longrightarrow> Inr a \<in> Field (r +c s)"
-unfolding Field_card_of csum_def by auto
-
-lemma nat_rec_0: "f = nat_rec f1 (%n rec. f2 n rec) \<Longrightarrow> f 0 = f1"
-by auto
-
-lemma nat_rec_Suc: "f = nat_rec f1 (%n rec. f2 n rec) \<Longrightarrow> f (Suc n) = f2 n (f n)"
-by auto
-
-lemma list_rec_Nil: "f = list_rec f1 (%x xs rec. f2 x xs rec) \<Longrightarrow> f [] = f1"
-by auto
-
-lemma list_rec_Cons: "f = list_rec f1 (%x xs rec. f2 x xs rec) \<Longrightarrow> f (x # xs) = f2 x xs (f xs)"
-by auto
-
-lemma not_arg_cong_Inr: "x \<noteq> y \<Longrightarrow> Inr x \<noteq> Inr y"
-by simp
-
-ML_file "Tools/bnf_gfp_util.ML"
-ML_file "Tools/bnf_gfp_tactics.ML"
-ML_file "Tools/bnf_gfp.ML"
-
-end
--- a/src/HOL/Codatatype/BNF_LFP.thy Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,228 +0,0 @@
-(* Title: HOL/BNF/BNF_LFP.thy
- Author: Dmitriy Traytel, TU Muenchen
- Copyright 2012
-
-Least fixed point operation on bounded natural functors.
-*)
-
-header {* Least Fixed Point Operation on Bounded Natural Functors *}
-
-theory BNF_LFP
-imports BNF_FP
-keywords
- "data_raw" :: thy_decl and
- "data" :: thy_decl
-begin
-
-lemma subset_emptyI: "(\<And>x. x \<in> A \<Longrightarrow> False) \<Longrightarrow> A \<subseteq> {}"
-by blast
-
-lemma image_Collect_subsetI:
- "(\<And>x. P x \<Longrightarrow> f x \<in> B) \<Longrightarrow> f ` {x. P x} \<subseteq> B"
-by blast
-
-lemma Collect_restrict: "{x. x \<in> X \<and> P x} \<subseteq> X"
-by auto
-
-lemma prop_restrict: "\<lbrakk>x \<in> Z; Z \<subseteq> {x. x \<in> X \<and> P x}\<rbrakk> \<Longrightarrow> P x"
-by auto
-
-lemma underS_I: "\<lbrakk>i \<noteq> j; (i, j) \<in> R\<rbrakk> \<Longrightarrow> i \<in> rel.underS R j"
-unfolding rel.underS_def by simp
-
-lemma underS_E: "i \<in> rel.underS R j \<Longrightarrow> i \<noteq> j \<and> (i, j) \<in> R"
-unfolding rel.underS_def by simp
-
-lemma underS_Field: "i \<in> rel.underS R j \<Longrightarrow> i \<in> Field R"
-unfolding rel.underS_def Field_def by auto
-
-lemma FieldI2: "(i, j) \<in> R \<Longrightarrow> j \<in> Field R"
-unfolding Field_def by auto
-
-lemma fst_convol': "fst (<f, g> x) = f x"
-using fst_convol unfolding convol_def by simp
-
-lemma snd_convol': "snd (<f, g> x) = g x"
-using snd_convol unfolding convol_def by simp
-
-lemma convol_expand_snd: "fst o f = g \<Longrightarrow> <g, snd o f> = f"
-unfolding convol_def by auto
-
-definition inver where
- "inver g f A = (ALL a : A. g (f a) = a)"
-
-lemma bij_betw_iff_ex:
- "bij_betw f A B = (EX g. g ` B = A \<and> inver g f A \<and> inver f g B)" (is "?L = ?R")
-proof (rule iffI)
- assume ?L
- hence f: "f ` A = B" and inj_f: "inj_on f A" unfolding bij_betw_def by auto
- let ?phi = "% b a. a : A \<and> f a = b"
- have "ALL b : B. EX a. ?phi b a" using f by blast
- then obtain g where g: "ALL b : B. g b : A \<and> f (g b) = b"
- using bchoice[of B ?phi] by blast
- hence gg: "ALL b : f ` A. g b : A \<and> f (g b) = b" using f by blast
- have gf: "inver g f A" unfolding inver_def
- by (metis (no_types) gg imageI[of _ A f] the_inv_into_f_f[OF inj_f])
- moreover have "g ` B \<le> A \<and> inver f g B" using g unfolding inver_def by blast
- moreover have "A \<le> g ` B"
- proof safe
- fix a assume a: "a : A"
- hence "f a : B" using f by auto
- moreover have "a = g (f a)" using a gf unfolding inver_def by auto
- ultimately show "a : g ` B" by blast
- qed
- ultimately show ?R by blast
-next
- assume ?R
- then obtain g where g: "g ` B = A \<and> inver g f A \<and> inver f g B" by blast
- show ?L unfolding bij_betw_def
- proof safe
- show "inj_on f A" unfolding inj_on_def
- proof safe
- fix a1 a2 assume a: "a1 : A" "a2 : A" and "f a1 = f a2"
- hence "g (f a1) = g (f a2)" by simp
- thus "a1 = a2" using a g unfolding inver_def by simp
- qed
- next
- fix a assume "a : A"
- then obtain b where b: "b : B" and a: "a = g b" using g by blast
- hence "b = f (g b)" using g unfolding inver_def by auto
- thus "f a : B" unfolding a using b by simp
- next
- fix b assume "b : B"
- hence "g b : A \<and> b = f (g b)" using g unfolding inver_def by auto
- thus "b : f ` A" by auto
- qed
-qed
-
-lemma bij_betw_ex_weakE:
- "\<lbrakk>bij_betw f A B\<rbrakk> \<Longrightarrow> \<exists>g. g ` B \<subseteq> A \<and> inver g f A \<and> inver f g B"
-by (auto simp only: bij_betw_iff_ex)
-
-lemma inver_surj: "\<lbrakk>g ` B \<subseteq> A; f ` A \<subseteq> B; inver g f A\<rbrakk> \<Longrightarrow> g ` B = A"
-unfolding inver_def by auto (rule rev_image_eqI, auto)
-
-lemma inver_mono: "\<lbrakk>A \<subseteq> B; inver f g B\<rbrakk> \<Longrightarrow> inver f g A"
-unfolding inver_def by auto
-
-lemma inver_pointfree: "inver f g A = (\<forall>a \<in> A. (f o g) a = a)"
-unfolding inver_def by simp
-
-lemma bij_betwE: "bij_betw f A B \<Longrightarrow> \<forall>a\<in>A. f a \<in> B"
-unfolding bij_betw_def by auto
-
-lemma bij_betw_imageE: "bij_betw f A B \<Longrightarrow> f ` A = B"
-unfolding bij_betw_def by auto
-
-lemma inverE: "\<lbrakk>inver f f' A; x \<in> A\<rbrakk> \<Longrightarrow> f (f' x) = x"
-unfolding inver_def by auto
-
-lemma bij_betw_inver1: "bij_betw f A B \<Longrightarrow> inver (inv_into A f) f A"
-unfolding bij_betw_def inver_def by auto
-
-lemma bij_betw_inver2: "bij_betw f A B \<Longrightarrow> inver f (inv_into A f) B"
-unfolding bij_betw_def inver_def by auto
-
-lemma bij_betwI: "\<lbrakk>bij_betw g B A; inver g f A; inver f g B\<rbrakk> \<Longrightarrow> bij_betw f A B"
-by (drule bij_betw_imageE, unfold bij_betw_iff_ex) blast
-
-lemma bij_betwI':
- "\<lbrakk>\<And>x y. \<lbrakk>x \<in> X; y \<in> X\<rbrakk> \<Longrightarrow> (f x = f y) = (x = y);
- \<And>x. x \<in> X \<Longrightarrow> f x \<in> Y;
- \<And>y. y \<in> Y \<Longrightarrow> \<exists>x \<in> X. y = f x\<rbrakk> \<Longrightarrow> bij_betw f X Y"
-unfolding bij_betw_def inj_on_def
-apply (rule conjI)
- apply blast
-by (erule thin_rl) blast
-
-lemma surj_fun_eq:
- assumes surj_on: "f ` X = UNIV" and eq_on: "\<forall>x \<in> X. (g1 o f) x = (g2 o f) x"
- shows "g1 = g2"
-proof (rule ext)
- fix y
- from surj_on obtain x where "x \<in> X" and "y = f x" by blast
- thus "g1 y = g2 y" using eq_on by simp
-qed
-
-lemma Card_order_wo_rel: "Card_order r \<Longrightarrow> wo_rel r"
-unfolding wo_rel_def card_order_on_def by blast
-
-lemma Cinfinite_limit: "\<lbrakk>x \<in> Field r; Cinfinite r\<rbrakk> \<Longrightarrow>
- \<exists>y \<in> Field r. x \<noteq> y \<and> (x, y) \<in> r"
-unfolding cinfinite_def by (auto simp add: infinite_Card_order_limit)
-
-lemma Card_order_trans:
- "\<lbrakk>Card_order r; x \<noteq> y; (x, y) \<in> r; y \<noteq> z; (y, z) \<in> r\<rbrakk> \<Longrightarrow> x \<noteq> z \<and> (x, z) \<in> r"
-unfolding card_order_on_def well_order_on_def linear_order_on_def
- partial_order_on_def preorder_on_def trans_def antisym_def by blast
-
-lemma Cinfinite_limit2:
- assumes x1: "x1 \<in> Field r" and x2: "x2 \<in> Field r" and r: "Cinfinite r"
- shows "\<exists>y \<in> Field r. (x1 \<noteq> y \<and> (x1, y) \<in> r) \<and> (x2 \<noteq> y \<and> (x2, y) \<in> r)"
-proof -
- from r have trans: "trans r" and total: "Total r" and antisym: "antisym r"
- unfolding card_order_on_def well_order_on_def linear_order_on_def
- partial_order_on_def preorder_on_def by auto
- obtain y1 where y1: "y1 \<in> Field r" "x1 \<noteq> y1" "(x1, y1) \<in> r"
- using Cinfinite_limit[OF x1 r] by blast
- obtain y2 where y2: "y2 \<in> Field r" "x2 \<noteq> y2" "(x2, y2) \<in> r"
- using Cinfinite_limit[OF x2 r] by blast
- show ?thesis
- proof (cases "y1 = y2")
- case True with y1 y2 show ?thesis by blast
- next
- case False
- with y1(1) y2(1) total have "(y1, y2) \<in> r \<or> (y2, y1) \<in> r"
- unfolding total_on_def by auto
- thus ?thesis
- proof
- assume *: "(y1, y2) \<in> r"
- with trans y1(3) have "(x1, y2) \<in> r" unfolding trans_def by blast
- with False y1 y2 * antisym show ?thesis by (cases "x1 = y2") (auto simp: antisym_def)
- next
- assume *: "(y2, y1) \<in> r"
- with trans y2(3) have "(x2, y1) \<in> r" unfolding trans_def by blast
- with False y1 y2 * antisym show ?thesis by (cases "x2 = y1") (auto simp: antisym_def)
- qed
- qed
-qed
-
-lemma Cinfinite_limit_finite: "\<lbrakk>finite X; X \<subseteq> Field r; Cinfinite r\<rbrakk>
- \<Longrightarrow> \<exists>y \<in> Field r. \<forall>x \<in> X. (x \<noteq> y \<and> (x, y) \<in> r)"
-proof (induct X rule: finite_induct)
- case empty thus ?case unfolding cinfinite_def using ex_in_conv[of "Field r"] finite.emptyI by auto
-next
- case (insert x X)
- then obtain y where y: "y \<in> Field r" "\<forall>x \<in> X. (x \<noteq> y \<and> (x, y) \<in> r)" by blast
- then obtain z where z: "z \<in> Field r" "x \<noteq> z \<and> (x, z) \<in> r" "y \<noteq> z \<and> (y, z) \<in> r"
- using Cinfinite_limit2[OF _ y(1) insert(5), of x] insert(4) by blast
- show ?case
- apply (intro bexI ballI)
- apply (erule insertE)
- apply hypsubst
- apply (rule z(2))
- using Card_order_trans[OF insert(5)[THEN conjunct2]] y(2) z(3)
- apply blast
- apply (rule z(1))
- done
-qed
-
-lemma insert_subsetI: "\<lbrakk>x \<in> A; X \<subseteq> A\<rbrakk> \<Longrightarrow> insert x X \<subseteq> A"
-by auto
-
-(*helps resolution*)
-lemma well_order_induct_imp:
- "wo_rel r \<Longrightarrow> (\<And>x. \<forall>y. y \<noteq> x \<and> (y, x) \<in> r \<longrightarrow> y \<in> Field r \<longrightarrow> P y \<Longrightarrow> x \<in> Field r \<longrightarrow> P x) \<Longrightarrow>
- x \<in> Field r \<longrightarrow> P x"
-by (erule wo_rel.well_order_induct)
-
-lemma meta_spec2:
- assumes "(\<And>x y. PROP P x y)"
- shows "PROP P x y"
-by (rule `(\<And>x y. PROP P x y)`)
-
-ML_file "Tools/bnf_lfp_util.ML"
-ML_file "Tools/bnf_lfp_tactics.ML"
-ML_file "Tools/bnf_lfp.ML"
-
-end
--- a/src/HOL/Codatatype/BNF_Util.thy Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,66 +0,0 @@
-(* Title: HOL/BNF/BNF_Util.thy
- Author: Dmitriy Traytel, TU Muenchen
- Author: Jasmin Blanchette, TU Muenchen
- Copyright 2012
-
-Library for bounded natural functors.
-*)
-
-header {* Library for Bounded Natural Functors *}
-
-theory BNF_Util
-imports "../Cardinals/Cardinal_Arithmetic"
-begin
-
-lemma subset_Collect_iff: "B \<subseteq> A \<Longrightarrow> (B \<subseteq> {x \<in> A. P x}) = (\<forall>x \<in> B. P x)"
-by blast
-
-lemma subset_CollectI: "B \<subseteq> A \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> Q x \<Longrightarrow> P x) \<Longrightarrow> ({x \<in> B. Q x} \<subseteq> {x \<in> A. P x})"
-by blast
-
-definition collect where
-"collect F x = (\<Union>f \<in> F. f x)"
-
-(* Weak pullbacks: *)
-definition wpull where
-"wpull A B1 B2 f1 f2 p1 p2 \<longleftrightarrow>
- (\<forall> b1 b2. b1 \<in> B1 \<and> b2 \<in> B2 \<and> f1 b1 = f2 b2 \<longrightarrow> (\<exists> a \<in> A. p1 a = b1 \<and> p2 a = b2))"
-
-(* Weak pseudo-pullbacks *)
-definition wppull where
-"wppull A B1 B2 f1 f2 e1 e2 p1 p2 \<longleftrightarrow>
- (\<forall> b1 b2. b1 \<in> B1 \<and> b2 \<in> B2 \<and> f1 b1 = f2 b2 \<longrightarrow>
- (\<exists> a \<in> A. e1 (p1 a) = e1 b1 \<and> e2 (p2 a) = e2 b2))"
-
-lemma fst_snd: "\<lbrakk>snd x = (y, z)\<rbrakk> \<Longrightarrow> fst (snd x) = y"
-by simp
-
-lemma snd_snd: "\<lbrakk>snd x = (y, z)\<rbrakk> \<Longrightarrow> snd (snd x) = z"
-by simp
-
-lemma fstI: "x = (y, z) \<Longrightarrow> fst x = y"
-by simp
-
-lemma sndI: "x = (y, z) \<Longrightarrow> snd x = z"
-by simp
-
-lemma bijI: "\<lbrakk>\<And>x y. (f x = f y) = (x = y); \<And>y. \<exists>x. y = f x\<rbrakk> \<Longrightarrow> bij f"
-unfolding bij_def inj_on_def by auto blast
-
-lemma pair_mem_Collect_split:
-"(\<lambda>x y. (x, y) \<in> {(x, y). P x y}) = P"
-by simp
-
-lemma Collect_pair_mem_eq: "{(x, y). (x, y) \<in> R} = R"
-by simp
-
-lemma Collect_fst_snd_mem_eq: "{p. (fst p, snd p) \<in> A} = A"
-by simp
-
-(* Operator: *)
-definition "Gr A f = {(a, f a) | a. a \<in> A}"
-
-ML_file "Tools/bnf_util.ML"
-ML_file "Tools/bnf_tactics.ML"
-
-end
--- a/src/HOL/Codatatype/BNF_Wrap.thy Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,28 +0,0 @@
-(* Title: HOL/BNF/BNF_Wrap.thy
- Author: Jasmin Blanchette, TU Muenchen
- Copyright 2012
-
-Wrapping datatypes.
-*)
-
-header {* Wrapping Datatypes *}
-
-theory BNF_Wrap
-imports BNF_Util
-keywords
- "wrap_data" :: thy_goal and
- "no_dests"
-begin
-
-lemma iffI_np: "\<lbrakk>x \<Longrightarrow> \<not> y; \<not> x \<Longrightarrow> y\<rbrakk> \<Longrightarrow> \<not> x \<longleftrightarrow> y"
-by (erule iffI) (erule contrapos_pn)
-
-lemma iff_contradict:
-"\<not> P \<Longrightarrow> P \<longleftrightarrow> Q \<Longrightarrow> Q \<Longrightarrow> R"
-"\<not> Q \<Longrightarrow> P \<longleftrightarrow> Q \<Longrightarrow> P \<Longrightarrow> R"
-by blast+
-
-ML_file "Tools/bnf_wrap_tactics.ML"
-ML_file "Tools/bnf_wrap.ML"
-
-end
--- a/src/HOL/Codatatype/Basic_BNFs.thy Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,417 +0,0 @@
-(* Title: HOL/BNF/Basic_BNFs.thy
- Author: Dmitriy Traytel, TU Muenchen
- Author: Andrei Popescu, TU Muenchen
- Author: Jasmin Blanchette, TU Muenchen
- Copyright 2012
-
-Registration of basic types as bounded natural functors.
-*)
-
-header {* Registration of Basic Types as Bounded Natural Functors *}
-
-theory Basic_BNFs
-imports BNF_Def
-begin
-
-lemma wpull_id: "wpull UNIV B1 B2 id id id id"
-unfolding wpull_def by simp
-
-lemmas natLeq_card_order = natLeq_Card_order[unfolded Field_natLeq]
-
-lemma ctwo_card_order: "card_order ctwo"
-using Card_order_ctwo by (unfold ctwo_def Field_card_of)
-
-lemma natLeq_cinfinite: "cinfinite natLeq"
-unfolding cinfinite_def Field_natLeq by (rule nat_infinite)
-
-bnf_def ID: "id :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" ["\<lambda>x. {x}"] "\<lambda>_:: 'a. natLeq" ["id :: 'a \<Rightarrow> 'a"]
- "\<lambda>x :: 'a \<Rightarrow> 'b \<Rightarrow> bool. x"
-apply auto
-apply (rule natLeq_card_order)
-apply (rule natLeq_cinfinite)
-apply (rule ordLess_imp_ordLeq[OF finite_ordLess_infinite[OF _ natLeq_Well_order]])
-apply (auto simp add: Field_card_of Field_natLeq card_of_well_order_on)[3]
-apply (rule ordLeq_transitive)
-apply (rule ordLeq_cexp1[of natLeq])
-apply (rule Cinfinite_Cnotzero)
-apply (rule conjI)
-apply (rule natLeq_cinfinite)
-apply (rule natLeq_Card_order)
-apply (rule card_of_Card_order)
-apply (rule cexp_mono1)
-apply (rule ordLeq_csum1)
-apply (rule card_of_Card_order)
-apply (rule disjI2)
-apply (rule cone_ordLeq_cexp)
-apply (rule ordLeq_transitive)
-apply (rule cone_ordLeq_ctwo)
-apply (rule ordLeq_csum2)
-apply (rule Card_order_ctwo)
-apply (rule natLeq_Card_order)
-apply (auto simp: Gr_def fun_eq_iff)
-done
-
-bnf_def DEADID: "id :: 'a \<Rightarrow> 'a" [] "\<lambda>_:: 'a. natLeq +c |UNIV :: 'a set|" ["SOME x :: 'a. True"]
- "op =::'a \<Rightarrow> 'a \<Rightarrow> bool"
-apply (auto simp add: wpull_id)
-apply (rule card_order_csum)
-apply (rule natLeq_card_order)
-apply (rule card_of_card_order_on)
-apply (rule cinfinite_csum)
-apply (rule disjI1)
-apply (rule natLeq_cinfinite)
-apply (rule ordLess_imp_ordLeq)
-apply (rule ordLess_ordLeq_trans)
-apply (rule ordLess_ctwo_cexp)
-apply (rule card_of_Card_order)
-apply (rule cexp_mono2'')
-apply (rule ordLeq_csum2)
-apply (rule card_of_Card_order)
-apply (rule ctwo_Cnotzero)
-apply (rule card_of_Card_order)
-apply (auto simp: Id_def Gr_def fun_eq_iff)
-done
-
-definition setl :: "'a + 'b \<Rightarrow> 'a set" where
-"setl x = (case x of Inl z => {z} | _ => {})"
-
-definition setr :: "'a + 'b \<Rightarrow> 'b set" where
-"setr x = (case x of Inr z => {z} | _ => {})"
-
-lemmas sum_set_defs = setl_def[abs_def] setr_def[abs_def]
-
-definition sum_rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('c \<Rightarrow> 'd \<Rightarrow> bool) \<Rightarrow> 'a + 'c \<Rightarrow> 'b + 'd \<Rightarrow> bool" where
-"sum_rel \<phi> \<psi> x y =
- (case x of Inl a1 \<Rightarrow> (case y of Inl a2 \<Rightarrow> \<phi> a1 a2 | Inr _ \<Rightarrow> False)
- | Inr b1 \<Rightarrow> (case y of Inl _ \<Rightarrow> False | Inr b2 \<Rightarrow> \<psi> b1 b2))"
-
-bnf_def sum_map [setl, setr] "\<lambda>_::'a + 'b. natLeq" [Inl, Inr] sum_rel
-proof -
- show "sum_map id id = id" by (rule sum_map.id)
-next
- fix f1 f2 g1 g2
- show "sum_map (g1 o f1) (g2 o f2) = sum_map g1 g2 o sum_map f1 f2"
- by (rule sum_map.comp[symmetric])
-next
- fix x f1 f2 g1 g2
- assume a1: "\<And>z. z \<in> setl x \<Longrightarrow> f1 z = g1 z" and
- a2: "\<And>z. z \<in> setr x \<Longrightarrow> f2 z = g2 z"
- thus "sum_map f1 f2 x = sum_map g1 g2 x"
- proof (cases x)
- case Inl thus ?thesis using a1 by (clarsimp simp: setl_def)
- next
- case Inr thus ?thesis using a2 by (clarsimp simp: setr_def)
- qed
-next
- fix f1 f2
- show "setl o sum_map f1 f2 = image f1 o setl"
- by (rule ext, unfold o_apply) (simp add: setl_def split: sum.split)
-next
- fix f1 f2
- show "setr o sum_map f1 f2 = image f2 o setr"
- by (rule ext, unfold o_apply) (simp add: setr_def split: sum.split)
-next
- show "card_order natLeq" by (rule natLeq_card_order)
-next
- show "cinfinite natLeq" by (rule natLeq_cinfinite)
-next
- fix x
- show "|setl x| \<le>o natLeq"
- apply (rule ordLess_imp_ordLeq)
- apply (rule finite_iff_ordLess_natLeq[THEN iffD1])
- by (simp add: setl_def split: sum.split)
-next
- fix x
- show "|setr x| \<le>o natLeq"
- apply (rule ordLess_imp_ordLeq)
- apply (rule finite_iff_ordLess_natLeq[THEN iffD1])
- by (simp add: setr_def split: sum.split)
-next
- fix A1 :: "'a set" and A2 :: "'b set"
- have in_alt: "{x. (case x of Inl z => {z} | _ => {}) \<subseteq> A1 \<and>
- (case x of Inr z => {z} | _ => {}) \<subseteq> A2} = A1 <+> A2" (is "?L = ?R")
- proof safe
- fix x :: "'a + 'b"
- assume "(case x of Inl z \<Rightarrow> {z} | _ \<Rightarrow> {}) \<subseteq> A1" "(case x of Inr z \<Rightarrow> {z} | _ \<Rightarrow> {}) \<subseteq> A2"
- hence "x \<in> Inl ` A1 \<or> x \<in> Inr ` A2" by (cases x) simp+
- thus "x \<in> A1 <+> A2" by blast
- qed (auto split: sum.split)
- show "|{x. setl x \<subseteq> A1 \<and> setr x \<subseteq> A2}| \<le>o
- (( |A1| +c |A2| ) +c ctwo) ^c natLeq"
- apply (rule ordIso_ordLeq_trans)
- apply (rule card_of_ordIso_subst)
- apply (unfold sum_set_defs)
- apply (rule in_alt)
- apply (rule ordIso_ordLeq_trans)
- apply (rule Plus_csum)
- apply (rule ordLeq_transitive)
- apply (rule ordLeq_csum1)
- apply (rule Card_order_csum)
- apply (rule ordLeq_cexp1)
- apply (rule conjI)
- using Field_natLeq UNIV_not_empty czeroE apply fast
- apply (rule natLeq_Card_order)
- by (rule Card_order_csum)
-next
- fix A1 A2 B11 B12 B21 B22 f11 f12 f21 f22 p11 p12 p21 p22
- assume "wpull A1 B11 B21 f11 f21 p11 p21" "wpull A2 B12 B22 f12 f22 p12 p22"
- hence
- pull1: "\<And>b1 b2. \<lbrakk>b1 \<in> B11; b2 \<in> B21; f11 b1 = f21 b2\<rbrakk> \<Longrightarrow> \<exists>a \<in> A1. p11 a = b1 \<and> p21 a = b2"
- and pull2: "\<And>b1 b2. \<lbrakk>b1 \<in> B12; b2 \<in> B22; f12 b1 = f22 b2\<rbrakk> \<Longrightarrow> \<exists>a \<in> A2. p12 a = b1 \<and> p22 a = b2"
- unfolding wpull_def by blast+
- show "wpull {x. setl x \<subseteq> A1 \<and> setr x \<subseteq> A2}
- {x. setl x \<subseteq> B11 \<and> setr x \<subseteq> B12} {x. setl x \<subseteq> B21 \<and> setr x \<subseteq> B22}
- (sum_map f11 f12) (sum_map f21 f22) (sum_map p11 p12) (sum_map p21 p22)"
- (is "wpull ?in ?in1 ?in2 ?mapf1 ?mapf2 ?mapp1 ?mapp2")
- proof (unfold wpull_def)
- { fix B1 B2
- assume *: "B1 \<in> ?in1" "B2 \<in> ?in2" "?mapf1 B1 = ?mapf2 B2"
- have "\<exists>A \<in> ?in. ?mapp1 A = B1 \<and> ?mapp2 A = B2"
- proof (cases B1)
- case (Inl b1)
- { fix b2 assume "B2 = Inr b2"
- with Inl *(3) have False by simp
- } then obtain b2 where Inl': "B2 = Inl b2" by (cases B2) (simp, blast)
- with Inl * have "b1 \<in> B11" "b2 \<in> B21" "f11 b1 = f21 b2"
- by (simp add: setl_def)+
- with pull1 obtain a where "a \<in> A1" "p11 a = b1" "p21 a = b2" by blast+
- with Inl Inl' have "Inl a \<in> ?in" "?mapp1 (Inl a) = B1 \<and> ?mapp2 (Inl a) = B2"
- by (simp add: sum_set_defs)+
- thus ?thesis by blast
- next
- case (Inr b1)
- { fix b2 assume "B2 = Inl b2"
- with Inr *(3) have False by simp
- } then obtain b2 where Inr': "B2 = Inr b2" by (cases B2) (simp, blast)
- with Inr * have "b1 \<in> B12" "b2 \<in> B22" "f12 b1 = f22 b2"
- by (simp add: sum_set_defs)+
- with pull2 obtain a where "a \<in> A2" "p12 a = b1" "p22 a = b2" by blast+
- with Inr Inr' have "Inr a \<in> ?in" "?mapp1 (Inr a) = B1 \<and> ?mapp2 (Inr a) = B2"
- by (simp add: sum_set_defs)+
- thus ?thesis by blast
- qed
- }
- thus "\<forall>B1 B2. B1 \<in> ?in1 \<and> B2 \<in> ?in2 \<and> ?mapf1 B1 = ?mapf2 B2 \<longrightarrow>
- (\<exists>A \<in> ?in. ?mapp1 A = B1 \<and> ?mapp2 A = B2)" by fastforce
- qed
-next
- fix R S
- show "{p. sum_rel (\<lambda>x y. (x, y) \<in> R) (\<lambda>x y. (x, y) \<in> S) (fst p) (snd p)} =
- (Gr {x. setl x \<subseteq> R \<and> setr x \<subseteq> S} (sum_map fst fst))\<inverse> O
- Gr {x. setl x \<subseteq> R \<and> setr x \<subseteq> S} (sum_map snd snd)"
- unfolding setl_def setr_def sum_rel_def Gr_def relcomp_unfold converse_unfold
- by (fastforce split: sum.splits)
-qed (auto simp: sum_set_defs)
-
-lemma singleton_ordLeq_ctwo_natLeq: "|{x}| \<le>o ctwo *c natLeq"
- apply (rule ordLeq_transitive)
- apply (rule ordLeq_cprod2)
- apply (rule ctwo_Cnotzero)
- apply (auto simp: Field_card_of intro: card_of_card_order_on)
- apply (rule cprod_mono2)
- apply (rule ordLess_imp_ordLeq)
- apply (unfold finite_iff_ordLess_natLeq[symmetric])
- by simp
-
-definition fsts :: "'a \<times> 'b \<Rightarrow> 'a set" where
-"fsts x = {fst x}"
-
-definition snds :: "'a \<times> 'b \<Rightarrow> 'b set" where
-"snds x = {snd x}"
-
-lemmas prod_set_defs = fsts_def[abs_def] snds_def[abs_def]
-
-definition prod_rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('c \<Rightarrow> 'd \<Rightarrow> bool) \<Rightarrow> 'a \<times> 'c \<Rightarrow> 'b \<times> 'd \<Rightarrow> bool" where
-"prod_rel \<phi> \<psi> p1 p2 = (case p1 of (a1, b1) \<Rightarrow> case p2 of (a2, b2) \<Rightarrow> \<phi> a1 a2 \<and> \<psi> b1 b2)"
-
-bnf_def map_pair [fsts, snds] "\<lambda>_::'a \<times> 'b. ctwo *c natLeq" [Pair] prod_rel
-proof (unfold prod_set_defs)
- show "map_pair id id = id" by (rule map_pair.id)
-next
- fix f1 f2 g1 g2
- show "map_pair (g1 o f1) (g2 o f2) = map_pair g1 g2 o map_pair f1 f2"
- by (rule map_pair.comp[symmetric])
-next
- fix x f1 f2 g1 g2
- assume "\<And>z. z \<in> {fst x} \<Longrightarrow> f1 z = g1 z" "\<And>z. z \<in> {snd x} \<Longrightarrow> f2 z = g2 z"
- thus "map_pair f1 f2 x = map_pair g1 g2 x" by (cases x) simp
-next
- fix f1 f2
- show "(\<lambda>x. {fst x}) o map_pair f1 f2 = image f1 o (\<lambda>x. {fst x})"
- by (rule ext, unfold o_apply) simp
-next
- fix f1 f2
- show "(\<lambda>x. {snd x}) o map_pair f1 f2 = image f2 o (\<lambda>x. {snd x})"
- by (rule ext, unfold o_apply) simp
-next
- show "card_order (ctwo *c natLeq)"
- apply (rule card_order_cprod)
- apply (rule ctwo_card_order)
- by (rule natLeq_card_order)
-next
- show "cinfinite (ctwo *c natLeq)"
- apply (rule cinfinite_cprod2)
- apply (rule ctwo_Cnotzero)
- apply (rule conjI[OF _ natLeq_Card_order])
- by (rule natLeq_cinfinite)
-next
- fix x
- show "|{fst x}| \<le>o ctwo *c natLeq"
- by (rule singleton_ordLeq_ctwo_natLeq)
-next
- fix x
- show "|{snd x}| \<le>o ctwo *c natLeq"
- by (rule singleton_ordLeq_ctwo_natLeq)
-next
- fix A1 :: "'a set" and A2 :: "'b set"
- have in_alt: "{x. {fst x} \<subseteq> A1 \<and> {snd x} \<subseteq> A2} = A1 \<times> A2" by auto
- show "|{x. {fst x} \<subseteq> A1 \<and> {snd x} \<subseteq> A2}| \<le>o
- ( ( |A1| +c |A2| ) +c ctwo) ^c (ctwo *c natLeq)"
- apply (rule ordIso_ordLeq_trans)
- apply (rule card_of_ordIso_subst)
- apply (rule in_alt)
- apply (rule ordIso_ordLeq_trans)
- apply (rule Times_cprod)
- apply (rule ordLeq_transitive)
- apply (rule cprod_csum_cexp)
- apply (rule cexp_mono)
- apply (rule ordLeq_csum1)
- apply (rule Card_order_csum)
- apply (rule ordLeq_cprod1)
- apply (rule Card_order_ctwo)
- apply (rule Cinfinite_Cnotzero)
- apply (rule conjI[OF _ natLeq_Card_order])
- apply (rule natLeq_cinfinite)
- apply (rule disjI2)
- apply (rule cone_ordLeq_cexp)
- apply (rule ordLeq_transitive)
- apply (rule cone_ordLeq_ctwo)
- apply (rule ordLeq_csum2)
- apply (rule Card_order_ctwo)
- apply (rule notE)
- apply (rule ctwo_not_czero)
- apply assumption
- by (rule Card_order_ctwo)
-next
- fix A1 A2 B11 B12 B21 B22 f11 f12 f21 f22 p11 p12 p21 p22
- assume "wpull A1 B11 B21 f11 f21 p11 p21" "wpull A2 B12 B22 f12 f22 p12 p22"
- thus "wpull {x. {fst x} \<subseteq> A1 \<and> {snd x} \<subseteq> A2}
- {x. {fst x} \<subseteq> B11 \<and> {snd x} \<subseteq> B12} {x. {fst x} \<subseteq> B21 \<and> {snd x} \<subseteq> B22}
- (map_pair f11 f12) (map_pair f21 f22) (map_pair p11 p12) (map_pair p21 p22)"
- unfolding wpull_def by simp fast
-next
- fix R S
- show "{p. prod_rel (\<lambda>x y. (x, y) \<in> R) (\<lambda>x y. (x, y) \<in> S) (fst p) (snd p)} =
- (Gr {x. {fst x} \<subseteq> R \<and> {snd x} \<subseteq> S} (map_pair fst fst))\<inverse> O
- Gr {x. {fst x} \<subseteq> R \<and> {snd x} \<subseteq> S} (map_pair snd snd)"
- unfolding prod_set_defs prod_rel_def Gr_def relcomp_unfold converse_unfold
- by auto
-qed simp+
-
-(* Categorical version of pullback: *)
-lemma wpull_cat:
-assumes p: "wpull A B1 B2 f1 f2 p1 p2"
-and c: "f1 o q1 = f2 o q2"
-and r: "range q1 \<subseteq> B1" "range q2 \<subseteq> B2"
-obtains h where "range h \<subseteq> A \<and> q1 = p1 o h \<and> q2 = p2 o h"
-proof-
- have *: "\<forall>d. \<exists>a \<in> A. p1 a = q1 d & p2 a = q2 d"
- proof safe
- fix d
- have "f1 (q1 d) = f2 (q2 d)" using c unfolding comp_def[abs_def] by (rule fun_cong)
- moreover
- have "q1 d : B1" "q2 d : B2" using r unfolding image_def by auto
- ultimately show "\<exists>a \<in> A. p1 a = q1 d \<and> p2 a = q2 d"
- using p unfolding wpull_def by auto
- qed
- then obtain h where "!! d. h d \<in> A & p1 (h d) = q1 d & p2 (h d) = q2 d" by metis
- thus ?thesis using that by fastforce
-qed
-
-lemma card_of_bounded_range:
- "|{f :: 'd \<Rightarrow> 'a. range f \<subseteq> B}| \<le>o |Func (UNIV :: 'd set) B|" (is "|?LHS| \<le>o |?RHS|")
-proof -
- let ?f = "\<lambda>f. %x. if f x \<in> B then Some (f x) else None"
- have "inj_on ?f ?LHS" unfolding inj_on_def
- proof (unfold fun_eq_iff, safe)
- fix g :: "'d \<Rightarrow> 'a" and f :: "'d \<Rightarrow> 'a" and x
- assume "range f \<subseteq> B" "range g \<subseteq> B" and eq: "\<forall>x. ?f f x = ?f g x"
- hence "f x \<in> B" "g x \<in> B" by auto
- with eq have "Some (f x) = Some (g x)" by metis
- thus "f x = g x" by simp
- qed
- moreover have "?f ` ?LHS \<subseteq> ?RHS" unfolding Func_def by fastforce
- ultimately show ?thesis using card_of_ordLeq by fast
-qed
-
-definition fun_rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('c \<Rightarrow> 'a) \<Rightarrow> ('c \<Rightarrow> 'b) \<Rightarrow> bool" where
-"fun_rel \<phi> f g = (\<forall>x. \<phi> (f x) (g x))"
-
-bnf_def "op \<circ>" [range] "\<lambda>_:: 'a \<Rightarrow> 'b. natLeq +c |UNIV :: 'a set|" ["%c x::'b::type. c::'a::type"]
- fun_rel
-proof
- fix f show "id \<circ> f = id f" by simp
-next
- fix f g show "op \<circ> (g \<circ> f) = op \<circ> g \<circ> op \<circ> f"
- unfolding comp_def[abs_def] ..
-next
- fix x f g
- assume "\<And>z. z \<in> range x \<Longrightarrow> f z = g z"
- thus "f \<circ> x = g \<circ> x" by auto
-next
- fix f show "range \<circ> op \<circ> f = op ` f \<circ> range"
- unfolding image_def comp_def[abs_def] by auto
-next
- show "card_order (natLeq +c |UNIV| )" (is "_ (_ +c ?U)")
- apply (rule card_order_csum)
- apply (rule natLeq_card_order)
- by (rule card_of_card_order_on)
-(* *)
- show "cinfinite (natLeq +c ?U)"
- apply (rule cinfinite_csum)
- apply (rule disjI1)
- by (rule natLeq_cinfinite)
-next
- fix f :: "'d => 'a"
- have "|range f| \<le>o | (UNIV::'d set) |" (is "_ \<le>o ?U") by (rule card_of_image)
- also have "?U \<le>o natLeq +c ?U" by (rule ordLeq_csum2) (rule card_of_Card_order)
- finally show "|range f| \<le>o natLeq +c ?U" .
-next
- fix B :: "'a set"
- have "|{f::'d => 'a. range f \<subseteq> B}| \<le>o |Func (UNIV :: 'd set) B|" by (rule card_of_bounded_range)
- also have "|Func (UNIV :: 'd set) B| =o |B| ^c |UNIV :: 'd set|"
- unfolding cexp_def Field_card_of by (rule card_of_refl)
- also have "|B| ^c |UNIV :: 'd set| \<le>o
- ( |B| +c ctwo) ^c (natLeq +c |UNIV :: 'd set| )"
- apply (rule cexp_mono)
- apply (rule ordLeq_csum1) apply (rule card_of_Card_order)
- apply (rule ordLeq_csum2) apply (rule card_of_Card_order)
- apply (rule disjI2) apply (rule cone_ordLeq_cexp)
- apply (rule ordLeq_transitive) apply (rule cone_ordLeq_ctwo) apply (rule ordLeq_csum2)
- apply (rule Card_order_ctwo)
- apply (rule notE) apply (rule conjunct1) apply (rule Cnotzero_UNIV) apply blast
- apply (rule card_of_Card_order)
- done
- finally
- show "|{f::'d => 'a. range f \<subseteq> B}| \<le>o
- ( |B| +c ctwo) ^c (natLeq +c |UNIV :: 'd set| )" .
-next
- fix A B1 B2 f1 f2 p1 p2 assume p: "wpull A B1 B2 f1 f2 p1 p2"
- show "wpull {h. range h \<subseteq> A} {g1. range g1 \<subseteq> B1} {g2. range g2 \<subseteq> B2}
- (op \<circ> f1) (op \<circ> f2) (op \<circ> p1) (op \<circ> p2)"
- unfolding wpull_def
- proof safe
- fix g1 g2 assume r: "range g1 \<subseteq> B1" "range g2 \<subseteq> B2"
- and c: "f1 \<circ> g1 = f2 \<circ> g2"
- show "\<exists>h \<in> {h. range h \<subseteq> A}. p1 \<circ> h = g1 \<and> p2 \<circ> h = g2"
- using wpull_cat[OF p c r] by simp metis
- qed
-next
- fix R
- show "{p. fun_rel (\<lambda>x y. (x, y) \<in> R) (fst p) (snd p)} =
- (Gr {x. range x \<subseteq> R} (op \<circ> fst))\<inverse> O Gr {x. range x \<subseteq> R} (op \<circ> snd)"
- unfolding fun_rel_def Gr_def relcomp_unfold converse_unfold
- by (auto intro!: exI dest!: in_mono)
-qed auto
-
-end
--- a/src/HOL/Codatatype/Countable_Set.thy Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,366 +0,0 @@
-(* Title: HOL/BNF/Countable_Set.thy
- Author: Andrei Popescu, TU Muenchen
- Copyright 2012
-
-(At most) countable sets.
-*)
-
-header {* (At Most) Countable Sets *}
-
-theory Countable_Set
-imports "../Cardinals/Cardinals" "~~/src/HOL/Library/Countable"
-begin
-
-
-subsection{* Basics *}
-
-definition "countable A \<equiv> |A| \<le>o natLeq"
-
-lemma countable_card_of_nat:
-"countable A \<longleftrightarrow> |A| \<le>o |UNIV::nat set|"
-unfolding countable_def using card_of_nat
-using ordLeq_ordIso_trans ordIso_symmetric by blast
-
-lemma countable_ex_to_nat:
-fixes A :: "'a set"
-shows "countable A \<longleftrightarrow> (\<exists> f::'a\<Rightarrow>nat. inj_on f A)"
-unfolding countable_card_of_nat card_of_ordLeq[symmetric] by auto
-
-lemma countable_or_card_of:
-assumes "countable A"
-shows "(finite A \<and> |A| <o |UNIV::nat set| ) \<or>
- (infinite A \<and> |A| =o |UNIV::nat set| )"
-apply (cases "finite A")
- apply(metis finite_iff_cardOf_nat)
- by (metis assms countable_card_of_nat infinite_iff_card_of_nat ordIso_iff_ordLeq)
-
-lemma countable_or:
-assumes "countable A"
-shows "(\<exists> f::'a\<Rightarrow>nat. finite A \<and> inj_on f A) \<or>
- (\<exists> f::'a\<Rightarrow>nat. infinite A \<and> bij_betw f A UNIV)"
-using countable_or_card_of[OF assms]
-by (metis assms card_of_ordIso countable_ex_to_nat)
-
-lemma countable_cases_card_of[elim, consumes 1, case_names Fin Inf]:
-assumes "countable A"
-and "\<lbrakk>finite A; |A| <o |UNIV::nat set|\<rbrakk> \<Longrightarrow> phi"
-and "\<lbrakk>infinite A; |A| =o |UNIV::nat set|\<rbrakk> \<Longrightarrow> phi"
-shows phi
-using assms countable_or_card_of by blast
-
-lemma countable_cases[elim, consumes 1, case_names Fin Inf]:
-assumes "countable A"
-and "\<And> f::'a\<Rightarrow>nat. \<lbrakk>finite A; inj_on f A\<rbrakk> \<Longrightarrow> phi"
-and "\<And> f::'a\<Rightarrow>nat. \<lbrakk>infinite A; bij_betw f A UNIV\<rbrakk> \<Longrightarrow> phi"
-shows phi
-using assms countable_or by metis
-
-definition toNat_pred :: "'a set \<Rightarrow> ('a \<Rightarrow> nat) \<Rightarrow> bool"
-where
-"toNat_pred (A::'a set) f \<equiv>
- (finite A \<and> inj_on f A) \<or> (infinite A \<and> bij_betw f A UNIV)"
-definition toNat where "toNat A \<equiv> SOME f. toNat_pred A f"
-
-lemma toNat_pred:
-assumes "countable A"
-shows "\<exists> f. toNat_pred A f"
-using assms countable_ex_to_nat toNat_pred_def by (cases rule: countable_cases) auto
-
-lemma toNat_pred_toNat:
-assumes "countable A"
-shows "toNat_pred A (toNat A)"
-unfolding toNat_def apply(rule someI_ex[of "toNat_pred A"])
-using toNat_pred[OF assms] .
-
-lemma bij_betw_toNat:
-assumes c: "countable A" and i: "infinite A"
-shows "bij_betw (toNat A) A (UNIV::nat set)"
-using toNat_pred_toNat[OF c] unfolding toNat_pred_def using i by auto
-
-lemma inj_on_toNat:
-assumes c: "countable A"
-shows "inj_on (toNat A) A"
-using c apply(cases rule: countable_cases)
-using bij_betw_toNat[OF c] toNat_pred_toNat[OF c]
-unfolding toNat_pred_def unfolding bij_betw_def by auto
-
-lemma toNat_inj[simp]:
-assumes c: "countable A" and a: "a \<in> A" and b: "b \<in> A"
-shows "toNat A a = toNat A b \<longleftrightarrow> a = b"
-using inj_on_toNat[OF c] using a b unfolding inj_on_def by auto
-
-lemma image_toNat:
-assumes c: "countable A" and i: "infinite A"
-shows "toNat A ` A = UNIV"
-using bij_betw_toNat[OF assms] unfolding bij_betw_def by simp
-
-lemma toNat_surj:
-assumes "countable A" and i: "infinite A"
-shows "\<exists> a. a \<in> A \<and> toNat A a = n"
-using image_toNat[OF assms]
-by (metis (no_types) image_iff iso_tuple_UNIV_I)
-
-definition
-"fromNat A n \<equiv>
- if n \<in> toNat A ` A then inv_into A (toNat A) n
- else (SOME a. a \<in> A)"
-
-lemma fromNat:
-assumes "A \<noteq> {}"
-shows "fromNat A n \<in> A"
-unfolding fromNat_def by (metis assms equals0I inv_into_into someI_ex)
-
-lemma toNat_fromNat[simp]:
-assumes "n \<in> toNat A ` A"
-shows "toNat A (fromNat A n) = n"
-by (metis assms f_inv_into_f fromNat_def)
-
-lemma infinite_toNat_fromNat[simp]:
-assumes c: "countable A" and i: "infinite A"
-shows "toNat A (fromNat A n) = n"
-apply(rule toNat_fromNat) using toNat_surj[OF assms]
-by (metis image_iff)
-
-lemma fromNat_toNat[simp]:
-assumes c: "countable A" and a: "a \<in> A"
-shows "fromNat A (toNat A a) = a"
-by (metis a c equals0D fromNat imageI toNat_fromNat toNat_inj)
-
-lemma fromNat_inj:
-assumes c: "countable A" and i: "infinite A"
-shows "fromNat A m = fromNat A n \<longleftrightarrow> m = n" (is "?L = ?R \<longleftrightarrow> ?K")
-proof-
- have "?L = ?R \<longleftrightarrow> toNat A ?L = toNat A ?R"
- unfolding toNat_inj[OF c fromNat[OF infinite_imp_nonempty[OF i]]
- fromNat[OF infinite_imp_nonempty[OF i]]] ..
- also have "... \<longleftrightarrow> ?K" using c i by simp
- finally show ?thesis .
-qed
-
-lemma fromNat_surj:
-assumes c: "countable A" and a: "a \<in> A"
-shows "\<exists> n. fromNat A n = a"
-apply(rule exI[of _ "toNat A a"]) using assms by simp
-
-lemma fromNat_image_incl:
-assumes "A \<noteq> {}"
-shows "fromNat A ` UNIV \<subseteq> A"
-using fromNat[OF assms] by auto
-
-lemma incl_fromNat_image:
-assumes "countable A"
-shows "A \<subseteq> fromNat A ` UNIV"
-unfolding image_def using fromNat_surj[OF assms] by auto
-
-lemma fromNat_image[simp]:
-assumes "A \<noteq> {}" and "countable A"
-shows "fromNat A ` UNIV = A"
-by (metis assms equalityI fromNat_image_incl incl_fromNat_image)
-
-lemma fromNat_inject[simp]:
-assumes A: "A \<noteq> {}" "countable A" and B: "B \<noteq> {}" "countable B"
-shows "fromNat A = fromNat B \<longleftrightarrow> A = B"
-by (metis assms fromNat_image)
-
-lemma inj_on_fromNat:
-"inj_on fromNat ({A. A \<noteq> {} \<and> countable A})"
-unfolding inj_on_def by auto
-
-
-subsection {* Preservation under the set theoretic operations *}
-
-lemma contable_empty[simp,intro]:
-"countable {}"
-by (metis countable_ex_to_nat inj_on_empty)
-
-lemma incl_countable:
-assumes "A \<subseteq> B" and "countable B"
-shows "countable A"
-by (metis assms countable_ex_to_nat subset_inj_on)
-
-lemma countable_diff:
-assumes "countable A"
-shows "countable (A - B)"
-by (metis Diff_subset assms incl_countable)
-
-lemma finite_countable[simp]:
-assumes "finite A"
-shows "countable A"
-by (metis assms countable_ex_to_nat finite_imp_inj_to_nat_seg)
-
-lemma countable_singl[simp]:
-"countable {a}"
-by simp
-
-lemma countable_insert[simp]:
-"countable (insert a A) \<longleftrightarrow> countable A"
-proof
- assume c: "countable A"
- thus "countable (insert a A)"
- apply (cases rule: countable_cases_card_of)
- apply (metis finite_countable finite_insert)
- unfolding countable_card_of_nat
- by (metis infinite_card_of_insert ordIso_imp_ordLeq ordIso_transitive)
-qed(insert incl_countable, metis incl_countable subset_insertI)
-
-lemma contable_IntL[simp]:
-assumes "countable A"
-shows "countable (A \<inter> B)"
-by (metis Int_lower1 assms incl_countable)
-
-lemma contable_IntR[simp]:
-assumes "countable B"
-shows "countable (A \<inter> B)"
-by (metis assms contable_IntL inf.commute)
-
-lemma countable_UN[simp]:
-assumes cI: "countable I" and cA: "\<And> i. i \<in> I \<Longrightarrow> countable (A i)"
-shows "countable (\<Union> i \<in> I. A i)"
-using assms unfolding countable_card_of_nat
-apply(intro card_of_UNION_ordLeq_infinite) by auto
-
-lemma contable_Un[simp]:
-"countable (A \<union> B) \<longleftrightarrow> countable A \<and> countable B"
-proof safe
- assume cA: "countable A" and cB: "countable B"
- let ?I = "{0,Suc 0}" let ?As = "\<lambda> i. case i of 0 \<Rightarrow> A|Suc 0 \<Rightarrow> B"
- have AB: "A \<union> B = (\<Union> i \<in> ?I. ?As i)" by simp
- show "countable (A \<union> B)" unfolding AB apply(rule countable_UN)
- using cA cB by auto
-qed (metis Un_upper1 incl_countable, metis Un_upper2 incl_countable)
-
-lemma countable_INT[simp]:
-assumes "i \<in> I" and "countable (A i)"
-shows "countable (\<Inter> i \<in> I. A i)"
-by (metis INF_insert assms contable_IntL insert_absorb)
-
-lemma countable_class[simp]:
-fixes A :: "('a::countable) set"
-shows "countable A"
-proof-
- have "inj_on to_nat A" by (metis inj_on_to_nat)
- thus ?thesis by (metis countable_ex_to_nat)
-qed
-
-lemma countable_image[simp]:
-assumes "countable A"
-shows "countable (f ` A)"
-using assms unfolding countable_card_of_nat
-by (metis card_of_image ordLeq_transitive)
-
-lemma countable_ordLeq:
-assumes "|A| \<le>o |B|" and "countable B"
-shows "countable A"
-using assms unfolding countable_card_of_nat by(rule ordLeq_transitive)
-
-lemma countable_ordLess:
-assumes AB: "|A| <o |B|" and B: "countable B"
-shows "countable A"
-using countable_ordLeq[OF ordLess_imp_ordLeq[OF AB] B] .
-
-lemma countable_vimage:
-assumes "B \<subseteq> range f" and "countable (f -` B)"
-shows "countable B"
-by (metis Int_absorb2 assms countable_image image_vimage_eq)
-
-lemma surj_countable_vimage:
-assumes s: "surj f" and c: "countable (f -` B)"
-shows "countable B"
-apply(rule countable_vimage[OF _ c]) using s by auto
-
-lemma countable_Collect[simp]:
-assumes "countable A"
-shows "countable {a \<in> A. \<phi> a}"
-by (metis Collect_conj_eq Int_absorb Int_commute Int_def assms contable_IntR)
-
-lemma countable_Plus[simp]:
-assumes A: "countable A" and B: "countable B"
-shows "countable (A <+> B)"
-proof-
- let ?U = "UNIV::nat set"
- have "|A| \<le>o |?U|" and "|B| \<le>o |?U|" using A B
- using card_of_nat[THEN ordIso_symmetric] ordLeq_ordIso_trans
- unfolding countable_def by blast+
- hence "|A <+> B| \<le>o |?U|" by (intro card_of_Plus_ordLeq_infinite) auto
- thus ?thesis using card_of_nat unfolding countable_def by(rule ordLeq_ordIso_trans)
-qed
-
-lemma countable_Times[simp]:
-assumes A: "countable A" and B: "countable B"
-shows "countable (A \<times> B)"
-proof-
- let ?U = "UNIV::nat set"
- have "|A| \<le>o |?U|" and "|B| \<le>o |?U|" using A B
- using card_of_nat[THEN ordIso_symmetric] ordLeq_ordIso_trans
- unfolding countable_def by blast+
- hence "|A \<times> B| \<le>o |?U|" by (intro card_of_Times_ordLeq_infinite) auto
- thus ?thesis using card_of_nat unfolding countable_def by(rule ordLeq_ordIso_trans)
-qed
-
-lemma ordLeq_countable:
-assumes "|A| \<le>o |B|" and "countable B"
-shows "countable A"
-using assms unfolding countable_def by(rule ordLeq_transitive)
-
-lemma ordLess_countable:
-assumes A: "|A| <o |B|" and B: "countable B"
-shows "countable A"
-by (rule ordLeq_countable[OF ordLess_imp_ordLeq[OF A] B])
-
-lemma countable_def2: "countable A \<longleftrightarrow> |A| \<le>o |UNIV :: nat set|"
-unfolding countable_def using card_of_nat[THEN ordIso_symmetric]
-by (metis (lifting) Field_card_of Field_natLeq card_of_mono2 card_of_nat
- countable_def ordIso_imp_ordLeq ordLeq_countable)
-
-
-subsection{* The type of countable sets *}
-
-typedef (open) 'a cset = "{A :: 'a set. countable A}"
-apply(rule exI[of _ "{}"]) by simp
-
-abbreviation rcset where "rcset \<equiv> Rep_cset"
-abbreviation acset where "acset \<equiv> Abs_cset"
-
-lemmas acset_rcset = Rep_cset_inverse
-declare acset_rcset[simp]
-
-lemma acset_surj:
-"\<exists> A. countable A \<and> acset A = C"
-apply(cases rule: Abs_cset_cases[of C]) by auto
-
-lemma rcset_acset[simp]:
-assumes "countable A"
-shows "rcset (acset A) = A"
-using Abs_cset_inverse assms by auto
-
-lemma acset_inj[simp]:
-assumes "countable A" and "countable B"
-shows "acset A = acset B \<longleftrightarrow> A = B"
-using assms Abs_cset_inject by auto
-
-lemma rcset[simp]:
-"countable (rcset C)"
-using Rep_cset by simp
-
-lemma rcset_inj[simp]:
-"rcset C = rcset D \<longleftrightarrow> C = D"
-by (metis acset_rcset)
-
-lemma rcset_surj:
-assumes "countable A"
-shows "\<exists> C. rcset C = A"
-apply(cases rule: Rep_cset_cases[of A])
-using assms by auto
-
-definition "cIn a C \<equiv> (a \<in> rcset C)"
-definition "cEmp \<equiv> acset {}"
-definition "cIns a C \<equiv> acset (insert a (rcset C))"
-abbreviation cSingl where "cSingl a \<equiv> cIns a cEmp"
-definition "cUn C D \<equiv> acset (rcset C \<union> rcset D)"
-definition "cInt C D \<equiv> acset (rcset C \<inter> rcset D)"
-definition "cDif C D \<equiv> acset (rcset C - rcset D)"
-definition "cIm f C \<equiv> acset (f ` rcset C)"
-definition "cVim f D \<equiv> acset (f -` rcset D)"
-(* TODO eventually: nice setup for these operations, copied from the set setup *)
-
-end
--- a/src/HOL/Codatatype/Equiv_Relations_More.thy Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,161 +0,0 @@
-(* Title: HOL/BNF/Equiv_Relations_More.thy
- Author: Andrei Popescu, TU Muenchen
- Copyright 2012
-
-Some preliminaries on equivalence relations and quotients.
-*)
-
-header {* Some Preliminaries on Equivalence Relations and Quotients *}
-
-theory Equiv_Relations_More
-imports Equiv_Relations Hilbert_Choice
-begin
-
-
-(* Recall the following constants and lemmas:
-
-term Eps
-term "A//r"
-lemmas equiv_def
-lemmas refl_on_def
- -- note that "reflexivity on" also assumes inclusion of the relation's field into r
-
-*)
-
-definition proj where "proj r x = r `` {x}"
-
-definition univ where "univ f X == f (Eps (%x. x \<in> X))"
-
-lemma proj_preserves:
-"x \<in> A \<Longrightarrow> proj r x \<in> A//r"
-unfolding proj_def by (rule quotientI)
-
-lemma proj_in_iff:
-assumes "equiv A r"
-shows "(proj r x \<in> A//r) = (x \<in> A)"
-apply(rule iffI, auto simp add: proj_preserves)
-unfolding proj_def quotient_def proof clarsimp
- fix y assume y: "y \<in> A" and "r `` {x} = r `` {y}"
- moreover have "y \<in> r `` {y}" using assms y unfolding equiv_def refl_on_def by blast
- ultimately have "(x,y) \<in> r" by blast
- thus "x \<in> A" using assms unfolding equiv_def refl_on_def by blast
-qed
-
-lemma proj_iff:
-"\<lbrakk>equiv A r; {x,y} \<subseteq> A\<rbrakk> \<Longrightarrow> (proj r x = proj r y) = ((x,y) \<in> r)"
-by (simp add: proj_def eq_equiv_class_iff)
-
-(*
-lemma in_proj: "\<lbrakk>equiv A r; x \<in> A\<rbrakk> \<Longrightarrow> x \<in> proj r x"
-unfolding proj_def equiv_def refl_on_def by blast
-*)
-
-lemma proj_image: "(proj r) ` A = A//r"
-unfolding proj_def[abs_def] quotient_def by blast
-
-lemma in_quotient_imp_non_empty:
-"\<lbrakk>equiv A r; X \<in> A//r\<rbrakk> \<Longrightarrow> X \<noteq> {}"
-unfolding quotient_def using equiv_class_self by fast
-
-lemma in_quotient_imp_in_rel:
-"\<lbrakk>equiv A r; X \<in> A//r; {x,y} \<subseteq> X\<rbrakk> \<Longrightarrow> (x,y) \<in> r"
-using quotient_eq_iff by fastforce
-
-lemma in_quotient_imp_closed:
-"\<lbrakk>equiv A r; X \<in> A//r; x \<in> X; (x,y) \<in> r\<rbrakk> \<Longrightarrow> y \<in> X"
-unfolding quotient_def equiv_def trans_def by blast
-
-lemma in_quotient_imp_subset:
-"\<lbrakk>equiv A r; X \<in> A//r\<rbrakk> \<Longrightarrow> X \<subseteq> A"
-using assms in_quotient_imp_in_rel equiv_type by fastforce
-
-lemma equiv_Eps_in:
-"\<lbrakk>equiv A r; X \<in> A//r\<rbrakk> \<Longrightarrow> Eps (%x. x \<in> X) \<in> X"
-apply (rule someI2_ex)
-using in_quotient_imp_non_empty by blast
-
-lemma equiv_Eps_preserves:
-assumes ECH: "equiv A r" and X: "X \<in> A//r"
-shows "Eps (%x. x \<in> X) \<in> A"
-apply (rule in_mono[rule_format])
- using assms apply (rule in_quotient_imp_subset)
-by (rule equiv_Eps_in) (rule assms)+
-
-lemma proj_Eps:
-assumes "equiv A r" and "X \<in> A//r"
-shows "proj r (Eps (%x. x \<in> X)) = X"
-unfolding proj_def proof auto
- fix x assume x: "x \<in> X"
- thus "(Eps (%x. x \<in> X), x) \<in> r" using assms equiv_Eps_in in_quotient_imp_in_rel by fast
-next
- fix x assume "(Eps (%x. x \<in> X),x) \<in> r"
- thus "x \<in> X" using in_quotient_imp_closed[OF assms equiv_Eps_in[OF assms]] by fast
-qed
-
-(*
-lemma Eps_proj:
-assumes "equiv A r" and "x \<in> A"
-shows "(Eps (%y. y \<in> proj r x), x) \<in> r"
-proof-
- have 1: "proj r x \<in> A//r" using assms proj_preserves by fastforce
- hence "Eps(%y. y \<in> proj r x) \<in> proj r x" using assms equiv_Eps_in by auto
- moreover have "x \<in> proj r x" using assms in_proj by fastforce
- ultimately show ?thesis using assms 1 in_quotient_imp_in_rel by fastforce
-qed
-
-lemma equiv_Eps_iff:
-assumes "equiv A r" and "{X,Y} \<subseteq> A//r"
-shows "((Eps (%x. x \<in> X),Eps (%y. y \<in> Y)) \<in> r) = (X = Y)"
-proof-
- have "Eps (%x. x \<in> X) \<in> X \<and> Eps (%y. y \<in> Y) \<in> Y" using assms equiv_Eps_in by auto
- thus ?thesis using assms quotient_eq_iff by fastforce
-qed
-
-lemma equiv_Eps_inj_on:
-assumes "equiv A r"
-shows "inj_on (%X. Eps (%x. x \<in> X)) (A//r)"
-unfolding inj_on_def proof clarify
- fix X Y assume X: "X \<in> A//r" and Y: "Y \<in> A//r" and Eps: "Eps (%x. x \<in> X) = Eps (%y. y \<in> Y)"
- hence "Eps (%x. x \<in> X) \<in> A" using assms equiv_Eps_preserves by auto
- hence "(Eps (%x. x \<in> X), Eps (%y. y \<in> Y)) \<in> r"
- using assms Eps unfolding quotient_def equiv_def refl_on_def by auto
- thus "X= Y" using X Y assms equiv_Eps_iff by auto
-qed
-*)
-
-lemma univ_commute:
-assumes ECH: "equiv A r" and RES: "f respects r" and x: "x \<in> A"
-shows "(univ f) (proj r x) = f x"
-unfolding univ_def proof -
- have prj: "proj r x \<in> A//r" using x proj_preserves by fast
- hence "Eps (%y. y \<in> proj r x) \<in> A" using ECH equiv_Eps_preserves by fast
- moreover have "proj r (Eps (%y. y \<in> proj r x)) = proj r x" using ECH prj proj_Eps by fast
- ultimately have "(x, Eps (%y. y \<in> proj r x)) \<in> r" using x ECH proj_iff by fast
- thus "f (Eps (%y. y \<in> proj r x)) = f x" using RES unfolding congruent_def by fastforce
-qed
-
-(*
-lemma univ_unique:
-assumes ECH: "equiv A r" and
- RES: "f respects r" and COM: "\<forall> x \<in> A. G (proj r x) = f x"
-shows "\<forall> X \<in> A//r. G X = univ f X"
-proof
- fix X assume "X \<in> A//r"
- then obtain x where x: "x \<in> A" and X: "X = proj r x" using ECH proj_image[of r A] by blast
- have "G X = f x" unfolding X using x COM by simp
- thus "G X = univ f X" unfolding X using ECH RES x univ_commute by fastforce
-qed
-*)
-
-lemma univ_preserves:
-assumes ECH: "equiv A r" and RES: "f respects r" and
- PRES: "\<forall> x \<in> A. f x \<in> B"
-shows "\<forall> X \<in> A//r. univ f X \<in> B"
-proof
- fix X assume "X \<in> A//r"
- then obtain x where x: "x \<in> A" and X: "X = proj r x" using ECH proj_image[of r A] by blast
- hence "univ f X = f x" using assms univ_commute by fastforce
- thus "univ f X \<in> B" using x PRES by simp
-qed
-
-end
--- a/src/HOL/Codatatype/Examples/HFset.thy Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,60 +0,0 @@
-(* Title: HOL/BNF/Examples/HFset.thy
- Author: Andrei Popescu, TU Muenchen
- Copyright 2012
-
-Hereditary sets.
-*)
-
-header {* Hereditary Sets *}
-
-theory HFset
-imports "../BNF"
-begin
-
-
-section {* Datatype definition *}
-
-data_raw hfset: 'hfset = "'hfset fset"
-
-
-section {* Customization of terms *}
-
-subsection{* Constructors *}
-
-definition "Fold hs \<equiv> hfset_ctor hs"
-
-lemma hfset_simps[simp]:
-"\<And>hs1 hs2. Fold hs1 = Fold hs2 \<longrightarrow> hs1 = hs2"
-unfolding Fold_def hfset.ctor_inject by auto
-
-theorem hfset_cases[elim, case_names Fold]:
-assumes Fold: "\<And> hs. h = Fold hs \<Longrightarrow> phi"
-shows phi
-using Fold unfolding Fold_def
-by (cases rule: hfset.ctor_exhaust[of h]) simp
-
-lemma hfset_induct[case_names Fold, induct type: hfset]:
-assumes Fold: "\<And> hs. (\<And> h. h |\<in>| hs \<Longrightarrow> phi h) \<Longrightarrow> phi (Fold hs)"
-shows "phi t"
-apply (induct rule: hfset.ctor_induct)
-using Fold unfolding Fold_def fset_fset_member mem_Collect_eq ..
-
-(* alternative induction principle, using fset: *)
-lemma hfset_induct_fset[case_names Fold, induct type: hfset]:
-assumes Fold: "\<And> hs. (\<And> h. h \<in> fset hs \<Longrightarrow> phi h) \<Longrightarrow> phi (Fold hs)"
-shows "phi t"
-apply (induct rule: hfset_induct)
-using Fold by (metis notin_fset)
-
-subsection{* Recursion and iteration (fold) *}
-
-lemma hfset_ctor_rec:
-"hfset_ctor_rec R (Fold hs) = R (map_fset <id, hfset_ctor_rec R> hs)"
-using hfset.ctor_recs unfolding Fold_def .
-
-(* The iterator has a simpler form: *)
-lemma hfset_ctor_fold:
-"hfset_ctor_fold R (Fold hs) = R (map_fset (hfset_ctor_fold R) hs)"
-using hfset.ctor_folds unfolding Fold_def .
-
-end
--- a/src/HOL/Codatatype/Examples/Infinite_Derivation_Trees/Gram_Lang.thy Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1366 +0,0 @@
-(* Title: HOL/BNF/Examples/Infinite_Derivation_Trees/Gram_Lang.thy
- Author: Andrei Popescu, TU Muenchen
- Copyright 2012
-
-Language of a grammar.
-*)
-
-header {* Language of a Grammar *}
-
-theory Gram_Lang
-imports Tree
-begin
-
-
-consts P :: "(N \<times> (T + N) set) set"
-axiomatization where
- finite_N: "finite (UNIV::N set)"
-and finite_in_P: "\<And> n tns. (n,tns) \<in> P \<longrightarrow> finite tns"
-and used: "\<And> n. \<exists> tns. (n,tns) \<in> P"
-
-
-subsection{* Tree basics: frontier, interior, etc. *}
-
-lemma Tree_cong:
-assumes "root tr = root tr'" and "cont tr = cont tr'"
-shows "tr = tr'"
-by (metis Node_root_cont assms)
-
-inductive finiteT where
-Node: "\<lbrakk>finite as; (finiteT^#) as\<rbrakk> \<Longrightarrow> finiteT (Node a as)"
-monos lift_mono
-
-lemma finiteT_induct[consumes 1, case_names Node, induct pred: finiteT]:
-assumes 1: "finiteT tr"
-and IH: "\<And>as n. \<lbrakk>finite as; (\<phi>^#) as\<rbrakk> \<Longrightarrow> \<phi> (Node n as)"
-shows "\<phi> tr"
-using 1 apply(induct rule: finiteT.induct)
-apply(rule IH) apply assumption apply(elim mono_lift) by simp
-
-
-(* Frontier *)
-
-inductive inFr :: "N set \<Rightarrow> Tree \<Rightarrow> T \<Rightarrow> bool" where
-Base: "\<lbrakk>root tr \<in> ns; Inl t \<in> cont tr\<rbrakk> \<Longrightarrow> inFr ns tr t"
-|
-Ind: "\<lbrakk>root tr \<in> ns; Inr tr1 \<in> cont tr; inFr ns tr1 t\<rbrakk> \<Longrightarrow> inFr ns tr t"
-
-definition "Fr ns tr \<equiv> {t. inFr ns tr t}"
-
-lemma inFr_root_in: "inFr ns tr t \<Longrightarrow> root tr \<in> ns"
-by (metis inFr.simps)
-
-lemma inFr_mono:
-assumes "inFr ns tr t" and "ns \<subseteq> ns'"
-shows "inFr ns' tr t"
-using assms apply(induct arbitrary: ns' rule: inFr.induct)
-using Base Ind by (metis inFr.simps set_mp)+
-
-lemma inFr_Ind_minus:
-assumes "inFr ns1 tr1 t" and "Inr tr1 \<in> cont tr"
-shows "inFr (insert (root tr) ns1) tr t"
-using assms apply(induct rule: inFr.induct)
- apply (metis inFr.simps insert_iff)
- by (metis inFr.simps inFr_mono insertI1 subset_insertI)
-
-(* alternative definition *)
-inductive inFr2 :: "N set \<Rightarrow> Tree \<Rightarrow> T \<Rightarrow> bool" where
-Base: "\<lbrakk>root tr \<in> ns; Inl t \<in> cont tr\<rbrakk> \<Longrightarrow> inFr2 ns tr t"
-|
-Ind: "\<lbrakk>Inr tr1 \<in> cont tr; inFr2 ns1 tr1 t\<rbrakk>
- \<Longrightarrow> inFr2 (insert (root tr) ns1) tr t"
-
-lemma inFr2_root_in: "inFr2 ns tr t \<Longrightarrow> root tr \<in> ns"
-apply(induct rule: inFr2.induct) by auto
-
-lemma inFr2_mono:
-assumes "inFr2 ns tr t" and "ns \<subseteq> ns'"
-shows "inFr2 ns' tr t"
-using assms apply(induct arbitrary: ns' rule: inFr2.induct)
-using Base Ind
-apply (metis subsetD) by (metis inFr2.simps insert_absorb insert_subset)
-
-lemma inFr2_Ind:
-assumes "inFr2 ns tr1 t" and "root tr \<in> ns" and "Inr tr1 \<in> cont tr"
-shows "inFr2 ns tr t"
-using assms apply(induct rule: inFr2.induct)
- apply (metis inFr2.simps insert_absorb)
- by (metis inFr2.simps insert_absorb)
-
-lemma inFr_inFr2:
-"inFr = inFr2"
-apply (rule ext)+ apply(safe)
- apply(erule inFr.induct)
- apply (metis (lifting) inFr2.Base)
- apply (metis (lifting) inFr2_Ind)
- apply(erule inFr2.induct)
- apply (metis (lifting) inFr.Base)
- apply (metis (lifting) inFr_Ind_minus)
-done
-
-lemma not_root_inFr:
-assumes "root tr \<notin> ns"
-shows "\<not> inFr ns tr t"
-by (metis assms inFr_root_in)
-
-theorem not_root_Fr:
-assumes "root tr \<notin> ns"
-shows "Fr ns tr = {}"
-using not_root_inFr[OF assms] unfolding Fr_def by auto
-
-
-(* Interior *)
-
-inductive inItr :: "N set \<Rightarrow> Tree \<Rightarrow> N \<Rightarrow> bool" where
-Base: "root tr \<in> ns \<Longrightarrow> inItr ns tr (root tr)"
-|
-Ind: "\<lbrakk>root tr \<in> ns; Inr tr1 \<in> cont tr; inItr ns tr1 n\<rbrakk> \<Longrightarrow> inItr ns tr n"
-
-definition "Itr ns tr \<equiv> {n. inItr ns tr n}"
-
-lemma inItr_root_in: "inItr ns tr n \<Longrightarrow> root tr \<in> ns"
-by (metis inItr.simps)
-
-lemma inItr_mono:
-assumes "inItr ns tr n" and "ns \<subseteq> ns'"
-shows "inItr ns' tr n"
-using assms apply(induct arbitrary: ns' rule: inItr.induct)
-using Base Ind by (metis inItr.simps set_mp)+
-
-
-(* The subtree relation *)
-
-inductive subtr where
-Refl: "root tr \<in> ns \<Longrightarrow> subtr ns tr tr"
-|
-Step: "\<lbrakk>root tr3 \<in> ns; subtr ns tr1 tr2; Inr tr2 \<in> cont tr3\<rbrakk> \<Longrightarrow> subtr ns tr1 tr3"
-
-lemma subtr_rootL_in:
-assumes "subtr ns tr1 tr2"
-shows "root tr1 \<in> ns"
-using assms apply(induct rule: subtr.induct) by auto
-
-lemma subtr_rootR_in:
-assumes "subtr ns tr1 tr2"
-shows "root tr2 \<in> ns"
-using assms apply(induct rule: subtr.induct) by auto
-
-lemmas subtr_roots_in = subtr_rootL_in subtr_rootR_in
-
-lemma subtr_mono:
-assumes "subtr ns tr1 tr2" and "ns \<subseteq> ns'"
-shows "subtr ns' tr1 tr2"
-using assms apply(induct arbitrary: ns' rule: subtr.induct)
-using Refl Step by (metis subtr.simps set_mp)+
-
-lemma subtr_trans_Un:
-assumes "subtr ns12 tr1 tr2" and "subtr ns23 tr2 tr3"
-shows "subtr (ns12 \<union> ns23) tr1 tr3"
-proof-
- have "subtr ns23 tr2 tr3 \<Longrightarrow>
- (\<forall> ns12 tr1. subtr ns12 tr1 tr2 \<longrightarrow> subtr (ns12 \<union> ns23) tr1 tr3)"
- apply(induct rule: subtr.induct, safe)
- apply (metis subtr_mono sup_commute sup_ge2)
- by (metis (lifting) Step UnI2)
- thus ?thesis using assms by auto
-qed
-
-lemma subtr_trans:
-assumes "subtr ns tr1 tr2" and "subtr ns tr2 tr3"
-shows "subtr ns tr1 tr3"
-using subtr_trans_Un[OF assms] by simp
-
-lemma subtr_StepL:
-assumes r: "root tr1 \<in> ns" and tr12: "Inr tr1 \<in> cont tr2" and s: "subtr ns tr2 tr3"
-shows "subtr ns tr1 tr3"
-apply(rule subtr_trans[OF _ s]) apply(rule Step[of tr2 ns tr1 tr1])
-by (metis assms subtr_rootL_in Refl)+
-
-(* alternative definition: *)
-inductive subtr2 where
-Refl: "root tr \<in> ns \<Longrightarrow> subtr2 ns tr tr"
-|
-Step: "\<lbrakk>root tr1 \<in> ns; Inr tr1 \<in> cont tr2; subtr2 ns tr2 tr3\<rbrakk> \<Longrightarrow> subtr2 ns tr1 tr3"
-
-lemma subtr2_rootL_in:
-assumes "subtr2 ns tr1 tr2"
-shows "root tr1 \<in> ns"
-using assms apply(induct rule: subtr2.induct) by auto
-
-lemma subtr2_rootR_in:
-assumes "subtr2 ns tr1 tr2"
-shows "root tr2 \<in> ns"
-using assms apply(induct rule: subtr2.induct) by auto
-
-lemmas subtr2_roots_in = subtr2_rootL_in subtr2_rootR_in
-
-lemma subtr2_mono:
-assumes "subtr2 ns tr1 tr2" and "ns \<subseteq> ns'"
-shows "subtr2 ns' tr1 tr2"
-using assms apply(induct arbitrary: ns' rule: subtr2.induct)
-using Refl Step by (metis subtr2.simps set_mp)+
-
-lemma subtr2_trans_Un:
-assumes "subtr2 ns12 tr1 tr2" and "subtr2 ns23 tr2 tr3"
-shows "subtr2 (ns12 \<union> ns23) tr1 tr3"
-proof-
- have "subtr2 ns12 tr1 tr2 \<Longrightarrow>
- (\<forall> ns23 tr3. subtr2 ns23 tr2 tr3 \<longrightarrow> subtr2 (ns12 \<union> ns23) tr1 tr3)"
- apply(induct rule: subtr2.induct, safe)
- apply (metis subtr2_mono sup_commute sup_ge2)
- by (metis Un_iff subtr2.simps)
- thus ?thesis using assms by auto
-qed
-
-lemma subtr2_trans:
-assumes "subtr2 ns tr1 tr2" and "subtr2 ns tr2 tr3"
-shows "subtr2 ns tr1 tr3"
-using subtr2_trans_Un[OF assms] by simp
-
-lemma subtr2_StepR:
-assumes r: "root tr3 \<in> ns" and tr23: "Inr tr2 \<in> cont tr3" and s: "subtr2 ns tr1 tr2"
-shows "subtr2 ns tr1 tr3"
-apply(rule subtr2_trans[OF s]) apply(rule Step[of _ _ tr3])
-by (metis assms subtr2_rootR_in Refl)+
-
-lemma subtr_subtr2:
-"subtr = subtr2"
-apply (rule ext)+ apply(safe)
- apply(erule subtr.induct)
- apply (metis (lifting) subtr2.Refl)
- apply (metis (lifting) subtr2_StepR)
- apply(erule subtr2.induct)
- apply (metis (lifting) subtr.Refl)
- apply (metis (lifting) subtr_StepL)
-done
-
-lemma subtr_inductL[consumes 1, case_names Refl Step]:
-assumes s: "subtr ns tr1 tr2" and Refl: "\<And>ns tr. \<phi> ns tr tr"
-and Step:
-"\<And>ns tr1 tr2 tr3.
- \<lbrakk>root tr1 \<in> ns; Inr tr1 \<in> cont tr2; subtr ns tr2 tr3; \<phi> ns tr2 tr3\<rbrakk> \<Longrightarrow> \<phi> ns tr1 tr3"
-shows "\<phi> ns tr1 tr2"
-using s unfolding subtr_subtr2 apply(rule subtr2.induct)
-using Refl Step unfolding subtr_subtr2 by auto
-
-lemma subtr_UNIV_inductL[consumes 1, case_names Refl Step]:
-assumes s: "subtr UNIV tr1 tr2" and Refl: "\<And>tr. \<phi> tr tr"
-and Step:
-"\<And>tr1 tr2 tr3.
- \<lbrakk>Inr tr1 \<in> cont tr2; subtr UNIV tr2 tr3; \<phi> tr2 tr3\<rbrakk> \<Longrightarrow> \<phi> tr1 tr3"
-shows "\<phi> tr1 tr2"
-using s apply(induct rule: subtr_inductL)
-apply(rule Refl) using Step subtr_mono by (metis subset_UNIV)
-
-(* Subtree versus frontier: *)
-lemma subtr_inFr:
-assumes "inFr ns tr t" and "subtr ns tr tr1"
-shows "inFr ns tr1 t"
-proof-
- have "subtr ns tr tr1 \<Longrightarrow> (\<forall> t. inFr ns tr t \<longrightarrow> inFr ns tr1 t)"
- apply(induct rule: subtr.induct, safe) by (metis inFr.Ind)
- thus ?thesis using assms by auto
-qed
-
-corollary Fr_subtr:
-"Fr ns tr = \<Union> {Fr ns tr' | tr'. subtr ns tr' tr}"
-unfolding Fr_def proof safe
- fix t assume t: "inFr ns tr t" hence "root tr \<in> ns" by (rule inFr_root_in)
- thus "t \<in> \<Union>{{t. inFr ns tr' t} |tr'. subtr ns tr' tr}"
- apply(intro UnionI[of "{t. inFr ns tr t}" _ t]) using t subtr.Refl by auto
-qed(metis subtr_inFr)
-
-lemma inFr_subtr:
-assumes "inFr ns tr t"
-shows "\<exists> tr'. subtr ns tr' tr \<and> Inl t \<in> cont tr'"
-using assms apply(induct rule: inFr.induct) apply safe
- apply (metis subtr.Refl)
- by (metis (lifting) subtr.Step)
-
-corollary Fr_subtr_cont:
-"Fr ns tr = \<Union> {Inl -` cont tr' | tr'. subtr ns tr' tr}"
-unfolding Fr_def
-apply safe
-apply (frule inFr_subtr)
-apply auto
-by (metis inFr.Base subtr_inFr subtr_rootL_in)
-
-(* Subtree versus interior: *)
-lemma subtr_inItr:
-assumes "inItr ns tr n" and "subtr ns tr tr1"
-shows "inItr ns tr1 n"
-proof-
- have "subtr ns tr tr1 \<Longrightarrow> (\<forall> t. inItr ns tr n \<longrightarrow> inItr ns tr1 n)"
- apply(induct rule: subtr.induct, safe) by (metis inItr.Ind)
- thus ?thesis using assms by auto
-qed
-
-corollary Itr_subtr:
-"Itr ns tr = \<Union> {Itr ns tr' | tr'. subtr ns tr' tr}"
-unfolding Itr_def apply safe
-apply (metis (lifting, mono_tags) UnionI inItr_root_in mem_Collect_eq subtr.Refl)
-by (metis subtr_inItr)
-
-lemma inItr_subtr:
-assumes "inItr ns tr n"
-shows "\<exists> tr'. subtr ns tr' tr \<and> root tr' = n"
-using assms apply(induct rule: inItr.induct) apply safe
- apply (metis subtr.Refl)
- by (metis (lifting) subtr.Step)
-
-corollary Itr_subtr_cont:
-"Itr ns tr = {root tr' | tr'. subtr ns tr' tr}"
-unfolding Itr_def apply safe
- apply (metis (lifting, mono_tags) UnionI inItr_subtr mem_Collect_eq vimageI2)
- by (metis inItr.Base subtr_inItr subtr_rootL_in)
-
-
-subsection{* The immediate subtree function *}
-
-(* production of: *)
-abbreviation "prodOf tr \<equiv> (id \<oplus> root) ` (cont tr)"
-(* subtree of: *)
-definition "subtrOf tr n \<equiv> SOME tr'. Inr tr' \<in> cont tr \<and> root tr' = n"
-
-lemma subtrOf:
-assumes n: "Inr n \<in> prodOf tr"
-shows "Inr (subtrOf tr n) \<in> cont tr \<and> root (subtrOf tr n) = n"
-proof-
- obtain tr' where "Inr tr' \<in> cont tr \<and> root tr' = n"
- using n unfolding image_def by (metis (lifting) Inr_oplus_elim assms)
- thus ?thesis unfolding subtrOf_def by(rule someI)
-qed
-
-lemmas Inr_subtrOf = subtrOf[THEN conjunct1]
-lemmas root_subtrOf[simp] = subtrOf[THEN conjunct2]
-
-lemma Inl_prodOf: "Inl -` (prodOf tr) = Inl -` (cont tr)"
-proof safe
- fix t ttr assume "Inl t = (id \<oplus> root) ttr" and "ttr \<in> cont tr"
- thus "t \<in> Inl -` cont tr" by(cases ttr, auto)
-next
- fix t assume "Inl t \<in> cont tr" thus "t \<in> Inl -` prodOf tr"
- by (metis (lifting) id_def image_iff sum_map.simps(1) vimageI2)
-qed
-
-lemma root_prodOf:
-assumes "Inr tr' \<in> cont tr"
-shows "Inr (root tr') \<in> prodOf tr"
-by (metis (lifting) assms image_iff sum_map.simps(2))
-
-
-subsection{* Derivation trees *}
-
-coinductive dtree where
-Tree: "\<lbrakk>(root tr, (id \<oplus> root) ` (cont tr)) \<in> P; inj_on root (Inr -` cont tr);
- lift dtree (cont tr)\<rbrakk> \<Longrightarrow> dtree tr"
-monos lift_mono
-
-(* destruction rules: *)
-lemma dtree_P:
-assumes "dtree tr"
-shows "(root tr, (id \<oplus> root) ` (cont tr)) \<in> P"
-using assms unfolding dtree.simps by auto
-
-lemma dtree_inj_on:
-assumes "dtree tr"
-shows "inj_on root (Inr -` cont tr)"
-using assms unfolding dtree.simps by auto
-
-lemma dtree_inj[simp]:
-assumes "dtree tr" and "Inr tr1 \<in> cont tr" and "Inr tr2 \<in> cont tr"
-shows "root tr1 = root tr2 \<longleftrightarrow> tr1 = tr2"
-using assms dtree_inj_on unfolding inj_on_def by auto
-
-lemma dtree_lift:
-assumes "dtree tr"
-shows "lift dtree (cont tr)"
-using assms unfolding dtree.simps by auto
-
-
-(* coinduction:*)
-lemma dtree_coind[elim, consumes 1, case_names Hyp]:
-assumes phi: "\<phi> tr"
-and Hyp:
-"\<And> tr. \<phi> tr \<Longrightarrow>
- (root tr, image (id \<oplus> root) (cont tr)) \<in> P \<and>
- inj_on root (Inr -` cont tr) \<and>
- lift (\<lambda> tr. \<phi> tr \<or> dtree tr) (cont tr)"
-shows "dtree tr"
-apply(rule dtree.coinduct[of \<phi> tr, OF phi])
-using Hyp by blast
-
-lemma dtree_raw_coind[elim, consumes 1, case_names Hyp]:
-assumes phi: "\<phi> tr"
-and Hyp:
-"\<And> tr. \<phi> tr \<Longrightarrow>
- (root tr, image (id \<oplus> root) (cont tr)) \<in> P \<and>
- inj_on root (Inr -` cont tr) \<and>
- lift \<phi> (cont tr)"
-shows "dtree tr"
-using phi apply(induct rule: dtree_coind)
-using Hyp mono_lift
-by (metis (mono_tags) mono_lift)
-
-lemma dtree_subtr_inj_on:
-assumes d: "dtree tr1" and s: "subtr ns tr tr1"
-shows "inj_on root (Inr -` cont tr)"
-using s d apply(induct rule: subtr.induct)
-apply (metis (lifting) dtree_inj_on) by (metis dtree_lift lift_def)
-
-lemma dtree_subtr_P:
-assumes d: "dtree tr1" and s: "subtr ns tr tr1"
-shows "(root tr, (id \<oplus> root) ` cont tr) \<in> P"
-using s d apply(induct rule: subtr.induct)
-apply (metis (lifting) dtree_P) by (metis dtree_lift lift_def)
-
-lemma subtrOf_root[simp]:
-assumes tr: "dtree tr" and cont: "Inr tr' \<in> cont tr"
-shows "subtrOf tr (root tr') = tr'"
-proof-
- have 0: "Inr (subtrOf tr (root tr')) \<in> cont tr" using Inr_subtrOf
- by (metis (lifting) cont root_prodOf)
- have "root (subtrOf tr (root tr')) = root tr'"
- using root_subtrOf by (metis (lifting) cont root_prodOf)
- thus ?thesis unfolding dtree_inj[OF tr 0 cont] .
-qed
-
-lemma surj_subtrOf:
-assumes "dtree tr" and 0: "Inr tr' \<in> cont tr"
-shows "\<exists> n. Inr n \<in> prodOf tr \<and> subtrOf tr n = tr'"
-apply(rule exI[of _ "root tr'"])
-using root_prodOf[OF 0] subtrOf_root[OF assms] by simp
-
-lemma dtree_subtr:
-assumes "dtree tr1" and "subtr ns tr tr1"
-shows "dtree tr"
-proof-
- have "(\<exists> ns tr1. dtree tr1 \<and> subtr ns tr tr1) \<Longrightarrow> dtree tr"
- proof (induct rule: dtree_raw_coind)
- case (Hyp tr)
- then obtain ns tr1 where tr1: "dtree tr1" and tr_tr1: "subtr ns tr tr1" by auto
- show ?case unfolding lift_def proof safe
- show "(root tr, (id \<oplus> root) ` cont tr) \<in> P" using dtree_subtr_P[OF tr1 tr_tr1] .
- next
- show "inj_on root (Inr -` cont tr)" using dtree_subtr_inj_on[OF tr1 tr_tr1] .
- next
- fix tr' assume tr': "Inr tr' \<in> cont tr"
- have tr_tr1: "subtr (ns \<union> {root tr'}) tr tr1" using subtr_mono[OF tr_tr1] by auto
- have "subtr (ns \<union> {root tr'}) tr' tr1" using subtr_StepL[OF _ tr' tr_tr1] by auto
- thus "\<exists>ns' tr1. dtree tr1 \<and> subtr ns' tr' tr1" using tr1 by blast
- qed
- qed
- thus ?thesis using assms by auto
-qed
-
-
-subsection{* Default trees *}
-
-(* Pick a left-hand side of a production for each nonterminal *)
-definition S where "S n \<equiv> SOME tns. (n,tns) \<in> P"
-
-lemma S_P: "(n, S n) \<in> P"
-using used unfolding S_def by(rule someI_ex)
-
-lemma finite_S: "finite (S n)"
-using S_P finite_in_P by auto
-
-
-(* The default tree of a nonterminal *)
-definition deftr :: "N \<Rightarrow> Tree" where
-"deftr \<equiv> unfold id S"
-
-lemma deftr_simps[simp]:
-"root (deftr n) = n"
-"cont (deftr n) = image (id \<oplus> deftr) (S n)"
-using unfold(1)[of id S n] unfold(2)[of S n id, OF finite_S]
-unfolding deftr_def by simp_all
-
-lemmas root_deftr = deftr_simps(1)
-lemmas cont_deftr = deftr_simps(2)
-
-lemma root_o_deftr[simp]: "root o deftr = id"
-by (rule ext, auto)
-
-lemma dtree_deftr: "dtree (deftr n)"
-proof-
- {fix tr assume "\<exists> n. tr = deftr n" hence "dtree tr"
- apply(induct rule: dtree_raw_coind) apply safe
- unfolding deftr_simps image_compose[symmetric] sum_map.comp id_o
- root_o_deftr sum_map.id image_id id_apply apply(rule S_P)
- unfolding inj_on_def lift_def by auto
- }
- thus ?thesis by auto
-qed
-
-
-subsection{* Hereditary substitution *}
-
-(* Auxiliary concept: The root-ommiting frontier: *)
-definition "inFrr ns tr t \<equiv> \<exists> tr'. Inr tr' \<in> cont tr \<and> inFr ns tr' t"
-definition "Frr ns tr \<equiv> {t. \<exists> tr'. Inr tr' \<in> cont tr \<and> t \<in> Fr ns tr'}"
-
-context
-fixes tr0 :: Tree
-begin
-
-definition "hsubst_r tr \<equiv> root tr"
-definition "hsubst_c tr \<equiv> if root tr = root tr0 then cont tr0 else cont tr"
-
-(* Hereditary substitution: *)
-definition hsubst :: "Tree \<Rightarrow> Tree" where
-"hsubst \<equiv> unfold hsubst_r hsubst_c"
-
-lemma finite_hsubst_c: "finite (hsubst_c n)"
-unfolding hsubst_c_def by (metis finite_cont)
-
-lemma root_hsubst[simp]: "root (hsubst tr) = root tr"
-using unfold(1)[of hsubst_r hsubst_c tr] unfolding hsubst_def hsubst_r_def by simp
-
-lemma root_o_subst[simp]: "root o hsubst = root"
-unfolding comp_def root_hsubst ..
-
-lemma cont_hsubst_eq[simp]:
-assumes "root tr = root tr0"
-shows "cont (hsubst tr) = (id \<oplus> hsubst) ` (cont tr0)"
-apply(subst id_o[symmetric, of id]) unfolding id_o
-using unfold(2)[of hsubst_c tr hsubst_r, OF finite_hsubst_c]
-unfolding hsubst_def hsubst_c_def using assms by simp
-
-lemma hsubst_eq:
-assumes "root tr = root tr0"
-shows "hsubst tr = hsubst tr0"
-apply(rule Tree_cong) using assms cont_hsubst_eq by auto
-
-lemma cont_hsubst_neq[simp]:
-assumes "root tr \<noteq> root tr0"
-shows "cont (hsubst tr) = (id \<oplus> hsubst) ` (cont tr)"
-apply(subst id_o[symmetric, of id]) unfolding id_o
-using unfold(2)[of hsubst_c tr hsubst_r, OF finite_hsubst_c]
-unfolding hsubst_def hsubst_c_def using assms by simp
-
-lemma Inl_cont_hsubst_eq[simp]:
-assumes "root tr = root tr0"
-shows "Inl -` cont (hsubst tr) = Inl -` (cont tr0)"
-unfolding cont_hsubst_eq[OF assms] by simp
-
-lemma Inr_cont_hsubst_eq[simp]:
-assumes "root tr = root tr0"
-shows "Inr -` cont (hsubst tr) = hsubst ` Inr -` cont tr0"
-unfolding cont_hsubst_eq[OF assms] by simp
-
-lemma Inl_cont_hsubst_neq[simp]:
-assumes "root tr \<noteq> root tr0"
-shows "Inl -` cont (hsubst tr) = Inl -` (cont tr)"
-unfolding cont_hsubst_neq[OF assms] by simp
-
-lemma Inr_cont_hsubst_neq[simp]:
-assumes "root tr \<noteq> root tr0"
-shows "Inr -` cont (hsubst tr) = hsubst ` Inr -` cont tr"
-unfolding cont_hsubst_neq[OF assms] by simp
-
-lemma dtree_hsubst:
-assumes tr0: "dtree tr0" and tr: "dtree tr"
-shows "dtree (hsubst tr)"
-proof-
- {fix tr1 have "(\<exists> tr. dtree tr \<and> tr1 = hsubst tr) \<Longrightarrow> dtree tr1"
- proof (induct rule: dtree_raw_coind)
- case (Hyp tr1) then obtain tr
- where dtr: "dtree tr" and tr1: "tr1 = hsubst tr" by auto
- show ?case unfolding lift_def tr1 proof safe
- show "(root (hsubst tr), prodOf (hsubst tr)) \<in> P"
- unfolding tr1 apply(cases "root tr = root tr0")
- using dtree_P[OF dtr] dtree_P[OF tr0]
- by (auto simp add: image_compose[symmetric] sum_map.comp)
- show "inj_on root (Inr -` cont (hsubst tr))"
- apply(cases "root tr = root tr0") using dtree_inj_on[OF dtr] dtree_inj_on[OF tr0]
- unfolding inj_on_def by (auto, blast)
- fix tr' assume "Inr tr' \<in> cont (hsubst tr)"
- thus "\<exists>tra. dtree tra \<and> tr' = hsubst tra"
- apply(cases "root tr = root tr0", simp_all)
- apply (metis dtree_lift lift_def tr0)
- by (metis dtr dtree_lift lift_def)
- qed
- qed
- }
- thus ?thesis using assms by blast
-qed
-
-lemma Frr: "Frr ns tr = {t. inFrr ns tr t}"
-unfolding inFrr_def Frr_def Fr_def by auto
-
-lemma inFr_hsubst_imp:
-assumes "inFr ns (hsubst tr) t"
-shows "t \<in> Inl -` (cont tr0) \<or> inFrr (ns - {root tr0}) tr0 t \<or>
- inFr (ns - {root tr0}) tr t"
-proof-
- {fix tr1
- have "inFr ns tr1 t \<Longrightarrow>
- (\<And> tr. tr1 = hsubst tr \<Longrightarrow> (t \<in> Inl -` (cont tr0) \<or> inFrr (ns - {root tr0}) tr0 t \<or>
- inFr (ns - {root tr0}) tr t))"
- proof(induct rule: inFr.induct)
- case (Base tr1 ns t tr)
- hence rtr: "root tr1 \<in> ns" and t_tr1: "Inl t \<in> cont tr1" and tr1: "tr1 = hsubst tr"
- by auto
- show ?case
- proof(cases "root tr1 = root tr0")
- case True
- hence "t \<in> Inl -` (cont tr0)" using t_tr1 unfolding tr1 by auto
- thus ?thesis by simp
- next
- case False
- hence "inFr (ns - {root tr0}) tr t" using t_tr1 unfolding tr1 apply simp
- by (metis Base.prems Diff_iff root_hsubst inFr.Base rtr singletonE)
- thus ?thesis by simp
- qed
- next
- case (Ind tr1 ns tr1' t) note IH = Ind(4)
- have rtr1: "root tr1 \<in> ns" and tr1'_tr1: "Inr tr1' \<in> cont tr1"
- and t_tr1': "inFr ns tr1' t" and tr1: "tr1 = hsubst tr" using Ind by auto
- have rtr1: "root tr1 = root tr" unfolding tr1 by simp
- show ?case
- proof(cases "root tr1 = root tr0")
- case True
- then obtain tr' where tr'_tr0: "Inr tr' \<in> cont tr0" and tr1': "tr1' = hsubst tr'"
- using tr1'_tr1 unfolding tr1 by auto
- show ?thesis using IH[OF tr1'] proof (elim disjE)
- assume "inFr (ns - {root tr0}) tr' t"
- thus ?thesis using tr'_tr0 unfolding inFrr_def by auto
- qed auto
- next
- case False
- then obtain tr' where tr'_tr: "Inr tr' \<in> cont tr" and tr1': "tr1' = hsubst tr'"
- using tr1'_tr1 unfolding tr1 by auto
- show ?thesis using IH[OF tr1'] proof (elim disjE)
- assume "inFr (ns - {root tr0}) tr' t"
- thus ?thesis using tr'_tr unfolding inFrr_def
- by (metis Diff_iff False Ind(1) empty_iff inFr2_Ind inFr_inFr2 insert_iff rtr1)
- qed auto
- qed
- qed
- }
- thus ?thesis using assms by auto
-qed
-
-lemma inFr_hsubst_notin:
-assumes "inFr ns tr t" and "root tr0 \<notin> ns"
-shows "inFr ns (hsubst tr) t"
-using assms apply(induct rule: inFr.induct)
-apply (metis Inl_cont_hsubst_neq inFr2.Base inFr_inFr2 root_hsubst vimageD vimageI2)
-by (metis (lifting) Inr_cont_hsubst_neq inFr.Ind rev_image_eqI root_hsubst vimageD vimageI2)
-
-lemma inFr_hsubst_minus:
-assumes "inFr (ns - {root tr0}) tr t"
-shows "inFr ns (hsubst tr) t"
-proof-
- have 1: "inFr (ns - {root tr0}) (hsubst tr) t"
- using inFr_hsubst_notin[OF assms] by simp
- show ?thesis using inFr_mono[OF 1] by auto
-qed
-
-lemma inFr_self_hsubst:
-assumes "root tr0 \<in> ns"
-shows
-"inFr ns (hsubst tr0) t \<longleftrightarrow>
- t \<in> Inl -` (cont tr0) \<or> inFrr (ns - {root tr0}) tr0 t"
-(is "?A \<longleftrightarrow> ?B \<or> ?C")
-apply(intro iffI)
-apply (metis inFr_hsubst_imp Diff_iff inFr_root_in insertI1) proof(elim disjE)
- assume ?B thus ?A apply(intro inFr.Base) using assms by auto
-next
- assume ?C then obtain tr where
- tr_tr0: "Inr tr \<in> cont tr0" and t_tr: "inFr (ns - {root tr0}) tr t"
- unfolding inFrr_def by auto
- def tr1 \<equiv> "hsubst tr"
- have 1: "inFr ns tr1 t" using t_tr unfolding tr1_def using inFr_hsubst_minus by auto
- have "Inr tr1 \<in> cont (hsubst tr0)" unfolding tr1_def using tr_tr0 by auto
- thus ?A using 1 inFr.Ind assms by (metis root_hsubst)
-qed
-
-theorem Fr_self_hsubst:
-assumes "root tr0 \<in> ns"
-shows "Fr ns (hsubst tr0) = Inl -` (cont tr0) \<union> Frr (ns - {root tr0}) tr0"
-using inFr_self_hsubst[OF assms] unfolding Frr Fr_def by auto
-
-end (* context *)
-
-
-subsection{* Regular trees *}
-
-hide_const regular
-
-definition "reg f tr \<equiv> \<forall> tr'. subtr UNIV tr' tr \<longrightarrow> tr' = f (root tr')"
-definition "regular tr \<equiv> \<exists> f. reg f tr"
-
-lemma reg_def2: "reg f tr \<longleftrightarrow> (\<forall> ns tr'. subtr ns tr' tr \<longrightarrow> tr' = f (root tr'))"
-unfolding reg_def using subtr_mono by (metis subset_UNIV)
-
-lemma regular_def2: "regular tr \<longleftrightarrow> (\<exists> f. reg f tr \<and> (\<forall> n. root (f n) = n))"
-unfolding regular_def proof safe
- fix f assume f: "reg f tr"
- def g \<equiv> "\<lambda> n. if inItr UNIV tr n then f n else deftr n"
- show "\<exists>g. reg g tr \<and> (\<forall>n. root (g n) = n)"
- apply(rule exI[of _ g])
- using f deftr_simps(1) unfolding g_def reg_def apply safe
- apply (metis (lifting) inItr.Base subtr_inItr subtr_rootL_in)
- by (metis (full_types) inItr_subtr subtr_subtr2)
-qed auto
-
-lemma reg_root:
-assumes "reg f tr"
-shows "f (root tr) = tr"
-using assms unfolding reg_def
-by (metis (lifting) iso_tuple_UNIV_I subtr.Refl)
-
-
-lemma reg_Inr_cont:
-assumes "reg f tr" and "Inr tr' \<in> cont tr"
-shows "reg f tr'"
-by (metis (lifting) assms iso_tuple_UNIV_I reg_def subtr.Step)
-
-lemma reg_subtr:
-assumes "reg f tr" and "subtr ns tr' tr"
-shows "reg f tr'"
-using assms unfolding reg_def using subtr_trans[of UNIV tr] UNIV_I
-by (metis UNIV_eq_I UnCI Un_upper1 iso_tuple_UNIV_I subtr_mono subtr_trans)
-
-lemma regular_subtr:
-assumes r: "regular tr" and s: "subtr ns tr' tr"
-shows "regular tr'"
-using r reg_subtr[OF _ s] unfolding regular_def by auto
-
-lemma subtr_deftr:
-assumes "subtr ns tr' (deftr n)"
-shows "tr' = deftr (root tr')"
-proof-
- {fix tr have "subtr ns tr' tr \<Longrightarrow> (\<forall> n. tr = deftr n \<longrightarrow> tr' = deftr (root tr'))"
- apply (induct rule: subtr.induct)
- proof(metis (lifting) deftr_simps(1), safe)
- fix tr3 ns tr1 tr2 n
- assume 1: "root (deftr n) \<in> ns" and 2: "subtr ns tr1 tr2"
- and IH: "\<forall>n. tr2 = deftr n \<longrightarrow> tr1 = deftr (root tr1)"
- and 3: "Inr tr2 \<in> cont (deftr n)"
- have "tr2 \<in> deftr ` UNIV"
- using 3 unfolding deftr_simps image_def
- by (metis (lifting, full_types) 3 CollectI Inr_oplus_iff cont_deftr
- iso_tuple_UNIV_I)
- then obtain n where "tr2 = deftr n" by auto
- thus "tr1 = deftr (root tr1)" using IH by auto
- qed
- }
- thus ?thesis using assms by auto
-qed
-
-lemma reg_deftr: "reg deftr (deftr n)"
-unfolding reg_def using subtr_deftr by auto
-
-lemma dtree_subtrOf_Union:
-assumes "dtree tr"
-shows "\<Union>{K tr' |tr'. Inr tr' \<in> cont tr} =
- \<Union>{K (subtrOf tr n) |n. Inr n \<in> prodOf tr}"
-unfolding Union_eq Bex_def mem_Collect_eq proof safe
- fix x xa tr'
- assume x: "x \<in> K tr'" and tr'_tr: "Inr tr' \<in> cont tr"
- show "\<exists>X. (\<exists>n. X = K (subtrOf tr n) \<and> Inr n \<in> prodOf tr) \<and> x \<in> X"
- apply(rule exI[of _ "K (subtrOf tr (root tr'))"]) apply(intro conjI)
- apply(rule exI[of _ "root tr'"]) apply (metis (lifting) root_prodOf tr'_tr)
- by (metis (lifting) assms subtrOf_root tr'_tr x)
-next
- fix x X n ttr
- assume x: "x \<in> K (subtrOf tr n)" and n: "Inr n = (id \<oplus> root) ttr" and ttr: "ttr \<in> cont tr"
- show "\<exists>X. (\<exists>tr'. X = K tr' \<and> Inr tr' \<in> cont tr) \<and> x \<in> X"
- apply(rule exI[of _ "K (subtrOf tr n)"]) apply(intro conjI)
- apply(rule exI[of _ "subtrOf tr n"]) apply (metis imageI n subtrOf ttr)
- using x .
-qed
-
-
-
-
-subsection {* Paths in a regular tree *}
-
-inductive path :: "(N \<Rightarrow> Tree) \<Rightarrow> N list \<Rightarrow> bool" for f where
-Base: "path f [n]"
-|
-Ind: "\<lbrakk>path f (n1 # nl); Inr (f n1) \<in> cont (f n)\<rbrakk>
- \<Longrightarrow> path f (n # n1 # nl)"
-
-lemma path_NE:
-assumes "path f nl"
-shows "nl \<noteq> Nil"
-using assms apply(induct rule: path.induct) by auto
-
-lemma path_post:
-assumes f: "path f (n # nl)" and nl: "nl \<noteq> []"
-shows "path f nl"
-proof-
- obtain n1 nl1 where nl: "nl = n1 # nl1" using nl by (cases nl, auto)
- show ?thesis using assms unfolding nl using path.simps by (metis (lifting) list.inject)
-qed
-
-lemma path_post_concat:
-assumes "path f (nl1 @ nl2)" and "nl2 \<noteq> Nil"
-shows "path f nl2"
-using assms apply (induct nl1)
-apply (metis append_Nil) by (metis Nil_is_append_conv append_Cons path_post)
-
-lemma path_concat:
-assumes "path f nl1" and "path f ((last nl1) # nl2)"
-shows "path f (nl1 @ nl2)"
-using assms apply(induct rule: path.induct) apply simp
-by (metis append_Cons last.simps list.simps(3) path.Ind)
-
-lemma path_distinct:
-assumes "path f nl"
-shows "\<exists> nl'. path f nl' \<and> hd nl' = hd nl \<and> last nl' = last nl \<and>
- set nl' \<subseteq> set nl \<and> distinct nl'"
-using assms proof(induct rule: length_induct)
- case (1 nl) hence p_nl: "path f nl" by simp
- then obtain n nl1 where nl: "nl = n # nl1" by (metis list.exhaust path_NE)
- show ?case
- proof(cases nl1)
- case Nil
- show ?thesis apply(rule exI[of _ nl]) using path.Base unfolding nl Nil by simp
- next
- case (Cons n1 nl2)
- hence p1: "path f nl1" by (metis list.simps nl p_nl path_post)
- show ?thesis
- proof(cases "n \<in> set nl1")
- case False
- obtain nl1' where p1': "path f nl1'" and hd_nl1': "hd nl1' = hd nl1" and
- l_nl1': "last nl1' = last nl1" and d_nl1': "distinct nl1'"
- and s_nl1': "set nl1' \<subseteq> set nl1"
- using 1(1)[THEN allE[of _ nl1]] p1 unfolding nl by auto
- obtain nl2' where nl1': "nl1' = n1 # nl2'" using path_NE[OF p1'] hd_nl1'
- unfolding Cons by(cases nl1', auto)
- show ?thesis apply(intro exI[of _ "n # nl1'"]) unfolding nl proof safe
- show "path f (n # nl1')" unfolding nl1'
- apply(rule path.Ind, metis nl1' p1')
- by (metis (lifting) Cons list.inject nl p1 p_nl path.simps path_NE)
- qed(insert l_nl1' Cons nl1' s_nl1' d_nl1' False, auto)
- next
- case True
- then obtain nl11 nl12 where nl1: "nl1 = nl11 @ n # nl12"
- by (metis split_list)
- have p12: "path f (n # nl12)"
- apply(rule path_post_concat[of _ "n # nl11"]) using p_nl[unfolded nl nl1] by auto
- obtain nl12' where p1': "path f nl12'" and hd_nl12': "hd nl12' = n" and
- l_nl12': "last nl12' = last (n # nl12)" and d_nl12': "distinct nl12'"
- and s_nl12': "set nl12' \<subseteq> {n} \<union> set nl12"
- using 1(1)[THEN allE[of _ "n # nl12"]] p12 unfolding nl nl1 by auto
- thus ?thesis apply(intro exI[of _ nl12']) unfolding nl nl1 by auto
- qed
- qed
-qed
-
-lemma path_subtr:
-assumes f: "\<And> n. root (f n) = n"
-and p: "path f nl"
-shows "subtr (set nl) (f (last nl)) (f (hd nl))"
-using p proof (induct rule: path.induct)
- case (Ind n1 nl n) let ?ns1 = "insert n1 (set nl)"
- have "path f (n1 # nl)"
- and "subtr ?ns1 (f (last (n1 # nl))) (f n1)"
- and fn1: "Inr (f n1) \<in> cont (f n)" using Ind by simp_all
- hence fn1_flast: "subtr (insert n ?ns1) (f (last (n1 # nl))) (f n1)"
- by (metis subset_insertI subtr_mono)
- have 1: "last (n # n1 # nl) = last (n1 # nl)" by auto
- have "subtr (insert n ?ns1) (f (last (n1 # nl))) (f n)"
- using f subtr.Step[OF _ fn1_flast fn1] by auto
- thus ?case unfolding 1 by simp
-qed (metis f hd.simps last_ConsL last_in_set not_Cons_self2 subtr.Refl)
-
-lemma reg_subtr_path_aux:
-assumes f: "reg f tr" and n: "subtr ns tr1 tr"
-shows "\<exists> nl. path f nl \<and> f (hd nl) = tr \<and> f (last nl) = tr1 \<and> set nl \<subseteq> ns"
-using n f proof(induct rule: subtr.induct)
- case (Refl tr ns)
- thus ?case
- apply(intro exI[of _ "[root tr]"]) apply simp by (metis (lifting) path.Base reg_root)
-next
- case (Step tr ns tr2 tr1)
- hence rtr: "root tr \<in> ns" and tr1_tr: "Inr tr1 \<in> cont tr"
- and tr2_tr1: "subtr ns tr2 tr1" and tr: "reg f tr" by auto
- have tr1: "reg f tr1" using reg_subtr[OF tr] rtr tr1_tr
- by (metis (lifting) Step.prems iso_tuple_UNIV_I reg_def subtr.Step)
- obtain nl where nl: "path f nl" and f_nl: "f (hd nl) = tr1"
- and last_nl: "f (last nl) = tr2" and set: "set nl \<subseteq> ns" using Step(3)[OF tr1] by auto
- have 0: "path f (root tr # nl)" apply (subst path.simps)
- using f_nl nl reg_root tr tr1_tr by (metis hd.simps neq_Nil_conv)
- show ?case apply(rule exI[of _ "(root tr) # nl"])
- using 0 reg_root tr last_nl nl path_NE rtr set by auto
-qed
-
-lemma reg_subtr_path:
-assumes f: "reg f tr" and n: "subtr ns tr1 tr"
-shows "\<exists> nl. distinct nl \<and> path f nl \<and> f (hd nl) = tr \<and> f (last nl) = tr1 \<and> set nl \<subseteq> ns"
-using reg_subtr_path_aux[OF assms] path_distinct[of f]
-by (metis (lifting) order_trans)
-
-lemma subtr_iff_path:
-assumes r: "reg f tr" and f: "\<And> n. root (f n) = n"
-shows "subtr ns tr1 tr \<longleftrightarrow>
- (\<exists> nl. distinct nl \<and> path f nl \<and> f (hd nl) = tr \<and> f (last nl) = tr1 \<and> set nl \<subseteq> ns)"
-proof safe
- fix nl assume p: "path f nl" and nl: "set nl \<subseteq> ns"
- have "subtr (set nl) (f (last nl)) (f (hd nl))"
- apply(rule path_subtr) using p f by simp_all
- thus "subtr ns (f (last nl)) (f (hd nl))"
- using subtr_mono nl by auto
-qed(insert reg_subtr_path[OF r], auto)
-
-lemma inFr_iff_path:
-assumes r: "reg f tr" and f: "\<And> n. root (f n) = n"
-shows
-"inFr ns tr t \<longleftrightarrow>
- (\<exists> nl tr1. distinct nl \<and> path f nl \<and> f (hd nl) = tr \<and> f (last nl) = tr1 \<and>
- set nl \<subseteq> ns \<and> Inl t \<in> cont tr1)"
-apply safe
-apply (metis (no_types) inFr_subtr r reg_subtr_path)
-by (metis f inFr.Base path_subtr subtr_inFr subtr_mono subtr_rootL_in)
-
-
-
-subsection{* The regular cut of a tree *}
-
-context fixes tr0 :: Tree
-begin
-
-(* Picking a subtree of a certain root: *)
-definition "pick n \<equiv> SOME tr. subtr UNIV tr tr0 \<and> root tr = n"
-
-lemma pick:
-assumes "inItr UNIV tr0 n"
-shows "subtr UNIV (pick n) tr0 \<and> root (pick n) = n"
-proof-
- have "\<exists> tr. subtr UNIV tr tr0 \<and> root tr = n"
- using assms by (metis (lifting) inItr_subtr)
- thus ?thesis unfolding pick_def by(rule someI_ex)
-qed
-
-lemmas subtr_pick = pick[THEN conjunct1]
-lemmas root_pick = pick[THEN conjunct2]
-
-lemma dtree_pick:
-assumes tr0: "dtree tr0" and n: "inItr UNIV tr0 n"
-shows "dtree (pick n)"
-using dtree_subtr[OF tr0 subtr_pick[OF n]] .
-
-definition "regOf_r n \<equiv> root (pick n)"
-definition "regOf_c n \<equiv> (id \<oplus> root) ` cont (pick n)"
-
-(* The regular tree of a function: *)
-definition regOf :: "N \<Rightarrow> Tree" where
-"regOf \<equiv> unfold regOf_r regOf_c"
-
-lemma finite_regOf_c: "finite (regOf_c n)"
-unfolding regOf_c_def by (metis finite_cont finite_imageI)
-
-lemma root_regOf_pick: "root (regOf n) = root (pick n)"
-using unfold(1)[of regOf_r regOf_c n] unfolding regOf_def regOf_r_def by simp
-
-lemma root_regOf[simp]:
-assumes "inItr UNIV tr0 n"
-shows "root (regOf n) = n"
-unfolding root_regOf_pick root_pick[OF assms] ..
-
-lemma cont_regOf[simp]:
-"cont (regOf n) = (id \<oplus> (regOf o root)) ` cont (pick n)"
-apply(subst id_o[symmetric, of id]) unfolding sum_map.comp[symmetric]
-unfolding image_compose unfolding regOf_c_def[symmetric]
-using unfold(2)[of regOf_c n regOf_r, OF finite_regOf_c]
-unfolding regOf_def ..
-
-lemma Inl_cont_regOf[simp]:
-"Inl -` (cont (regOf n)) = Inl -` (cont (pick n))"
-unfolding cont_regOf by simp
-
-lemma Inr_cont_regOf:
-"Inr -` (cont (regOf n)) = (regOf \<circ> root) ` (Inr -` cont (pick n))"
-unfolding cont_regOf by simp
-
-lemma subtr_regOf:
-assumes n: "inItr UNIV tr0 n" and "subtr UNIV tr1 (regOf n)"
-shows "\<exists> n1. inItr UNIV tr0 n1 \<and> tr1 = regOf n1"
-proof-
- {fix tr ns assume "subtr UNIV tr1 tr"
- hence "tr = regOf n \<longrightarrow> (\<exists> n1. inItr UNIV tr0 n1 \<and> tr1 = regOf n1)"
- proof (induct rule: subtr_UNIV_inductL)
- case (Step tr2 tr1 tr)
- show ?case proof
- assume "tr = regOf n"
- then obtain n1 where tr2: "Inr tr2 \<in> cont tr1"
- and tr1_tr: "subtr UNIV tr1 tr" and n1: "inItr UNIV tr0 n1" and tr1: "tr1 = regOf n1"
- using Step by auto
- obtain tr2' where tr2: "tr2 = regOf (root tr2')"
- and tr2': "Inr tr2' \<in> cont (pick n1)"
- using tr2 Inr_cont_regOf[of n1]
- unfolding tr1 image_def o_def using vimage_eq by auto
- have "inItr UNIV tr0 (root tr2')"
- using inItr.Base inItr.Ind n1 pick subtr_inItr tr2' by (metis iso_tuple_UNIV_I)
- thus "\<exists>n2. inItr UNIV tr0 n2 \<and> tr2 = regOf n2" using tr2 by blast
- qed
- qed(insert n, auto)
- }
- thus ?thesis using assms by auto
-qed
-
-lemma root_regOf_root:
-assumes n: "inItr UNIV tr0 n" and t_tr: "t_tr \<in> cont (pick n)"
-shows "(id \<oplus> (root \<circ> regOf \<circ> root)) t_tr = (id \<oplus> root) t_tr"
-using assms apply(cases t_tr)
- apply (metis (lifting) sum_map.simps(1))
- using pick regOf_def regOf_r_def unfold(1)
- inItr.Base o_apply subtr_StepL subtr_inItr sum_map.simps(2)
- by (metis UNIV_I)
-
-lemma regOf_P:
-assumes tr0: "dtree tr0" and n: "inItr UNIV tr0 n"
-shows "(n, (id \<oplus> root) ` cont (regOf n)) \<in> P" (is "?L \<in> P")
-proof-
- have "?L = (n, (id \<oplus> root) ` cont (pick n))"
- unfolding cont_regOf image_compose[symmetric] sum_map.comp id_o o_assoc
- unfolding Pair_eq apply(rule conjI[OF refl]) apply(rule image_cong[OF refl])
- by(rule root_regOf_root[OF n])
- moreover have "... \<in> P" by (metis (lifting) dtree_pick root_pick dtree_P n tr0)
- ultimately show ?thesis by simp
-qed
-
-lemma dtree_regOf:
-assumes tr0: "dtree tr0" and "inItr UNIV tr0 n"
-shows "dtree (regOf n)"
-proof-
- {fix tr have "\<exists> n. inItr UNIV tr0 n \<and> tr = regOf n \<Longrightarrow> dtree tr"
- proof (induct rule: dtree_raw_coind)
- case (Hyp tr)
- then obtain n where n: "inItr UNIV tr0 n" and tr: "tr = regOf n" by auto
- show ?case unfolding lift_def apply safe
- apply (metis (lifting) regOf_P root_regOf n tr tr0)
- unfolding tr Inr_cont_regOf unfolding inj_on_def apply clarsimp using root_regOf
- apply (metis UNIV_I inItr.Base n pick subtr2.simps subtr_inItr subtr_subtr2)
- by (metis n subtr.Refl subtr_StepL subtr_regOf tr UNIV_I)
- qed
- }
- thus ?thesis using assms by blast
-qed
-
-(* The regular cut of a tree: *)
-definition "rcut \<equiv> regOf (root tr0)"
-
-theorem reg_rcut: "reg regOf rcut"
-unfolding reg_def rcut_def
-by (metis inItr.Base root_regOf subtr_regOf UNIV_I)
-
-lemma rcut_reg:
-assumes "reg regOf tr0"
-shows "rcut = tr0"
-using assms unfolding rcut_def reg_def by (metis subtr.Refl UNIV_I)
-
-theorem rcut_eq: "rcut = tr0 \<longleftrightarrow> reg regOf tr0"
-using reg_rcut rcut_reg by metis
-
-theorem regular_rcut: "regular rcut"
-using reg_rcut unfolding regular_def by blast
-
-theorem Fr_rcut: "Fr UNIV rcut \<subseteq> Fr UNIV tr0"
-proof safe
- fix t assume "t \<in> Fr UNIV rcut"
- then obtain tr where t: "Inl t \<in> cont tr" and tr: "subtr UNIV tr (regOf (root tr0))"
- using Fr_subtr[of UNIV "regOf (root tr0)"] unfolding rcut_def
- by (metis (full_types) Fr_def inFr_subtr mem_Collect_eq)
- obtain n where n: "inItr UNIV tr0 n" and tr: "tr = regOf n" using tr
- by (metis (lifting) inItr.Base subtr_regOf UNIV_I)
- have "Inl t \<in> cont (pick n)" using t using Inl_cont_regOf[of n] unfolding tr
- by (metis (lifting) vimageD vimageI2)
- moreover have "subtr UNIV (pick n) tr0" using subtr_pick[OF n] ..
- ultimately show "t \<in> Fr UNIV tr0" unfolding Fr_subtr_cont by auto
-qed
-
-theorem dtree_rcut:
-assumes "dtree tr0"
-shows "dtree rcut"
-unfolding rcut_def using dtree_regOf[OF assms inItr.Base] by simp
-
-theorem root_rcut[simp]: "root rcut = root tr0"
-unfolding rcut_def
-by (metis (lifting) root_regOf inItr.Base reg_def reg_root subtr_rootR_in)
-
-end (* context *)
-
-
-subsection{* Recursive description of the regular tree frontiers *}
-
-lemma regular_inFr:
-assumes r: "regular tr" and In: "root tr \<in> ns"
-and t: "inFr ns tr t"
-shows "t \<in> Inl -` (cont tr) \<or>
- (\<exists> tr'. Inr tr' \<in> cont tr \<and> inFr (ns - {root tr}) tr' t)"
-(is "?L \<or> ?R")
-proof-
- obtain f where r: "reg f tr" and f: "\<And>n. root (f n) = n"
- using r unfolding regular_def2 by auto
- obtain nl tr1 where d_nl: "distinct nl" and p: "path f nl" and hd_nl: "f (hd nl) = tr"
- and l_nl: "f (last nl) = tr1" and s_nl: "set nl \<subseteq> ns" and t_tr1: "Inl t \<in> cont tr1"
- using t unfolding inFr_iff_path[OF r f] by auto
- obtain n nl1 where nl: "nl = n # nl1" by (metis (lifting) p path.simps)
- hence f_n: "f n = tr" using hd_nl by simp
- have n_nl1: "n \<notin> set nl1" using d_nl unfolding nl by auto
- show ?thesis
- proof(cases nl1)
- case Nil hence "tr = tr1" using f_n l_nl unfolding nl by simp
- hence ?L using t_tr1 by simp thus ?thesis by simp
- next
- case (Cons n1 nl2) note nl1 = Cons
- have 1: "last nl1 = last nl" "hd nl1 = n1" unfolding nl nl1 by simp_all
- have p1: "path f nl1" and n1_tr: "Inr (f n1) \<in> cont tr"
- using path.simps[of f nl] p f_n unfolding nl nl1 by auto
- have r1: "reg f (f n1)" using reg_Inr_cont[OF r n1_tr] .
- have 0: "inFr (set nl1) (f n1) t" unfolding inFr_iff_path[OF r1 f]
- apply(intro exI[of _ nl1], intro exI[of _ tr1])
- using d_nl unfolding 1 l_nl unfolding nl using p1 t_tr1 by auto
- have root_tr: "root tr = n" by (metis f f_n)
- have "inFr (ns - {root tr}) (f n1) t" apply(rule inFr_mono[OF 0])
- using s_nl unfolding root_tr unfolding nl using n_nl1 by auto
- thus ?thesis using n1_tr by auto
- qed
-qed
-
-theorem regular_Fr:
-assumes r: "regular tr" and In: "root tr \<in> ns"
-shows "Fr ns tr =
- Inl -` (cont tr) \<union>
- \<Union> {Fr (ns - {root tr}) tr' | tr'. Inr tr' \<in> cont tr}"
-unfolding Fr_def
-using In inFr.Base regular_inFr[OF assms] apply safe
-apply (simp, metis (full_types) UnionI mem_Collect_eq)
-apply simp
-by (simp, metis (lifting) inFr_Ind_minus insert_Diff)
-
-
-subsection{* The generated languages *}
-
-(* The (possibly inifinite tree) generated language *)
-definition "L ns n \<equiv> {Fr ns tr | tr. dtree tr \<and> root tr = n}"
-
-(* The regular-tree generated language *)
-definition "Lr ns n \<equiv> {Fr ns tr | tr. dtree tr \<and> root tr = n \<and> regular tr}"
-
-theorem L_rec_notin:
-assumes "n \<notin> ns"
-shows "L ns n = {{}}"
-using assms unfolding L_def apply safe
- using not_root_Fr apply force
- apply(rule exI[of _ "deftr n"])
- by (metis (no_types) dtree_deftr not_root_Fr root_deftr)
-
-theorem Lr_rec_notin:
-assumes "n \<notin> ns"
-shows "Lr ns n = {{}}"
-using assms unfolding Lr_def apply safe
- using not_root_Fr apply force
- apply(rule exI[of _ "deftr n"])
- by (metis (no_types) regular_def dtree_deftr not_root_Fr reg_deftr root_deftr)
-
-lemma dtree_subtrOf:
-assumes "dtree tr" and "Inr n \<in> prodOf tr"
-shows "dtree (subtrOf tr n)"
-by (metis assms dtree_lift lift_def subtrOf)
-
-theorem Lr_rec_in:
-assumes n: "n \<in> ns"
-shows "Lr ns n \<subseteq>
-{Inl -` tns \<union> (\<Union> {K n' | n'. Inr n' \<in> tns}) | tns K.
- (n,tns) \<in> P \<and>
- (\<forall> n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> Lr (ns - {n}) n')}"
-(is "Lr ns n \<subseteq> {?F tns K | tns K. (n,tns) \<in> P \<and> ?\<phi> tns K}")
-proof safe
- fix ts assume "ts \<in> Lr ns n"
- then obtain tr where dtr: "dtree tr" and r: "root tr = n" and tr: "regular tr"
- and ts: "ts = Fr ns tr" unfolding Lr_def by auto
- def tns \<equiv> "(id \<oplus> root) ` (cont tr)"
- def K \<equiv> "\<lambda> n'. Fr (ns - {n}) (subtrOf tr n')"
- show "\<exists>tns K. ts = ?F tns K \<and> (n, tns) \<in> P \<and> ?\<phi> tns K"
- apply(rule exI[of _ tns], rule exI[of _ K]) proof(intro conjI allI impI)
- show "ts = Inl -` tns \<union> \<Union>{K n' |n'. Inr n' \<in> tns}"
- unfolding ts regular_Fr[OF tr n[unfolded r[symmetric]]]
- unfolding tns_def K_def r[symmetric]
- unfolding Inl_prodOf dtree_subtrOf_Union[OF dtr] ..
- show "(n, tns) \<in> P" unfolding tns_def r[symmetric] using dtree_P[OF dtr] .
- fix n' assume "Inr n' \<in> tns" thus "K n' \<in> Lr (ns - {n}) n'"
- unfolding K_def Lr_def mem_Collect_eq apply(intro exI[of _ "subtrOf tr n'"])
- using dtr tr apply(intro conjI refl) unfolding tns_def
- apply(erule dtree_subtrOf[OF dtr])
- apply (metis subtrOf)
- by (metis Inr_subtrOf UNIV_I regular_subtr subtr.simps)
- qed
-qed
-
-lemma hsubst_aux:
-fixes n ftr tns
-assumes n: "n \<in> ns" and tns: "finite tns" and
-1: "\<And> n'. Inr n' \<in> tns \<Longrightarrow> dtree (ftr n')"
-defines "tr \<equiv> Node n ((id \<oplus> ftr) ` tns)" defines "tr' \<equiv> hsubst tr tr"
-shows "Fr ns tr' = Inl -` tns \<union> \<Union>{Fr (ns - {n}) (ftr n') |n'. Inr n' \<in> tns}"
-(is "_ = ?B") proof-
- have rtr: "root tr = n" and ctr: "cont tr = (id \<oplus> ftr) ` tns"
- unfolding tr_def using tns by auto
- have Frr: "Frr (ns - {n}) tr = \<Union>{Fr (ns - {n}) (ftr n') |n'. Inr n' \<in> tns}"
- unfolding Frr_def ctr by auto
- have "Fr ns tr' = Inl -` (cont tr) \<union> Frr (ns - {n}) tr"
- using Fr_self_hsubst[OF n[unfolded rtr[symmetric]]] unfolding tr'_def rtr ..
- also have "... = ?B" unfolding ctr Frr by simp
- finally show ?thesis .
-qed
-
-theorem L_rec_in:
-assumes n: "n \<in> ns"
-shows "
-{Inl -` tns \<union> (\<Union> {K n' | n'. Inr n' \<in> tns}) | tns K.
- (n,tns) \<in> P \<and>
- (\<forall> n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> L (ns - {n}) n')}
- \<subseteq> L ns n"
-proof safe
- fix tns K
- assume P: "(n, tns) \<in> P" and 0: "\<forall>n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> L (ns - {n}) n'"
- {fix n' assume "Inr n' \<in> tns"
- hence "K n' \<in> L (ns - {n}) n'" using 0 by auto
- hence "\<exists> tr'. K n' = Fr (ns - {n}) tr' \<and> dtree tr' \<and> root tr' = n'"
- unfolding L_def mem_Collect_eq by auto
- }
- then obtain ftr where 0: "\<And> n'. Inr n' \<in> tns \<Longrightarrow>
- K n' = Fr (ns - {n}) (ftr n') \<and> dtree (ftr n') \<and> root (ftr n') = n'"
- by metis
- def tr \<equiv> "Node n ((id \<oplus> ftr) ` tns)" def tr' \<equiv> "hsubst tr tr"
- have rtr: "root tr = n" and ctr: "cont tr = (id \<oplus> ftr) ` tns"
- unfolding tr_def by (simp, metis P cont_Node finite_imageI finite_in_P)
- have prtr: "prodOf tr = tns" apply(rule Inl_Inr_image_cong)
- unfolding ctr apply simp apply simp apply safe
- using 0 unfolding image_def apply force apply simp by (metis 0 vimageI2)
- have 1: "{K n' |n'. Inr n' \<in> tns} = {Fr (ns - {n}) (ftr n') |n'. Inr n' \<in> tns}"
- using 0 by auto
- have dtr: "dtree tr" apply(rule dtree.Tree)
- apply (metis (lifting) P prtr rtr)
- unfolding inj_on_def ctr lift_def using 0 by auto
- hence dtr': "dtree tr'" unfolding tr'_def by (metis dtree_hsubst)
- have tns: "finite tns" using finite_in_P P by simp
- have "Inl -` tns \<union> \<Union>{Fr (ns - {n}) (ftr n') |n'. Inr n' \<in> tns} \<in> L ns n"
- unfolding L_def mem_Collect_eq apply(intro exI[of _ tr'] conjI)
- using dtr' 0 hsubst_aux[OF assms tns, of ftr] unfolding tr_def tr'_def by auto
- thus "Inl -` tns \<union> \<Union>{K n' |n'. Inr n' \<in> tns} \<in> L ns n" unfolding 1 .
-qed
-
-lemma card_N: "(n::N) \<in> ns \<Longrightarrow> card (ns - {n}) < card ns"
-by (metis finite_N Diff_UNIV Diff_infinite_finite card_Diff1_less finite.emptyI)
-
-function LL where
-"LL ns n =
- (if n \<notin> ns then {{}} else
- {Inl -` tns \<union> (\<Union> {K n' | n'. Inr n' \<in> tns}) | tns K.
- (n,tns) \<in> P \<and>
- (\<forall> n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> LL (ns - {n}) n')})"
-by(pat_completeness, auto)
-termination apply(relation "inv_image (measure card) fst")
-using card_N by auto
-
-declare LL.simps[code] (* TODO: Does code generation for LL work? *)
-declare LL.simps[simp del]
-
-theorem Lr_LL: "Lr ns n \<subseteq> LL ns n"
-proof (induct ns arbitrary: n rule: measure_induct[of card])
- case (1 ns n) show ?case proof(cases "n \<in> ns")
- case False thus ?thesis unfolding Lr_rec_notin[OF False] by (simp add: LL.simps)
- next
- case True show ?thesis apply(rule subset_trans)
- using Lr_rec_in[OF True] apply assumption
- unfolding LL.simps[of ns n] using True 1 card_N proof clarsimp
- fix tns K
- assume "n \<in> ns" hence c: "card (ns - {n}) < card ns" using card_N by blast
- assume "(n, tns) \<in> P"
- and "\<forall>n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> Lr (ns - {n}) n'"
- thus "\<exists>tnsa Ka.
- Inl -` tns \<union> \<Union>{K n' |n'. Inr n' \<in> tns} =
- Inl -` tnsa \<union> \<Union>{Ka n' |n'. Inr n' \<in> tnsa} \<and>
- (n, tnsa) \<in> P \<and> (\<forall>n'. Inr n' \<in> tnsa \<longrightarrow> Ka n' \<in> LL (ns - {n}) n')"
- apply(intro exI[of _ tns] exI[of _ K]) using c 1 by auto
- qed
- qed
-qed
-
-theorem LL_L: "LL ns n \<subseteq> L ns n"
-proof (induct ns arbitrary: n rule: measure_induct[of card])
- case (1 ns n) show ?case proof(cases "n \<in> ns")
- case False thus ?thesis unfolding L_rec_notin[OF False] by (simp add: LL.simps)
- next
- case True show ?thesis apply(rule subset_trans)
- prefer 2 using L_rec_in[OF True] apply assumption
- unfolding LL.simps[of ns n] using True 1 card_N proof clarsimp
- fix tns K
- assume "n \<in> ns" hence c: "card (ns - {n}) < card ns" using card_N by blast
- assume "(n, tns) \<in> P"
- and "\<forall>n'. Inr n' \<in> tns \<longrightarrow> K n' \<in> LL (ns - {n}) n'"
- thus "\<exists>tnsa Ka.
- Inl -` tns \<union> \<Union>{K n' |n'. Inr n' \<in> tns} =
- Inl -` tnsa \<union> \<Union>{Ka n' |n'. Inr n' \<in> tnsa} \<and>
- (n, tnsa) \<in> P \<and> (\<forall>n'. Inr n' \<in> tnsa \<longrightarrow> Ka n' \<in> L (ns - {n}) n')"
- apply(intro exI[of _ tns] exI[of _ K]) using c 1 by auto
- qed
- qed
-qed
-
-(* The subsumpsion relation between languages *)
-definition "subs L1 L2 \<equiv> \<forall> ts2 \<in> L2. \<exists> ts1 \<in> L1. ts1 \<subseteq> ts2"
-
-lemma incl_subs[simp]: "L2 \<subseteq> L1 \<Longrightarrow> subs L1 L2"
-unfolding subs_def by auto
-
-lemma subs_refl[simp]: "subs L1 L1" unfolding subs_def by auto
-
-lemma subs_trans: "\<lbrakk>subs L1 L2; subs L2 L3\<rbrakk> \<Longrightarrow> subs L1 L3"
-unfolding subs_def by (metis subset_trans)
-
-(* Language equivalence *)
-definition "leqv L1 L2 \<equiv> subs L1 L2 \<and> subs L2 L1"
-
-lemma subs_leqv[simp]: "leqv L1 L2 \<Longrightarrow> subs L1 L2"
-unfolding leqv_def by auto
-
-lemma subs_leqv_sym[simp]: "leqv L1 L2 \<Longrightarrow> subs L2 L1"
-unfolding leqv_def by auto
-
-lemma leqv_refl[simp]: "leqv L1 L1" unfolding leqv_def by auto
-
-lemma leqv_trans:
-assumes 12: "leqv L1 L2" and 23: "leqv L2 L3"
-shows "leqv L1 L3"
-using assms unfolding leqv_def by (metis (lifting) subs_trans)
-
-lemma leqv_sym: "leqv L1 L2 \<Longrightarrow> leqv L2 L1"
-unfolding leqv_def by auto
-
-lemma leqv_Sym: "leqv L1 L2 \<longleftrightarrow> leqv L2 L1"
-unfolding leqv_def by auto
-
-lemma Lr_incl_L: "Lr ns ts \<subseteq> L ns ts"
-unfolding Lr_def L_def by auto
-
-lemma Lr_subs_L: "subs (Lr UNIV ts) (L UNIV ts)"
-unfolding subs_def proof safe
- fix ts2 assume "ts2 \<in> L UNIV ts"
- then obtain tr where ts2: "ts2 = Fr UNIV tr" and dtr: "dtree tr" and rtr: "root tr = ts"
- unfolding L_def by auto
- thus "\<exists>ts1\<in>Lr UNIV ts. ts1 \<subseteq> ts2"
- apply(intro bexI[of _ "Fr UNIV (rcut tr)"])
- unfolding Lr_def L_def using Fr_rcut dtree_rcut root_rcut regular_rcut by auto
-qed
-
-theorem Lr_leqv_L: "leqv (Lr UNIV ts) (L UNIV ts)"
-using Lr_subs_L unfolding leqv_def by (metis (lifting) Lr_incl_L incl_subs)
-
-theorem LL_leqv_L: "leqv (LL UNIV ts) (L UNIV ts)"
-by (metis (lifting) LL_L Lr_LL Lr_subs_L incl_subs leqv_def subs_trans)
-
-theorem LL_leqv_Lr: "leqv (LL UNIV ts) (Lr UNIV ts)"
-using Lr_leqv_L LL_leqv_L by (metis leqv_Sym leqv_trans)
-
-
-end
--- a/src/HOL/Codatatype/Examples/Infinite_Derivation_Trees/Parallel.thy Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,152 +0,0 @@
-(* Title: HOL/BNF/Examples/Infinite_Derivation_Trees/Parallel.thy
- Author: Andrei Popescu, TU Muenchen
- Copyright 2012
-
-Parallel composition.
-*)
-
-header {* Parallel Composition *}
-
-theory Parallel
-imports Tree
-begin
-
-
-consts Nplus :: "N \<Rightarrow> N \<Rightarrow> N" (infixl "+" 60)
-
-axiomatization where
- Nplus_comm: "(a::N) + b = b + (a::N)"
-and Nplus_assoc: "((a::N) + b) + c = a + (b + c)"
-
-
-
-section{* Parallel composition *}
-
-fun par_r where "par_r (tr1,tr2) = root tr1 + root tr2"
-fun par_c where
-"par_c (tr1,tr2) =
- Inl ` (Inl -` (cont tr1 \<union> cont tr2)) \<union>
- Inr ` (Inr -` cont tr1 \<times> Inr -` cont tr2)"
-
-declare par_r.simps[simp del] declare par_c.simps[simp del]
-
-definition par :: "Tree \<times> Tree \<Rightarrow> Tree" where
-"par \<equiv> unfold par_r par_c"
-
-abbreviation par_abbr (infixr "\<parallel>" 80) where "tr1 \<parallel> tr2 \<equiv> par (tr1, tr2)"
-
-lemma finite_par_c: "finite (par_c (tr1, tr2))"
-unfolding par_c.simps apply(rule finite_UnI)
- apply (metis finite_Un finite_cont finite_imageI finite_vimageI inj_Inl)
- apply(intro finite_imageI finite_cartesian_product finite_vimageI)
- using finite_cont by auto
-
-lemma root_par: "root (tr1 \<parallel> tr2) = root tr1 + root tr2"
-using unfold(1)[of par_r par_c "(tr1,tr2)"] unfolding par_def par_r.simps by simp
-
-lemma cont_par:
-"cont (tr1 \<parallel> tr2) = (id \<oplus> par) ` par_c (tr1,tr2)"
-using unfold(2)[of par_c "(tr1,tr2)" par_r, OF finite_par_c]
-unfolding par_def ..
-
-lemma Inl_cont_par[simp]:
-"Inl -` (cont (tr1 \<parallel> tr2)) = Inl -` (cont tr1 \<union> cont tr2)"
-unfolding cont_par par_c.simps by auto
-
-lemma Inr_cont_par[simp]:
-"Inr -` (cont (tr1 \<parallel> tr2)) = par ` (Inr -` cont tr1 \<times> Inr -` cont tr2)"
-unfolding cont_par par_c.simps by auto
-
-lemma Inl_in_cont_par:
-"Inl t \<in> cont (tr1 \<parallel> tr2) \<longleftrightarrow> (Inl t \<in> cont tr1 \<or> Inl t \<in> cont tr2)"
-using Inl_cont_par[of tr1 tr2] unfolding vimage_def by auto
-
-lemma Inr_in_cont_par:
-"Inr t \<in> cont (tr1 \<parallel> tr2) \<longleftrightarrow> (t \<in> par ` (Inr -` cont tr1 \<times> Inr -` cont tr2))"
-using Inr_cont_par[of tr1 tr2] unfolding vimage_def by auto
-
-
-section{* =-coinductive proofs *}
-
-(* Detailed proofs of commutativity and associativity: *)
-theorem par_com: "tr1 \<parallel> tr2 = tr2 \<parallel> tr1"
-proof-
- let ?\<phi> = "\<lambda> trA trB. \<exists> tr1 tr2. trA = tr1 \<parallel> tr2 \<and> trB = tr2 \<parallel> tr1"
- {fix trA trB
- assume "?\<phi> trA trB" hence "trA = trB"
- proof (induct rule: Tree_coind, safe)
- fix tr1 tr2
- show "root (tr1 \<parallel> tr2) = root (tr2 \<parallel> tr1)"
- unfolding root_par by (rule Nplus_comm)
- next
- fix tr1 tr2 :: Tree
- let ?trA = "tr1 \<parallel> tr2" let ?trB = "tr2 \<parallel> tr1"
- show "(?\<phi> ^#2) (cont ?trA) (cont ?trB)"
- unfolding lift2_def proof(intro conjI allI impI)
- fix n show "Inl n \<in> cont (tr1 \<parallel> tr2) \<longleftrightarrow> Inl n \<in> cont (tr2 \<parallel> tr1)"
- unfolding Inl_in_cont_par by auto
- next
- fix trA' assume "Inr trA' \<in> cont ?trA"
- then obtain tr1' tr2' where "trA' = tr1' \<parallel> tr2'"
- and "Inr tr1' \<in> cont tr1" and "Inr tr2' \<in> cont tr2"
- unfolding Inr_in_cont_par by auto
- thus "\<exists> trB'. Inr trB' \<in> cont ?trB \<and> ?\<phi> trA' trB'"
- apply(intro exI[of _ "tr2' \<parallel> tr1'"]) unfolding Inr_in_cont_par by auto
- next
- fix trB' assume "Inr trB' \<in> cont ?trB"
- then obtain tr1' tr2' where "trB' = tr2' \<parallel> tr1'"
- and "Inr tr1' \<in> cont tr1" and "Inr tr2' \<in> cont tr2"
- unfolding Inr_in_cont_par by auto
- thus "\<exists> trA'. Inr trA' \<in> cont ?trA \<and> ?\<phi> trA' trB'"
- apply(intro exI[of _ "tr1' \<parallel> tr2'"]) unfolding Inr_in_cont_par by auto
- qed
- qed
- }
- thus ?thesis by blast
-qed
-
-theorem par_assoc: "(tr1 \<parallel> tr2) \<parallel> tr3 = tr1 \<parallel> (tr2 \<parallel> tr3)"
-proof-
- let ?\<phi> =
- "\<lambda> trA trB. \<exists> tr1 tr2 tr3. trA = (tr1 \<parallel> tr2) \<parallel> tr3 \<and>
- trB = tr1 \<parallel> (tr2 \<parallel> tr3)"
- {fix trA trB
- assume "?\<phi> trA trB" hence "trA = trB"
- proof (induct rule: Tree_coind, safe)
- fix tr1 tr2 tr3
- show "root ((tr1 \<parallel> tr2) \<parallel> tr3) = root (tr1 \<parallel> (tr2 \<parallel> tr3))"
- unfolding root_par by (rule Nplus_assoc)
- next
- fix tr1 tr2 tr3
- let ?trA = "(tr1 \<parallel> tr2) \<parallel> tr3" let ?trB = "tr1 \<parallel> (tr2 \<parallel> tr3)"
- show "(?\<phi> ^#2) (cont ?trA) (cont ?trB)"
- unfolding lift2_def proof(intro conjI allI impI)
- fix n show "Inl n \<in> (cont ?trA) \<longleftrightarrow> Inl n \<in> (cont ?trB)"
- unfolding Inl_in_cont_par by simp
- next
- fix trA' assume "Inr trA' \<in> cont ?trA"
- then obtain tr1' tr2' tr3' where "trA' = (tr1' \<parallel> tr2') \<parallel> tr3'"
- and "Inr tr1' \<in> cont tr1" and "Inr tr2' \<in> cont tr2"
- and "Inr tr3' \<in> cont tr3" unfolding Inr_in_cont_par by auto
- thus "\<exists> trB'. Inr trB' \<in> cont ?trB \<and> ?\<phi> trA' trB'"
- apply(intro exI[of _ "tr1' \<parallel> (tr2' \<parallel> tr3')"])
- unfolding Inr_in_cont_par by auto
- next
- fix trB' assume "Inr trB' \<in> cont ?trB"
- then obtain tr1' tr2' tr3' where "trB' = tr1' \<parallel> (tr2' \<parallel> tr3')"
- and "Inr tr1' \<in> cont tr1" and "Inr tr2' \<in> cont tr2"
- and "Inr tr3' \<in> cont tr3" unfolding Inr_in_cont_par by auto
- thus "\<exists> trA'. Inr trA' \<in> cont ?trA \<and> ?\<phi> trA' trB'"
- apply(intro exI[of _ "(tr1' \<parallel> tr2') \<parallel> tr3'"])
- unfolding Inr_in_cont_par by auto
- qed
- qed
- }
- thus ?thesis by blast
-qed
-
-
-
-
-
-end
\ No newline at end of file
--- a/src/HOL/Codatatype/Examples/Infinite_Derivation_Trees/Prelim.thy Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,67 +0,0 @@
-(* Title: HOL/BNF/Examples/Infinite_Derivation_Trees/Prelim.thy
- Author: Andrei Popescu, TU Muenchen
- Copyright 2012
-
-Preliminaries.
-*)
-
-header {* Preliminaries *}
-
-theory Prelim
-imports "../../BNF"
-begin
-
-declare fset_to_fset[simp]
-
-lemma fst_snd_convol_o[simp]: "<fst o s, snd o s> = s"
-apply(rule ext) by (simp add: convol_def)
-
-abbreviation sm_abbrev (infix "\<oplus>" 60)
-where "f \<oplus> g \<equiv> Sum_Type.sum_map f g"
-
-lemma sum_map_InlD: "(f \<oplus> g) z = Inl x \<Longrightarrow> \<exists>y. z = Inl y \<and> f y = x"
-by (cases z) auto
-
-lemma sum_map_InrD: "(f \<oplus> g) z = Inr x \<Longrightarrow> \<exists>y. z = Inr y \<and> g y = x"
-by (cases z) auto
-
-abbreviation sum_case_abbrev ("[[_,_]]" 800)
-where "[[f,g]] \<equiv> Sum_Type.sum_case f g"
-
-lemma inj_Inl[simp]: "inj Inl" unfolding inj_on_def by auto
-lemma inj_Inr[simp]: "inj Inr" unfolding inj_on_def by auto
-
-lemma Inl_oplus_elim:
-assumes "Inl tr \<in> (id \<oplus> f) ` tns"
-shows "Inl tr \<in> tns"
-using assms apply clarify by (case_tac x, auto)
-
-lemma Inl_oplus_iff[simp]: "Inl tr \<in> (id \<oplus> f) ` tns \<longleftrightarrow> Inl tr \<in> tns"
-using Inl_oplus_elim
-by (metis id_def image_iff sum_map.simps(1))
-
-lemma Inl_m_oplus[simp]: "Inl -` (id \<oplus> f) ` tns = Inl -` tns"
-using Inl_oplus_iff unfolding vimage_def by auto
-
-lemma Inr_oplus_elim:
-assumes "Inr tr \<in> (id \<oplus> f) ` tns"
-shows "\<exists> n. Inr n \<in> tns \<and> f n = tr"
-using assms apply clarify by (case_tac x, auto)
-
-lemma Inr_oplus_iff[simp]:
-"Inr tr \<in> (id \<oplus> f) ` tns \<longleftrightarrow> (\<exists> n. Inr n \<in> tns \<and> f n = tr)"
-apply (rule iffI)
- apply (metis Inr_oplus_elim)
-by (metis image_iff sum_map.simps(2))
-
-lemma Inr_m_oplus[simp]: "Inr -` (id \<oplus> f) ` tns = f ` (Inr -` tns)"
-using Inr_oplus_iff unfolding vimage_def by auto
-
-lemma Inl_Inr_image_cong:
-assumes "Inl -` A = Inl -` B" and "Inr -` A = Inr -` B"
-shows "A = B"
-apply safe using assms apply(case_tac x, auto) by(case_tac x, auto)
-
-
-
-end
\ No newline at end of file
--- a/src/HOL/Codatatype/Examples/Infinite_Derivation_Trees/Tree.thy Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,326 +0,0 @@
-(* Title: HOL/BNF/Examples/Infinite_Derivation_Trees/Tree.thy
- Author: Andrei Popescu, TU Muenchen
- Copyright 2012
-
-Trees with nonterminal internal nodes and terminal leaves.
-*)
-
-header {* Trees with Nonterminal Internal Nodes and Terminal Leaves *}
-
-theory Tree
-imports Prelim
-begin
-
-hide_fact (open) Quotient_Product.prod_rel_def
-
-typedecl N typedecl T
-
-codata_raw Tree: 'Tree = "N \<times> (T + 'Tree) fset"
-
-
-section {* Sugar notations for Tree *}
-
-subsection{* Setup for map, set, rel *}
-
-(* These should be eventually inferred from compositionality *)
-
-lemma pre_Tree_map:
-"pre_Tree_map f (n, as) = (n, map_fset (id \<oplus> f) as)"
-unfolding pre_Tree_map_def id_apply
-sum_map_def by simp
-
-lemma pre_Tree_map':
-"pre_Tree_map f n_as = (fst n_as, map_fset (id \<oplus> f) (snd n_as))"
-using pre_Tree_map by(cases n_as, simp)
-
-
-definition
-"llift2 \<phi> as1 as2 \<longleftrightarrow>
- (\<forall> n. Inl n \<in> fset as1 \<longleftrightarrow> Inl n \<in> fset as2) \<and>
- (\<forall> tr1. Inr tr1 \<in> fset as1 \<longrightarrow> (\<exists> tr2. Inr tr2 \<in> fset as2 \<and> \<phi> tr1 tr2)) \<and>
- (\<forall> tr2. Inr tr2 \<in> fset as2 \<longrightarrow> (\<exists> tr1. Inr tr1 \<in> fset as1 \<and> \<phi> tr1 tr2))"
-
-lemma pre_Tree_rel: "pre_Tree_rel \<phi> (n1,as1) (n2,as2) \<longleftrightarrow> n1 = n2 \<and> llift2 \<phi> as1 as2"
-unfolding llift2_def pre_Tree_rel_def sum_rel_def[abs_def] prod_rel_def fset_rel_def split_conv
-apply (auto split: sum.splits)
-apply (metis sumE)
-apply (metis sumE)
-apply (metis sumE)
-apply (metis sumE)
-apply (metis sumE sum.simps(1,2,4))
-apply (metis sumE sum.simps(1,2,4))
-done
-
-
-subsection{* Constructors *}
-
-definition NNode :: "N \<Rightarrow> (T + Tree)fset \<Rightarrow> Tree"
-where "NNode n as \<equiv> Tree_ctor (n,as)"
-
-lemmas ctor_defs = NNode_def
-
-
-subsection {* Pre-selectors *}
-
-(* These are mere auxiliaries *)
-
-definition "asNNode tr \<equiv> SOME n_as. NNode (fst n_as) (snd n_as) = tr"
-lemmas pre_sel_defs = asNNode_def
-
-
-subsection {* Selectors *}
-
-(* One for each pair (constructor, constructor argument) *)
-
-(* For NNode: *)
-definition root :: "Tree \<Rightarrow> N" where "root tr = fst (asNNode tr)"
-definition ccont :: "Tree \<Rightarrow> (T + Tree)fset" where "ccont tr = snd (asNNode tr)"
-
-lemmas sel_defs = root_def ccont_def
-
-
-subsection {* Basic properties *}
-
-(* Constructors versus selectors *)
-lemma NNode_surj: "\<exists> n as. NNode n as = tr"
-unfolding NNode_def
-by (metis Tree.ctor_dtor pair_collapse)
-
-lemma NNode_asNNode:
-"NNode (fst (asNNode tr)) (snd (asNNode tr)) = tr"
-proof-
- obtain n as where "NNode n as = tr" using NNode_surj[of tr] by blast
- hence "NNode (fst (n,as)) (snd (n,as)) = tr" by simp
- thus ?thesis unfolding asNNode_def by(rule someI)
-qed
-
-theorem NNode_root_ccont[simp]:
-"NNode (root tr) (ccont tr) = tr"
-using NNode_asNNode unfolding root_def ccont_def .
-
-(* Constructors *)
-theorem TTree_simps[simp]:
-"NNode n as = NNode n' as' \<longleftrightarrow> n = n' \<and> as = as'"
-unfolding ctor_defs Tree.ctor_inject by auto
-
-theorem TTree_cases[elim, case_names NNode Choice]:
-assumes NNode: "\<And> n as. tr = NNode n as \<Longrightarrow> phi"
-shows phi
-proof(cases rule: Tree.ctor_exhaust[of tr])
- fix x assume "tr = Tree_ctor x"
- thus ?thesis
- apply(cases x)
- using NNode unfolding ctor_defs apply blast
- done
-qed
-
-(* Constructors versus selectors *)
-theorem TTree_sel_ctor[simp]:
-"root (NNode n as) = n"
-"ccont (NNode n as) = as"
-unfolding root_def ccont_def
-by (metis (no_types) NNode_asNNode TTree_simps)+
-
-
-subsection{* Coinduction *}
-
-theorem TTree_coind_Node[elim, consumes 1, case_names NNode, induct pred: "HOL.eq"]:
-assumes phi: "\<phi> tr1 tr2" and
-NNode: "\<And> n1 n2 as1 as2.
- \<lbrakk>\<phi> (NNode n1 as1) (NNode n2 as2)\<rbrakk> \<Longrightarrow>
- n1 = n2 \<and> llift2 \<phi> as1 as2"
-shows "tr1 = tr2"
-apply(rule mp[OF Tree.rel_coinduct[of \<phi> tr1 tr2] phi]) proof clarify
- fix tr1 tr2 assume \<phi>: "\<phi> tr1 tr2"
- show "pre_Tree_rel \<phi> (Tree_dtor tr1) (Tree_dtor tr2)"
- apply(cases rule: Tree.ctor_exhaust[of tr1], cases rule: Tree.ctor_exhaust[of tr2])
- apply (simp add: Tree.dtor_ctor)
- apply(case_tac x, case_tac xa, simp)
- unfolding pre_Tree_rel apply(rule NNode) using \<phi> unfolding NNode_def by simp
-qed
-
-theorem TTree_coind[elim, consumes 1, case_names LLift]:
-assumes phi: "\<phi> tr1 tr2" and
-LLift: "\<And> tr1 tr2. \<phi> tr1 tr2 \<Longrightarrow>
- root tr1 = root tr2 \<and> llift2 \<phi> (ccont tr1) (ccont tr2)"
-shows "tr1 = tr2"
-using phi apply(induct rule: TTree_coind_Node)
-using LLift by (metis TTree_sel_ctor)
-
-
-
-subsection {* Coiteration *}
-
-(* Preliminaries: *)
-declare Tree.dtor_ctor[simp]
-declare Tree.ctor_dtor[simp]
-
-lemma Tree_dtor_NNode[simp]:
-"Tree_dtor (NNode n as) = (n,as)"
-unfolding NNode_def Tree.dtor_ctor ..
-
-lemma Tree_dtor_root_ccont:
-"Tree_dtor tr = (root tr, ccont tr)"
-unfolding root_def ccont_def
-by (metis (lifting) NNode_asNNode Tree_dtor_NNode)
-
-(* Coiteration *)
-definition TTree_unfold ::
-"('b \<Rightarrow> N) \<Rightarrow> ('b \<Rightarrow> (T + 'b) fset) \<Rightarrow> 'b \<Rightarrow> Tree"
-where "TTree_unfold rt ct \<equiv> Tree_dtor_unfold <rt,ct>"
-
-lemma Tree_unfold_unfold:
-"Tree_dtor_unfold s = TTree_unfold (fst o s) (snd o s)"
-apply(rule ext)
-unfolding TTree_unfold_def by simp
-
-theorem TTree_unfold:
-"root (TTree_unfold rt ct b) = rt b"
-"ccont (TTree_unfold rt ct b) = map_fset (id \<oplus> TTree_unfold rt ct) (ct b)"
-using Tree.dtor_unfolds[of "<rt,ct>" b] unfolding Tree_unfold_unfold fst_convol snd_convol
-unfolding pre_Tree_map' fst_convol' snd_convol'
-unfolding Tree_dtor_root_ccont by simp_all
-
-(* Corecursion, stronger than coiteration (unfold) *)
-definition TTree_corec ::
-"('b \<Rightarrow> N) \<Rightarrow> ('b \<Rightarrow> (T + (Tree + 'b)) fset) \<Rightarrow> 'b \<Rightarrow> Tree"
-where "TTree_corec rt ct \<equiv> Tree_dtor_corec <rt,ct>"
-
-lemma Tree_dtor_corec_corec:
-"Tree_dtor_corec s = TTree_corec (fst o s) (snd o s)"
-apply(rule ext)
-unfolding TTree_corec_def by simp
-
-theorem TTree_corec:
-"root (TTree_corec rt ct b) = rt b"
-"ccont (TTree_corec rt ct b) = map_fset (id \<oplus> ([[id, TTree_corec rt ct]]) ) (ct b)"
-using Tree.dtor_corecs[of "<rt,ct>" b] unfolding Tree_dtor_corec_corec fst_convol snd_convol
-unfolding pre_Tree_map' fst_convol' snd_convol'
-unfolding Tree_dtor_root_ccont by simp_all
-
-
-subsection{* The characteristic theorems transported from fset to set *}
-
-definition "Node n as \<equiv> NNode n (the_inv fset as)"
-definition "cont \<equiv> fset o ccont"
-definition "unfold rt ct \<equiv> TTree_unfold rt (the_inv fset o ct)"
-definition "corec rt ct \<equiv> TTree_corec rt (the_inv fset o ct)"
-
-definition lift ("_ ^#" 200) where
-"lift \<phi> as \<longleftrightarrow> (\<forall> tr. Inr tr \<in> as \<longrightarrow> \<phi> tr)"
-
-definition lift2 ("_ ^#2" 200) where
-"lift2 \<phi> as1 as2 \<longleftrightarrow>
- (\<forall> n. Inl n \<in> as1 \<longleftrightarrow> Inl n \<in> as2) \<and>
- (\<forall> tr1. Inr tr1 \<in> as1 \<longrightarrow> (\<exists> tr2. Inr tr2 \<in> as2 \<and> \<phi> tr1 tr2)) \<and>
- (\<forall> tr2. Inr tr2 \<in> as2 \<longrightarrow> (\<exists> tr1. Inr tr1 \<in> as1 \<and> \<phi> tr1 tr2))"
-
-definition liftS ("_ ^#s" 200) where
-"liftS trs = {as. Inr -` as \<subseteq> trs}"
-
-lemma lift2_llift2:
-"\<lbrakk>finite as1; finite as2\<rbrakk> \<Longrightarrow>
- lift2 \<phi> as1 as2 \<longleftrightarrow> llift2 \<phi> (the_inv fset as1) (the_inv fset as2)"
-unfolding lift2_def llift2_def by auto
-
-lemma llift2_lift2:
-"llift2 \<phi> as1 as2 \<longleftrightarrow> lift2 \<phi> (fset as1) (fset as2)"
-using lift2_llift2 by (metis finite_fset fset_cong fset_to_fset)
-
-lemma mono_lift:
-assumes "(\<phi>^#) as"
-and "\<And> tr. \<phi> tr \<Longrightarrow> \<phi>' tr"
-shows "(\<phi>'^#) as"
-using assms unfolding lift_def[abs_def] by blast
-
-lemma mono_liftS:
-assumes "trs1 \<subseteq> trs2 "
-shows "(trs1 ^#s) \<subseteq> (trs2 ^#s)"
-using assms unfolding liftS_def[abs_def] by blast
-
-lemma lift_mono:
-assumes "\<phi> \<le> \<phi>'"
-shows "(\<phi>^#) \<le> (\<phi>'^#)"
-using assms unfolding lift_def[abs_def] by blast
-
-lemma mono_lift2:
-assumes "(\<phi>^#2) as1 as2"
-and "\<And> tr1 tr2. \<phi> tr1 tr2 \<Longrightarrow> \<phi>' tr1 tr2"
-shows "(\<phi>'^#2) as1 as2"
-using assms unfolding lift2_def[abs_def] by blast
-
-lemma lift2_mono:
-assumes "\<phi> \<le> \<phi>'"
-shows "(\<phi>^#2) \<le> (\<phi>'^#2)"
-using assms unfolding lift2_def[abs_def] by blast
-
-lemma finite_cont[simp]: "finite (cont tr)"
-unfolding cont_def by auto
-
-theorem Node_root_cont[simp]:
-"Node (root tr) (cont tr) = tr"
-using NNode_root_ccont unfolding Node_def cont_def
-by (metis cont_def finite_cont fset_cong fset_to_fset o_def)
-
-theorem Tree_simps[simp]:
-assumes "finite as" and "finite as'"
-shows "Node n as = Node n' as' \<longleftrightarrow> n = n' \<and> as = as'"
-using assms TTree_simps unfolding Node_def
-by (metis fset_to_fset)
-
-theorem Tree_cases[elim, case_names Node Choice]:
-assumes Node: "\<And> n as. \<lbrakk>finite as; tr = Node n as\<rbrakk> \<Longrightarrow> phi"
-shows phi
-apply(cases rule: TTree_cases[of tr])
-using Node unfolding Node_def
-by (metis Node Node_root_cont finite_cont)
-
-theorem Tree_sel_ctor[simp]:
-"root (Node n as) = n"
-"finite as \<Longrightarrow> cont (Node n as) = as"
-unfolding Node_def cont_def by auto
-
-theorems root_Node = Tree_sel_ctor(1)
-theorems cont_Node = Tree_sel_ctor(2)
-
-theorem Tree_coind_Node[elim, consumes 1, case_names Node]:
-assumes phi: "\<phi> tr1 tr2" and
-Node:
-"\<And> n1 n2 as1 as2.
- \<lbrakk>finite as1; finite as2; \<phi> (Node n1 as1) (Node n2 as2)\<rbrakk>
- \<Longrightarrow> n1 = n2 \<and> (\<phi>^#2) as1 as2"
-shows "tr1 = tr2"
-using phi apply(induct rule: TTree_coind_Node)
-unfolding llift2_lift2 apply(rule Node)
-unfolding Node_def
-apply (metis finite_fset)
-apply (metis finite_fset)
-by (metis finite_fset fset_cong fset_to_fset)
-
-theorem Tree_coind[elim, consumes 1, case_names Lift, induct pred: "HOL.eq"]:
-assumes phi: "\<phi> tr1 tr2" and
-Lift: "\<And> tr1 tr2. \<phi> tr1 tr2 \<Longrightarrow>
- root tr1 = root tr2 \<and> (\<phi>^#2) (cont tr1) (cont tr2)"
-shows "tr1 = tr2"
-using phi apply(induct rule: TTree_coind)
-unfolding llift2_lift2 apply(rule Lift[unfolded cont_def comp_def]) .
-
-theorem unfold:
-"root (unfold rt ct b) = rt b"
-"finite (ct b) \<Longrightarrow> cont (unfold rt ct b) = image (id \<oplus> unfold rt ct) (ct b)"
-using TTree_unfold[of rt "the_inv fset \<circ> ct" b] unfolding unfold_def
-apply - apply metis
-unfolding cont_def comp_def
-by (metis (no_types) fset_to_fset map_fset_image)
-
-
-theorem corec:
-"root (corec rt ct b) = rt b"
-"finite (ct b) \<Longrightarrow> cont (corec rt ct b) = image (id \<oplus> ([[id, corec rt ct]])) (ct b)"
-using TTree_corec[of rt "the_inv fset \<circ> ct" b] unfolding corec_def
-apply - apply metis
-unfolding cont_def comp_def
-by (metis (no_types) fset_to_fset map_fset_image)
-
-
-end
--- a/src/HOL/Codatatype/Examples/Lambda_Term.thy Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,259 +0,0 @@
-(* Title: HOL/BNF/Examples/Lambda_Term.thy
- Author: Dmitriy Traytel, TU Muenchen
- Author: Andrei Popescu, TU Muenchen
- Copyright 2012
-
-Lambda-terms.
-*)
-
-header {* Lambda-Terms *}
-
-theory Lambda_Term
-imports "../BNF"
-begin
-
-
-section {* Datatype definition *}
-
-data_raw trm: 'trm = "'a + 'trm \<times> 'trm + 'a \<times> 'trm + ('a \<times> 'trm) fset \<times> 'trm"
-
-
-section {* Customization of terms *}
-
-subsection{* Set and map *}
-
-lemma pre_trm_set2_Lt: "pre_trm_set2 (Inr (Inr (Inr (xts, t)))) = snd ` (fset xts) \<union> {t}"
-unfolding pre_trm_set2_def sum_set_defs prod_set_defs collect_def[abs_def]
-by auto
-
-lemma pre_trm_set2_Var: "\<And>x. pre_trm_set2 (Inl x) = {}"
-and pre_trm_set2_App:
-"\<And>t1 t2. pre_trm_set2 (Inr (Inl t1t2)) = {fst t1t2, snd t1t2}"
-and pre_trm_set2_Lam:
-"\<And>x t. pre_trm_set2 (Inr (Inr (Inl (x, t)))) = {t}"
-unfolding pre_trm_set2_def sum_set_defs prod_set_defs collect_def[abs_def]
-by auto
-
-lemma pre_trm_map:
-"\<And> a1. pre_trm_map f1 f2 (Inl a1) = Inl (f1 a1)"
-"\<And> a2 b2. pre_trm_map f1 f2 (Inr (Inl (a2,b2))) = Inr (Inl (f2 a2, f2 b2))"
-"\<And> a1 a2. pre_trm_map f1 f2 (Inr (Inr (Inl (a1,a2)))) = Inr (Inr (Inl (f1 a1, f2 a2)))"
-"\<And> a1a2s a2.
- pre_trm_map f1 f2 (Inr (Inr (Inr (a1a2s, a2)))) =
- Inr (Inr (Inr (map_fset (\<lambda> (a1', a2'). (f1 a1', f2 a2')) a1a2s, f2 a2)))"
-unfolding pre_trm_map_def collect_def[abs_def] map_pair_def by auto
-
-
-subsection{* Constructors *}
-
-definition "Var x \<equiv> trm_ctor (Inl x)"
-definition "App t1 t2 \<equiv> trm_ctor (Inr (Inl (t1,t2)))"
-definition "Lam x t \<equiv> trm_ctor (Inr (Inr (Inl (x,t))))"
-definition "Lt xts t \<equiv> trm_ctor (Inr (Inr (Inr (xts,t))))"
-
-lemmas ctor_defs = Var_def App_def Lam_def Lt_def
-
-theorem trm_simps[simp]:
-"\<And>x y. Var x = Var y \<longleftrightarrow> x = y"
-"\<And>t1 t2 t1' t2'. App t1 t2 = App t1' t2' \<longleftrightarrow> t1 = t1' \<and> t2 = t2'"
-"\<And>x x' t t'. Lam x t = Lam x' t' \<longleftrightarrow> x = x' \<and> t = t'"
-"\<And> xts xts' t t'. Lt xts t = Lt xts' t' \<longleftrightarrow> xts = xts' \<and> t = t'"
-(* *)
-"\<And> x t1 t2. Var x \<noteq> App t1 t2" "\<And>x y t. Var x \<noteq> Lam y t" "\<And> x xts t. Var x \<noteq> Lt xts t"
-"\<And> t1 t2 x t. App t1 t2 \<noteq> Lam x t" "\<And> t1 t2 xts t. App t1 t2 \<noteq> Lt xts t"
-"\<And>x t xts t1. Lam x t \<noteq> Lt xts t1"
-unfolding ctor_defs trm.ctor_inject by auto
-
-theorem trm_cases[elim, case_names Var App Lam Lt]:
-assumes Var: "\<And> x. t = Var x \<Longrightarrow> phi"
-and App: "\<And> t1 t2. t = App t1 t2 \<Longrightarrow> phi"
-and Lam: "\<And> x t1. t = Lam x t1 \<Longrightarrow> phi"
-and Lt: "\<And> xts t1. t = Lt xts t1 \<Longrightarrow> phi"
-shows phi
-proof(cases rule: trm.ctor_exhaust[of t])
- fix x assume "t = trm_ctor x"
- thus ?thesis
- apply(cases x) using Var unfolding ctor_defs apply blast
- apply(case_tac b) using App unfolding ctor_defs apply(case_tac a, blast)
- apply(case_tac ba) using Lam unfolding ctor_defs apply(case_tac a, blast)
- apply(case_tac bb) using Lt unfolding ctor_defs by blast
-qed
-
-lemma trm_induct[case_names Var App Lam Lt, induct type: trm]:
-assumes Var: "\<And> (x::'a). phi (Var x)"
-and App: "\<And> t1 t2. \<lbrakk>phi t1; phi t2\<rbrakk> \<Longrightarrow> phi (App t1 t2)"
-and Lam: "\<And> x t. phi t \<Longrightarrow> phi (Lam x t)"
-and Lt: "\<And> xts t. \<lbrakk>\<And> x1 t1. (x1,t1) |\<in>| xts \<Longrightarrow> phi t1; phi t\<rbrakk> \<Longrightarrow> phi (Lt xts t)"
-shows "phi t"
-proof(induct rule: trm.ctor_induct)
- fix u :: "'a + 'a trm \<times> 'a trm + 'a \<times> 'a trm + ('a \<times> 'a trm) fset \<times> 'a trm"
- assume IH: "\<And>t. t \<in> pre_trm_set2 u \<Longrightarrow> phi t"
- show "phi (trm_ctor u)"
- proof(cases u)
- case (Inl x)
- show ?thesis using Var unfolding Var_def Inl .
- next
- case (Inr uu) note Inr1 = Inr
- show ?thesis
- proof(cases uu)
- case (Inl t1t2)
- obtain t1 t2 where t1t2: "t1t2 = (t1,t2)" by (cases t1t2, blast)
- show ?thesis unfolding Inr1 Inl t1t2 App_def[symmetric] apply(rule App)
- using IH unfolding Inr1 Inl pre_trm_set2_App t1t2 fst_conv snd_conv by blast+
- next
- case (Inr uuu) note Inr2 = Inr
- show ?thesis
- proof(cases uuu)
- case (Inl xt)
- obtain x t where xt: "xt = (x,t)" by (cases xt, blast)
- show ?thesis unfolding Inr1 Inr2 Inl xt Lam_def[symmetric] apply(rule Lam)
- using IH unfolding Inr1 Inr2 Inl pre_trm_set2_Lam xt by blast
- next
- case (Inr xts_t)
- obtain xts t where xts_t: "xts_t = (xts,t)" by (cases xts_t, blast)
- show ?thesis unfolding Inr1 Inr2 Inr xts_t Lt_def[symmetric] apply(rule Lt) using IH
- unfolding Inr1 Inr2 Inr pre_trm_set2_Lt xts_t fset_fset_member image_def by auto
- qed
- qed
- qed
-qed
-
-
-subsection{* Recursion and iteration (fold) *}
-
-definition
-"sumJoin4 f1 f2 f3 f4 \<equiv>
-\<lambda> k. (case k of
- Inl x1 \<Rightarrow> f1 x1
-|Inr k1 \<Rightarrow> (case k1 of
- Inl ((s2,a2),(t2,b2)) \<Rightarrow> f2 s2 a2 t2 b2
-|Inr k2 \<Rightarrow> (case k2 of Inl (x3,(t3,b3)) \<Rightarrow> f3 x3 t3 b3
-|Inr (xts,(t4,b4)) \<Rightarrow> f4 xts t4 b4)))"
-
-lemma sumJoin4_simps[simp]:
-"\<And>x. sumJoin4 var app lam lt (Inl x) = var x"
-"\<And> t1 a1 t2 a2. sumJoin4 var app lam lt (Inr (Inl ((t1,a1),(t2,a2)))) = app t1 a1 t2 a2"
-"\<And> x t a. sumJoin4 var app lam lt (Inr (Inr (Inl (x,(t,a))))) = lam x t a"
-"\<And> xtas t a. sumJoin4 var app lam lt (Inr (Inr (Inr (xtas,(t,a))))) = lt xtas t a"
-unfolding sumJoin4_def by auto
-
-definition "trmrec var app lam lt \<equiv> trm_ctor_rec (sumJoin4 var app lam lt)"
-
-lemma trmrec_Var[simp]:
-"trmrec var app lam lt (Var x) = var x"
-unfolding trmrec_def Var_def trm.ctor_recs pre_trm_map(1) by simp
-
-lemma trmrec_App[simp]:
-"trmrec var app lam lt (App t1 t2) =
- app t1 (trmrec var app lam lt t1) t2 (trmrec var app lam lt t2)"
-unfolding trmrec_def App_def trm.ctor_recs pre_trm_map(2) convol_def by simp
-
-lemma trmrec_Lam[simp]:
-"trmrec var app lam lt (Lam x t) = lam x t (trmrec var app lam lt t)"
-unfolding trmrec_def Lam_def trm.ctor_recs pre_trm_map(3) convol_def by simp
-
-lemma trmrec_Lt[simp]:
-"trmrec var app lam lt (Lt xts t) =
- lt (map_fset (\<lambda> (x,t). (x,t,trmrec var app lam lt t)) xts) t (trmrec var app lam lt t)"
-unfolding trmrec_def Lt_def trm.ctor_recs pre_trm_map(4) convol_def by simp
-
-definition
-"sumJoinI4 f1 f2 f3 f4 \<equiv>
-\<lambda> k. (case k of
- Inl x1 \<Rightarrow> f1 x1
-|Inr k1 \<Rightarrow> (case k1 of
- Inl (a2,b2) \<Rightarrow> f2 a2 b2
-|Inr k2 \<Rightarrow> (case k2 of Inl (x3,b3) \<Rightarrow> f3 x3 b3
-|Inr (xts,b4) \<Rightarrow> f4 xts b4)))"
-
-lemma sumJoinI4_simps[simp]:
-"\<And>x. sumJoinI4 var app lam lt (Inl x) = var x"
-"\<And> a1 a2. sumJoinI4 var app lam lt (Inr (Inl (a1,a2))) = app a1 a2"
-"\<And> x a. sumJoinI4 var app lam lt (Inr (Inr (Inl (x,a)))) = lam x a"
-"\<And> xtas a. sumJoinI4 var app lam lt (Inr (Inr (Inr (xtas,a)))) = lt xtas a"
-unfolding sumJoinI4_def by auto
-
-(* The iterator has a simpler, hence more manageable type. *)
-definition "trmfold var app lam lt \<equiv> trm_ctor_fold (sumJoinI4 var app lam lt)"
-
-lemma trmfold_Var[simp]:
-"trmfold var app lam lt (Var x) = var x"
-unfolding trmfold_def Var_def trm.ctor_folds pre_trm_map(1) by simp
-
-lemma trmfold_App[simp]:
-"trmfold var app lam lt (App t1 t2) =
- app (trmfold var app lam lt t1) (trmfold var app lam lt t2)"
-unfolding trmfold_def App_def trm.ctor_folds pre_trm_map(2) by simp
-
-lemma trmfold_Lam[simp]:
-"trmfold var app lam lt (Lam x t) = lam x (trmfold var app lam lt t)"
-unfolding trmfold_def Lam_def trm.ctor_folds pre_trm_map(3) by simp
-
-lemma trmfold_Lt[simp]:
-"trmfold var app lam lt (Lt xts t) =
- lt (map_fset (\<lambda> (x,t). (x,trmfold var app lam lt t)) xts) (trmfold var app lam lt t)"
-unfolding trmfold_def Lt_def trm.ctor_folds pre_trm_map(4) by simp
-
-
-subsection{* Example: The set of all variables varsOf and free variables fvarsOf of a term: *}
-
-definition "varsOf = trmfold
-(\<lambda> x. {x})
-(\<lambda> X1 X2. X1 \<union> X2)
-(\<lambda> x X. X \<union> {x})
-(\<lambda> xXs Y. Y \<union> (\<Union> { {x} \<union> X | x X. (x,X) |\<in>| xXs}))"
-
-lemma varsOf_simps[simp]:
-"varsOf (Var x) = {x}"
-"varsOf (App t1 t2) = varsOf t1 \<union> varsOf t2"
-"varsOf (Lam x t) = varsOf t \<union> {x}"
-"varsOf (Lt xts t) =
- varsOf t \<union> (\<Union> { {x} \<union> X | x X. (x,X) |\<in>| map_fset (\<lambda> (x,t1). (x,varsOf t1)) xts})"
-unfolding varsOf_def by simp_all
-
-definition "fvarsOf = trmfold
-(\<lambda> x. {x})
-(\<lambda> X1 X2. X1 \<union> X2)
-(\<lambda> x X. X - {x})
-(\<lambda> xtXs Y. Y - {x | x X. (x,X) |\<in>| xtXs} \<union> (\<Union> {X | x X. (x,X) |\<in>| xtXs}))"
-
-lemma fvarsOf_simps[simp]:
-"fvarsOf (Var x) = {x}"
-"fvarsOf (App t1 t2) = fvarsOf t1 \<union> fvarsOf t2"
-"fvarsOf (Lam x t) = fvarsOf t - {x}"
-"fvarsOf (Lt xts t) =
- fvarsOf t
- - {x | x X. (x,X) |\<in>| map_fset (\<lambda> (x,t1). (x,fvarsOf t1)) xts}
- \<union> (\<Union> {X | x X. (x,X) |\<in>| map_fset (\<lambda> (x,t1). (x,fvarsOf t1)) xts})"
-unfolding fvarsOf_def by simp_all
-
-lemma diff_Un_incl_triv: "\<lbrakk>A \<subseteq> D; C \<subseteq> E\<rbrakk> \<Longrightarrow> A - B \<union> C \<subseteq> D \<union> E" by blast
-
-lemma in_map_fset_iff:
-"(x, X) |\<in>| map_fset (\<lambda>(x, t1). (x, f t1)) xts \<longleftrightarrow>
- (\<exists> t1. (x,t1) |\<in>| xts \<and> X = f t1)"
-unfolding map_fset_def2_raw in_fset fset_afset unfolding fset_def2_raw by auto
-
-lemma fvarsOf_varsOf: "fvarsOf t \<subseteq> varsOf t"
-proof induct
- case (Lt xts t)
- thus ?case unfolding fvarsOf_simps varsOf_simps
- proof (elim diff_Un_incl_triv)
- show
- "\<Union>{X | x X. (x, X) |\<in>| map_fset (\<lambda>(x, t1). (x, fvarsOf t1)) xts}
- \<subseteq> \<Union>{{x} \<union> X |x X. (x, X) |\<in>| map_fset (\<lambda>(x, t1). (x, varsOf t1)) xts}"
- (is "_ \<subseteq> \<Union> ?L")
- proof(rule Sup_mono, safe)
- fix a x X
- assume "(x, X) |\<in>| map_fset (\<lambda>(x, t1). (x, fvarsOf t1)) xts"
- then obtain t1 where x_t1: "(x,t1) |\<in>| xts" and X: "X = fvarsOf t1"
- unfolding in_map_fset_iff by auto
- let ?Y = "varsOf t1"
- have x_Y: "(x,?Y) |\<in>| map_fset (\<lambda>(x, t1). (x, varsOf t1)) xts"
- using x_t1 unfolding in_map_fset_iff by auto
- show "\<exists> Y \<in> ?L. X \<subseteq> Y" unfolding X using Lt(1) x_Y x_t1 by auto
- qed
- qed
-qed auto
-
-end
--- a/src/HOL/Codatatype/Examples/ListF.thy Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,171 +0,0 @@
-(* Title: HOL/BNF/Examples/ListF.thy
- Author: Dmitriy Traytel, TU Muenchen
- Author: Andrei Popescu, TU Muenchen
- Copyright 2012
-
-Finite lists.
-*)
-
-header {* Finite Lists *}
-
-theory ListF
-imports "../BNF"
-begin
-
-data_raw listF: 'list = "unit + 'a \<times> 'list"
-
-definition "NilF = listF_ctor (Inl ())"
-definition "Conss a as \<equiv> listF_ctor (Inr (a, as))"
-
-lemma listF_map_NilF[simp]: "listF_map f NilF = NilF"
-unfolding listF_map_def pre_listF_map_def NilF_def listF.ctor_folds by simp
-
-lemma listF_map_Conss[simp]:
- "listF_map f (Conss x xs) = Conss (f x) (listF_map f xs)"
-unfolding listF_map_def pre_listF_map_def Conss_def listF.ctor_folds by simp
-
-lemma listF_set_NilF[simp]: "listF_set NilF = {}"
-unfolding listF_set_def NilF_def listF.ctor_folds pre_listF_set1_def pre_listF_set2_def
- sum_set_defs pre_listF_map_def collect_def[abs_def] by simp
-
-lemma listF_set_Conss[simp]: "listF_set (Conss x xs) = {x} \<union> listF_set xs"
-unfolding listF_set_def Conss_def listF.ctor_folds pre_listF_set1_def pre_listF_set2_def
- sum_set_defs prod_set_defs pre_listF_map_def collect_def[abs_def] by simp
-
-lemma fold_sum_case_NilF: "listF_ctor_fold (sum_case f g) NilF = f ()"
-unfolding NilF_def listF.ctor_folds pre_listF_map_def by simp
-
-
-lemma fold_sum_case_Conss:
- "listF_ctor_fold (sum_case f g) (Conss y ys) = g (y, listF_ctor_fold (sum_case f g) ys)"
-unfolding Conss_def listF.ctor_folds pre_listF_map_def by simp
-
-(* familiar induction principle *)
-lemma listF_induct:
- fixes xs :: "'a listF"
- assumes IB: "P NilF" and IH: "\<And>x xs. P xs \<Longrightarrow> P (Conss x xs)"
- shows "P xs"
-proof (rule listF.ctor_induct)
- fix xs :: "unit + 'a \<times> 'a listF"
- assume raw_IH: "\<And>a. a \<in> pre_listF_set2 xs \<Longrightarrow> P a"
- show "P (listF_ctor xs)"
- proof (cases xs)
- case (Inl a) with IB show ?thesis unfolding NilF_def by simp
- next
- case (Inr b)
- then obtain y ys where yys: "listF_ctor xs = Conss y ys"
- unfolding Conss_def listF.ctor_inject by (blast intro: prod.exhaust)
- hence "ys \<in> pre_listF_set2 xs"
- unfolding pre_listF_set2_def Conss_def listF.ctor_inject sum_set_defs prod_set_defs
- collect_def[abs_def] by simp
- with raw_IH have "P ys" by blast
- with IH have "P (Conss y ys)" by blast
- with yys show ?thesis by simp
- qed
-qed
-
-rep_datatype NilF Conss
-by (blast intro: listF_induct) (auto simp add: NilF_def Conss_def listF.ctor_inject)
-
-definition Singll ("[[_]]") where
- [simp]: "Singll a \<equiv> Conss a NilF"
-
-definition appendd (infixr "@@" 65) where
- "appendd \<equiv> listF_ctor_fold (sum_case (\<lambda> _. id) (\<lambda> (a,f) bs. Conss a (f bs)))"
-
-definition "lrev \<equiv> listF_ctor_fold (sum_case (\<lambda> _. NilF) (\<lambda> (b,bs). bs @@ [[b]]))"
-
-lemma lrev_NilF[simp]: "lrev NilF = NilF"
-unfolding lrev_def by (simp add: fold_sum_case_NilF)
-
-lemma lrev_Conss[simp]: "lrev (Conss y ys) = lrev ys @@ [[y]]"
-unfolding lrev_def by (simp add: fold_sum_case_Conss)
-
-lemma NilF_appendd[simp]: "NilF @@ ys = ys"
-unfolding appendd_def by (simp add: fold_sum_case_NilF)
-
-lemma Conss_append[simp]: "Conss x xs @@ ys = Conss x (xs @@ ys)"
-unfolding appendd_def by (simp add: fold_sum_case_Conss)
-
-lemma appendd_NilF[simp]: "xs @@ NilF = xs"
-by (rule listF_induct) auto
-
-lemma appendd_assoc[simp]: "(xs @@ ys) @@ zs = xs @@ ys @@ zs"
-by (rule listF_induct) auto
-
-lemma lrev_appendd[simp]: "lrev (xs @@ ys) = lrev ys @@ lrev xs"
-by (rule listF_induct[of _ xs]) auto
-
-lemma listF_map_appendd[simp]:
- "listF_map f (xs @@ ys) = listF_map f xs @@ listF_map f ys"
-by (rule listF_induct[of _ xs]) auto
-
-lemma lrev_listF_map[simp]: "lrev (listF_map f xs) = listF_map f (lrev xs)"
-by (rule listF_induct[of _ xs]) auto
-
-lemma lrev_lrev[simp]: "lrev (lrev as) = as"
-by (rule listF_induct) auto
-
-fun lengthh where
- "lengthh NilF = 0"
-| "lengthh (Conss x xs) = Suc (lengthh xs)"
-
-fun nthh where
- "nthh (Conss x xs) 0 = x"
-| "nthh (Conss x xs) (Suc n) = nthh xs n"
-| "nthh xs i = undefined"
-
-lemma lengthh_listF_map[simp]: "lengthh (listF_map f xs) = lengthh xs"
-by (rule listF_induct[of _ xs]) auto
-
-lemma nthh_listF_map[simp]:
- "i < lengthh xs \<Longrightarrow> nthh (listF_map f xs) i = f (nthh xs i)"
-by (induct rule: nthh.induct) auto
-
-lemma nthh_listF_set[simp]: "i < lengthh xs \<Longrightarrow> nthh xs i \<in> listF_set xs"
-by (induct rule: nthh.induct) auto
-
-lemma NilF_iff[iff]: "(lengthh xs = 0) = (xs = NilF)"
-by (induct xs) auto
-
-lemma Conss_iff[iff]:
- "(lengthh xs = Suc n) = (\<exists>y ys. xs = Conss y ys \<and> lengthh ys = n)"
-by (induct xs) auto
-
-lemma Conss_iff'[iff]:
- "(Suc n = lengthh xs) = (\<exists>y ys. xs = Conss y ys \<and> lengthh ys = n)"
-by (induct xs) (simp, simp, blast)
-
-lemma listF_induct2: "\<lbrakk>lengthh xs = lengthh ys; P NilF NilF;
- \<And>x xs y ys. P xs ys \<Longrightarrow> P (Conss x xs) (Conss y ys)\<rbrakk> \<Longrightarrow> P xs ys"
-by (induct xs arbitrary: ys rule: listF_induct) auto
-
-fun zipp where
- "zipp NilF NilF = NilF"
-| "zipp (Conss x xs) (Conss y ys) = Conss (x, y) (zipp xs ys)"
-| "zipp xs ys = undefined"
-
-lemma listF_map_fst_zip[simp]:
- "lengthh xs = lengthh ys \<Longrightarrow> listF_map fst (zipp xs ys) = xs"
-by (erule listF_induct2) auto
-
-lemma listF_map_snd_zip[simp]:
- "lengthh xs = lengthh ys \<Longrightarrow> listF_map snd (zipp xs ys) = ys"
-by (erule listF_induct2) auto
-
-lemma lengthh_zip[simp]:
- "lengthh xs = lengthh ys \<Longrightarrow> lengthh (zipp xs ys) = lengthh xs"
-by (erule listF_induct2) auto
-
-lemma nthh_zip[simp]:
- assumes *: "lengthh xs = lengthh ys"
- shows "i < lengthh xs \<Longrightarrow> nthh (zipp xs ys) i = (nthh xs i, nthh ys i)"
-proof (induct arbitrary: i rule: listF_induct2[OF *])
- case (2 x xs y ys) thus ?case by (induct i) auto
-qed simp
-
-lemma list_set_nthh[simp]:
- "(x \<in> listF_set xs) \<Longrightarrow> (\<exists>i < lengthh xs. nthh xs i = x)"
-by (induct xs) (auto, induct rule: nthh.induct, auto)
-
-end
--- a/src/HOL/Codatatype/Examples/Misc_Codata.thy Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,110 +0,0 @@
-(* Title: HOL/BNF/Examples/Misc_Data.thy
- Author: Dmitriy Traytel, TU Muenchen
- Author: Andrei Popescu, TU Muenchen
- Copyright 2012
-
-Miscellaneous codatatype declarations.
-*)
-
-header {* Miscellaneous Codatatype Declarations *}
-
-theory Misc_Codata
-imports "../BNF"
-begin
-
-codata simple = X1 | X2 | X3 | X4
-
-codata simple' = X1' unit | X2' unit | X3' unit | X4' unit
-
-codata 'a stream = Stream 'a "'a stream"
-
-codata 'a mylist = MyNil | MyCons 'a "'a mylist"
-
-codata ('b, 'c, 'd, 'e) some_passive =
- SP1 "('b, 'c, 'd, 'e) some_passive" | SP2 'b | SP3 'c | SP4 'd | SP5 'e
-
-codata lambda =
- Var string |
- App lambda lambda |
- Abs string lambda |
- Let "(string \<times> lambda) fset" lambda
-
-codata 'a par_lambda =
- PVar 'a |
- PApp "'a par_lambda" "'a par_lambda" |
- PAbs 'a "'a par_lambda" |
- PLet "('a \<times> 'a par_lambda) fset" "'a par_lambda"
-
-(*
- ('a, 'b1, 'b2) F1 = 'a * 'b1 + 'a * 'b2
- ('a, 'b1, 'b2) F2 = unit + 'b1 * 'b2
-*)
-
-codata 'a J1 = J11 'a "'a J1" | J12 'a "'a J2"
- and 'a J2 = J21 | J22 "'a J1" "'a J2"
-
-codata 'a tree = TEmpty | TNode 'a "'a forest"
- and 'a forest = FNil | FCons "'a tree" "'a forest"
-
-codata 'a tree' = TEmpty' | TNode' "'a branch" "'a branch"
- and 'a branch = Branch 'a "'a tree'"
-
-codata ('a, 'b) exp = Term "('a, 'b) trm" | Sum "('a, 'b) trm" "('a, 'b) exp"
- and ('a, 'b) trm = Factor "('a, 'b) factor" | Prod "('a, 'b) factor" "('a, 'b) trm"
- and ('a, 'b) factor = C 'a | V 'b | Paren "('a, 'b) exp"
-
-codata ('a, 'b, 'c) some_killing =
- SK "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b, 'c) some_killing + ('a, 'b, 'c) in_here"
- and ('a, 'b, 'c) in_here =
- IH1 'b 'a | IH2 'c
-
-codata_raw some_killing': 'a = "'b \<Rightarrow> 'd \<Rightarrow> ('a + 'c)"
-and in_here': 'c = "'d + 'e"
-
-codata_raw some_killing'': 'a = "'b \<Rightarrow> 'c"
-and in_here'': 'c = "'d \<times> 'b + 'e"
-
-codata ('b, 'c) less_killing = LK "'b \<Rightarrow> 'c"
-
-codata 'b cps = CPS1 'b | CPS2 "'b \<Rightarrow> 'b cps"
-
-codata ('b1, 'b2, 'b3, 'b4, 'b5, 'b6, 'b7, 'b8, 'b9) fun_rhs =
- FR "'b1 \<Rightarrow> 'b2 \<Rightarrow> 'b3 \<Rightarrow> 'b4 \<Rightarrow> 'b5 \<Rightarrow> 'b6 \<Rightarrow> 'b7 \<Rightarrow> 'b8 \<Rightarrow> 'b9 \<Rightarrow>
- ('b1, 'b2, 'b3, 'b4, 'b5, 'b6, 'b7, 'b8, 'b9) fun_rhs"
-
-codata ('b1, 'b2, 'b3, 'b4, 'b5, 'b6, 'b7, 'b8, 'b9, 'b10, 'b11, 'b12, 'b13, 'b14, 'b15, 'b16, 'b17,
- 'b18, 'b19, 'b20) fun_rhs' =
- FR' "'b1 \<Rightarrow> 'b2 \<Rightarrow> 'b3 \<Rightarrow> 'b4 \<Rightarrow> 'b5 \<Rightarrow> 'b6 \<Rightarrow> 'b7 \<Rightarrow> 'b8 \<Rightarrow> 'b9 \<Rightarrow> 'b10 \<Rightarrow> 'b11 \<Rightarrow> 'b12 \<Rightarrow> 'b13 \<Rightarrow> 'b14 \<Rightarrow>
- 'b15 \<Rightarrow> 'b16 \<Rightarrow> 'b17 \<Rightarrow> 'b18 \<Rightarrow> 'b19 \<Rightarrow> 'b20 \<Rightarrow>
- ('b1, 'b2, 'b3, 'b4, 'b5, 'b6, 'b7, 'b8, 'b9, 'b10, 'b11, 'b12, 'b13, 'b14, 'b15, 'b16, 'b17,
- 'b18, 'b19, 'b20) fun_rhs'"
-
-codata ('a, 'b, 'c) wit3_F1 = W1 'a "('a, 'b, 'c) wit3_F1" "('a, 'b, 'c) wit3_F2"
- and ('a, 'b, 'c) wit3_F2 = W2 'b "('a, 'b, 'c) wit3_F2"
- and ('a, 'b, 'c) wit3_F3 = W31 'a 'b "('a, 'b, 'c) wit3_F1" | W32 'c 'a 'b "('a, 'b, 'c) wit3_F1"
-
-codata ('c, 'e, 'g) coind_wit1 =
- CW1 'c "('c, 'e, 'g) coind_wit1" "('c, 'e, 'g) ind_wit" "('c, 'e, 'g) coind_wit2"
- and ('c, 'e, 'g) coind_wit2 =
- CW21 "('c, 'e, 'g) coind_wit2" 'e | CW22 'c 'g
- and ('c, 'e, 'g) ind_wit =
- IW1 | IW2 'c
-
-codata ('b, 'a) bar = BAR "'a \<Rightarrow> 'b"
-codata ('a, 'b, 'c, 'd) foo = FOO "'d + 'b \<Rightarrow> 'c + 'a"
-
-codata 'a dead_foo = A
-codata ('a, 'b) use_dead_foo = Y "'a" "'b dead_foo"
-
-(* SLOW, MEMORY-HUNGRY
-codata ('a, 'c) D1 = A1 "('a, 'c) D2" | B1 "'a list"
- and ('a, 'c) D2 = A2 "('a, 'c) D3" | B2 "'c list"
- and ('a, 'c) D3 = A3 "('a, 'c) D3" | B3 "('a, 'c) D4" | C3 "('a, 'c) D4" "('a, 'c) D5"
- and ('a, 'c) D4 = A4 "('a, 'c) D5" | B4 "'a list list list"
- and ('a, 'c) D5 = A5 "('a, 'c) D6"
- and ('a, 'c) D6 = A6 "('a, 'c) D7"
- and ('a, 'c) D7 = A7 "('a, 'c) D8"
- and ('a, 'c) D8 = A8 "('a, 'c) D1 list"
-*)
-
-end
--- a/src/HOL/Codatatype/Examples/Misc_Data.thy Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,154 +0,0 @@
-(* Title: HOL/BNF/Examples/Misc_Data.thy
- Author: Dmitriy Traytel, TU Muenchen
- Author: Andrei Popescu, TU Muenchen
- Copyright 2012
-
-Miscellaneous datatype declarations.
-*)
-
-header {* Miscellaneous Datatype Declarations *}
-
-theory Misc_Data
-imports "../BNF"
-begin
-
-data simple = X1 | X2 | X3 | X4
-
-data simple' = X1' unit | X2' unit | X3' unit | X4' unit
-
-data 'a mylist = MyNil | MyCons 'a "'a mylist"
-
-data ('b, 'c, 'd, 'e) some_passive =
- SP1 "('b, 'c, 'd, 'e) some_passive" | SP2 'b | SP3 'c | SP4 'd | SP5 'e
-
-data lambda =
- Var string |
- App lambda lambda |
- Abs string lambda |
- Let "(string \<times> lambda) fset" lambda
-
-data 'a par_lambda =
- PVar 'a |
- PApp "'a par_lambda" "'a par_lambda" |
- PAbs 'a "'a par_lambda" |
- PLet "('a \<times> 'a par_lambda) fset" "'a par_lambda"
-
-(*
- ('a, 'b1, 'b2) F1 = 'a * 'b1 + 'a * 'b2
- ('a, 'b1, 'b2) F2 = unit + 'b1 * 'b2
-*)
-
-data 'a I1 = I11 'a "'a I1" | I12 'a "'a I2"
- and 'a I2 = I21 | I22 "'a I1" "'a I2"
-
-data 'a tree = TEmpty | TNode 'a "'a forest"
- and 'a forest = FNil | FCons "'a tree" "'a forest"
-
-data 'a tree' = TEmpty' | TNode' "'a branch" "'a branch"
- and 'a branch = Branch 'a "'a tree'"
-
-data ('a, 'b) exp = Term "('a, 'b) trm" | Sum "('a, 'b) trm" "('a, 'b) exp"
- and ('a, 'b) trm = Factor "('a, 'b) factor" | Prod "('a, 'b) factor" "('a, 'b) trm"
- and ('a, 'b) factor = C 'a | V 'b | Paren "('a, 'b) exp"
-
-data ('a, 'b, 'c) some_killing =
- SK "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b, 'c) some_killing + ('a, 'b, 'c) in_here"
- and ('a, 'b, 'c) in_here =
- IH1 'b 'a | IH2 'c
-
-data 'b nofail1 = NF11 "'b nofail1" 'b | NF12 'b
-data 'b nofail2 = NF2 "('b nofail2 \<times> 'b \<times> 'b nofail2 \<times> 'b) list"
-data 'b nofail3 = NF3 'b "('b nofail3 \<times> 'b \<times> 'b nofail3 \<times> 'b) fset"
-data 'b nofail4 = NF4 "('b nofail4 \<times> ('b nofail4 \<times> 'b \<times> 'b nofail4 \<times> 'b) fset) list"
-
-(*
-data 'b fail = F "'b fail" 'b "'b fail" "'b list"
-data 'b fail = F "'b fail" 'b "'b fail" 'b
-data 'b fail = F1 "'b fail" 'b | F2 "'b fail"
-data 'b fail = F "'b fail" 'b
-*)
-
-data l1 = L1 "l2 list"
- and l2 = L21 "l1 fset" | L22 l2
-
-data kk1 = KK1 kk2
- and kk2 = KK2 kk3
- and kk3 = KK3 "kk1 list"
-
-data t1 = T11 t3 | T12 t2
- and t2 = T2 t1
- and t3 = T3
-
-data t1' = T11' t2' | T12' t3'
- and t2' = T2' t1'
- and t3' = T3'
-
-(*
-data fail1 = F1 fail2
- and fail2 = F2 fail3
- and fail3 = F3 fail1
-
-data fail1 = F1 "fail2 list" fail2
- and fail2 = F2 "fail2 fset" fail3
- and fail3 = F3 fail1
-
-data fail1 = F1 "fail2 list" fail2
- and fail2 = F2 "fail1 fset" fail1
-*)
-
-(* SLOW
-data ('a, 'c) D1 = A1 "('a, 'c) D2" | B1 "'a list"
- and ('a, 'c) D2 = A2 "('a, 'c) D3" | B2 "'c list"
- and ('a, 'c) D3 = A3 "('a, 'c) D3" | B3 "('a, 'c) D4" | C3 "('a, 'c) D4" "('a, 'c) D5"
- and ('a, 'c) D4 = A4 "('a, 'c) D5" | B4 "'a list list list"
- and ('a, 'c) D5 = A5 "('a, 'c) D6"
- and ('a, 'c) D6 = A6 "('a, 'c) D7"
- and ('a, 'c) D7 = A7 "('a, 'c) D8"
- and ('a, 'c) D8 = A8 "('a, 'c) D1 list"
-
-(*time comparison*)
-datatype ('a, 'c) D1' = A1' "('a, 'c) D2'" | B1' "'a list"
- and ('a, 'c) D2' = A2' "('a, 'c) D3'" | B2' "'c list"
- and ('a, 'c) D3' = A3' "('a, 'c) D3'" | B3' "('a, 'c) D4'" | C3' "('a, 'c) D4'" "('a, 'c) D5'"
- and ('a, 'c) D4' = A4' "('a, 'c) D5'" | B4' "'a list list list"
- and ('a, 'c) D5' = A5' "('a, 'c) D6'"
- and ('a, 'c) D6' = A6' "('a, 'c) D7'"
- and ('a, 'c) D7' = A7' "('a, 'c) D8'"
- and ('a, 'c) D8' = A8' "('a, 'c) D1' list"
-*)
-
-(* fail:
-data tt1 = TT11 tt2 tt3 | TT12 tt2 tt4
- and tt2 = TT2
- and tt3 = TT3 tt4
- and tt4 = TT4 tt1
-*)
-
-data k1 = K11 k2 k3 | K12 k2 k4
- and k2 = K2
- and k3 = K3 k4
- and k4 = K4
-
-data tt1 = TT11 tt3 tt2 | TT12 tt2 tt4
- and tt2 = TT2
- and tt3 = TT3 tt1
- and tt4 = TT4
-
-(* SLOW
-data s1 = S11 s2 s3 s4 | S12 s3 | S13 s2 s6 | S14 s4 s2 | S15 s2 s2
- and s2 = S21 s7 s5 | S22 s5 s4 s6
- and s3 = S31 s1 s7 s2 | S32 s3 s3 | S33 s4 s5
- and s4 = S4 s5
- and s5 = S5
- and s6 = S61 s6 | S62 s1 s2 | S63 s6
- and s7 = S71 s8 | S72 s5
- and s8 = S8 nat
-*)
-
-data ('a, 'b) bar = Bar "'b \<Rightarrow> 'a"
-data ('a, 'b, 'c, 'd) foo = Foo "'d + 'b \<Rightarrow> 'c + 'a"
-
-data 'a dead_foo = A
-data ('a, 'b) use_dead_foo = Y "'a" "'b dead_foo"
-
-end
--- a/src/HOL/Codatatype/Examples/Process.thy Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,367 +0,0 @@
-(* Title: HOL/BNF/Examples/Process.thy
- Author: Andrei Popescu, TU Muenchen
- Copyright 2012
-
-Processes.
-*)
-
-header {* Processes *}
-
-theory Process
-imports "../BNF"
-begin
-
-hide_fact (open) Quotient_Product.prod_rel_def
-
-codata 'a process =
- isAction: Action (prefOf: 'a) (contOf: "'a process") |
- isChoice: Choice (ch1Of: "'a process") (ch2Of: "'a process")
-
-(* Read: prefix of, continuation of, choice 1 of, choice 2 of *)
-
-section {* Customization *}
-
-subsection {* Basic properties *}
-
-declare
- pre_process_rel_def[simp]
- sum_rel_def[simp]
- prod_rel_def[simp]
-
-(* Constructors versus discriminators *)
-theorem isAction_isChoice:
-"isAction p \<or> isChoice p"
-by (rule process.disc_exhaust) auto
-
-theorem not_isAction_isChoice: "\<not> (isAction p \<and> isChoice p)"
-by (cases rule: process.exhaust[of p]) auto
-
-
-subsection{* Coinduction *}
-
-theorem process_coind[elim, consumes 1, case_names iss Action Choice, induct pred: "HOL.eq"]:
-assumes phi: "\<phi> p p'" and
-iss: "\<And>p p'. \<phi> p p' \<Longrightarrow> (isAction p \<longleftrightarrow> isAction p') \<and> (isChoice p \<longleftrightarrow> isChoice p')" and
-Act: "\<And> a a' p p'. \<phi> (Action a p) (Action a' p') \<Longrightarrow> a = a' \<and> \<phi> p p'" and
-Ch: "\<And> p q p' q'. \<phi> (Choice p q) (Choice p' q') \<Longrightarrow> \<phi> p p' \<and> \<phi> q q'"
-shows "p = p'"
-proof(intro mp[OF process.rel_coinduct, of \<phi>, OF _ phi], clarify)
- fix p p' assume \<phi>: "\<phi> p p'"
- show "pre_process_rel (op =) \<phi> (process_dtor p) (process_dtor p')"
- proof(cases rule: process.exhaust[of p])
- case (Action a q) note p = Action
- hence "isAction p'" using iss[OF \<phi>] by (cases rule: process.exhaust[of p'], auto)
- then obtain a' q' where p': "p' = Action a' q'" by (cases rule: process.exhaust[of p'], auto)
- have 0: "a = a' \<and> \<phi> q q'" using Act[OF \<phi>[unfolded p p']] .
- have dtor: "process_dtor p = Inl (a,q)" "process_dtor p' = Inl (a',q')"
- unfolding p p' Action_def process.dtor_ctor by simp_all
- show ?thesis using 0 unfolding dtor by simp
- next
- case (Choice p1 p2) note p = Choice
- hence "isChoice p'" using iss[OF \<phi>] by (cases rule: process.exhaust[of p'], auto)
- then obtain p1' p2' where p': "p' = Choice p1' p2'"
- by (cases rule: process.exhaust[of p'], auto)
- have 0: "\<phi> p1 p1' \<and> \<phi> p2 p2'" using Ch[OF \<phi>[unfolded p p']] .
- have dtor: "process_dtor p = Inr (p1,p2)" "process_dtor p' = Inr (p1',p2')"
- unfolding p p' Choice_def process.dtor_ctor by simp_all
- show ?thesis using 0 unfolding dtor by simp
- qed
-qed
-
-(* Stronger coinduction, up to equality: *)
-theorem process_strong_coind[elim, consumes 1, case_names iss Action Choice]:
-assumes phi: "\<phi> p p'" and
-iss: "\<And>p p'. \<phi> p p' \<Longrightarrow> (isAction p \<longleftrightarrow> isAction p') \<and> (isChoice p \<longleftrightarrow> isChoice p')" and
-Act: "\<And> a a' p p'. \<phi> (Action a p) (Action a' p') \<Longrightarrow> a = a' \<and> (\<phi> p p' \<or> p = p')" and
-Ch: "\<And> p q p' q'. \<phi> (Choice p q) (Choice p' q') \<Longrightarrow> (\<phi> p p' \<or> p = p') \<and> (\<phi> q q' \<or> q = q')"
-shows "p = p'"
-proof(intro mp[OF process.rel_strong_coinduct, of \<phi>, OF _ phi], clarify)
- fix p p' assume \<phi>: "\<phi> p p'"
- show "pre_process_rel (op =) (\<lambda>a b. \<phi> a b \<or> a = b) (process_dtor p) (process_dtor p')"
- proof(cases rule: process.exhaust[of p])
- case (Action a q) note p = Action
- hence "isAction p'" using iss[OF \<phi>] by (cases rule: process.exhaust[of p'], auto)
- then obtain a' q' where p': "p' = Action a' q'" by (cases rule: process.exhaust[of p'], auto)
- have 0: "a = a' \<and> (\<phi> q q' \<or> q = q')" using Act[OF \<phi>[unfolded p p']] .
- have dtor: "process_dtor p = Inl (a,q)" "process_dtor p' = Inl (a',q')"
- unfolding p p' Action_def process.dtor_ctor by simp_all
- show ?thesis using 0 unfolding dtor by simp
- next
- case (Choice p1 p2) note p = Choice
- hence "isChoice p'" using iss[OF \<phi>] by (cases rule: process.exhaust[of p'], auto)
- then obtain p1' p2' where p': "p' = Choice p1' p2'"
- by (cases rule: process.exhaust[of p'], auto)
- have 0: "(\<phi> p1 p1' \<or> p1 = p1') \<and> (\<phi> p2 p2' \<or> p2 = p2')" using Ch[OF \<phi>[unfolded p p']] .
- have dtor: "process_dtor p = Inr (p1,p2)" "process_dtor p' = Inr (p1',p2')"
- unfolding p p' Choice_def process.dtor_ctor by simp_all
- show ?thesis using 0 unfolding dtor by simp
- qed
-qed
-
-
-subsection {* Coiteration (unfold) *}
-
-
-section{* Coinductive definition of the notion of trace *}
-
-(* Say we have a type of streams: *)
-
-typedecl 'a stream
-
-consts Ccons :: "'a \<Rightarrow> 'a stream \<Rightarrow> 'a stream"
-
-(* Use the existing coinductive package (distinct from our
-new codatatype package, but highly compatible with it): *)
-
-coinductive trace where
-"trace p as \<Longrightarrow> trace (Action a p) (Ccons a as)"
-|
-"trace p as \<or> trace q as \<Longrightarrow> trace (Choice p q) as"
-
-
-section{* Examples of corecursive definitions: *}
-
-subsection{* Single-guard fixpoint definition *}
-
-definition
-"BX \<equiv>
- process_unfold
- (\<lambda> P. True)
- (\<lambda> P. ''a'')
- (\<lambda> P. P)
- undefined
- undefined
- ()"
-
-lemma BX: "BX = Action ''a'' BX"
-unfolding BX_def
-using process.unfolds(1)[of "\<lambda> P. True" "()" "\<lambda> P. ''a''" "\<lambda> P. P"] by simp
-
-
-subsection{* Multi-guard fixpoint definitions, simulated with auxiliary arguments *}
-
-datatype x_y_ax = x | y | ax
-
-definition "isA \<equiv> \<lambda> K. case K of x \<Rightarrow> False |y \<Rightarrow> True |ax \<Rightarrow> True"
-definition "pr \<equiv> \<lambda> K. case K of x \<Rightarrow> undefined |y \<Rightarrow> ''b'' |ax \<Rightarrow> ''a''"
-definition "co \<equiv> \<lambda> K. case K of x \<Rightarrow> undefined |y \<Rightarrow> x |ax \<Rightarrow> x"
-lemmas Action_defs = isA_def pr_def co_def
-
-definition "c1 \<equiv> \<lambda> K. case K of x \<Rightarrow> ax |y \<Rightarrow> undefined |ax \<Rightarrow> undefined"
-definition "c2 \<equiv> \<lambda> K. case K of x \<Rightarrow> y |y \<Rightarrow> undefined |ax \<Rightarrow> undefined"
-lemmas Choice_defs = c1_def c2_def
-
-definition "F \<equiv> process_unfold isA pr co c1 c2"
-definition "X = F x" definition "Y = F y" definition "AX = F ax"
-
-lemma X_Y_AX: "X = Choice AX Y" "Y = Action ''b'' X" "AX = Action ''a'' X"
-unfolding X_def Y_def AX_def F_def
-using process.unfolds(2)[of isA x "pr" co c1 c2]
- process.unfolds(1)[of isA y "pr" co c1 c2]
- process.unfolds(1)[of isA ax "pr" co c1 c2]
-unfolding Action_defs Choice_defs by simp_all
-
-(* end product: *)
-lemma X_AX:
-"X = Choice AX (Action ''b'' X)"
-"AX = Action ''a'' X"
-using X_Y_AX by simp_all
-
-
-
-section{* Case study: Multi-guard fixpoint definitions, without auxiliary arguments *}
-
-hide_const x y ax X Y AX
-
-(* Process terms *)
-datatype ('a,'pvar) process_term =
- VAR 'pvar |
- PROC "'a process" |
- ACT 'a "('a,'pvar) process_term" | CH "('a,'pvar) process_term" "('a,'pvar) process_term"
-
-(* below, sys represents a system of equations *)
-fun isACT where
-"isACT sys (VAR X) =
- (case sys X of ACT a T \<Rightarrow> True |PROC p \<Rightarrow> isAction p |_ \<Rightarrow> False)"
-|
-"isACT sys (PROC p) = isAction p"
-|
-"isACT sys (ACT a T) = True"
-|
-"isACT sys (CH T1 T2) = False"
-
-fun PREF where
-"PREF sys (VAR X) =
- (case sys X of ACT a T \<Rightarrow> a | PROC p \<Rightarrow> prefOf p)"
-|
-"PREF sys (PROC p) = prefOf p"
-|
-"PREF sys (ACT a T) = a"
-
-fun CONT where
-"CONT sys (VAR X) =
- (case sys X of ACT a T \<Rightarrow> T | PROC p \<Rightarrow> PROC (contOf p))"
-|
-"CONT sys (PROC p) = PROC (contOf p)"
-|
-"CONT sys (ACT a T) = T"
-
-fun CH1 where
-"CH1 sys (VAR X) =
- (case sys X of CH T1 T2 \<Rightarrow> T1 |PROC p \<Rightarrow> PROC (ch1Of p))"
-|
-"CH1 sys (PROC p) = PROC (ch1Of p)"
-|
-"CH1 sys (CH T1 T2) = T1"
-
-fun CH2 where
-"CH2 sys (VAR X) =
- (case sys X of CH T1 T2 \<Rightarrow> T2 |PROC p \<Rightarrow> PROC (ch2Of p))"
-|
-"CH2 sys (PROC p) = PROC (ch2Of p)"
-|
-"CH2 sys (CH T1 T2) = T2"
-
-definition "guarded sys \<equiv> \<forall> X Y. sys X \<noteq> VAR Y"
-
-definition
-"solution sys \<equiv>
- process_unfold
- (isACT sys)
- (PREF sys)
- (CONT sys)
- (CH1 sys)
- (CH2 sys)"
-
-lemma solution_Action:
-assumes "isACT sys T"
-shows "solution sys T = Action (PREF sys T) (solution sys (CONT sys T))"
-unfolding solution_def
-using process.unfolds(1)[of "isACT sys" T "PREF sys" "CONT sys" "CH1 sys" "CH2 sys"]
- assms by simp
-
-lemma solution_Choice:
-assumes "\<not> isACT sys T"
-shows "solution sys T = Choice (solution sys (CH1 sys T)) (solution sys (CH2 sys T))"
-unfolding solution_def
-using process.unfolds(2)[of "isACT sys" T "PREF sys" "CONT sys" "CH1 sys" "CH2 sys"]
- assms by simp
-
-lemma isACT_VAR:
-assumes g: "guarded sys"
-shows "isACT sys (VAR X) \<longleftrightarrow> isACT sys (sys X)"
-using g unfolding guarded_def by (cases "sys X") auto
-
-lemma solution_VAR:
-assumes g: "guarded sys"
-shows "solution sys (VAR X) = solution sys (sys X)"
-proof(cases "isACT sys (VAR X)")
- case True
- hence T: "isACT sys (sys X)" unfolding isACT_VAR[OF g] .
- show ?thesis
- unfolding solution_Action[OF T] using solution_Action[of sys "VAR X"] True g
- unfolding guarded_def by (cases "sys X", auto)
-next
- case False note FFalse = False
- hence TT: "\<not> isACT sys (sys X)" unfolding isACT_VAR[OF g] .
- show ?thesis
- unfolding solution_Choice[OF TT] using solution_Choice[of sys "VAR X"] FFalse g
- unfolding guarded_def by (cases "sys X", auto)
-qed
-
-lemma solution_PROC[simp]:
-"solution sys (PROC p) = p"
-proof-
- {fix q assume "q = solution sys (PROC p)"
- hence "p = q"
- proof(induct rule: process_coind)
- case (iss p p')
- from isAction_isChoice[of p] show ?case
- proof
- assume p: "isAction p"
- hence 0: "isACT sys (PROC p)" by simp
- thus ?thesis using iss not_isAction_isChoice
- unfolding solution_Action[OF 0] by auto
- next
- assume "isChoice p"
- hence 0: "\<not> isACT sys (PROC p)"
- using not_isAction_isChoice by auto
- thus ?thesis using iss isAction_isChoice
- unfolding solution_Choice[OF 0] by auto
- qed
- next
- case (Action a a' p p')
- hence 0: "isACT sys (PROC (Action a p))" by simp
- show ?case using Action unfolding solution_Action[OF 0] by simp
- next
- case (Choice p q p' q')
- hence 0: "\<not> isACT sys (PROC (Choice p q))" using not_isAction_isChoice by auto
- show ?case using Choice unfolding solution_Choice[OF 0] by simp
- qed
- }
- thus ?thesis by metis
-qed
-
-lemma solution_ACT[simp]:
-"solution sys (ACT a T) = Action a (solution sys T)"
-by (metis CONT.simps(3) PREF.simps(3) isACT.simps(3) solution_Action)
-
-lemma solution_CH[simp]:
-"solution sys (CH T1 T2) = Choice (solution sys T1) (solution sys T2)"
-by (metis CH1.simps(3) CH2.simps(3) isACT.simps(4) solution_Choice)
-
-
-(* Example: *)
-
-fun sys where
-"sys 0 = CH (VAR (Suc 0)) (ACT ''b'' (VAR 0))"
-|
-"sys (Suc 0) = ACT ''a'' (VAR 0)"
-| (* dummy guarded term for variables outside the system: *)
-"sys X = ACT ''a'' (VAR 0)"
-
-lemma guarded_sys:
-"guarded sys"
-unfolding guarded_def proof (intro allI)
- fix X Y show "sys X \<noteq> VAR Y" by (cases X, simp, case_tac nat, auto)
-qed
-
-(* the actual processes: *)
-definition "x \<equiv> solution sys (VAR 0)"
-definition "ax \<equiv> solution sys (VAR (Suc 0))"
-
-(* end product: *)
-lemma x_ax:
-"x = Choice ax (Action ''b'' x)"
-"ax = Action ''a'' x"
-unfolding x_def ax_def by (subst solution_VAR[OF guarded_sys], simp)+
-
-
-(* Thanks to the inclusion of processes as process terms, one can
-also consider parametrized systems of equations---here, x is a (semantic)
-process parameter: *)
-
-fun sys' where
-"sys' 0 = CH (PROC x) (ACT ''b'' (VAR 0))"
-|
-"sys' (Suc 0) = CH (ACT ''a'' (VAR 0)) (PROC x)"
-| (* dummy guarded term : *)
-"sys' X = ACT ''a'' (VAR 0)"
-
-lemma guarded_sys':
-"guarded sys'"
-unfolding guarded_def proof (intro allI)
- fix X Y show "sys' X \<noteq> VAR Y" by (cases X, simp, case_tac nat, auto)
-qed
-
-(* the actual processes: *)
-definition "y \<equiv> solution sys' (VAR 0)"
-definition "ay \<equiv> solution sys' (VAR (Suc 0))"
-
-(* end product: *)
-lemma y_ay:
-"y = Choice x (Action ''b'' y)"
-"ay = Choice (Action ''a'' y) x"
-unfolding y_def ay_def by (subst solution_VAR[OF guarded_sys'], simp)+
-
-end
--- a/src/HOL/Codatatype/Examples/Stream.thy Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,157 +0,0 @@
-(* Title: HOL/BNF/Examples/Stream.thy
- Author: Dmitriy Traytel, TU Muenchen
- Author: Andrei Popescu, TU Muenchen
- Copyright 2012
-
-Infinite streams.
-*)
-
-header {* Infinite Streams *}
-
-theory Stream
-imports TreeFI
-begin
-
-hide_const (open) Quotient_Product.prod_rel
-hide_fact (open) Quotient_Product.prod_rel_def
-
-codata_raw stream: 's = "'a \<times> 's"
-
-(* selectors for streams *)
-definition "hdd as \<equiv> fst (stream_dtor as)"
-definition "tll as \<equiv> snd (stream_dtor as)"
-
-lemma unfold_pair_fun_hdd[simp]: "hdd (stream_dtor_unfold (f \<odot> g) t) = f t"
-unfolding hdd_def pair_fun_def stream.dtor_unfolds by simp
-
-lemma unfold_pair_fun_tll[simp]: "tll (stream_dtor_unfold (f \<odot> g) t) =
- stream_dtor_unfold (f \<odot> g) (g t)"
-unfolding tll_def pair_fun_def stream.dtor_unfolds by simp
-
-(* infinite trees: *)
-coinductive infiniteTr where
-"\<lbrakk>tr' \<in> listF_set (sub tr); infiniteTr tr'\<rbrakk> \<Longrightarrow> infiniteTr tr"
-
-lemma infiniteTr_strong_coind[consumes 1, case_names sub]:
-assumes *: "phi tr" and
-**: "\<And> tr. phi tr \<Longrightarrow> \<exists> tr' \<in> listF_set (sub tr). phi tr' \<or> infiniteTr tr'"
-shows "infiniteTr tr"
-using assms by (elim infiniteTr.coinduct) blast
-
-lemma infiniteTr_coind[consumes 1, case_names sub, induct pred: infiniteTr]:
-assumes *: "phi tr" and
-**: "\<And> tr. phi tr \<Longrightarrow> \<exists> tr' \<in> listF_set (sub tr). phi tr'"
-shows "infiniteTr tr"
-using assms by (elim infiniteTr.coinduct) blast
-
-lemma infiniteTr_sub[simp]:
-"infiniteTr tr \<Longrightarrow> (\<exists> tr' \<in> listF_set (sub tr). infiniteTr tr')"
-by (erule infiniteTr.cases) blast
-
-definition "konigPath \<equiv> stream_dtor_unfold
- (lab \<odot> (\<lambda>tr. SOME tr'. tr' \<in> listF_set (sub tr) \<and> infiniteTr tr'))"
-
-lemma hdd_simps1[simp]: "hdd (konigPath t) = lab t"
-unfolding konigPath_def by simp
-
-lemma tll_simps2[simp]: "tll (konigPath t) =
- konigPath (SOME tr. tr \<in> listF_set (sub t) \<and> infiniteTr tr)"
-unfolding konigPath_def by simp
-
-(* proper paths in trees: *)
-coinductive properPath where
-"\<lbrakk>hdd as = lab tr; tr' \<in> listF_set (sub tr); properPath (tll as) tr'\<rbrakk> \<Longrightarrow>
- properPath as tr"
-
-lemma properPath_strong_coind[consumes 1, case_names hdd_lab sub]:
-assumes *: "phi as tr" and
-**: "\<And> as tr. phi as tr \<Longrightarrow> hdd as = lab tr" and
-***: "\<And> as tr.
- phi as tr \<Longrightarrow>
- \<exists> tr' \<in> listF_set (sub tr). phi (tll as) tr' \<or> properPath (tll as) tr'"
-shows "properPath as tr"
-using assms by (elim properPath.coinduct) blast
-
-lemma properPath_coind[consumes 1, case_names hdd_lab sub, induct pred: properPath]:
-assumes *: "phi as tr" and
-**: "\<And> as tr. phi as tr \<Longrightarrow> hdd as = lab tr" and
-***: "\<And> as tr.
- phi as tr \<Longrightarrow>
- \<exists> tr' \<in> listF_set (sub tr). phi (tll as) tr'"
-shows "properPath as tr"
-using properPath_strong_coind[of phi, OF * **] *** by blast
-
-lemma properPath_hdd_lab:
-"properPath as tr \<Longrightarrow> hdd as = lab tr"
-by (erule properPath.cases) blast
-
-lemma properPath_sub:
-"properPath as tr \<Longrightarrow>
- \<exists> tr' \<in> listF_set (sub tr). phi (tll as) tr' \<or> properPath (tll as) tr'"
-by (erule properPath.cases) blast
-
-(* prove the following by coinduction *)
-theorem Konig:
- assumes "infiniteTr tr"
- shows "properPath (konigPath tr) tr"
-proof-
- {fix as
- assume "infiniteTr tr \<and> as = konigPath tr" hence "properPath as tr"
- proof (induct rule: properPath_coind, safe)
- fix t
- let ?t = "SOME t'. t' \<in> listF_set (sub t) \<and> infiniteTr t'"
- assume "infiniteTr t"
- hence "\<exists>t' \<in> listF_set (sub t). infiniteTr t'" by simp
- hence "\<exists>t'. t' \<in> listF_set (sub t) \<and> infiniteTr t'" by blast
- hence "?t \<in> listF_set (sub t) \<and> infiniteTr ?t" by (elim someI_ex)
- moreover have "tll (konigPath t) = konigPath ?t" by simp
- ultimately show "\<exists>t' \<in> listF_set (sub t).
- infiniteTr t' \<and> tll (konigPath t) = konigPath t'" by blast
- qed simp
- }
- thus ?thesis using assms by blast
-qed
-
-(* some more stream theorems *)
-
-lemma stream_map[simp]: "stream_map f = stream_dtor_unfold (f o hdd \<odot> tll)"
-unfolding stream_map_def pair_fun_def hdd_def[abs_def] tll_def[abs_def]
- map_pair_def o_def prod_case_beta by simp
-
-lemma prod_rel[simp]: "prod_rel \<phi>1 \<phi>2 a b = (\<phi>1 (fst a) (fst b) \<and> \<phi>2 (snd a) (snd b))"
-unfolding prod_rel_def by auto
-
-lemmas stream_coind =
- mp[OF stream.rel_coinduct, unfolded prod_rel[abs_def], folded hdd_def tll_def]
-
-definition plus :: "nat stream \<Rightarrow> nat stream \<Rightarrow> nat stream" (infixr "\<oplus>" 66) where
- [simp]: "plus xs ys =
- stream_dtor_unfold ((%(xs, ys). hdd xs + hdd ys) \<odot> (%(xs, ys). (tll xs, tll ys))) (xs, ys)"
-
-definition scalar :: "nat \<Rightarrow> nat stream \<Rightarrow> nat stream" (infixr "\<cdot>" 68) where
- [simp]: "scalar n = stream_map (\<lambda>x. n * x)"
-
-definition ones :: "nat stream" where [simp]: "ones = stream_dtor_unfold ((%x. 1) \<odot> id) ()"
-definition twos :: "nat stream" where [simp]: "twos = stream_dtor_unfold ((%x. 2) \<odot> id) ()"
-definition ns :: "nat \<Rightarrow> nat stream" where [simp]: "ns n = scalar n ones"
-
-lemma "ones \<oplus> ones = twos"
-by (intro stream_coind[where P="%x1 x2. \<exists>x. x1 = ones \<oplus> ones \<and> x2 = twos"]) auto
-
-lemma "n \<cdot> twos = ns (2 * n)"
-by (intro stream_coind[where P="%x1 x2. \<exists>n. x1 = n \<cdot> twos \<and> x2 = ns (2 * n)"]) force+
-
-lemma prod_scalar: "(n * m) \<cdot> xs = n \<cdot> m \<cdot> xs"
-by (intro stream_coind[where P="%x1 x2. \<exists>n m xs. x1 = (n * m) \<cdot> xs \<and> x2 = n \<cdot> m \<cdot> xs"]) force+
-
-lemma scalar_plus: "n \<cdot> (xs \<oplus> ys) = n \<cdot> xs \<oplus> n \<cdot> ys"
-by (intro stream_coind[where P="%x1 x2. \<exists>n xs ys. x1 = n \<cdot> (xs \<oplus> ys) \<and> x2 = n \<cdot> xs \<oplus> n \<cdot> ys"])
- (force simp: add_mult_distrib2)+
-
-lemma plus_comm: "xs \<oplus> ys = ys \<oplus> xs"
-by (intro stream_coind[where P="%x1 x2. \<exists>xs ys. x1 = xs \<oplus> ys \<and> x2 = ys \<oplus> xs"]) force+
-
-lemma plus_assoc: "(xs \<oplus> ys) \<oplus> zs = xs \<oplus> ys \<oplus> zs"
-by (intro stream_coind[where P="%x1 x2. \<exists>xs ys zs. x1 = (xs \<oplus> ys) \<oplus> zs \<and> x2 = xs \<oplus> ys \<oplus> zs"]) force+
-
-end
--- a/src/HOL/Codatatype/Examples/TreeFI.thy Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,83 +0,0 @@
-(* Title: HOL/BNF/Examples/TreeFI.thy
- Author: Dmitriy Traytel, TU Muenchen
- Author: Andrei Popescu, TU Muenchen
- Copyright 2012
-
-Finitely branching possibly infinite trees.
-*)
-
-header {* Finitely Branching Possibly Infinite Trees *}
-
-theory TreeFI
-imports ListF
-begin
-
-hide_const (open) Sublist.sub
-
-codata_raw treeFI: 'tree = "'a \<times> 'tree listF"
-
-lemma pre_treeFI_listF_set[simp]: "pre_treeFI_set2 (i, xs) = listF_set xs"
-unfolding pre_treeFI_set2_def collect_def[abs_def] prod_set_defs
-by (auto simp add: listF.set_natural')
-
-(* selectors for trees *)
-definition "lab tr \<equiv> fst (treeFI_dtor tr)"
-definition "sub tr \<equiv> snd (treeFI_dtor tr)"
-
-lemma dtor[simp]: "treeFI_dtor tr = (lab tr, sub tr)"
-unfolding lab_def sub_def by simp
-
-definition pair_fun (infixr "\<odot>" 50) where
- "f \<odot> g \<equiv> \<lambda>x. (f x, g x)"
-
-lemma unfold_pair_fun_lab: "lab (treeFI_dtor_unfold (f \<odot> g) t) = f t"
-unfolding lab_def pair_fun_def treeFI.dtor_unfolds pre_treeFI_map_def by simp
-
-lemma unfold_pair_fun_sub: "sub (treeFI_dtor_unfold (f \<odot> g) t) = listF_map (treeFI_dtor_unfold (f \<odot> g)) (g t)"
-unfolding sub_def pair_fun_def treeFI.dtor_unfolds pre_treeFI_map_def by simp
-
-(* Tree reverse:*)
-definition "trev \<equiv> treeFI_dtor_unfold (lab \<odot> lrev o sub)"
-
-lemma trev_simps1[simp]: "lab (trev t) = lab t"
-unfolding trev_def by (simp add: unfold_pair_fun_lab)
-
-lemma trev_simps2[simp]: "sub (trev t) = listF_map trev (lrev (sub t))"
-unfolding trev_def by (simp add: unfold_pair_fun_sub)
-
-lemma treeFI_coinduct:
-assumes *: "phi x y"
-and step: "\<And>a b. phi a b \<Longrightarrow>
- lab a = lab b \<and>
- lengthh (sub a) = lengthh (sub b) \<and>
- (\<forall>i < lengthh (sub a). phi (nthh (sub a) i) (nthh (sub b) i))"
-shows "x = y"
-proof (rule mp[OF treeFI.dtor_coinduct, of phi, OF _ *])
- fix a b :: "'a treeFI"
- let ?zs = "zipp (sub a) (sub b)"
- let ?z = "(lab a, ?zs)"
- assume "phi a b"
- with step have step': "lab a = lab b" "lengthh (sub a) = lengthh (sub b)"
- "\<forall>i < lengthh (sub a). phi (nthh (sub a) i) (nthh (sub b) i)" by auto
- hence "pre_treeFI_map id fst ?z = treeFI_dtor a" "pre_treeFI_map id snd ?z = treeFI_dtor b"
- unfolding pre_treeFI_map_def by auto
- moreover have "\<forall>(x, y) \<in> pre_treeFI_set2 ?z. phi x y"
- proof safe
- fix z1 z2
- assume "(z1, z2) \<in> pre_treeFI_set2 ?z"
- hence "(z1, z2) \<in> listF_set ?zs" by auto
- hence "\<exists>i < lengthh ?zs. nthh ?zs i = (z1, z2)" by auto
- with step'(2) obtain i where "i < lengthh (sub a)"
- "nthh (sub a) i = z1" "nthh (sub b) i = z2" by auto
- with step'(3) show "phi z1 z2" by auto
- qed
- ultimately show "\<exists>z.
- (pre_treeFI_map id fst z = treeFI_dtor a \<and>
- pre_treeFI_map id snd z = treeFI_dtor b) \<and>
- (\<forall>x y. (x, y) \<in> pre_treeFI_set2 z \<longrightarrow> phi x y)" by blast
-qed
-
-lemma trev_trev: "trev (trev tr) = tr"
-by (rule treeFI_coinduct[of "%a b. trev (trev b) = a"]) auto
-
-end
--- a/src/HOL/Codatatype/Examples/TreeFsetI.thy Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,59 +0,0 @@
-(* Title: HOL/BNF/Examples/TreeFsetI.thy
- Author: Dmitriy Traytel, TU Muenchen
- Author: Andrei Popescu, TU Muenchen
- Copyright 2012
-
-Finitely branching possibly infinite trees, with sets of children.
-*)
-
-header {* Finitely Branching Possibly Infinite Trees, with Sets of Children *}
-
-theory TreeFsetI
-imports "../BNF"
-begin
-
-hide_const (open) Sublist.sub
-hide_fact (open) Quotient_Product.prod_rel_def
-
-definition pair_fun (infixr "\<odot>" 50) where
- "f \<odot> g \<equiv> \<lambda>x. (f x, g x)"
-
-codata_raw treeFsetI: 't = "'a \<times> 't fset"
-
-(* selectors for trees *)
-definition "lab t \<equiv> fst (treeFsetI_dtor t)"
-definition "sub t \<equiv> snd (treeFsetI_dtor t)"
-
-lemma dtor[simp]: "treeFsetI_dtor t = (lab t, sub t)"
-unfolding lab_def sub_def by simp
-
-lemma unfold_pair_fun_lab: "lab (treeFsetI_dtor_unfold (f \<odot> g) t) = f t"
-unfolding lab_def pair_fun_def treeFsetI.dtor_unfolds pre_treeFsetI_map_def by simp
-
-lemma unfold_pair_fun_sub: "sub (treeFsetI_dtor_unfold (f \<odot> g) t) = map_fset (treeFsetI_dtor_unfold (f \<odot> g)) (g t)"
-unfolding sub_def pair_fun_def treeFsetI.dtor_unfolds pre_treeFsetI_map_def by simp
-
-(* tree map (contrived example): *)
-definition "tmap f \<equiv> treeFsetI_dtor_unfold (f o lab \<odot> sub)"
-
-lemma tmap_simps1[simp]: "lab (tmap f t) = f (lab t)"
-unfolding tmap_def by (simp add: unfold_pair_fun_lab)
-
-lemma trev_simps2[simp]: "sub (tmap f t) = map_fset (tmap f) (sub t)"
-unfolding tmap_def by (simp add: unfold_pair_fun_sub)
-
-lemma pre_treeFsetI_rel[simp]: "pre_treeFsetI_rel R1 R2 a b = (R1 (fst a) (fst b) \<and>
- (\<forall>t \<in> fset (snd a). (\<exists>u \<in> fset (snd b). R2 t u)) \<and>
- (\<forall>t \<in> fset (snd b). (\<exists>u \<in> fset (snd a). R2 u t)))"
-apply (cases a)
-apply (cases b)
-apply (simp add: pre_treeFsetI_rel_def prod_rel_def fset_rel_def)
-done
-
-lemmas treeFsetI_coind = mp[OF treeFsetI.rel_coinduct]
-
-lemma "tmap (f o g) x = tmap f (tmap g x)"
-by (intro treeFsetI_coind[where P="%x1 x2. \<exists>x. x1 = tmap (f o g) x \<and> x2 = tmap f (tmap g x)"])
- force+
-
-end
--- a/src/HOL/Codatatype/More_BNFs.thy Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1511 +0,0 @@
-(* Title: HOL/BNF/More_BNFs.thy
- Author: Dmitriy Traytel, TU Muenchen
- Author: Andrei Popescu, TU Muenchen
- Author: Andreas Lochbihler, Karlsruhe Institute of Technology
- Author: Jasmin Blanchette, TU Muenchen
- Copyright 2012
-
-Registration of various types as bounded natural functors.
-*)
-
-header {* Registration of Various Types as Bounded Natural Functors *}
-
-theory More_BNFs
-imports
- BNF_LFP
- BNF_GFP
- "~~/src/HOL/Quotient_Examples/FSet"
- "~~/src/HOL/Library/Multiset"
- Countable_Set
-begin
-
-lemma option_rec_conv_option_case: "option_rec = option_case"
-by (simp add: fun_eq_iff split: option.split)
-
-definition option_rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a option \<Rightarrow> 'b option \<Rightarrow> bool" where
-"option_rel R x_opt y_opt =
- (case (x_opt, y_opt) of
- (None, None) \<Rightarrow> True
- | (Some x, Some y) \<Rightarrow> R x y
- | _ \<Rightarrow> False)"
-
-bnf_def Option.map [Option.set] "\<lambda>_::'a option. natLeq" ["None"] option_rel
-proof -
- show "Option.map id = id" by (simp add: fun_eq_iff Option.map_def split: option.split)
-next
- fix f g
- show "Option.map (g \<circ> f) = Option.map g \<circ> Option.map f"
- by (auto simp add: fun_eq_iff Option.map_def split: option.split)
-next
- fix f g x
- assume "\<And>z. z \<in> Option.set x \<Longrightarrow> f z = g z"
- thus "Option.map f x = Option.map g x"
- by (simp cong: Option.map_cong)
-next
- fix f
- show "Option.set \<circ> Option.map f = op ` f \<circ> Option.set"
- by fastforce
-next
- show "card_order natLeq" by (rule natLeq_card_order)
-next
- show "cinfinite natLeq" by (rule natLeq_cinfinite)
-next
- fix x
- show "|Option.set x| \<le>o natLeq"
- by (cases x) (simp_all add: ordLess_imp_ordLeq finite_iff_ordLess_natLeq[symmetric])
-next
- fix A
- have unfold: "{x. Option.set x \<subseteq> A} = Some ` A \<union> {None}"
- by (auto simp add: option_rec_conv_option_case Option.set_def split: option.split_asm)
- show "|{x. Option.set x \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq"
- apply (rule ordIso_ordLeq_trans)
- apply (rule card_of_ordIso_subst[OF unfold])
- apply (rule ordLeq_transitive)
- apply (rule Un_csum)
- apply (rule ordLeq_transitive)
- apply (rule csum_mono)
- apply (rule card_of_image)
- apply (rule ordIso_ordLeq_trans)
- apply (rule single_cone)
- apply (rule cone_ordLeq_ctwo)
- apply (rule ordLeq_cexp1)
- apply (simp_all add: natLeq_cinfinite natLeq_Card_order cinfinite_not_czero Card_order_csum)
- done
-next
- fix A B1 B2 f1 f2 p1 p2
- assume wpull: "wpull A B1 B2 f1 f2 p1 p2"
- show "wpull {x. Option.set x \<subseteq> A} {x. Option.set x \<subseteq> B1} {x. Option.set x \<subseteq> B2}
- (Option.map f1) (Option.map f2) (Option.map p1) (Option.map p2)"
- (is "wpull ?A ?B1 ?B2 ?f1 ?f2 ?p1 ?p2")
- unfolding wpull_def
- proof (intro strip, elim conjE)
- fix b1 b2
- assume "b1 \<in> ?B1" "b2 \<in> ?B2" "?f1 b1 = ?f2 b2"
- thus "\<exists>a \<in> ?A. ?p1 a = b1 \<and> ?p2 a = b2" using wpull
- unfolding wpull_def by (cases b2) (auto 4 5)
- qed
-next
- fix z
- assume "z \<in> Option.set None"
- thus False by simp
-next
- fix R
- show "{p. option_rel (\<lambda>x y. (x, y) \<in> R) (fst p) (snd p)} =
- (Gr {x. Option.set x \<subseteq> R} (Option.map fst))\<inverse> O Gr {x. Option.set x \<subseteq> R} (Option.map snd)"
- unfolding option_rel_def Gr_def relcomp_unfold converse_unfold
- by (auto simp: trans[OF eq_commute option_map_is_None] trans[OF eq_commute option_map_eq_Some]
- split: option.splits) blast
-qed
-
-lemma card_of_list_in:
- "|{xs. set xs \<subseteq> A}| \<le>o |Pfunc (UNIV :: nat set) A|" (is "|?LHS| \<le>o |?RHS|")
-proof -
- let ?f = "%xs. %i. if i < length xs \<and> set xs \<subseteq> A then Some (nth xs i) else None"
- have "inj_on ?f ?LHS" unfolding inj_on_def fun_eq_iff
- proof safe
- fix xs :: "'a list" and ys :: "'a list"
- assume su: "set xs \<subseteq> A" "set ys \<subseteq> A" and eq: "\<forall>i. ?f xs i = ?f ys i"
- hence *: "length xs = length ys"
- by (metis linorder_cases option.simps(2) order_less_irrefl)
- thus "xs = ys" by (rule nth_equalityI) (metis * eq su option.inject)
- qed
- moreover have "?f ` ?LHS \<subseteq> ?RHS" unfolding Pfunc_def by fastforce
- ultimately show ?thesis using card_of_ordLeq by blast
-qed
-
-lemma list_in_empty: "A = {} \<Longrightarrow> {x. set x \<subseteq> A} = {[]}"
-by simp
-
-lemma card_of_Func: "|Func A B| =o |B| ^c |A|"
-unfolding cexp_def Field_card_of by (rule card_of_refl)
-
-lemma not_emp_czero_notIn_ordIso_Card_order:
-"A \<noteq> {} \<Longrightarrow> ( |A|, czero) \<notin> ordIso \<and> Card_order |A|"
- apply (rule conjI)
- apply (metis Field_card_of czeroE)
- by (rule card_of_Card_order)
-
-lemma list_in_bd: "|{x. set x \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq"
-proof -
- fix A :: "'a set"
- show "|{x. set x \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq"
- proof (cases "A = {}")
- case False thus ?thesis
- apply -
- apply (rule ordLeq_transitive)
- apply (rule card_of_list_in)
- apply (rule ordLeq_transitive)
- apply (erule card_of_Pfunc_Pow_Func)
- apply (rule ordIso_ordLeq_trans)
- apply (rule Times_cprod)
- apply (rule cprod_cinfinite_bound)
- apply (rule ordIso_ordLeq_trans)
- apply (rule Pow_cexp_ctwo)
- apply (rule ordIso_ordLeq_trans)
- apply (rule cexp_cong2)
- apply (rule card_of_nat)
- apply (rule Card_order_ctwo)
- apply (rule card_of_Card_order)
- apply (rule natLeq_Card_order)
- apply (rule disjI1)
- apply (rule ctwo_Cnotzero)
- apply (rule cexp_mono1)
- apply (rule ordLeq_csum2)
- apply (rule Card_order_ctwo)
- apply (rule disjI1)
- apply (rule ctwo_Cnotzero)
- apply (rule natLeq_Card_order)
- apply (rule ordIso_ordLeq_trans)
- apply (rule card_of_Func)
- apply (rule ordIso_ordLeq_trans)
- apply (rule cexp_cong2)
- apply (rule card_of_nat)
- apply (rule card_of_Card_order)
- apply (rule card_of_Card_order)
- apply (rule natLeq_Card_order)
- apply (rule disjI1)
- apply (erule not_emp_czero_notIn_ordIso_Card_order)
- apply (rule cexp_mono1)
- apply (rule ordLeq_csum1)
- apply (rule card_of_Card_order)
- apply (rule disjI1)
- apply (erule not_emp_czero_notIn_ordIso_Card_order)
- apply (rule natLeq_Card_order)
- apply (rule card_of_Card_order)
- apply (rule card_of_Card_order)
- apply (rule Cinfinite_cexp)
- apply (rule ordLeq_csum2)
- apply (rule Card_order_ctwo)
- apply (rule conjI)
- apply (rule natLeq_cinfinite)
- by (rule natLeq_Card_order)
- next
- case True thus ?thesis
- apply -
- apply (rule ordIso_ordLeq_trans)
- apply (rule card_of_ordIso_subst)
- apply (erule list_in_empty)
- apply (rule ordIso_ordLeq_trans)
- apply (rule single_cone)
- apply (rule cone_ordLeq_cexp)
- apply (rule ordLeq_transitive)
- apply (rule cone_ordLeq_ctwo)
- apply (rule ordLeq_csum2)
- by (rule Card_order_ctwo)
- qed
-qed
-
-bnf_def map [set] "\<lambda>_::'a list. natLeq" ["[]"]
-proof -
- show "map id = id" by (rule List.map.id)
-next
- fix f g
- show "map (g o f) = map g o map f" by (rule List.map.comp[symmetric])
-next
- fix x f g
- assume "\<And>z. z \<in> set x \<Longrightarrow> f z = g z"
- thus "map f x = map g x" by simp
-next
- fix f
- show "set o map f = image f o set" by (rule ext, unfold o_apply, rule set_map)
-next
- show "card_order natLeq" by (rule natLeq_card_order)
-next
- show "cinfinite natLeq" by (rule natLeq_cinfinite)
-next
- fix x
- show "|set x| \<le>o natLeq"
- apply (rule ordLess_imp_ordLeq)
- apply (rule finite_ordLess_infinite[OF _ natLeq_Well_order])
- unfolding Field_natLeq Field_card_of by (auto simp: card_of_well_order_on)
-next
- fix A :: "'a set"
- show "|{x. set x \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq" by (rule list_in_bd)
-next
- fix A B1 B2 f1 f2 p1 p2
- assume "wpull A B1 B2 f1 f2 p1 p2"
- hence pull: "\<And>b1 b2. b1 \<in> B1 \<and> b2 \<in> B2 \<and> f1 b1 = f2 b2 \<Longrightarrow> \<exists>a \<in> A. p1 a = b1 \<and> p2 a = b2"
- unfolding wpull_def by auto
- show "wpull {x. set x \<subseteq> A} {x. set x \<subseteq> B1} {x. set x \<subseteq> B2} (map f1) (map f2) (map p1) (map p2)"
- (is "wpull ?A ?B1 ?B2 _ _ _ _")
- proof (unfold wpull_def)
- { fix as bs assume *: "as \<in> ?B1" "bs \<in> ?B2" "map f1 as = map f2 bs"
- hence "length as = length bs" by (metis length_map)
- hence "\<exists>zs \<in> ?A. map p1 zs = as \<and> map p2 zs = bs" using *
- proof (induct as bs rule: list_induct2)
- case (Cons a as b bs)
- hence "a \<in> B1" "b \<in> B2" "f1 a = f2 b" by auto
- with pull obtain z where "z \<in> A" "p1 z = a" "p2 z = b" by blast
- moreover
- from Cons obtain zs where "zs \<in> ?A" "map p1 zs = as" "map p2 zs = bs" by auto
- ultimately have "z # zs \<in> ?A" "map p1 (z # zs) = a # as \<and> map p2 (z # zs) = b # bs" by auto
- thus ?case by (rule_tac x = "z # zs" in bexI)
- qed simp
- }
- thus "\<forall>as bs. as \<in> ?B1 \<and> bs \<in> ?B2 \<and> map f1 as = map f2 bs \<longrightarrow>
- (\<exists>zs \<in> ?A. map p1 zs = as \<and> map p2 zs = bs)" by blast
- qed
-qed simp+
-
-(* Finite sets *)
-abbreviation afset where "afset \<equiv> abs_fset"
-abbreviation rfset where "rfset \<equiv> rep_fset"
-
-lemma fset_fset_member:
-"fset A = {a. a |\<in>| A}"
-unfolding fset_def fset_member_def by auto
-
-lemma afset_rfset:
-"afset (rfset x) = x"
-by (rule Quotient_fset[unfolded Quotient_def, THEN conjunct1, rule_format])
-
-lemma afset_rfset_id:
-"afset o rfset = id"
-unfolding comp_def afset_rfset id_def ..
-
-lemma rfset:
-"rfset A = rfset B \<longleftrightarrow> A = B"
-by (metis afset_rfset)
-
-lemma afset_set:
-"afset as = afset bs \<longleftrightarrow> set as = set bs"
-using Quotient_fset unfolding Quotient_def list_eq_def by auto
-
-lemma surj_afset:
-"\<exists> as. A = afset as"
-by (metis afset_rfset)
-
-lemma fset_def2:
-"fset = set o rfset"
-unfolding fset_def map_fun_def[abs_def] by simp
-
-lemma fset_def2_raw:
-"fset A = set (rfset A)"
-unfolding fset_def2 by simp
-
-lemma fset_comp_afset:
-"fset o afset = set"
-unfolding fset_def2 comp_def apply(rule ext)
-unfolding afset_set[symmetric] afset_rfset ..
-
-lemma fset_afset:
-"fset (afset as) = set as"
-unfolding fset_comp_afset[symmetric] by simp
-
-lemma set_rfset_afset:
-"set (rfset (afset as)) = set as"
-unfolding afset_set[symmetric] afset_rfset ..
-
-lemma map_fset_comp_afset:
-"(map_fset f) o afset = afset o (map f)"
-unfolding map_fset_def map_fun_def[abs_def] comp_def apply(rule ext)
-unfolding afset_set set_map set_rfset_afset id_apply ..
-
-lemma map_fset_afset:
-"(map_fset f) (afset as) = afset (map f as)"
-using map_fset_comp_afset unfolding comp_def fun_eq_iff by auto
-
-lemma fset_map_fset:
-"fset (map_fset f A) = (image f) (fset A)"
-apply(subst afset_rfset[symmetric, of A])
-unfolding map_fset_afset fset_afset set_map
-unfolding fset_def2_raw ..
-
-lemma map_fset_def2:
-"map_fset f = afset o (map f) o rfset"
-unfolding map_fset_def map_fun_def[abs_def] by simp
-
-lemma map_fset_def2_raw:
-"map_fset f A = afset (map f (rfset A))"
-unfolding map_fset_def2 by simp
-
-lemma finite_ex_fset:
-assumes "finite A"
-shows "\<exists> B. fset B = A"
-by (metis assms finite_list fset_afset)
-
-lemma wpull_image:
-assumes "wpull A B1 B2 f1 f2 p1 p2"
-shows "wpull (Pow A) (Pow B1) (Pow B2) (image f1) (image f2) (image p1) (image p2)"
-unfolding wpull_def Pow_def Bex_def mem_Collect_eq proof clarify
- fix Y1 Y2 assume Y1: "Y1 \<subseteq> B1" and Y2: "Y2 \<subseteq> B2" and EQ: "f1 ` Y1 = f2 ` Y2"
- def X \<equiv> "{a \<in> A. p1 a \<in> Y1 \<and> p2 a \<in> Y2}"
- show "\<exists>X\<subseteq>A. p1 ` X = Y1 \<and> p2 ` X = Y2"
- proof (rule exI[of _ X], intro conjI)
- show "p1 ` X = Y1"
- proof
- show "Y1 \<subseteq> p1 ` X"
- proof safe
- fix y1 assume y1: "y1 \<in> Y1"
- then obtain y2 where y2: "y2 \<in> Y2" and eq: "f1 y1 = f2 y2" using EQ by auto
- then obtain x where "x \<in> A" and "p1 x = y1" and "p2 x = y2"
- using assms y1 Y1 Y2 unfolding wpull_def by blast
- thus "y1 \<in> p1 ` X" unfolding X_def using y1 y2 by auto
- qed
- qed(unfold X_def, auto)
- show "p2 ` X = Y2"
- proof
- show "Y2 \<subseteq> p2 ` X"
- proof safe
- fix y2 assume y2: "y2 \<in> Y2"
- then obtain y1 where y1: "y1 \<in> Y1" and eq: "f1 y1 = f2 y2" using EQ by force
- then obtain x where "x \<in> A" and "p1 x = y1" and "p2 x = y2"
- using assms y2 Y1 Y2 unfolding wpull_def by blast
- thus "y2 \<in> p2 ` X" unfolding X_def using y1 y2 by auto
- qed
- qed(unfold X_def, auto)
- qed(unfold X_def, auto)
-qed
-
-lemma fset_to_fset: "finite A \<Longrightarrow> fset (the_inv fset A) = A"
-by (rule f_the_inv_into_f) (auto simp: inj_on_def fset_cong dest!: finite_ex_fset)
-
-definition fset_rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> 'b fset \<Rightarrow> bool" where
-"fset_rel R a b \<longleftrightarrow>
- (\<forall>t \<in> fset a. \<exists>u \<in> fset b. R t u) \<and>
- (\<forall>t \<in> fset b. \<exists>u \<in> fset a. R u t)"
-
-lemma fset_rel_aux:
-"(\<forall>t \<in> fset a. \<exists>u \<in> fset b. R t u) \<and> (\<forall>u \<in> fset b. \<exists>t \<in> fset a. R t u) \<longleftrightarrow>
- (a, b) \<in> (Gr {a. fset a \<subseteq> {(a, b). R a b}} (map_fset fst))\<inverse> O
- Gr {a. fset a \<subseteq> {(a, b). R a b}} (map_fset snd)" (is "?L = ?R")
-proof
- assume ?L
- def R' \<equiv> "the_inv fset (Collect (split R) \<inter> (fset a \<times> fset b))" (is "the_inv fset ?L'")
- have "finite ?L'" by (intro finite_Int[OF disjI2] finite_cartesian_product) auto
- hence *: "fset R' = ?L'" unfolding R'_def by (intro fset_to_fset)
- show ?R unfolding Gr_def relcomp_unfold converse_unfold
- proof (intro CollectI prod_caseI exI conjI)
- from * show "(R', a) = (R', map_fset fst R')" using conjunct1[OF `?L`]
- by (auto simp add: fset_cong[symmetric] image_def Int_def split: prod.splits)
- from * show "(R', b) = (R', map_fset snd R')" using conjunct2[OF `?L`]
- by (auto simp add: fset_cong[symmetric] image_def Int_def split: prod.splits)
- qed (auto simp add: *)
-next
- assume ?R thus ?L unfolding Gr_def relcomp_unfold converse_unfold
- apply (simp add: subset_eq Ball_def)
- apply (rule conjI)
- apply (clarsimp, metis snd_conv)
- by (clarsimp, metis fst_conv)
-qed
-
-bnf_def map_fset [fset] "\<lambda>_::'a fset. natLeq" ["{||}"] fset_rel
-proof -
- show "map_fset id = id"
- unfolding map_fset_def2 map_id o_id afset_rfset_id ..
-next
- fix f g
- show "map_fset (g o f) = map_fset g o map_fset f"
- unfolding map_fset_def2 map.comp[symmetric] comp_def apply(rule ext)
- unfolding afset_set set_map fset_def2_raw[symmetric] image_image[symmetric]
- unfolding map_fset_afset[symmetric] map_fset_image afset_rfset
- by (rule refl)
-next
- fix x f g
- assume "\<And>z. z \<in> fset x \<Longrightarrow> f z = g z"
- hence "map f (rfset x) = map g (rfset x)"
- apply(intro map_cong) unfolding fset_def2_raw by auto
- thus "map_fset f x = map_fset g x" unfolding map_fset_def2_raw
- by (rule arg_cong)
-next
- fix f
- show "fset o map_fset f = image f o fset"
- unfolding comp_def fset_map_fset ..
-next
- show "card_order natLeq" by (rule natLeq_card_order)
-next
- show "cinfinite natLeq" by (rule natLeq_cinfinite)
-next
- fix x
- show "|fset x| \<le>o natLeq"
- unfolding fset_def2_raw
- apply (rule ordLess_imp_ordLeq)
- apply (rule finite_iff_ordLess_natLeq[THEN iffD1])
- by (rule finite_set)
-next
- fix A :: "'a set"
- have "|{x. fset x \<subseteq> A}| \<le>o |afset ` {as. set as \<subseteq> A}|"
- apply(rule card_of_mono1) unfolding fset_def2_raw apply auto
- apply (rule image_eqI)
- by (auto simp: afset_rfset)
- also have "|afset ` {as. set as \<subseteq> A}| \<le>o |{as. set as \<subseteq> A}|" using card_of_image .
- also have "|{as. set as \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq" by (rule list_in_bd)
- finally show "|{x. fset x \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq" .
-next
- fix A B1 B2 f1 f2 p1 p2
- assume wp: "wpull A B1 B2 f1 f2 p1 p2"
- hence "wpull (Pow A) (Pow B1) (Pow B2) (image f1) (image f2) (image p1) (image p2)"
- by (rule wpull_image)
- show "wpull {x. fset x \<subseteq> A} {x. fset x \<subseteq> B1} {x. fset x \<subseteq> B2}
- (map_fset f1) (map_fset f2) (map_fset p1) (map_fset p2)"
- unfolding wpull_def Pow_def Bex_def mem_Collect_eq proof clarify
- fix y1 y2
- assume Y1: "fset y1 \<subseteq> B1" and Y2: "fset y2 \<subseteq> B2"
- assume "map_fset f1 y1 = map_fset f2 y2"
- hence EQ: "f1 ` (fset y1) = f2 ` (fset y2)" unfolding map_fset_def2_raw
- unfolding afset_set set_map fset_def2_raw .
- with Y1 Y2 obtain X where X: "X \<subseteq> A"
- and Y1: "p1 ` X = fset y1" and Y2: "p2 ` X = fset y2"
- using wpull_image[OF wp] unfolding wpull_def Pow_def
- unfolding Bex_def mem_Collect_eq apply -
- apply(erule allE[of _ "fset y1"], erule allE[of _ "fset y2"]) by auto
- have "\<forall> y1' \<in> fset y1. \<exists> x. x \<in> X \<and> y1' = p1 x" using Y1 by auto
- then obtain q1 where q1: "\<forall> y1' \<in> fset y1. q1 y1' \<in> X \<and> y1' = p1 (q1 y1')" by metis
- have "\<forall> y2' \<in> fset y2. \<exists> x. x \<in> X \<and> y2' = p2 x" using Y2 by auto
- then obtain q2 where q2: "\<forall> y2' \<in> fset y2. q2 y2' \<in> X \<and> y2' = p2 (q2 y2')" by metis
- def X' \<equiv> "q1 ` (fset y1) \<union> q2 ` (fset y2)"
- have X': "X' \<subseteq> A" and Y1: "p1 ` X' = fset y1" and Y2: "p2 ` X' = fset y2"
- using X Y1 Y2 q1 q2 unfolding X'_def by auto
- have fX': "finite X'" unfolding X'_def by simp
- then obtain x where X'eq: "X' = fset x" by (auto dest: finite_ex_fset)
- show "\<exists>x. fset x \<subseteq> A \<and> map_fset p1 x = y1 \<and> map_fset p2 x = y2"
- apply(intro exI[of _ "x"]) using X' Y1 Y2
- unfolding X'eq map_fset_def2_raw fset_def2_raw set_map[symmetric]
- afset_set[symmetric] afset_rfset by simp
- qed
-next
- fix R
- show "{p. fset_rel (\<lambda>x y. (x, y) \<in> R) (fst p) (snd p)} =
- (Gr {x. fset x \<subseteq> R} (map_fset fst))\<inverse> O Gr {x. fset x \<subseteq> R} (map_fset snd)"
- unfolding fset_rel_def fset_rel_aux by simp
-qed auto
-
-(* Countable sets *)
-
-lemma card_of_countable_sets_range:
-fixes A :: "'a set"
-shows "|{X. X \<subseteq> A \<and> countable X \<and> X \<noteq> {}}| \<le>o |{f::nat \<Rightarrow> 'a. range f \<subseteq> A}|"
-apply(rule card_of_ordLeqI[of fromNat]) using inj_on_fromNat
-unfolding inj_on_def by auto
-
-lemma card_of_countable_sets_Func:
-"|{X. X \<subseteq> A \<and> countable X \<and> X \<noteq> {}}| \<le>o |A| ^c natLeq"
-using card_of_countable_sets_range card_of_Func_UNIV[THEN ordIso_symmetric]
-unfolding cexp_def Field_natLeq Field_card_of
-by (rule ordLeq_ordIso_trans)
-
-lemma ordLeq_countable_subsets:
-"|A| \<le>o |{X. X \<subseteq> A \<and> countable X}|"
-apply (rule card_of_ordLeqI[of "\<lambda> a. {a}"]) unfolding inj_on_def by auto
-
-lemma finite_countable_subset:
-"finite {X. X \<subseteq> A \<and> countable X} \<longleftrightarrow> finite A"
-apply default
- apply (erule contrapos_pp)
- apply (rule card_of_ordLeq_infinite)
- apply (rule ordLeq_countable_subsets)
- apply assumption
-apply (rule finite_Collect_conjI)
-apply (rule disjI1)
-by (erule finite_Collect_subsets)
-
-lemma card_of_countable_sets:
-"|{X. X \<subseteq> A \<and> countable X}| \<le>o ( |A| +c ctwo) ^c natLeq"
-(is "|?L| \<le>o _")
-proof(cases "finite A")
- let ?R = "Func (UNIV::nat set) (A <+> (UNIV::bool set))"
- case True hence "finite ?L" by simp
- moreover have "infinite ?R"
- apply(rule infinite_Func[of _ "Inr True" "Inr False"]) by auto
- ultimately show ?thesis unfolding cexp_def csum_def ctwo_def Field_natLeq Field_card_of
- apply(intro ordLess_imp_ordLeq) by (rule finite_ordLess_infinite2)
-next
- case False
- hence "|{X. X \<subseteq> A \<and> countable X}| =o |{X. X \<subseteq> A \<and> countable X} - {{}}|"
- by (intro card_of_infinite_diff_finitte finite.emptyI finite.insertI ordIso_symmetric)
- (unfold finite_countable_subset)
- also have "|{X. X \<subseteq> A \<and> countable X} - {{}}| \<le>o |A| ^c natLeq"
- using card_of_countable_sets_Func[of A] unfolding set_diff_eq by auto
- also have "|A| ^c natLeq \<le>o ( |A| +c ctwo) ^c natLeq"
- apply(rule cexp_mono1_cone_ordLeq)
- apply(rule ordLeq_csum1, rule card_of_Card_order)
- apply (rule cone_ordLeq_cexp)
- apply (rule cone_ordLeq_Cnotzero)
- using csum_Cnotzero2 ctwo_Cnotzero apply blast
- by (rule natLeq_Card_order)
- finally show ?thesis .
-qed
-
-lemma rcset_to_rcset: "countable A \<Longrightarrow> rcset (the_inv rcset A) = A"
-apply (rule f_the_inv_into_f)
-unfolding inj_on_def rcset_inj using rcset_surj by auto
-
-lemma Collect_Int_Times:
-"{(x, y). R x y} \<inter> A \<times> B = {(x, y). R x y \<and> x \<in> A \<and> y \<in> B}"
-by auto
-
-lemma rcset_natural': "rcset (cIm f x) = f ` rcset x"
-unfolding cIm_def[abs_def] by simp
-
-definition cset_rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a cset \<Rightarrow> 'b cset \<Rightarrow> bool" where
-"cset_rel R a b \<longleftrightarrow>
- (\<forall>t \<in> rcset a. \<exists>u \<in> rcset b. R t u) \<and>
- (\<forall>t \<in> rcset b. \<exists>u \<in> rcset a. R u t)"
-
-lemma cset_rel_aux:
-"(\<forall>t \<in> rcset a. \<exists>u \<in> rcset b. R t u) \<and> (\<forall>t \<in> rcset b. \<exists>u \<in> rcset a. R u t) \<longleftrightarrow>
- (a, b) \<in> (Gr {x. rcset x \<subseteq> {(a, b). R a b}} (cIm fst))\<inverse> O
- Gr {x. rcset x \<subseteq> {(a, b). R a b}} (cIm snd)" (is "?L = ?R")
-proof
- assume ?L
- def R' \<equiv> "the_inv rcset (Collect (split R) \<inter> (rcset a \<times> rcset b))"
- (is "the_inv rcset ?L'")
- have "countable ?L'" by auto
- hence *: "rcset R' = ?L'" unfolding R'_def using fset_to_fset by (intro rcset_to_rcset)
- show ?R unfolding Gr_def relcomp_unfold converse_unfold
- proof (intro CollectI prod_caseI exI conjI)
- have "rcset a = fst ` ({(x, y). R x y} \<inter> rcset a \<times> rcset b)" (is "_ = ?A")
- using conjunct1[OF `?L`] unfolding image_def by (auto simp add: Collect_Int_Times)
- hence "a = acset ?A" by (metis acset_rcset)
- thus "(R', a) = (R', cIm fst R')" unfolding cIm_def * by auto
- have "rcset b = snd ` ({(x, y). R x y} \<inter> rcset a \<times> rcset b)" (is "_ = ?B")
- using conjunct2[OF `?L`] unfolding image_def by (auto simp add: Collect_Int_Times)
- hence "b = acset ?B" by (metis acset_rcset)
- thus "(R', b) = (R', cIm snd R')" unfolding cIm_def * by auto
- qed (auto simp add: *)
-next
- assume ?R thus ?L unfolding Gr_def relcomp_unfold converse_unfold
- apply (simp add: subset_eq Ball_def)
- apply (rule conjI)
- apply (clarsimp, metis (lifting, no_types) rcset_natural' image_iff surjective_pairing)
- apply (clarsimp)
- by (metis Domain.intros Range.simps rcset_natural' fst_eq_Domain snd_eq_Range)
-qed
-
-bnf_def cIm [rcset] "\<lambda>_::'a cset. natLeq" ["cEmp"] cset_rel
-proof -
- show "cIm id = id" unfolding cIm_def[abs_def] id_def by auto
-next
- fix f g show "cIm (g \<circ> f) = cIm g \<circ> cIm f"
- unfolding cIm_def[abs_def] apply(rule ext) unfolding comp_def by auto
-next
- fix C f g assume eq: "\<And>a. a \<in> rcset C \<Longrightarrow> f a = g a"
- thus "cIm f C = cIm g C"
- unfolding cIm_def[abs_def] unfolding image_def by auto
-next
- fix f show "rcset \<circ> cIm f = op ` f \<circ> rcset" unfolding cIm_def[abs_def] by auto
-next
- show "card_order natLeq" by (rule natLeq_card_order)
-next
- show "cinfinite natLeq" by (rule natLeq_cinfinite)
-next
- fix C show "|rcset C| \<le>o natLeq" using rcset unfolding countable_def .
-next
- fix A :: "'a set"
- have "|{Z. rcset Z \<subseteq> A}| \<le>o |acset ` {X. X \<subseteq> A \<and> countable X}|"
- apply(rule card_of_mono1) unfolding Pow_def image_def
- proof (rule Collect_mono, clarsimp)
- fix x
- assume "rcset x \<subseteq> A"
- hence "rcset x \<subseteq> A \<and> countable (rcset x) \<and> x = acset (rcset x)"
- using acset_rcset[of x] rcset[of x] by force
- thus "\<exists>y \<subseteq> A. countable y \<and> x = acset y" by blast
- qed
- also have "|acset ` {X. X \<subseteq> A \<and> countable X}| \<le>o |{X. X \<subseteq> A \<and> countable X}|"
- using card_of_image .
- also have "|{X. X \<subseteq> A \<and> countable X}| \<le>o ( |A| +c ctwo) ^c natLeq"
- using card_of_countable_sets .
- finally show "|{Z. rcset Z \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq" .
-next
- fix A B1 B2 f1 f2 p1 p2
- assume wp: "wpull A B1 B2 f1 f2 p1 p2"
- show "wpull {x. rcset x \<subseteq> A} {x. rcset x \<subseteq> B1} {x. rcset x \<subseteq> B2}
- (cIm f1) (cIm f2) (cIm p1) (cIm p2)"
- unfolding wpull_def proof safe
- fix y1 y2
- assume Y1: "rcset y1 \<subseteq> B1" and Y2: "rcset y2 \<subseteq> B2"
- assume "cIm f1 y1 = cIm f2 y2"
- hence EQ: "f1 ` (rcset y1) = f2 ` (rcset y2)"
- unfolding cIm_def by auto
- with Y1 Y2 obtain X where X: "X \<subseteq> A"
- and Y1: "p1 ` X = rcset y1" and Y2: "p2 ` X = rcset y2"
- using wpull_image[OF wp] unfolding wpull_def Pow_def
- unfolding Bex_def mem_Collect_eq apply -
- apply(erule allE[of _ "rcset y1"], erule allE[of _ "rcset y2"]) by auto
- have "\<forall> y1' \<in> rcset y1. \<exists> x. x \<in> X \<and> y1' = p1 x" using Y1 by auto
- then obtain q1 where q1: "\<forall> y1' \<in> rcset y1. q1 y1' \<in> X \<and> y1' = p1 (q1 y1')" by metis
- have "\<forall> y2' \<in> rcset y2. \<exists> x. x \<in> X \<and> y2' = p2 x" using Y2 by auto
- then obtain q2 where q2: "\<forall> y2' \<in> rcset y2. q2 y2' \<in> X \<and> y2' = p2 (q2 y2')" by metis
- def X' \<equiv> "q1 ` (rcset y1) \<union> q2 ` (rcset y2)"
- have X': "X' \<subseteq> A" and Y1: "p1 ` X' = rcset y1" and Y2: "p2 ` X' = rcset y2"
- using X Y1 Y2 q1 q2 unfolding X'_def by fast+
- have fX': "countable X'" unfolding X'_def by simp
- then obtain x where X'eq: "X' = rcset x" by (metis rcset_acset)
- show "\<exists>x\<in>{x. rcset x \<subseteq> A}. cIm p1 x = y1 \<and> cIm p2 x = y2"
- apply(intro bexI[of _ "x"]) using X' Y1 Y2 unfolding X'eq cIm_def by auto
- qed
-next
- fix R
- show "{p. cset_rel (\<lambda>x y. (x, y) \<in> R) (fst p) (snd p)} =
- (Gr {x. rcset x \<subseteq> R} (cIm fst))\<inverse> O Gr {x. rcset x \<subseteq> R} (cIm snd)"
- unfolding cset_rel_def cset_rel_aux by simp
-qed (unfold cEmp_def, auto)
-
-
-(* Multisets *)
-
-(* The cardinal of a mutiset: this, and the following basic lemmas about it,
-should eventually go into Multiset.thy *)
-definition "mcard M \<equiv> setsum (count M) {a. count M a > 0}"
-
-lemma mcard_emp[simp]: "mcard {#} = 0"
-unfolding mcard_def by auto
-
-lemma mcard_emp_iff[simp]: "mcard M = 0 \<longleftrightarrow> M = {#}"
-unfolding mcard_def apply safe
- apply simp_all
- by (metis multi_count_eq zero_multiset.rep_eq)
-
-lemma mcard_singl[simp]: "mcard {#a#} = Suc 0"
-unfolding mcard_def by auto
-
-lemma mcard_Plus[simp]: "mcard (M + N) = mcard M + mcard N"
-proof-
- have "setsum (count M) {a. 0 < count M a + count N a} =
- setsum (count M) {a. a \<in># M}"
- apply(rule setsum_mono_zero_cong_right) by auto
- moreover
- have "setsum (count N) {a. 0 < count M a + count N a} =
- setsum (count N) {a. a \<in># N}"
- apply(rule setsum_mono_zero_cong_right) by auto
- ultimately show ?thesis
- unfolding mcard_def count_union[THEN ext] comm_monoid_add_class.setsum.F_fun_f by simp
-qed
-
-lemma setsum_gt_0_iff:
-fixes f :: "'a \<Rightarrow> nat" assumes "finite A"
-shows "setsum f A > 0 \<longleftrightarrow> (\<exists> a \<in> A. f a > 0)"
-(is "?L \<longleftrightarrow> ?R")
-proof-
- have "?L \<longleftrightarrow> \<not> setsum f A = 0" by fast
- also have "... \<longleftrightarrow> (\<exists> a \<in> A. f a \<noteq> 0)" using assms by simp
- also have "... \<longleftrightarrow> ?R" by simp
- finally show ?thesis .
-qed
-
-(* *)
-definition mmap :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> nat) \<Rightarrow> 'b \<Rightarrow> nat" where
-"mmap h f b = setsum f {a. h a = b \<and> f a > 0}"
-
-lemma mmap_id: "mmap id = id"
-proof (rule ext)+
- fix f a show "mmap id f a = id f a"
- proof(cases "f a = 0")
- case False
- hence 1: "{aa. aa = a \<and> 0 < f aa} = {a}" by auto
- show ?thesis by (simp add: mmap_def id_apply 1)
- qed(unfold mmap_def, auto)
-qed
-
-lemma inj_on_setsum_inv:
-assumes f: "f \<in> multiset"
-and 1: "(0::nat) < setsum f {a. h a = b' \<and> 0 < f a}" (is "0 < setsum f ?A'")
-and 2: "{a. h a = b \<and> 0 < f a} = {a. h a = b' \<and> 0 < f a}" (is "?A = ?A'")
-shows "b = b'"
-proof-
- have "finite ?A'" using f unfolding multiset_def by auto
- hence "?A' \<noteq> {}" using 1 setsum_gt_0_iff by auto
- thus ?thesis using 2 by auto
-qed
-
-lemma mmap_comp:
-fixes h1 :: "'a \<Rightarrow> 'b" and h2 :: "'b \<Rightarrow> 'c"
-assumes f: "f \<in> multiset"
-shows "mmap (h2 o h1) f = (mmap h2 o mmap h1) f"
-unfolding mmap_def[abs_def] comp_def proof(rule ext)+
- fix c :: 'c
- let ?A = "{a. h2 (h1 a) = c \<and> 0 < f a}"
- let ?As = "\<lambda> b. {a. h1 a = b \<and> 0 < f a}"
- let ?B = "{b. h2 b = c \<and> 0 < setsum f (?As b)}"
- have 0: "{?As b | b. b \<in> ?B} = ?As ` ?B" by auto
- have "\<And> b. finite (?As b)" using f unfolding multiset_def by simp
- hence "?B = {b. h2 b = c \<and> ?As b \<noteq> {}}" using setsum_gt_0_iff by auto
- hence A: "?A = \<Union> {?As b | b. b \<in> ?B}" by auto
- have "setsum f ?A = setsum (setsum f) {?As b | b. b \<in> ?B}"
- unfolding A apply(rule setsum_Union_disjoint)
- using f unfolding multiset_def by auto
- also have "... = setsum (setsum f) (?As ` ?B)" unfolding 0 ..
- also have "... = setsum (setsum f o ?As) ?B" apply(rule setsum_reindex)
- unfolding inj_on_def apply auto using inj_on_setsum_inv[OF f, of h1] by blast
- also have "... = setsum (\<lambda> b. setsum f (?As b)) ?B" unfolding comp_def ..
- finally show "setsum f ?A = setsum (\<lambda> b. setsum f (?As b)) ?B" .
-qed
-
-lemma mmap_comp1:
-fixes h1 :: "'a \<Rightarrow> 'b" and h2 :: "'b \<Rightarrow> 'c"
-assumes "f \<in> multiset"
-shows "mmap (\<lambda> a. h2 (h1 a)) f = mmap h2 (mmap h1 f)"
-using mmap_comp[OF assms] unfolding comp_def by auto
-
-lemma mmap:
-assumes "f \<in> multiset"
-shows "mmap h f \<in> multiset"
-using assms unfolding mmap_def[abs_def] multiset_def proof safe
- assume fin: "finite {a. 0 < f a}" (is "finite ?A")
- show "finite {b. 0 < setsum f {a. h a = b \<and> 0 < f a}}"
- (is "finite {b. 0 < setsum f (?As b)}")
- proof- let ?B = "{b. 0 < setsum f (?As b)}"
- have "\<And> b. finite (?As b)" using assms unfolding multiset_def by simp
- hence B: "?B = {b. ?As b \<noteq> {}}" using setsum_gt_0_iff by auto
- hence "?B \<subseteq> h ` ?A" by auto
- thus ?thesis using finite_surj[OF fin] by auto
- qed
-qed
-
-lemma mmap_cong:
-assumes "\<And>a. a \<in># M \<Longrightarrow> f a = g a"
-shows "mmap f (count M) = mmap g (count M)"
-using assms unfolding mmap_def[abs_def] by (intro ext, intro setsum_cong) auto
-
-abbreviation supp where "supp f \<equiv> {a. f a > 0}"
-
-lemma mmap_image_comp:
-assumes f: "f \<in> multiset"
-shows "(supp o mmap h) f = (image h o supp) f"
-unfolding mmap_def[abs_def] comp_def proof-
- have "\<And> b. finite {a. h a = b \<and> 0 < f a}" (is "\<And> b. finite (?As b)")
- using f unfolding multiset_def by auto
- thus "{b. 0 < setsum f (?As b)} = h ` {a. 0 < f a}"
- using setsum_gt_0_iff by auto
-qed
-
-lemma mmap_image:
-assumes f: "f \<in> multiset"
-shows "supp (mmap h f) = h ` (supp f)"
-using mmap_image_comp[OF assms] unfolding comp_def .
-
-lemma set_of_Abs_multiset:
-assumes f: "f \<in> multiset"
-shows "set_of (Abs_multiset f) = supp f"
-using assms unfolding set_of_def by (auto simp: Abs_multiset_inverse)
-
-lemma supp_count:
-"supp (count M) = set_of M"
-using assms unfolding set_of_def by auto
-
-lemma multiset_of_surj:
-"multiset_of ` {as. set as \<subseteq> A} = {M. set_of M \<subseteq> A}"
-proof safe
- fix M assume M: "set_of M \<subseteq> A"
- obtain as where eq: "M = multiset_of as" using surj_multiset_of unfolding surj_def by auto
- hence "set as \<subseteq> A" using M by auto
- thus "M \<in> multiset_of ` {as. set as \<subseteq> A}" using eq by auto
-next
- show "\<And>x xa xb. \<lbrakk>set xa \<subseteq> A; xb \<in> set_of (multiset_of xa)\<rbrakk> \<Longrightarrow> xb \<in> A"
- by (erule set_mp) (unfold set_of_multiset_of)
-qed
-
-lemma card_of_set_of:
-"|{M. set_of M \<subseteq> A}| \<le>o |{as. set as \<subseteq> A}|"
-apply(rule card_of_ordLeqI2[of _ multiset_of]) using multiset_of_surj by auto
-
-lemma nat_sum_induct:
-assumes "\<And>n1 n2. (\<And> m1 m2. m1 + m2 < n1 + n2 \<Longrightarrow> phi m1 m2) \<Longrightarrow> phi n1 n2"
-shows "phi (n1::nat) (n2::nat)"
-proof-
- let ?chi = "\<lambda> n1n2 :: nat * nat. phi (fst n1n2) (snd n1n2)"
- have "?chi (n1,n2)"
- apply(induct rule: measure_induct[of "\<lambda> n1n2. fst n1n2 + snd n1n2" ?chi])
- using assms by (metis fstI sndI)
- thus ?thesis by simp
-qed
-
-lemma matrix_count:
-fixes ct1 ct2 :: "nat \<Rightarrow> nat"
-assumes "setsum ct1 {..<Suc n1} = setsum ct2 {..<Suc n2}"
-shows
-"\<exists> ct. (\<forall> i1 \<le> n1. setsum (\<lambda> i2. ct i1 i2) {..<Suc n2} = ct1 i1) \<and>
- (\<forall> i2 \<le> n2. setsum (\<lambda> i1. ct i1 i2) {..<Suc n1} = ct2 i2)"
-(is "?phi ct1 ct2 n1 n2")
-proof-
- have "\<forall> ct1 ct2 :: nat \<Rightarrow> nat.
- setsum ct1 {..<Suc n1} = setsum ct2 {..<Suc n2} \<longrightarrow> ?phi ct1 ct2 n1 n2"
- proof(induct rule: nat_sum_induct[of
-"\<lambda> n1 n2. \<forall> ct1 ct2 :: nat \<Rightarrow> nat.
- setsum ct1 {..<Suc n1} = setsum ct2 {..<Suc n2} \<longrightarrow> ?phi ct1 ct2 n1 n2"],
- clarify)
- fix n1 n2 :: nat and ct1 ct2 :: "nat \<Rightarrow> nat"
- assume IH: "\<And> m1 m2. m1 + m2 < n1 + n2 \<Longrightarrow>
- \<forall> dt1 dt2 :: nat \<Rightarrow> nat.
- setsum dt1 {..<Suc m1} = setsum dt2 {..<Suc m2} \<longrightarrow> ?phi dt1 dt2 m1 m2"
- and ss: "setsum ct1 {..<Suc n1} = setsum ct2 {..<Suc n2}"
- show "?phi ct1 ct2 n1 n2"
- proof(cases n1)
- case 0 note n1 = 0
- show ?thesis
- proof(cases n2)
- case 0 note n2 = 0
- let ?ct = "\<lambda> i1 i2. ct2 0"
- show ?thesis apply(rule exI[of _ ?ct]) using n1 n2 ss by simp
- next
- case (Suc m2) note n2 = Suc
- let ?ct = "\<lambda> i1 i2. ct2 i2"
- show ?thesis apply(rule exI[of _ ?ct]) using n1 n2 ss by auto
- qed
- next
- case (Suc m1) note n1 = Suc
- show ?thesis
- proof(cases n2)
- case 0 note n2 = 0
- let ?ct = "\<lambda> i1 i2. ct1 i1"
- show ?thesis apply(rule exI[of _ ?ct]) using n1 n2 ss by auto
- next
- case (Suc m2) note n2 = Suc
- show ?thesis
- proof(cases "ct1 n1 \<le> ct2 n2")
- case True
- def dt2 \<equiv> "\<lambda> i2. if i2 = n2 then ct2 i2 - ct1 n1 else ct2 i2"
- have "setsum ct1 {..<Suc m1} = setsum dt2 {..<Suc n2}"
- unfolding dt2_def using ss n1 True by auto
- hence "?phi ct1 dt2 m1 n2" using IH[of m1 n2] n1 by simp
- then obtain dt where
- 1: "\<And> i1. i1 \<le> m1 \<Longrightarrow> setsum (\<lambda> i2. dt i1 i2) {..<Suc n2} = ct1 i1" and
- 2: "\<And> i2. i2 \<le> n2 \<Longrightarrow> setsum (\<lambda> i1. dt i1 i2) {..<Suc m1} = dt2 i2" by auto
- let ?ct = "\<lambda> i1 i2. if i1 = n1 then (if i2 = n2 then ct1 n1 else 0)
- else dt i1 i2"
- show ?thesis apply(rule exI[of _ ?ct])
- using n1 n2 1 2 True unfolding dt2_def by simp
- next
- case False
- hence False: "ct2 n2 < ct1 n1" by simp
- def dt1 \<equiv> "\<lambda> i1. if i1 = n1 then ct1 i1 - ct2 n2 else ct1 i1"
- have "setsum dt1 {..<Suc n1} = setsum ct2 {..<Suc m2}"
- unfolding dt1_def using ss n2 False by auto
- hence "?phi dt1 ct2 n1 m2" using IH[of n1 m2] n2 by simp
- then obtain dt where
- 1: "\<And> i1. i1 \<le> n1 \<Longrightarrow> setsum (\<lambda> i2. dt i1 i2) {..<Suc m2} = dt1 i1" and
- 2: "\<And> i2. i2 \<le> m2 \<Longrightarrow> setsum (\<lambda> i1. dt i1 i2) {..<Suc n1} = ct2 i2" by force
- let ?ct = "\<lambda> i1 i2. if i2 = n2 then (if i1 = n1 then ct2 n2 else 0)
- else dt i1 i2"
- show ?thesis apply(rule exI[of _ ?ct])
- using n1 n2 1 2 False unfolding dt1_def by simp
- qed
- qed
- qed
- qed
- thus ?thesis using assms by auto
-qed
-
-definition
-"inj2 u B1 B2 \<equiv>
- \<forall> b1 b1' b2 b2'. {b1,b1'} \<subseteq> B1 \<and> {b2,b2'} \<subseteq> B2 \<and> u b1 b2 = u b1' b2'
- \<longrightarrow> b1 = b1' \<and> b2 = b2'"
-
-lemma matrix_setsum_finite:
-assumes B1: "B1 \<noteq> {}" "finite B1" and B2: "B2 \<noteq> {}" "finite B2" and u: "inj2 u B1 B2"
-and ss: "setsum N1 B1 = setsum N2 B2"
-shows "\<exists> M :: 'a \<Rightarrow> nat.
- (\<forall> b1 \<in> B1. setsum (\<lambda> b2. M (u b1 b2)) B2 = N1 b1) \<and>
- (\<forall> b2 \<in> B2. setsum (\<lambda> b1. M (u b1 b2)) B1 = N2 b2)"
-proof-
- obtain n1 where "card B1 = Suc n1" using B1 by (metis card_insert finite.simps)
- then obtain e1 where e1: "bij_betw e1 {..<Suc n1} B1"
- using ex_bij_betw_finite_nat[OF B1(2)] by (metis atLeast0LessThan bij_betw_the_inv_into)
- hence e1_inj: "inj_on e1 {..<Suc n1}" and e1_surj: "e1 ` {..<Suc n1} = B1"
- unfolding bij_betw_def by auto
- def f1 \<equiv> "inv_into {..<Suc n1} e1"
- have f1: "bij_betw f1 B1 {..<Suc n1}"
- and f1e1[simp]: "\<And> i1. i1 < Suc n1 \<Longrightarrow> f1 (e1 i1) = i1"
- and e1f1[simp]: "\<And> b1. b1 \<in> B1 \<Longrightarrow> e1 (f1 b1) = b1" unfolding f1_def
- apply (metis bij_betw_inv_into e1, metis bij_betw_inv_into_left e1 lessThan_iff)
- by (metis e1_surj f_inv_into_f)
- (* *)
- obtain n2 where "card B2 = Suc n2" using B2 by (metis card_insert finite.simps)
- then obtain e2 where e2: "bij_betw e2 {..<Suc n2} B2"
- using ex_bij_betw_finite_nat[OF B2(2)] by (metis atLeast0LessThan bij_betw_the_inv_into)
- hence e2_inj: "inj_on e2 {..<Suc n2}" and e2_surj: "e2 ` {..<Suc n2} = B2"
- unfolding bij_betw_def by auto
- def f2 \<equiv> "inv_into {..<Suc n2} e2"
- have f2: "bij_betw f2 B2 {..<Suc n2}"
- and f2e2[simp]: "\<And> i2. i2 < Suc n2 \<Longrightarrow> f2 (e2 i2) = i2"
- and e2f2[simp]: "\<And> b2. b2 \<in> B2 \<Longrightarrow> e2 (f2 b2) = b2" unfolding f2_def
- apply (metis bij_betw_inv_into e2, metis bij_betw_inv_into_left e2 lessThan_iff)
- by (metis e2_surj f_inv_into_f)
- (* *)
- let ?ct1 = "N1 o e1" let ?ct2 = "N2 o e2"
- have ss: "setsum ?ct1 {..<Suc n1} = setsum ?ct2 {..<Suc n2}"
- unfolding setsum_reindex[OF e1_inj, symmetric] setsum_reindex[OF e2_inj, symmetric]
- e1_surj e2_surj using ss .
- obtain ct where
- ct1: "\<And> i1. i1 \<le> n1 \<Longrightarrow> setsum (\<lambda> i2. ct i1 i2) {..<Suc n2} = ?ct1 i1" and
- ct2: "\<And> i2. i2 \<le> n2 \<Longrightarrow> setsum (\<lambda> i1. ct i1 i2) {..<Suc n1} = ?ct2 i2"
- using matrix_count[OF ss] by blast
- (* *)
- def A \<equiv> "{u b1 b2 | b1 b2. b1 \<in> B1 \<and> b2 \<in> B2}"
- have "\<forall> a \<in> A. \<exists> b1b2 \<in> B1 <*> B2. u (fst b1b2) (snd b1b2) = a"
- unfolding A_def Ball_def mem_Collect_eq by auto
- then obtain h1h2 where h12:
- "\<And>a. a \<in> A \<Longrightarrow> u (fst (h1h2 a)) (snd (h1h2 a)) = a \<and> h1h2 a \<in> B1 <*> B2" by metis
- def h1 \<equiv> "fst o h1h2" def h2 \<equiv> "snd o h1h2"
- have h12[simp]: "\<And>a. a \<in> A \<Longrightarrow> u (h1 a) (h2 a) = a"
- "\<And> a. a \<in> A \<Longrightarrow> h1 a \<in> B1" "\<And> a. a \<in> A \<Longrightarrow> h2 a \<in> B2"
- using h12 unfolding h1_def h2_def by force+
- {fix b1 b2 assume b1: "b1 \<in> B1" and b2: "b2 \<in> B2"
- hence inA: "u b1 b2 \<in> A" unfolding A_def by auto
- hence "u b1 b2 = u (h1 (u b1 b2)) (h2 (u b1 b2))" by auto
- moreover have "h1 (u b1 b2) \<in> B1" "h2 (u b1 b2) \<in> B2" using inA by auto
- ultimately have "h1 (u b1 b2) = b1 \<and> h2 (u b1 b2) = b2"
- using u b1 b2 unfolding inj2_def by fastforce
- }
- hence h1[simp]: "\<And> b1 b2. \<lbrakk>b1 \<in> B1; b2 \<in> B2\<rbrakk> \<Longrightarrow> h1 (u b1 b2) = b1" and
- h2[simp]: "\<And> b1 b2. \<lbrakk>b1 \<in> B1; b2 \<in> B2\<rbrakk> \<Longrightarrow> h2 (u b1 b2) = b2" by auto
- def M \<equiv> "\<lambda> a. ct (f1 (h1 a)) (f2 (h2 a))"
- show ?thesis
- apply(rule exI[of _ M]) proof safe
- fix b1 assume b1: "b1 \<in> B1"
- hence f1b1: "f1 b1 \<le> n1" using f1 unfolding bij_betw_def
- by (metis bij_betwE f1 lessThan_iff less_Suc_eq_le)
- have "(\<Sum>b2\<in>B2. M (u b1 b2)) = (\<Sum>i2<Suc n2. ct (f1 b1) (f2 (e2 i2)))"
- unfolding e2_surj[symmetric] setsum_reindex[OF e2_inj]
- unfolding M_def comp_def apply(intro setsum_cong) apply force
- by (metis e2_surj b1 h1 h2 imageI)
- also have "... = N1 b1" using b1 ct1[OF f1b1] by simp
- finally show "(\<Sum>b2\<in>B2. M (u b1 b2)) = N1 b1" .
- next
- fix b2 assume b2: "b2 \<in> B2"
- hence f2b2: "f2 b2 \<le> n2" using f2 unfolding bij_betw_def
- by (metis bij_betwE f2 lessThan_iff less_Suc_eq_le)
- have "(\<Sum>b1\<in>B1. M (u b1 b2)) = (\<Sum>i1<Suc n1. ct (f1 (e1 i1)) (f2 b2))"
- unfolding e1_surj[symmetric] setsum_reindex[OF e1_inj]
- unfolding M_def comp_def apply(intro setsum_cong) apply force
- by (metis e1_surj b2 h1 h2 imageI)
- also have "... = N2 b2" using b2 ct2[OF f2b2] by simp
- finally show "(\<Sum>b1\<in>B1. M (u b1 b2)) = N2 b2" .
- qed
-qed
-
-lemma supp_vimage_mmap:
-assumes "M \<in> multiset"
-shows "supp M \<subseteq> f -` (supp (mmap f M))"
-using assms by (auto simp: mmap_image)
-
-lemma mmap_ge_0:
-assumes "M \<in> multiset"
-shows "0 < mmap f M b \<longleftrightarrow> (\<exists>a. 0 < M a \<and> f a = b)"
-proof-
- have f: "finite {a. f a = b \<and> 0 < M a}" using assms unfolding multiset_def by auto
- show ?thesis unfolding mmap_def setsum_gt_0_iff[OF f] by auto
-qed
-
-lemma finite_twosets:
-assumes "finite B1" and "finite B2"
-shows "finite {u b1 b2 |b1 b2. b1 \<in> B1 \<and> b2 \<in> B2}" (is "finite ?A")
-proof-
- have A: "?A = (\<lambda> b1b2. u (fst b1b2) (snd b1b2)) ` (B1 <*> B2)" by force
- show ?thesis unfolding A using finite_cartesian_product[OF assms] by auto
-qed
-
-lemma wp_mmap:
-fixes A :: "'a set" and B1 :: "'b1 set" and B2 :: "'b2 set"
-assumes wp: "wpull A B1 B2 f1 f2 p1 p2"
-shows
-"wpull {M. M \<in> multiset \<and> supp M \<subseteq> A}
- {N1. N1 \<in> multiset \<and> supp N1 \<subseteq> B1} {N2. N2 \<in> multiset \<and> supp N2 \<subseteq> B2}
- (mmap f1) (mmap f2) (mmap p1) (mmap p2)"
-unfolding wpull_def proof (safe, unfold Bex_def mem_Collect_eq)
- fix N1 :: "'b1 \<Rightarrow> nat" and N2 :: "'b2 \<Rightarrow> nat"
- assume mmap': "mmap f1 N1 = mmap f2 N2"
- and N1[simp]: "N1 \<in> multiset" "supp N1 \<subseteq> B1"
- and N2[simp]: "N2 \<in> multiset" "supp N2 \<subseteq> B2"
- have mN1[simp]: "mmap f1 N1 \<in> multiset" using N1 by (auto simp: mmap)
- have mN2[simp]: "mmap f2 N2 \<in> multiset" using N2 by (auto simp: mmap)
- def P \<equiv> "mmap f1 N1"
- have P1: "P = mmap f1 N1" and P2: "P = mmap f2 N2" unfolding P_def using mmap' by auto
- note P = P1 P2
- have P_mult[simp]: "P \<in> multiset" unfolding P_def using N1 by auto
- have fin_N1[simp]: "finite (supp N1)" using N1(1) unfolding multiset_def by auto
- have fin_N2[simp]: "finite (supp N2)" using N2(1) unfolding multiset_def by auto
- have fin_P[simp]: "finite (supp P)" using P_mult unfolding multiset_def by auto
- (* *)
- def set1 \<equiv> "\<lambda> c. {b1 \<in> supp N1. f1 b1 = c}"
- have set1[simp]: "\<And> c b1. b1 \<in> set1 c \<Longrightarrow> f1 b1 = c" unfolding set1_def by auto
- have fin_set1: "\<And> c. c \<in> supp P \<Longrightarrow> finite (set1 c)"
- using N1(1) unfolding set1_def multiset_def by auto
- have set1_NE: "\<And> c. c \<in> supp P \<Longrightarrow> set1 c \<noteq> {}"
- unfolding set1_def P1 mmap_ge_0[OF N1(1)] by auto
- have supp_N1_set1: "supp N1 = (\<Union> c \<in> supp P. set1 c)"
- using supp_vimage_mmap[OF N1(1), of f1] unfolding set1_def P1 by auto
- hence set1_inclN1: "\<And>c. c \<in> supp P \<Longrightarrow> set1 c \<subseteq> supp N1" by auto
- hence set1_incl: "\<And> c. c \<in> supp P \<Longrightarrow> set1 c \<subseteq> B1" using N1(2) by blast
- have set1_disj: "\<And> c c'. c \<noteq> c' \<Longrightarrow> set1 c \<inter> set1 c' = {}"
- unfolding set1_def by auto
- have setsum_set1: "\<And> c. setsum N1 (set1 c) = P c"
- unfolding P1 set1_def mmap_def apply(rule setsum_cong) by auto
- (* *)
- def set2 \<equiv> "\<lambda> c. {b2 \<in> supp N2. f2 b2 = c}"
- have set2[simp]: "\<And> c b2. b2 \<in> set2 c \<Longrightarrow> f2 b2 = c" unfolding set2_def by auto
- have fin_set2: "\<And> c. c \<in> supp P \<Longrightarrow> finite (set2 c)"
- using N2(1) unfolding set2_def multiset_def by auto
- have set2_NE: "\<And> c. c \<in> supp P \<Longrightarrow> set2 c \<noteq> {}"
- unfolding set2_def P2 mmap_ge_0[OF N2(1)] by auto
- have supp_N2_set2: "supp N2 = (\<Union> c \<in> supp P. set2 c)"
- using supp_vimage_mmap[OF N2(1), of f2] unfolding set2_def P2 by auto
- hence set2_inclN2: "\<And>c. c \<in> supp P \<Longrightarrow> set2 c \<subseteq> supp N2" by auto
- hence set2_incl: "\<And> c. c \<in> supp P \<Longrightarrow> set2 c \<subseteq> B2" using N2(2) by blast
- have set2_disj: "\<And> c c'. c \<noteq> c' \<Longrightarrow> set2 c \<inter> set2 c' = {}"
- unfolding set2_def by auto
- have setsum_set2: "\<And> c. setsum N2 (set2 c) = P c"
- unfolding P2 set2_def mmap_def apply(rule setsum_cong) by auto
- (* *)
- have ss: "\<And> c. c \<in> supp P \<Longrightarrow> setsum N1 (set1 c) = setsum N2 (set2 c)"
- unfolding setsum_set1 setsum_set2 ..
- have "\<forall> c \<in> supp P. \<forall> b1b2 \<in> (set1 c) \<times> (set2 c).
- \<exists> a \<in> A. p1 a = fst b1b2 \<and> p2 a = snd b1b2"
- using wp set1_incl set2_incl unfolding wpull_def Ball_def mem_Collect_eq
- by simp (metis set1 set2 set_rev_mp)
- then obtain uu where uu:
- "\<forall> c \<in> supp P. \<forall> b1b2 \<in> (set1 c) \<times> (set2 c).
- uu c b1b2 \<in> A \<and> p1 (uu c b1b2) = fst b1b2 \<and> p2 (uu c b1b2) = snd b1b2" by metis
- def u \<equiv> "\<lambda> c b1 b2. uu c (b1,b2)"
- have u[simp]:
- "\<And> c b1 b2. \<lbrakk>c \<in> supp P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk> \<Longrightarrow> u c b1 b2 \<in> A"
- "\<And> c b1 b2. \<lbrakk>c \<in> supp P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk> \<Longrightarrow> p1 (u c b1 b2) = b1"
- "\<And> c b1 b2. \<lbrakk>c \<in> supp P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk> \<Longrightarrow> p2 (u c b1 b2) = b2"
- using uu unfolding u_def by auto
- {fix c assume c: "c \<in> supp P"
- have "inj2 (u c) (set1 c) (set2 c)" unfolding inj2_def proof clarify
- fix b1 b1' b2 b2'
- assume "{b1, b1'} \<subseteq> set1 c" "{b2, b2'} \<subseteq> set2 c" and 0: "u c b1 b2 = u c b1' b2'"
- hence "p1 (u c b1 b2) = b1 \<and> p2 (u c b1 b2) = b2 \<and>
- p1 (u c b1' b2') = b1' \<and> p2 (u c b1' b2') = b2'"
- using u(2)[OF c] u(3)[OF c] by simp metis
- thus "b1 = b1' \<and> b2 = b2'" using 0 by auto
- qed
- } note inj = this
- def sset \<equiv> "\<lambda> c. {u c b1 b2 | b1 b2. b1 \<in> set1 c \<and> b2 \<in> set2 c}"
- have fin_sset[simp]: "\<And> c. c \<in> supp P \<Longrightarrow> finite (sset c)" unfolding sset_def
- using fin_set1 fin_set2 finite_twosets by blast
- have sset_A: "\<And> c. c \<in> supp P \<Longrightarrow> sset c \<subseteq> A" unfolding sset_def by auto
- {fix c a assume c: "c \<in> supp P" and ac: "a \<in> sset c"
- then obtain b1 b2 where b1: "b1 \<in> set1 c" and b2: "b2 \<in> set2 c"
- and a: "a = u c b1 b2" unfolding sset_def by auto
- have "p1 a \<in> set1 c" and p2a: "p2 a \<in> set2 c"
- using ac a b1 b2 c u(2) u(3) by simp+
- hence "u c (p1 a) (p2 a) = a" unfolding a using b1 b2 inj[OF c]
- unfolding inj2_def by (metis c u(2) u(3))
- } note u_p12[simp] = this
- {fix c a assume c: "c \<in> supp P" and ac: "a \<in> sset c"
- hence "p1 a \<in> set1 c" unfolding sset_def by auto
- }note p1[simp] = this
- {fix c a assume c: "c \<in> supp P" and ac: "a \<in> sset c"
- hence "p2 a \<in> set2 c" unfolding sset_def by auto
- }note p2[simp] = this
- (* *)
- {fix c assume c: "c \<in> supp P"
- hence "\<exists> M. (\<forall> b1 \<in> set1 c. setsum (\<lambda> b2. M (u c b1 b2)) (set2 c) = N1 b1) \<and>
- (\<forall> b2 \<in> set2 c. setsum (\<lambda> b1. M (u c b1 b2)) (set1 c) = N2 b2)"
- unfolding sset_def
- using matrix_setsum_finite[OF set1_NE[OF c] fin_set1[OF c]
- set2_NE[OF c] fin_set2[OF c] inj[OF c] ss[OF c]] by auto
- }
- then obtain Ms where
- ss1: "\<And> c b1. \<lbrakk>c \<in> supp P; b1 \<in> set1 c\<rbrakk> \<Longrightarrow>
- setsum (\<lambda> b2. Ms c (u c b1 b2)) (set2 c) = N1 b1" and
- ss2: "\<And> c b2. \<lbrakk>c \<in> supp P; b2 \<in> set2 c\<rbrakk> \<Longrightarrow>
- setsum (\<lambda> b1. Ms c (u c b1 b2)) (set1 c) = N2 b2"
- by metis
- def SET \<equiv> "\<Union> c \<in> supp P. sset c"
- have fin_SET[simp]: "finite SET" unfolding SET_def apply(rule finite_UN_I) by auto
- have SET_A: "SET \<subseteq> A" unfolding SET_def using sset_A by auto
- have u_SET[simp]: "\<And> c b1 b2. \<lbrakk>c \<in> supp P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk> \<Longrightarrow> u c b1 b2 \<in> SET"
- unfolding SET_def sset_def by blast
- {fix c a assume c: "c \<in> supp P" and a: "a \<in> SET" and p1a: "p1 a \<in> set1 c"
- then obtain c' where c': "c' \<in> supp P" and ac': "a \<in> sset c'"
- unfolding SET_def by auto
- hence "p1 a \<in> set1 c'" unfolding sset_def by auto
- hence eq: "c = c'" using p1a c c' set1_disj by auto
- hence "a \<in> sset c" using ac' by simp
- } note p1_rev = this
- {fix c a assume c: "c \<in> supp P" and a: "a \<in> SET" and p2a: "p2 a \<in> set2 c"
- then obtain c' where c': "c' \<in> supp P" and ac': "a \<in> sset c'"
- unfolding SET_def by auto
- hence "p2 a \<in> set2 c'" unfolding sset_def by auto
- hence eq: "c = c'" using p2a c c' set2_disj by auto
- hence "a \<in> sset c" using ac' by simp
- } note p2_rev = this
- (* *)
- have "\<forall> a \<in> SET. \<exists> c \<in> supp P. a \<in> sset c" unfolding SET_def by auto
- then obtain h where h: "\<forall> a \<in> SET. h a \<in> supp P \<and> a \<in> sset (h a)" by metis
- have h_u[simp]: "\<And> c b1 b2. \<lbrakk>c \<in> supp P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk>
- \<Longrightarrow> h (u c b1 b2) = c"
- by (metis h p2 set2 u(3) u_SET)
- have h_u1: "\<And> c b1 b2. \<lbrakk>c \<in> supp P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk>
- \<Longrightarrow> h (u c b1 b2) = f1 b1"
- using h unfolding sset_def by auto
- have h_u2: "\<And> c b1 b2. \<lbrakk>c \<in> supp P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk>
- \<Longrightarrow> h (u c b1 b2) = f2 b2"
- using h unfolding sset_def by auto
- def M \<equiv> "\<lambda> a. if a \<in> SET \<and> p1 a \<in> supp N1 \<and> p2 a \<in> supp N2 then Ms (h a) a else 0"
- have sM: "supp M \<subseteq> SET" "supp M \<subseteq> p1 -` (supp N1)" "supp M \<subseteq> p2 -` (supp N2)"
- unfolding M_def by auto
- show "\<exists>M. (M \<in> multiset \<and> supp M \<subseteq> A) \<and> mmap p1 M = N1 \<and> mmap p2 M = N2"
- proof(rule exI[of _ M], safe)
- show "M \<in> multiset"
- unfolding multiset_def using finite_subset[OF sM(1) fin_SET] by simp
- next
- fix a assume "0 < M a"
- thus "a \<in> A" unfolding M_def using SET_A by (cases "a \<in> SET") auto
- next
- show "mmap p1 M = N1"
- unfolding mmap_def[abs_def] proof(rule ext)
- fix b1
- let ?K = "{a. p1 a = b1 \<and> 0 < M a}"
- show "setsum M ?K = N1 b1"
- proof(cases "b1 \<in> supp N1")
- case False
- hence "?K = {}" using sM(2) by auto
- thus ?thesis using False by auto
- next
- case True
- def c \<equiv> "f1 b1"
- have c: "c \<in> supp P" and b1: "b1 \<in> set1 c"
- unfolding set1_def c_def P1 using True by (auto simp: mmap_image)
- have "setsum M ?K = setsum M {a. p1 a = b1 \<and> a \<in> SET}"
- apply(rule setsum_mono_zero_cong_left) unfolding M_def by auto
- also have "... = setsum M ((\<lambda> b2. u c b1 b2) ` (set2 c))"
- apply(rule setsum_cong) using c b1 proof safe
- fix a assume p1a: "p1 a \<in> set1 c" and "0 < P c" and "a \<in> SET"
- hence ac: "a \<in> sset c" using p1_rev by auto
- hence "a = u c (p1 a) (p2 a)" using c by auto
- moreover have "p2 a \<in> set2 c" using ac c by auto
- ultimately show "a \<in> u c (p1 a) ` set2 c" by auto
- next
- fix b2 assume b1: "b1 \<in> set1 c" and b2: "b2 \<in> set2 c"
- hence "u c b1 b2 \<in> SET" using c by auto
- qed auto
- also have "... = setsum (\<lambda> b2. M (u c b1 b2)) (set2 c)"
- unfolding comp_def[symmetric] apply(rule setsum_reindex)
- using inj unfolding inj_on_def inj2_def using b1 c u(3) by blast
- also have "... = N1 b1" unfolding ss1[OF c b1, symmetric]
- apply(rule setsum_cong[OF refl]) unfolding M_def
- using True h_u[OF c b1] set2_def u(2,3)[OF c b1] u_SET[OF c b1] by fastforce
- finally show ?thesis .
- qed
- qed
- next
- show "mmap p2 M = N2"
- unfolding mmap_def[abs_def] proof(rule ext)
- fix b2
- let ?K = "{a. p2 a = b2 \<and> 0 < M a}"
- show "setsum M ?K = N2 b2"
- proof(cases "b2 \<in> supp N2")
- case False
- hence "?K = {}" using sM(3) by auto
- thus ?thesis using False by auto
- next
- case True
- def c \<equiv> "f2 b2"
- have c: "c \<in> supp P" and b2: "b2 \<in> set2 c"
- unfolding set2_def c_def P2 using True by (auto simp: mmap_image)
- have "setsum M ?K = setsum M {a. p2 a = b2 \<and> a \<in> SET}"
- apply(rule setsum_mono_zero_cong_left) unfolding M_def by auto
- also have "... = setsum M ((\<lambda> b1. u c b1 b2) ` (set1 c))"
- apply(rule setsum_cong) using c b2 proof safe
- fix a assume p2a: "p2 a \<in> set2 c" and "0 < P c" and "a \<in> SET"
- hence ac: "a \<in> sset c" using p2_rev by auto
- hence "a = u c (p1 a) (p2 a)" using c by auto
- moreover have "p1 a \<in> set1 c" using ac c by auto
- ultimately show "a \<in> (\<lambda>b1. u c b1 (p2 a)) ` set1 c" by auto
- next
- fix b2 assume b1: "b1 \<in> set1 c" and b2: "b2 \<in> set2 c"
- hence "u c b1 b2 \<in> SET" using c by auto
- qed auto
- also have "... = setsum (M o (\<lambda> b1. u c b1 b2)) (set1 c)"
- apply(rule setsum_reindex)
- using inj unfolding inj_on_def inj2_def using b2 c u(2) by blast
- also have "... = setsum (\<lambda> b1. M (u c b1 b2)) (set1 c)"
- unfolding comp_def[symmetric] by simp
- also have "... = N2 b2" unfolding ss2[OF c b2, symmetric]
- apply(rule setsum_cong[OF refl]) unfolding M_def set2_def
- using True h_u1[OF c _ b2] u(2,3)[OF c _ b2] u_SET[OF c _ b2]
- unfolding set1_def by fastforce
- finally show ?thesis .
- qed
- qed
- qed
-qed
-
-definition multiset_map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a multiset \<Rightarrow> 'b multiset" where
-"multiset_map h = Abs_multiset \<circ> mmap h \<circ> count"
-
-bnf_def multiset_map [set_of] "\<lambda>_::'a multiset. natLeq" ["{#}"]
-unfolding multiset_map_def
-proof -
- show "Abs_multiset \<circ> mmap id \<circ> count = id" unfolding mmap_id by (auto simp: count_inverse)
-next
- fix f g
- show "Abs_multiset \<circ> mmap (g \<circ> f) \<circ> count =
- Abs_multiset \<circ> mmap g \<circ> count \<circ> (Abs_multiset \<circ> mmap f \<circ> count)"
- unfolding comp_def apply(rule ext)
- by (auto simp: Abs_multiset_inverse count mmap_comp1 mmap)
-next
- fix M f g assume eq: "\<And>a. a \<in> set_of M \<Longrightarrow> f a = g a"
- thus "(Abs_multiset \<circ> mmap f \<circ> count) M = (Abs_multiset \<circ> mmap g \<circ> count) M" apply auto
- unfolding cIm_def[abs_def] image_def
- by (auto intro!: mmap_cong simp: Abs_multiset_inject count mmap)
-next
- fix f show "set_of \<circ> (Abs_multiset \<circ> mmap f \<circ> count) = op ` f \<circ> set_of"
- by (auto simp: count mmap mmap_image set_of_Abs_multiset supp_count)
-next
- show "card_order natLeq" by (rule natLeq_card_order)
-next
- show "cinfinite natLeq" by (rule natLeq_cinfinite)
-next
- fix M show "|set_of M| \<le>o natLeq"
- apply(rule ordLess_imp_ordLeq)
- unfolding finite_iff_ordLess_natLeq[symmetric] using finite_set_of .
-next
- fix A :: "'a set"
- have "|{M. set_of M \<subseteq> A}| \<le>o |{as. set as \<subseteq> A}|" using card_of_set_of .
- also have "|{as. set as \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq"
- by (rule list_in_bd)
- finally show "|{M. set_of M \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq" .
-next
- fix A B1 B2 f1 f2 p1 p2
- let ?map = "\<lambda> f. Abs_multiset \<circ> mmap f \<circ> count"
- assume wp: "wpull A B1 B2 f1 f2 p1 p2"
- show "wpull {x. set_of x \<subseteq> A} {x. set_of x \<subseteq> B1} {x. set_of x \<subseteq> B2}
- (?map f1) (?map f2) (?map p1) (?map p2)"
- unfolding wpull_def proof safe
- fix y1 y2
- assume y1: "set_of y1 \<subseteq> B1" and y2: "set_of y2 \<subseteq> B2"
- and m: "?map f1 y1 = ?map f2 y2"
- def N1 \<equiv> "count y1" def N2 \<equiv> "count y2"
- have "N1 \<in> multiset \<and> supp N1 \<subseteq> B1" and "N2 \<in> multiset \<and> supp N2 \<subseteq> B2"
- and "mmap f1 N1 = mmap f2 N2"
- using y1 y2 m unfolding N1_def N2_def
- by (auto simp: Abs_multiset_inject count mmap)
- then obtain M where M: "M \<in> multiset \<and> supp M \<subseteq> A"
- and N1: "mmap p1 M = N1" and N2: "mmap p2 M = N2"
- using wp_mmap[OF wp] unfolding wpull_def by auto
- def x \<equiv> "Abs_multiset M"
- show "\<exists>x\<in>{x. set_of x \<subseteq> A}. ?map p1 x = y1 \<and> ?map p2 x = y2"
- apply(intro bexI[of _ x]) using M N1 N2 unfolding N1_def N2_def x_def
- by (auto simp: count_inverse Abs_multiset_inverse)
- qed
-qed (unfold set_of_empty, auto)
-
-inductive multiset_rel' where
-Zero: "multiset_rel' R {#} {#}"
-|
-Plus: "\<lbrakk>R a b; multiset_rel' R M N\<rbrakk> \<Longrightarrow> multiset_rel' R (M + {#a#}) (N + {#b#})"
-
-lemma multiset_map_Zero_iff[simp]: "multiset_map f M = {#} \<longleftrightarrow> M = {#}"
-by (metis image_is_empty multiset.set_natural' set_of_eq_empty_iff)
-
-lemma multiset_map_Zero[simp]: "multiset_map f {#} = {#}" by simp
-
-lemma multiset_rel_Zero: "multiset_rel R {#} {#}"
-unfolding multiset_rel_def Gr_def relcomp_unfold by auto
-
-declare multiset.count[simp]
-declare mmap[simp]
-declare Abs_multiset_inverse[simp]
-declare multiset.count_inverse[simp]
-declare union_preserves_multiset[simp]
-
-lemma mmap_Plus[simp]:
-assumes "K \<in> multiset" and "L \<in> multiset"
-shows "mmap f (\<lambda>a. K a + L a) a = mmap f K a + mmap f L a"
-proof-
- have "{aa. f aa = a \<and> (0 < K aa \<or> 0 < L aa)} \<subseteq>
- {aa. 0 < K aa} \<union> {aa. 0 < L aa}" (is "?C \<subseteq> ?A \<union> ?B") by auto
- moreover have "finite (?A \<union> ?B)" apply(rule finite_UnI)
- using assms unfolding multiset_def by auto
- ultimately have C: "finite ?C" using finite_subset by blast
- have "setsum K {aa. f aa = a \<and> 0 < K aa} = setsum K {aa. f aa = a \<and> 0 < K aa + L aa}"
- apply(rule setsum_mono_zero_cong_left) using C by auto
- moreover
- have "setsum L {aa. f aa = a \<and> 0 < L aa} = setsum L {aa. f aa = a \<and> 0 < K aa + L aa}"
- apply(rule setsum_mono_zero_cong_left) using C by auto
- ultimately show ?thesis
- unfolding mmap_def unfolding comm_monoid_add_class.setsum.F_fun_f by auto
-qed
-
-lemma multiset_map_Plus[simp]:
-"multiset_map f (M1 + M2) = multiset_map f M1 + multiset_map f M2"
-unfolding multiset_map_def
-apply(subst multiset.count_inject[symmetric])
-unfolding plus_multiset.rep_eq comp_def by auto
-
-lemma multiset_map_singl[simp]: "multiset_map f {#a#} = {#f a#}"
-proof-
- have 0: "\<And> b. card {aa. a = aa \<and> (a = aa \<longrightarrow> f aa = b)} =
- (if b = f a then 1 else 0)" by auto
- thus ?thesis
- unfolding multiset_map_def comp_def mmap_def[abs_def] map_fun_def
- by (simp, simp add: single_def)
-qed
-
-lemma multiset_rel_Plus:
-assumes ab: "R a b" and MN: "multiset_rel R M N"
-shows "multiset_rel R (M + {#a#}) (N + {#b#})"
-proof-
- {fix y assume "R a b" and "set_of y \<subseteq> {(x, y). R x y}"
- hence "\<exists>ya. multiset_map fst y + {#a#} = multiset_map fst ya \<and>
- multiset_map snd y + {#b#} = multiset_map snd ya \<and>
- set_of ya \<subseteq> {(x, y). R x y}"
- apply(intro exI[of _ "y + {#(a,b)#}"]) by auto
- }
- thus ?thesis
- using assms
- unfolding multiset_rel_def Gr_def relcomp_unfold by force
-qed
-
-lemma multiset_rel'_imp_multiset_rel:
-"multiset_rel' R M N \<Longrightarrow> multiset_rel R M N"
-apply(induct rule: multiset_rel'.induct)
-using multiset_rel_Zero multiset_rel_Plus by auto
-
-lemma mcard_multiset_map[simp]: "mcard (multiset_map f M) = mcard M"
-proof-
- def A \<equiv> "\<lambda> b. {a. f a = b \<and> a \<in># M}"
- let ?B = "{b. 0 < setsum (count M) (A b)}"
- have "{b. \<exists>a. f a = b \<and> a \<in># M} \<subseteq> f ` {a. a \<in># M}" by auto
- moreover have "finite (f ` {a. a \<in># M})" apply(rule finite_imageI)
- using finite_Collect_mem .
- ultimately have fin: "finite {b. \<exists>a. f a = b \<and> a \<in># M}" by(rule finite_subset)
- have i: "inj_on A ?B" unfolding inj_on_def A_def apply clarsimp
- by (metis (lifting, mono_tags) mem_Collect_eq rel_simps(54)
- setsum_gt_0_iff setsum_infinite)
- have 0: "\<And> b. 0 < setsum (count M) (A b) \<longleftrightarrow> (\<exists> a \<in> A b. count M a > 0)"
- apply safe
- apply (metis less_not_refl setsum_gt_0_iff setsum_infinite)
- by (metis A_def finite_Collect_conjI finite_Collect_mem setsum_gt_0_iff)
- hence AB: "A ` ?B = {A b | b. \<exists> a \<in> A b. count M a > 0}" by auto
-
- have "setsum (\<lambda> x. setsum (count M) (A x)) ?B = setsum (setsum (count M) o A) ?B"
- unfolding comp_def ..
- also have "... = (\<Sum>x\<in> A ` ?B. setsum (count M) x)"
- unfolding comm_monoid_add_class.setsum_reindex[OF i, symmetric] ..
- also have "... = setsum (count M) (\<Union>x\<in>A ` {b. 0 < setsum (count M) (A b)}. x)"
- (is "_ = setsum (count M) ?J")
- apply(rule comm_monoid_add_class.setsum_UN_disjoint[symmetric])
- using 0 fin unfolding A_def by (auto intro!: finite_imageI)
- also have "?J = {a. a \<in># M}" unfolding AB unfolding A_def by auto
- finally have "setsum (\<lambda> x. setsum (count M) (A x)) ?B =
- setsum (count M) {a. a \<in># M}" .
- thus ?thesis unfolding A_def mcard_def multiset_map_def by (simp add: mmap_def)
-qed
-
-lemma multiset_rel_mcard:
-assumes "multiset_rel R M N"
-shows "mcard M = mcard N"
-using assms unfolding multiset_rel_def relcomp_unfold Gr_def by auto
-
-lemma multiset_induct2[case_names empty addL addR]:
-assumes empty: "P {#} {#}"
-and addL: "\<And>M N a. P M N \<Longrightarrow> P (M + {#a#}) N"
-and addR: "\<And>M N a. P M N \<Longrightarrow> P M (N + {#a#})"
-shows "P M N"
-apply(induct N rule: multiset_induct)
- apply(induct M rule: multiset_induct, rule empty, erule addL)
- apply(induct M rule: multiset_induct, erule addR, erule addR)
-done
-
-lemma multiset_induct2_mcard[consumes 1, case_names empty add]:
-assumes c: "mcard M = mcard N"
-and empty: "P {#} {#}"
-and add: "\<And>M N a b. P M N \<Longrightarrow> P (M + {#a#}) (N + {#b#})"
-shows "P M N"
-using c proof(induct M arbitrary: N rule: measure_induct_rule[of mcard])
- case (less M) show ?case
- proof(cases "M = {#}")
- case True hence "N = {#}" using less.prems by auto
- thus ?thesis using True empty by auto
- next
- case False then obtain M1 a where M: "M = M1 + {#a#}" by (metis multi_nonempty_split)
- have "N \<noteq> {#}" using False less.prems by auto
- then obtain N1 b where N: "N = N1 + {#b#}" by (metis multi_nonempty_split)
- have "mcard M1 = mcard N1" using less.prems unfolding M N by auto
- thus ?thesis using M N less.hyps add by auto
- qed
-qed
-
-lemma msed_map_invL:
-assumes "multiset_map f (M + {#a#}) = N"
-shows "\<exists> N1. N = N1 + {#f a#} \<and> multiset_map f M = N1"
-proof-
- have "f a \<in># N"
- using assms multiset.set_natural'[of f "M + {#a#}"] by auto
- then obtain N1 where N: "N = N1 + {#f a#}" using multi_member_split by metis
- have "multiset_map f M = N1" using assms unfolding N by simp
- thus ?thesis using N by blast
-qed
-
-lemma msed_map_invR:
-assumes "multiset_map f M = N + {#b#}"
-shows "\<exists> M1 a. M = M1 + {#a#} \<and> f a = b \<and> multiset_map f M1 = N"
-proof-
- obtain a where a: "a \<in># M" and fa: "f a = b"
- using multiset.set_natural'[of f M] unfolding assms
- by (metis image_iff mem_set_of_iff union_single_eq_member)
- then obtain M1 where M: "M = M1 + {#a#}" using multi_member_split by metis
- have "multiset_map f M1 = N" using assms unfolding M fa[symmetric] by simp
- thus ?thesis using M fa by blast
-qed
-
-lemma msed_rel_invL:
-assumes "multiset_rel R (M + {#a#}) N"
-shows "\<exists> N1 b. N = N1 + {#b#} \<and> R a b \<and> multiset_rel R M N1"
-proof-
- obtain K where KM: "multiset_map fst K = M + {#a#}"
- and KN: "multiset_map snd K = N" and sK: "set_of K \<subseteq> {(a, b). R a b}"
- using assms
- unfolding multiset_rel_def Gr_def relcomp_unfold by auto
- obtain K1 ab where K: "K = K1 + {#ab#}" and a: "fst ab = a"
- and K1M: "multiset_map fst K1 = M" using msed_map_invR[OF KM] by auto
- obtain N1 where N: "N = N1 + {#snd ab#}" and K1N1: "multiset_map snd K1 = N1"
- using msed_map_invL[OF KN[unfolded K]] by auto
- have Rab: "R a (snd ab)" using sK a unfolding K by auto
- have "multiset_rel R M N1" using sK K1M K1N1
- unfolding K multiset_rel_def Gr_def relcomp_unfold by auto
- thus ?thesis using N Rab by auto
-qed
-
-lemma msed_rel_invR:
-assumes "multiset_rel R M (N + {#b#})"
-shows "\<exists> M1 a. M = M1 + {#a#} \<and> R a b \<and> multiset_rel R M1 N"
-proof-
- obtain K where KN: "multiset_map snd K = N + {#b#}"
- and KM: "multiset_map fst K = M" and sK: "set_of K \<subseteq> {(a, b). R a b}"
- using assms
- unfolding multiset_rel_def Gr_def relcomp_unfold by auto
- obtain K1 ab where K: "K = K1 + {#ab#}" and b: "snd ab = b"
- and K1N: "multiset_map snd K1 = N" using msed_map_invR[OF KN] by auto
- obtain M1 where M: "M = M1 + {#fst ab#}" and K1M1: "multiset_map fst K1 = M1"
- using msed_map_invL[OF KM[unfolded K]] by auto
- have Rab: "R (fst ab) b" using sK b unfolding K by auto
- have "multiset_rel R M1 N" using sK K1N K1M1
- unfolding K multiset_rel_def Gr_def relcomp_unfold by auto
- thus ?thesis using M Rab by auto
-qed
-
-lemma multiset_rel_imp_multiset_rel':
-assumes "multiset_rel R M N"
-shows "multiset_rel' R M N"
-using assms proof(induct M arbitrary: N rule: measure_induct_rule[of mcard])
- case (less M)
- have c: "mcard M = mcard N" using multiset_rel_mcard[OF less.prems] .
- show ?case
- proof(cases "M = {#}")
- case True hence "N = {#}" using c by simp
- thus ?thesis using True multiset_rel'.Zero by auto
- next
- case False then obtain M1 a where M: "M = M1 + {#a#}" by (metis multi_nonempty_split)
- obtain N1 b where N: "N = N1 + {#b#}" and R: "R a b" and ms: "multiset_rel R M1 N1"
- using msed_rel_invL[OF less.prems[unfolded M]] by auto
- have "multiset_rel' R M1 N1" using less.hyps[of M1 N1] ms unfolding M by simp
- thus ?thesis using multiset_rel'.Plus[of R a b, OF R] unfolding M N by simp
- qed
-qed
-
-lemma multiset_rel_multiset_rel':
-"multiset_rel R M N = multiset_rel' R M N"
-using multiset_rel_imp_multiset_rel' multiset_rel'_imp_multiset_rel by auto
-
-(* The main end product for multiset_rel: inductive characterization *)
-theorems multiset_rel_induct[case_names empty add, induct pred: multiset_rel] =
- multiset_rel'.induct[unfolded multiset_rel_multiset_rel'[symmetric]]
-
-end
--- a/src/HOL/Codatatype/README.html Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,54 +0,0 @@
-<!DOCTYPE HTML PUBLIC "-//W3C//DTD HTML 4.01 Transitional//EN" "http://www.w3.org/TR/html4/loose.dtd">
-
-<html>
-
-<head>
- <meta http-equiv="Content-Type" content="text/html; charset=iso-8859-1">
- <title>BNF Package</title>
-</head>
-
-<body>
-
-<h3><i>BNF</i>: A (co)datatype package based on bounded natural functors
-(BNFs)</h3>
-
-<p>
-The <i>BNF</i> package provides a fully modular framework for constructing
-inductive and coinductive datatypes in HOL, with support for mixed mutual and
-nested (co)recursion. Mixed (co)recursion enables type definitions involving
-both datatypes and codatatypes, such as the type of finitely branching trees of
-possibly infinite depth. The framework draws heavily from category theory.
-
-<p>
-The package is described in the following paper:
-
-<ul>
- <li><a href="http://www21.in.tum.de/~traytel/papers/codatatypes/index.html">Foundational, Compositional (Co)datatypes for Higher-Order Logic—Category Theory Applied to Theorem Proving</a>, <br>
- Dmitriy Traytel, Andrei Popescu, and Jasmin Christian Blanchette.<br>
- <i>Logic in Computer Science (LICS 2012)</i>, 2012.
-</ul>
-
-<p>
-The main entry point for applications is <tt>BNF.thy</tt>. The <tt>Examples</tt>
-directory contains various examples of (co)datatypes, including the examples
-from the paper.
-
-<p>
-The key notion underlying the package is that of a <i>bounded natural functor</i>
-(<i>BNF</i>)—an enriched type constructor satisfying specific properties
-preserved by interesting categorical operations (composition, least fixed point,
-and greatest fixed point). The <tt>Basic_BNFs.thy</tt> and <tt>More_BNFs.thy</tt>
-files register various basic types, notably for sums, products, function spaces,
-finite sets, multisets, and countable sets. Custom BNFs can be registered as well.
-
-<p>
-<b>Warning:</b> The package is under development. Please contact any nonempty
-subset of
-<a href="mailto:traytel@in.tum.de">the</a>
-<a href="mailto:popescua@in.tum.de">above</a>
-<a href="mailto:blanchette@in.tum.de">authors</a>
-if you have questions or comments.
-
-</body>
-
-</html>
--- a/src/HOL/Codatatype/Tools/bnf_comp.ML Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,726 +0,0 @@
-(* Title: HOL/BNF/Tools/bnf_comp.ML
- Author: Dmitriy Traytel, TU Muenchen
- Author: Jasmin Blanchette, TU Muenchen
- Copyright 2012
-
-Composition of bounded natural functors.
-*)
-
-signature BNF_COMP =
-sig
- type unfold_set
- val empty_unfolds: unfold_set
- val map_unfolds_of: unfold_set -> thm list
- val rel_unfolds_of: unfold_set -> thm list
- val set_unfoldss_of: unfold_set -> thm list list
- val srel_unfolds_of: unfold_set -> thm list
-
- val bnf_of_typ: BNF_Def.const_policy -> (binding -> binding) ->
- ((string * sort) list list -> (string * sort) list) -> typ -> unfold_set * Proof.context ->
- (BNF_Def.BNF * (typ list * typ list)) * (unfold_set * Proof.context)
- val default_comp_sort: (string * sort) list list -> (string * sort) list
- val normalize_bnfs: (int -> binding -> binding) -> ''a list list -> ''a list ->
- (''a list list -> ''a list) -> BNF_Def.BNF list -> unfold_set -> Proof.context ->
- (int list list * ''a list) * (BNF_Def.BNF list * (unfold_set * Proof.context))
- val seal_bnf: unfold_set -> binding -> typ list -> BNF_Def.BNF -> Proof.context ->
- (BNF_Def.BNF * typ list) * local_theory
-end;
-
-structure BNF_Comp : BNF_COMP =
-struct
-
-open BNF_Def
-open BNF_Util
-open BNF_Tactics
-open BNF_Comp_Tactics
-
-type unfold_set = {
- map_unfolds: thm list,
- set_unfoldss: thm list list,
- rel_unfolds: thm list,
- srel_unfolds: thm list
-};
-
-val empty_unfolds = {map_unfolds = [], set_unfoldss = [], rel_unfolds = [], srel_unfolds = []};
-
-fun add_to_thms thms new = thms |> not (Thm.is_reflexive new) ? insert Thm.eq_thm new;
-fun adds_to_thms thms news = insert (eq_set Thm.eq_thm) (no_reflexive news) thms;
-
-fun add_to_unfolds map sets rel srel
- {map_unfolds, set_unfoldss, rel_unfolds, srel_unfolds} =
- {map_unfolds = add_to_thms map_unfolds map,
- set_unfoldss = adds_to_thms set_unfoldss sets,
- rel_unfolds = add_to_thms rel_unfolds rel,
- srel_unfolds = add_to_thms srel_unfolds srel};
-
-fun add_bnf_to_unfolds bnf =
- add_to_unfolds (map_def_of_bnf bnf) (set_defs_of_bnf bnf) (rel_def_of_bnf bnf)
- (srel_def_of_bnf bnf);
-
-val map_unfolds_of = #map_unfolds;
-val set_unfoldss_of = #set_unfoldss;
-val rel_unfolds_of = #rel_unfolds;
-val srel_unfolds_of = #srel_unfolds;
-
-val bdTN = "bdT";
-
-fun mk_killN n = "_kill" ^ string_of_int n;
-fun mk_liftN n = "_lift" ^ string_of_int n;
-fun mk_permuteN src dest =
- "_permute_" ^ implode (map string_of_int src) ^ "_" ^ implode (map string_of_int dest);
-
-(*copied from Envir.expand_term_free*)
-fun expand_term_const defs =
- let
- val eqs = map ((fn ((x, U), u) => (x, (U, u))) o apfst dest_Const) defs;
- val get = fn Const (x, _) => AList.lookup (op =) eqs x | _ => NONE;
- in Envir.expand_term get end;
-
-fun clean_compose_bnf const_policy qualify b outer inners (unfold_set, lthy) =
- let
- val olive = live_of_bnf outer;
- val onwits = nwits_of_bnf outer;
- val odead = dead_of_bnf outer;
- val inner = hd inners;
- val ilive = live_of_bnf inner;
- val ideads = map dead_of_bnf inners;
- val inwitss = map nwits_of_bnf inners;
-
- (* TODO: check olive = length inners > 0,
- forall inner from inners. ilive = live,
- forall inner from inners. idead = dead *)
-
- val (oDs, lthy1) = apfst (map TFree)
- (Variable.invent_types (replicate odead HOLogic.typeS) lthy);
- val (Dss, lthy2) = apfst (map (map TFree))
- (fold_map Variable.invent_types (map (fn n => replicate n HOLogic.typeS) ideads) lthy1);
- val (Ass, lthy3) = apfst (replicate ilive o map TFree)
- (Variable.invent_types (replicate ilive HOLogic.typeS) lthy2);
- val As = if ilive > 0 then hd Ass else [];
- val Ass_repl = replicate olive As;
- val (Bs, _(*lthy4*)) = apfst (map TFree)
- (Variable.invent_types (replicate ilive HOLogic.typeS) lthy3);
- val Bss_repl = replicate olive Bs;
-
- val ((((fs', Qs'), Asets), xs), _(*names_lthy*)) = lthy
- |> apfst snd o mk_Frees' "f" (map2 (curry (op -->)) As Bs)
- ||>> apfst snd o mk_Frees' "Q" (map2 mk_pred2T As Bs)
- ||>> mk_Frees "A" (map HOLogic.mk_setT As)
- ||>> mk_Frees "x" As;
-
- val CAs = map3 mk_T_of_bnf Dss Ass_repl inners;
- val CCA = mk_T_of_bnf oDs CAs outer;
- val CBs = map3 mk_T_of_bnf Dss Bss_repl inners;
- val outer_sets = mk_sets_of_bnf (replicate olive oDs) (replicate olive CAs) outer;
- val inner_setss = map3 mk_sets_of_bnf (map (replicate ilive) Dss) (replicate olive Ass) inners;
- val inner_bds = map3 mk_bd_of_bnf Dss Ass_repl inners;
- val outer_bd = mk_bd_of_bnf oDs CAs outer;
-
- (*%f1 ... fn. outer.map (inner_1.map f1 ... fn) ... (inner_m.map f1 ... fn)*)
- val mapx = fold_rev Term.abs fs'
- (Term.list_comb (mk_map_of_bnf oDs CAs CBs outer,
- map2 (fn Ds => (fn f => Term.list_comb (f, map Bound (ilive - 1 downto 0))) o
- mk_map_of_bnf Ds As Bs) Dss inners));
- (*%Q1 ... Qn. outer.rel (inner_1.rel Q1 ... Qn) ... (inner_m.rel Q1 ... Qn)*)
- val rel = fold_rev Term.abs Qs'
- (Term.list_comb (mk_rel_of_bnf oDs CAs CBs outer,
- map2 (fn Ds => (fn f => Term.list_comb (f, map Bound (ilive - 1 downto 0))) o
- mk_rel_of_bnf Ds As Bs) Dss inners));
-
- (*Union o collect {outer.set_1 ... outer.set_m} o outer.map inner_1.set_i ... inner_m.set_i*)
- (*Union o collect {image inner_1.set_i o outer.set_1 ... image inner_m.set_i o outer.set_m}*)
- fun mk_set i =
- let
- val (setTs, T) = `(replicate olive o HOLogic.mk_setT) (nth As i);
- val outer_set = mk_collect
- (mk_sets_of_bnf (replicate olive oDs) (replicate olive setTs) outer)
- (mk_T_of_bnf oDs setTs outer --> HOLogic.mk_setT T);
- val inner_sets = map (fn sets => nth sets i) inner_setss;
- val outer_map = mk_map_of_bnf oDs CAs setTs outer;
- val map_inner_sets = Term.list_comb (outer_map, inner_sets);
- val collect_image = mk_collect
- (map2 (fn f => fn set => HOLogic.mk_comp (mk_image f, set)) inner_sets outer_sets)
- (CCA --> HOLogic.mk_setT T);
- in
- (Library.foldl1 HOLogic.mk_comp [mk_Union T, outer_set, map_inner_sets],
- HOLogic.mk_comp (mk_Union T, collect_image))
- end;
-
- val (sets, sets_alt) = map_split mk_set (0 upto ilive - 1);
-
- (*(inner_1.bd +c ... +c inner_m.bd) *c outer.bd*)
- val bd = Term.absdummy CCA (mk_cprod (Library.foldr1 (uncurry mk_csum) inner_bds) outer_bd);
-
- fun map_id_tac {context = ctxt, ...} =
- let
- (*order the theorems by reverse size to prevent bad interaction with nonconfluent rewrite
- rules*)
- val thms = (map map_id_of_bnf inners
- |> map (`(Term.size_of_term o Thm.prop_of))
- |> sort (rev_order o int_ord o pairself fst)
- |> map snd) @ [map_id_of_bnf outer];
- in
- (EVERY' (map (fn thm => subst_tac ctxt [thm]) thms) THEN' rtac refl) 1
- end;
-
- fun map_comp_tac _ =
- mk_comp_map_comp_tac (map_comp_of_bnf outer) (map_cong_of_bnf outer)
- (map map_comp_of_bnf inners);
-
- fun mk_single_set_natural_tac i _ =
- mk_comp_set_natural_tac (map_comp_of_bnf outer) (map_cong_of_bnf outer)
- (collect_set_natural_of_bnf outer)
- (map ((fn thms => nth thms i) o set_natural_of_bnf) inners);
-
- val set_natural_tacs = map mk_single_set_natural_tac (0 upto ilive - 1);
-
- fun bd_card_order_tac _ =
- mk_comp_bd_card_order_tac (map bd_card_order_of_bnf inners) (bd_card_order_of_bnf outer);
-
- fun bd_cinfinite_tac _ =
- mk_comp_bd_cinfinite_tac (bd_cinfinite_of_bnf inner) (bd_cinfinite_of_bnf outer);
-
- val set_alt_thms =
- if ! quick_and_dirty then
- []
- else
- map (fn goal =>
- Skip_Proof.prove lthy [] [] goal
- (fn {context, ...} => (mk_comp_set_alt_tac context (collect_set_natural_of_bnf outer)))
- |> Thm.close_derivation)
- (map2 (curry (HOLogic.mk_Trueprop o HOLogic.mk_eq)) sets sets_alt);
-
- fun map_cong_tac _ =
- mk_comp_map_cong_tac set_alt_thms (map_cong_of_bnf outer) (map map_cong_of_bnf inners);
-
- val set_bd_tacs =
- if ! quick_and_dirty then
- replicate (length set_alt_thms) (K all_tac)
- else
- let
- val outer_set_bds = set_bd_of_bnf outer;
- val inner_set_bdss = map set_bd_of_bnf inners;
- val inner_bd_Card_orders = map bd_Card_order_of_bnf inners;
- fun single_set_bd_thm i j =
- @{thm comp_single_set_bd} OF [nth inner_bd_Card_orders j, nth (nth inner_set_bdss j) i,
- nth outer_set_bds j]
- val single_set_bd_thmss =
- map ((fn f => map f (0 upto olive - 1)) o single_set_bd_thm) (0 upto ilive - 1);
- in
- map2 (fn set_alt => fn single_set_bds => fn {context, ...} =>
- mk_comp_set_bd_tac context set_alt single_set_bds)
- set_alt_thms single_set_bd_thmss
- end;
-
- val in_alt_thm =
- let
- val inx = mk_in Asets sets CCA;
- val in_alt = mk_in (map2 (mk_in Asets) inner_setss CAs) outer_sets CCA;
- val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt));
- in
- Skip_Proof.prove lthy [] [] goal
- (fn {context, ...} => mk_comp_in_alt_tac context set_alt_thms)
- |> Thm.close_derivation
- end;
-
- fun in_bd_tac _ =
- mk_comp_in_bd_tac in_alt_thm (map in_bd_of_bnf inners) (in_bd_of_bnf outer)
- (map bd_Cinfinite_of_bnf inners) (bd_Card_order_of_bnf outer);
-
- fun map_wpull_tac _ =
- mk_map_wpull_tac in_alt_thm (map map_wpull_of_bnf inners) (map_wpull_of_bnf outer);
-
- fun srel_O_Gr_tac _ =
- let
- val basic_thms = @{thms mem_Collect_eq fst_conv snd_conv}; (*TODO: tune*)
- val outer_srel_Gr = srel_Gr_of_bnf outer RS sym;
- val outer_srel_cong = srel_cong_of_bnf outer;
- val thm =
- (trans OF [in_alt_thm RS @{thm subst_rel_def},
- trans OF [@{thm arg_cong2[of _ _ _ _ relcomp]} OF
- [trans OF [outer_srel_Gr RS @{thm arg_cong[of _ _ converse]},
- srel_converse_of_bnf outer RS sym], outer_srel_Gr],
- trans OF [srel_O_of_bnf outer RS sym, outer_srel_cong OF
- (map (fn bnf => srel_O_Gr_of_bnf bnf RS sym) inners)]]] RS sym)
- |> unfold_thms lthy (basic_thms @ srel_def_of_bnf outer :: map srel_def_of_bnf inners);
- in
- unfold_thms_tac lthy basic_thms THEN rtac thm 1
- end;
-
- val tacs = zip_axioms map_id_tac map_comp_tac map_cong_tac set_natural_tacs bd_card_order_tac
- bd_cinfinite_tac set_bd_tacs in_bd_tac map_wpull_tac srel_O_Gr_tac;
-
- val outer_wits = mk_wits_of_bnf (replicate onwits oDs) (replicate onwits CAs) outer;
-
- val inner_witss = map (map (fn (I, wit) => Term.list_comb (wit, map (nth xs) I)))
- (map3 (fn Ds => fn n => mk_wits_of_bnf (replicate n Ds) (replicate n As))
- Dss inwitss inners);
-
- val inner_witsss = map (map (nth inner_witss) o fst) outer_wits;
-
- val wits = (inner_witsss, (map (single o snd) outer_wits))
- |-> map2 (fold (map_product (fn iwit => fn owit => owit $ iwit)))
- |> flat
- |> map (`(fn t => Term.add_frees t []))
- |> minimize_wits
- |> map (fn (frees, t) => fold absfree frees t);
-
- fun wit_tac {context = ctxt, ...} =
- mk_comp_wit_tac ctxt (wit_thms_of_bnf outer) (collect_set_natural_of_bnf outer)
- (maps wit_thms_of_bnf inners);
-
- val (bnf', lthy') =
- bnf_def const_policy (K Derive_Few_Facts) qualify tacs wit_tac (SOME (oDs @ flat Dss))
- (((((b, mapx), sets), bd), wits), SOME rel) lthy;
- in
- (bnf', (add_bnf_to_unfolds bnf' unfold_set, lthy'))
- end;
-
-(* Killing live variables *)
-
-fun kill_bnf qualify n bnf (unfold_set, lthy) = if n = 0 then (bnf, (unfold_set, lthy)) else
- let
- val b = Binding.suffix_name (mk_killN n) (name_of_bnf bnf);
- val live = live_of_bnf bnf;
- val dead = dead_of_bnf bnf;
- val nwits = nwits_of_bnf bnf;
-
- (* TODO: check 0 < n <= live *)
-
- val (Ds, lthy1) = apfst (map TFree)
- (Variable.invent_types (replicate dead HOLogic.typeS) lthy);
- val ((killedAs, As), lthy2) = apfst (`(take n) o map TFree)
- (Variable.invent_types (replicate live HOLogic.typeS) lthy1);
- val (Bs, _(*lthy3*)) = apfst (append killedAs o map TFree)
- (Variable.invent_types (replicate (live - n) HOLogic.typeS) lthy2);
-
- val ((Asets, lives), _(*names_lthy*)) = lthy
- |> mk_Frees "A" (map HOLogic.mk_setT (drop n As))
- ||>> mk_Frees "x" (drop n As);
- val xs = map (fn T => HOLogic.choice_const T $ absdummy T @{term True}) killedAs @ lives;
-
- val T = mk_T_of_bnf Ds As bnf;
-
- (*bnf.map id ... id*)
- val mapx = Term.list_comb (mk_map_of_bnf Ds As Bs bnf, map HOLogic.id_const killedAs);
- (*bnf.rel (op =) ... (op =)*)
- val rel = Term.list_comb (mk_rel_of_bnf Ds As Bs bnf, map HOLogic.eq_const killedAs);
-
- val bnf_sets = mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf;
- val sets = drop n bnf_sets;
-
- (*(|UNIV :: A1 set| +c ... +c |UNIV :: An set|) *c bnf.bd*)
- val bnf_bd = mk_bd_of_bnf Ds As bnf;
- val bd = mk_cprod
- (Library.foldr1 (uncurry mk_csum) (map (mk_card_of o HOLogic.mk_UNIV) killedAs)) bnf_bd;
-
- fun map_id_tac _ = rtac (map_id_of_bnf bnf) 1;
- fun map_comp_tac {context, ...} =
- unfold_thms_tac context ((map_comp_of_bnf bnf RS sym) :: @{thms o_assoc id_o o_id}) THEN
- rtac refl 1;
- fun map_cong_tac {context, ...} =
- mk_kill_map_cong_tac context n (live - n) (map_cong_of_bnf bnf);
- val set_natural_tacs = map (fn thm => fn _ => rtac thm 1) (drop n (set_natural_of_bnf bnf));
- fun bd_card_order_tac _ = mk_kill_bd_card_order_tac n (bd_card_order_of_bnf bnf);
- fun bd_cinfinite_tac _ = mk_kill_bd_cinfinite_tac (bd_Cinfinite_of_bnf bnf);
- val set_bd_tacs =
- map (fn thm => fn _ => mk_kill_set_bd_tac (bd_Card_order_of_bnf bnf) thm)
- (drop n (set_bd_of_bnf bnf));
-
- val in_alt_thm =
- let
- val inx = mk_in Asets sets T;
- val in_alt = mk_in (map HOLogic.mk_UNIV killedAs @ Asets) bnf_sets T;
- val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt));
- in
- Skip_Proof.prove lthy [] [] goal (K kill_in_alt_tac) |> Thm.close_derivation
- end;
-
- fun in_bd_tac _ =
- mk_kill_in_bd_tac n (live > n) in_alt_thm (in_bd_of_bnf bnf) (bd_Card_order_of_bnf bnf)
- (bd_Cinfinite_of_bnf bnf) (bd_Cnotzero_of_bnf bnf);
- fun map_wpull_tac _ = mk_map_wpull_tac in_alt_thm [] (map_wpull_of_bnf bnf);
-
- fun srel_O_Gr_tac _ =
- let
- val srel_Gr = srel_Gr_of_bnf bnf RS sym
- val thm =
- (trans OF [in_alt_thm RS @{thm subst_rel_def},
- trans OF [@{thm arg_cong2[of _ _ _ _ relcomp]} OF
- [trans OF [srel_Gr RS @{thm arg_cong[of _ _ converse]},
- srel_converse_of_bnf bnf RS sym], srel_Gr],
- trans OF [srel_O_of_bnf bnf RS sym, srel_cong_of_bnf bnf OF
- (replicate n @{thm trans[OF Gr_UNIV_id[OF refl] Id_alt[symmetric]]} @
- replicate (live - n) @{thm Gr_fst_snd})]]] RS sym)
- |> unfold_thms lthy (srel_def_of_bnf bnf :: @{thms Id_def' mem_Collect_eq split_conv});
- in
- rtac thm 1
- end;
-
- val tacs = zip_axioms map_id_tac map_comp_tac map_cong_tac set_natural_tacs bd_card_order_tac
- bd_cinfinite_tac set_bd_tacs in_bd_tac map_wpull_tac srel_O_Gr_tac;
-
- val bnf_wits = mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf;
-
- val wits = map (fn t => fold absfree (Term.add_frees t []) t)
- (map (fn (I, wit) => Term.list_comb (wit, map (nth xs) I)) bnf_wits);
-
- fun wit_tac _ = mk_simple_wit_tac (wit_thms_of_bnf bnf);
-
- val (bnf', lthy') =
- bnf_def Smart_Inline (K Derive_Few_Facts) qualify tacs wit_tac (SOME (killedAs @ Ds))
- (((((b, mapx), sets), Term.absdummy T bd), wits), SOME rel) lthy;
- in
- (bnf', (add_bnf_to_unfolds bnf' unfold_set, lthy'))
- end;
-
-(* Adding dummy live variables *)
-
-fun lift_bnf qualify n bnf (unfold_set, lthy) = if n = 0 then (bnf, (unfold_set, lthy)) else
- let
- val b = Binding.suffix_name (mk_liftN n) (name_of_bnf bnf);
- val live = live_of_bnf bnf;
- val dead = dead_of_bnf bnf;
- val nwits = nwits_of_bnf bnf;
-
- (* TODO: check 0 < n *)
-
- val (Ds, lthy1) = apfst (map TFree)
- (Variable.invent_types (replicate dead HOLogic.typeS) lthy);
- val ((newAs, As), lthy2) = apfst (chop n o map TFree)
- (Variable.invent_types (replicate (n + live) HOLogic.typeS) lthy1);
- val ((newBs, Bs), _(*lthy3*)) = apfst (chop n o map TFree)
- (Variable.invent_types (replicate (n + live) HOLogic.typeS) lthy2);
-
- val (Asets, _(*names_lthy*)) = lthy
- |> mk_Frees "A" (map HOLogic.mk_setT (newAs @ As));
-
- val T = mk_T_of_bnf Ds As bnf;
-
- (*%f1 ... fn. bnf.map*)
- val mapx =
- fold_rev Term.absdummy (map2 (curry (op -->)) newAs newBs) (mk_map_of_bnf Ds As Bs bnf);
- (*%Q1 ... Qn. bnf.rel*)
- val rel = fold_rev Term.absdummy (map2 mk_pred2T newAs newBs) (mk_rel_of_bnf Ds As Bs bnf);
-
- val bnf_sets = mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf;
- val sets = map (fn A => absdummy T (HOLogic.mk_set A [])) newAs @ bnf_sets;
-
- val bd = mk_bd_of_bnf Ds As bnf;
-
- fun map_id_tac _ = rtac (map_id_of_bnf bnf) 1;
- fun map_comp_tac {context, ...} =
- unfold_thms_tac context ((map_comp_of_bnf bnf RS sym) :: @{thms o_assoc id_o o_id}) THEN
- rtac refl 1;
- fun map_cong_tac {context, ...} =
- rtac (map_cong_of_bnf bnf) 1 THEN REPEAT_DETERM_N live (Goal.assume_rule_tac context 1);
- val set_natural_tacs =
- if ! quick_and_dirty then
- replicate (n + live) (K all_tac)
- else
- replicate n (K empty_natural_tac) @
- map (fn thm => fn _ => rtac thm 1) (set_natural_of_bnf bnf);
- fun bd_card_order_tac _ = rtac (bd_card_order_of_bnf bnf) 1;
- fun bd_cinfinite_tac _ = rtac (bd_cinfinite_of_bnf bnf) 1;
- val set_bd_tacs =
- if ! quick_and_dirty then
- replicate (n + live) (K all_tac)
- else
- replicate n (K (mk_lift_set_bd_tac (bd_Card_order_of_bnf bnf))) @
- (map (fn thm => fn _ => rtac thm 1) (set_bd_of_bnf bnf));
-
- val in_alt_thm =
- let
- val inx = mk_in Asets sets T;
- val in_alt = mk_in (drop n Asets) bnf_sets T;
- val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt));
- in
- Skip_Proof.prove lthy [] [] goal (K lift_in_alt_tac) |> Thm.close_derivation
- end;
-
- fun in_bd_tac _ = mk_lift_in_bd_tac n in_alt_thm (in_bd_of_bnf bnf) (bd_Card_order_of_bnf bnf);
- fun map_wpull_tac _ = mk_map_wpull_tac in_alt_thm [] (map_wpull_of_bnf bnf);
-
- fun srel_O_Gr_tac _ =
- mk_simple_srel_O_Gr_tac lthy (srel_def_of_bnf bnf) (srel_O_Gr_of_bnf bnf) in_alt_thm;
-
- val tacs = zip_axioms map_id_tac map_comp_tac map_cong_tac set_natural_tacs bd_card_order_tac
- bd_cinfinite_tac set_bd_tacs in_bd_tac map_wpull_tac srel_O_Gr_tac;
-
- val wits = map snd (mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf);
-
- fun wit_tac _ = mk_simple_wit_tac (wit_thms_of_bnf bnf);
-
- val (bnf', lthy') =
- bnf_def Smart_Inline (K Derive_Few_Facts) qualify tacs wit_tac (SOME Ds)
- (((((b, mapx), sets), Term.absdummy T bd), wits), SOME rel) lthy;
-
- in
- (bnf', (add_bnf_to_unfolds bnf' unfold_set, lthy'))
- end;
-
-(* Changing the order of live variables *)
-
-fun permute_bnf qualify src dest bnf (unfold_set, lthy) =
- if src = dest then (bnf, (unfold_set, lthy)) else
- let
- val b = Binding.suffix_name (mk_permuteN src dest) (name_of_bnf bnf);
- val live = live_of_bnf bnf;
- val dead = dead_of_bnf bnf;
- val nwits = nwits_of_bnf bnf;
- fun permute xs = mk_permute src dest xs;
- fun permute_rev xs = mk_permute dest src xs;
-
- val (Ds, lthy1) = apfst (map TFree)
- (Variable.invent_types (replicate dead HOLogic.typeS) lthy);
- val (As, lthy2) = apfst (map TFree)
- (Variable.invent_types (replicate live HOLogic.typeS) lthy1);
- val (Bs, _(*lthy3*)) = apfst (map TFree)
- (Variable.invent_types (replicate live HOLogic.typeS) lthy2);
-
- val (Asets, _(*names_lthy*)) = lthy
- |> mk_Frees "A" (map HOLogic.mk_setT (permute As));
-
- val T = mk_T_of_bnf Ds As bnf;
-
- (*%f(1) ... f(n). bnf.map f\<sigma>(1) ... f\<sigma>(n)*)
- val mapx = fold_rev Term.absdummy (permute (map2 (curry op -->) As Bs))
- (Term.list_comb (mk_map_of_bnf Ds As Bs bnf, permute_rev (map Bound (live - 1 downto 0))));
- (*%Q(1) ... Q(n). bnf.rel Q\<sigma>(1) ... Q\<sigma>(n)*)
- val rel = fold_rev Term.absdummy (permute (map2 mk_pred2T As Bs))
- (Term.list_comb (mk_rel_of_bnf Ds As Bs bnf, permute_rev (map Bound (live - 1 downto 0))));
-
- val bnf_sets = mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf;
- val sets = permute bnf_sets;
-
- val bd = mk_bd_of_bnf Ds As bnf;
-
- fun map_id_tac _ = rtac (map_id_of_bnf bnf) 1;
- fun map_comp_tac _ = rtac (map_comp_of_bnf bnf) 1;
- fun map_cong_tac {context, ...} =
- rtac (map_cong_of_bnf bnf) 1 THEN REPEAT_DETERM_N live (Goal.assume_rule_tac context 1);
- val set_natural_tacs = permute (map (fn thm => fn _ => rtac thm 1) (set_natural_of_bnf bnf));
- fun bd_card_order_tac _ = rtac (bd_card_order_of_bnf bnf) 1;
- fun bd_cinfinite_tac _ = rtac (bd_cinfinite_of_bnf bnf) 1;
- val set_bd_tacs = permute (map (fn thm => fn _ => rtac thm 1) (set_bd_of_bnf bnf));
-
- val in_alt_thm =
- let
- val inx = mk_in Asets sets T;
- val in_alt = mk_in (permute_rev Asets) bnf_sets T;
- val goal = fold_rev Logic.all Asets (mk_Trueprop_eq (inx, in_alt));
- in
- Skip_Proof.prove lthy [] [] goal (K (mk_permute_in_alt_tac src dest))
- |> Thm.close_derivation
- end;
-
- fun in_bd_tac _ =
- mk_permute_in_bd_tac src dest in_alt_thm (in_bd_of_bnf bnf) (bd_Card_order_of_bnf bnf);
- fun map_wpull_tac _ = mk_map_wpull_tac in_alt_thm [] (map_wpull_of_bnf bnf);
-
- fun srel_O_Gr_tac _ =
- mk_simple_srel_O_Gr_tac lthy (srel_def_of_bnf bnf) (srel_O_Gr_of_bnf bnf) in_alt_thm;
-
- val tacs = zip_axioms map_id_tac map_comp_tac map_cong_tac set_natural_tacs bd_card_order_tac
- bd_cinfinite_tac set_bd_tacs in_bd_tac map_wpull_tac srel_O_Gr_tac;
-
- val wits = map snd (mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf);
-
- fun wit_tac _ = mk_simple_wit_tac (wit_thms_of_bnf bnf);
-
- val (bnf', lthy') =
- bnf_def Smart_Inline (K Derive_Few_Facts) qualify tacs wit_tac (SOME Ds)
- (((((b, mapx), sets), Term.absdummy T bd), wits), SOME rel) lthy;
- in
- (bnf', (add_bnf_to_unfolds bnf' unfold_set, lthy'))
- end;
-
-(* Composition pipeline *)
-
-fun permute_and_kill qualify n src dest bnf =
- bnf
- |> permute_bnf qualify src dest
- #> uncurry (kill_bnf qualify n);
-
-fun lift_and_permute qualify n src dest bnf =
- bnf
- |> lift_bnf qualify n
- #> uncurry (permute_bnf qualify src dest);
-
-fun normalize_bnfs qualify Ass Ds sort bnfs unfold_set lthy =
- let
- val before_kill_src = map (fn As => 0 upto (length As - 1)) Ass;
- val kill_poss = map (find_indices Ds) Ass;
- val live_poss = map2 (subtract (op =)) kill_poss before_kill_src;
- val before_kill_dest = map2 append kill_poss live_poss;
- val kill_ns = map length kill_poss;
- val (inners', (unfold_set', lthy')) =
- fold_map5 (fn i => permute_and_kill (qualify i))
- (if length bnfs = 1 then [0] else (1 upto length bnfs))
- kill_ns before_kill_src before_kill_dest bnfs (unfold_set, lthy);
-
- val Ass' = map2 (map o nth) Ass live_poss;
- val As = sort Ass';
- val after_lift_dest = replicate (length Ass') (0 upto (length As - 1));
- val old_poss = map (map (fn x => find_index (fn y => x = y) As)) Ass';
- val new_poss = map2 (subtract (op =)) old_poss after_lift_dest;
- val after_lift_src = map2 append new_poss old_poss;
- val lift_ns = map (fn xs => length As - length xs) Ass';
- in
- ((kill_poss, As), fold_map5 (fn i => lift_and_permute (qualify i))
- (if length bnfs = 1 then [0] else (1 upto length bnfs))
- lift_ns after_lift_src after_lift_dest inners' (unfold_set', lthy'))
- end;
-
-fun default_comp_sort Ass =
- Library.sort (Term_Ord.typ_ord o pairself TFree) (fold (fold (insert (op =))) Ass []);
-
-fun compose_bnf const_policy qualify sort outer inners oDs Dss tfreess (unfold_set, lthy) =
- let
- val b = name_of_bnf outer;
-
- val Ass = map (map Term.dest_TFree) tfreess;
- val Ds = fold (fold Term.add_tfreesT) (oDs :: Dss) [];
-
- val ((kill_poss, As), (inners', (unfold_set', lthy'))) =
- normalize_bnfs qualify Ass Ds sort inners unfold_set lthy;
-
- val Ds = oDs @ flat (map3 (append oo map o nth) tfreess kill_poss Dss);
- val As = map TFree As;
- in
- apfst (rpair (Ds, As))
- (clean_compose_bnf const_policy (qualify 0) b outer inners' (unfold_set', lthy'))
- end;
-
-(* Hide the type of the bound (optimization) and unfold the definitions (nicer to the user) *)
-
-fun seal_bnf unfold_set b Ds bnf lthy =
- let
- val live = live_of_bnf bnf;
- val nwits = nwits_of_bnf bnf;
-
- val (As, lthy1) = apfst (map TFree)
- (Variable.invent_types (replicate live HOLogic.typeS) (fold Variable.declare_typ Ds lthy));
- val (Bs, _) = apfst (map TFree)
- (Variable.invent_types (replicate live HOLogic.typeS) lthy1);
-
- val map_unfolds = map_unfolds_of unfold_set;
- val set_unfoldss = set_unfoldss_of unfold_set;
- val rel_unfolds = rel_unfolds_of unfold_set;
- val srel_unfolds = srel_unfolds_of unfold_set;
-
- val expand_maps =
- fold expand_term_const (map (single o Logic.dest_equals o Thm.prop_of) map_unfolds);
- val expand_sets =
- fold expand_term_const (map (map (Logic.dest_equals o Thm.prop_of)) set_unfoldss);
- val expand_rels =
- fold expand_term_const (map (single o Logic.dest_equals o Thm.prop_of) rel_unfolds);
- val unfold_maps = fold (unfold_thms lthy o single) map_unfolds;
- val unfold_sets = fold (unfold_thms lthy) set_unfoldss;
- val unfold_rels = unfold_thms lthy rel_unfolds;
- val unfold_srels = unfold_thms lthy srel_unfolds;
- val unfold_all = unfold_sets o unfold_maps o unfold_rels o unfold_srels;
- val bnf_map = expand_maps (mk_map_of_bnf Ds As Bs bnf);
- val bnf_sets = map (expand_maps o expand_sets)
- (mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf);
- val bnf_bd = mk_bd_of_bnf Ds As bnf;
- val bnf_rel = expand_rels (mk_rel_of_bnf Ds As Bs bnf);
- val T = mk_T_of_bnf Ds As bnf;
-
- (*bd should only depend on dead type variables!*)
- val bd_repT = fst (dest_relT (fastype_of bnf_bd));
- val bdT_bind = Binding.suffix_name ("_" ^ bdTN) b;
- val params = fold Term.add_tfreesT Ds [];
- val deads = map TFree params;
-
- val ((bdT_name, (bdT_glob_info, bdT_loc_info)), lthy) =
- typedef false NONE (bdT_bind, params, NoSyn)
- (HOLogic.mk_UNIV bd_repT) NONE (EVERY' [rtac exI, rtac UNIV_I] 1) lthy;
-
- val bnf_bd' = mk_dir_image bnf_bd
- (Const (#Abs_name bdT_glob_info, bd_repT --> Type (bdT_name, deads)))
-
- val Abs_bdT_inj = mk_Abs_inj_thm (#Abs_inject bdT_loc_info);
- val Abs_bdT_bij = mk_Abs_bij_thm lthy Abs_bdT_inj (#Abs_cases bdT_loc_info);
-
- val bd_ordIso = @{thm dir_image} OF [Abs_bdT_inj, bd_Card_order_of_bnf bnf];
- val bd_card_order =
- @{thm card_order_dir_image} OF [Abs_bdT_bij, bd_card_order_of_bnf bnf];
- val bd_cinfinite =
- (@{thm Cinfinite_cong} OF [bd_ordIso, bd_Cinfinite_of_bnf bnf]) RS conjunct1;
-
- val set_bds =
- map (fn thm => @{thm ordLeq_ordIso_trans} OF [thm, bd_ordIso]) (set_bd_of_bnf bnf);
- val in_bd =
- @{thm ordLeq_ordIso_trans} OF [in_bd_of_bnf bnf,
- @{thm cexp_cong2_Cnotzero} OF [bd_ordIso, if live = 0 then
- @{thm ctwo_Cnotzero} else @{thm ctwo_Cnotzero} RS @{thm csum_Cnotzero2},
- bd_Card_order_of_bnf bnf]];
-
- fun mk_tac thm {context = ctxt, prems = _} =
- (rtac (unfold_all thm) THEN'
- SOLVE o REPEAT_DETERM o (atac ORELSE' Goal.assume_rule_tac ctxt)) 1;
-
- val tacs = zip_axioms (mk_tac (map_id_of_bnf bnf)) (mk_tac (map_comp_of_bnf bnf))
- (mk_tac (map_cong_of_bnf bnf)) (map mk_tac (set_natural_of_bnf bnf))
- (K (rtac bd_card_order 1)) (K (rtac bd_cinfinite 1)) (map mk_tac set_bds) (mk_tac in_bd)
- (mk_tac (map_wpull_of_bnf bnf))
- (mk_tac (unfold_thms lthy [srel_def_of_bnf bnf] (srel_O_Gr_of_bnf bnf)));
-
- val bnf_wits = map snd (mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf);
-
- fun wit_tac _ = mk_simple_wit_tac (map unfold_all (wit_thms_of_bnf bnf));
-
- val policy = user_policy Derive_All_Facts;
-
- val (bnf', lthy') = bnf_def Hardly_Inline policy I tacs wit_tac (SOME deads)
- (((((b, bnf_map), bnf_sets), Term.absdummy T bnf_bd'), bnf_wits), SOME bnf_rel) lthy;
- in
- ((bnf', deads), lthy')
- end;
-
-val ID_bnf = the (bnf_of @{context} "Basic_BNFs.ID");
-val DEADID_bnf = the (bnf_of @{context} "Basic_BNFs.DEADID");
-
-fun bnf_of_typ _ _ _ (T as TFree _) accum = ((ID_bnf, ([], [T])), accum)
- | bnf_of_typ _ _ _ (TVar _) _ = error "Unexpected schematic variable"
- | bnf_of_typ const_policy qualify' sort (T as Type (C, Ts)) (unfold_set, lthy) =
- let
- val tfrees = Term.add_tfreesT T [];
- val bnf_opt = if null tfrees then NONE else bnf_of lthy C;
- in
- (case bnf_opt of
- NONE => ((DEADID_bnf, ([T], [])), (unfold_set, lthy))
- | SOME bnf =>
- if forall (can Term.dest_TFree) Ts andalso length Ts = length tfrees then
- let
- val T' = T_of_bnf bnf;
- val deads = deads_of_bnf bnf;
- val lives = lives_of_bnf bnf;
- val tvars' = Term.add_tvarsT T' [];
- val deads_lives =
- pairself (map (Term.typ_subst_TVars (map fst tvars' ~~ map TFree tfrees)))
- (deads, lives);
- in ((bnf, deads_lives), (unfold_set, lthy)) end
- else
- let
- val name = Long_Name.base_name C;
- fun qualify i =
- let val namei = name ^ nonzero_string_of_int i;
- in qualify' o Binding.qualify true namei end;
- val odead = dead_of_bnf bnf;
- val olive = live_of_bnf bnf;
- val oDs_pos = find_indices [TFree ("dead", [])] (snd (Term.dest_Type
- (mk_T_of_bnf (replicate odead (TFree ("dead", []))) (replicate olive dummyT) bnf)));
- val oDs = map (nth Ts) oDs_pos;
- val Ts' = map (nth Ts) (subtract (op =) oDs_pos (0 upto length Ts - 1));
- val ((inners, (Dss, Ass)), (unfold_set', lthy')) =
- apfst (apsnd split_list o split_list)
- (fold_map2 (fn i => bnf_of_typ Smart_Inline (qualify i) sort)
- (if length Ts' = 1 then [0] else (1 upto length Ts')) Ts' (unfold_set, lthy));
- in
- compose_bnf const_policy qualify sort bnf inners oDs Dss Ass (unfold_set', lthy')
- end)
- end;
-
-end;
--- a/src/HOL/Codatatype/Tools/bnf_comp_tactics.ML Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,416 +0,0 @@
-(* Title: HOL/BNF/Tools/bnf_comp_tactics.ML
- Author: Dmitriy Traytel, TU Muenchen
- Author: Jasmin Blanchette, TU Muenchen
- Copyright 2012
-
-Tactics for composition of bounded natural functors.
-*)
-
-signature BNF_COMP_TACTICS =
-sig
- val mk_comp_bd_card_order_tac: thm list -> thm -> tactic
- val mk_comp_bd_cinfinite_tac: thm -> thm -> tactic
- val mk_comp_in_alt_tac: Proof.context -> thm list -> tactic
- val mk_comp_in_bd_tac: thm -> thm list -> thm -> thm list -> thm -> tactic
- val mk_comp_map_comp_tac: thm -> thm -> thm list -> tactic
- val mk_comp_map_cong_tac: thm list -> thm -> thm list -> tactic
- val mk_comp_set_alt_tac: Proof.context -> thm -> tactic
- val mk_comp_set_bd_tac: Proof.context -> thm -> thm list -> tactic
- val mk_comp_set_natural_tac: thm -> thm -> thm -> thm list -> tactic
- val mk_comp_wit_tac: Proof.context -> thm list -> thm -> thm list -> tactic
-
- val mk_kill_bd_card_order_tac: int -> thm -> tactic
- val mk_kill_bd_cinfinite_tac: thm -> tactic
- val kill_in_alt_tac: tactic
- val mk_kill_in_bd_tac: int -> bool -> thm -> thm -> thm -> thm -> thm -> tactic
- val mk_kill_map_cong_tac: Proof.context -> int -> int -> thm -> tactic
- val mk_kill_set_bd_tac: thm -> thm -> tactic
-
- val empty_natural_tac: tactic
- val lift_in_alt_tac: tactic
- val mk_lift_in_bd_tac: int -> thm -> thm -> thm -> tactic
- val mk_lift_set_bd_tac: thm -> tactic
-
- val mk_permute_in_alt_tac: ''a list -> ''a list -> tactic
- val mk_permute_in_bd_tac: ''a list -> ''a list -> thm -> thm -> thm -> tactic
-
- val mk_map_wpull_tac: thm -> thm list -> thm -> tactic
- val mk_simple_srel_O_Gr_tac: Proof.context -> thm -> thm -> thm -> tactic
- val mk_simple_wit_tac: thm list -> tactic
-end;
-
-structure BNF_Comp_Tactics : BNF_COMP_TACTICS =
-struct
-
-open BNF_Util
-open BNF_Tactics
-
-val Card_order_csum = @{thm Card_order_csum};
-val Card_order_ctwo = @{thm Card_order_ctwo};
-val Cnotzero_UNIV = @{thm Cnotzero_UNIV};
-val arg_cong_Union = @{thm arg_cong[of _ _ Union]};
-val card_of_Card_order = @{thm card_of_Card_order};
-val csum_Cnotzero1 = @{thm csum_Cnotzero1};
-val csum_Cnotzero2 = @{thm csum_Cnotzero2};
-val ctwo_Cnotzero = @{thm ctwo_Cnotzero};
-val o_eq_dest_lhs = @{thm o_eq_dest_lhs};
-val ordIso_transitive = @{thm ordIso_transitive};
-val ordLeq_csum2 = @{thm ordLeq_csum2};
-val trans_image_cong_o_apply = @{thm trans[OF image_cong[OF o_apply refl]]};
-val trans_o_apply = @{thm trans[OF o_apply]};
-
-
-
-(* Composition *)
-
-fun mk_comp_set_alt_tac ctxt collect_set_natural =
- unfold_thms_tac ctxt @{thms sym[OF o_assoc]} THEN
- unfold_thms_tac ctxt [collect_set_natural RS sym] THEN
- rtac refl 1;
-
-fun mk_comp_map_comp_tac Gmap_comp Gmap_cong map_comps =
- EVERY' ([rtac ext, rtac sym, rtac trans_o_apply,
- rtac (Gmap_comp RS sym RS o_eq_dest_lhs RS trans), rtac Gmap_cong] @
- map (fn thm => rtac (thm RS sym RS fun_cong)) map_comps) 1;
-
-fun mk_comp_set_natural_tac Gmap_comp Gmap_cong Gset_natural set_naturals =
- EVERY' ([rtac ext] @
- replicate 3 (rtac trans_o_apply) @
- [rtac (arg_cong_Union RS trans),
- rtac (@{thm arg_cong2[of _ _ _ _ collect, OF refl]} RS trans),
- rtac (Gmap_comp RS sym RS o_eq_dest_lhs RS trans),
- rtac Gmap_cong] @
- map (fn thm => rtac (thm RS fun_cong)) set_naturals @
- [rtac (Gset_natural RS o_eq_dest_lhs), rtac sym, rtac trans_o_apply,
- rtac trans_image_cong_o_apply, rtac trans_image_cong_o_apply,
- rtac (@{thm image_cong} OF [Gset_natural RS o_eq_dest_lhs RS arg_cong_Union, refl] RS trans),
- rtac @{thm trans[OF pointfreeE[OF Union_natural[symmetric]]]}, rtac arg_cong_Union,
- rtac @{thm trans[OF o_eq_dest_lhs[OF image_o_collect[symmetric]]]},
- rtac @{thm fun_cong[OF arg_cong[of _ _ collect]]}] @
- [REPEAT_DETERM_N (length set_naturals) o EVERY' [rtac @{thm trans[OF image_insert]},
- rtac @{thm arg_cong2[of _ _ _ _ insert]}, rtac ext, rtac trans_o_apply,
- rtac trans_image_cong_o_apply, rtac @{thm trans[OF image_image]},
- rtac @{thm sym[OF trans[OF o_apply]]}, rtac @{thm image_cong[OF refl o_apply]}],
- rtac @{thm image_empty}]) 1;
-
-fun mk_comp_map_cong_tac comp_set_alts map_cong map_congs =
- let
- val n = length comp_set_alts;
- in
- (if n = 0 then rtac refl 1
- else rtac map_cong 1 THEN
- EVERY' (map_index (fn (i, map_cong) =>
- rtac map_cong THEN' EVERY' (map_index (fn (k, set_alt) =>
- EVERY' [select_prem_tac n (dtac @{thm meta_spec}) (k + 1), etac @{thm meta_mp},
- rtac (equalityD2 RS set_mp), rtac (set_alt RS fun_cong RS trans),
- rtac trans_o_apply, rtac (@{thm collect_def} RS arg_cong_Union),
- rtac @{thm UnionI}, rtac @{thm UN_I}, REPEAT_DETERM_N i o rtac @{thm insertI2},
- rtac @{thm insertI1}, rtac (o_apply RS equalityD2 RS set_mp),
- etac @{thm imageI}, atac])
- comp_set_alts))
- map_congs) 1)
- end;
-
-fun mk_comp_bd_card_order_tac Fbd_card_orders Gbd_card_order =
- let
- val (card_orders, last_card_order) = split_last Fbd_card_orders;
- fun gen_before thm = rtac @{thm card_order_csum} THEN' rtac thm;
- in
- (rtac @{thm card_order_cprod} THEN'
- WRAP' gen_before (K (K all_tac)) card_orders (rtac last_card_order) THEN'
- rtac Gbd_card_order) 1
- end;
-
-fun mk_comp_bd_cinfinite_tac Fbd_cinfinite Gbd_cinfinite =
- (rtac @{thm cinfinite_cprod} THEN'
- ((K (TRY ((rtac @{thm cinfinite_csum} THEN' rtac disjI1) 1)) THEN'
- ((rtac @{thm cinfinite_csum} THEN' rtac disjI1 THEN' rtac Fbd_cinfinite) ORELSE'
- rtac Fbd_cinfinite)) ORELSE'
- rtac Fbd_cinfinite) THEN'
- rtac Gbd_cinfinite) 1;
-
-fun mk_comp_set_bd_tac ctxt comp_set_alt Gset_Fset_bds =
- let
- val (bds, last_bd) = split_last Gset_Fset_bds;
- fun gen_before bd =
- rtac ctrans THEN' rtac @{thm Un_csum} THEN'
- rtac ctrans THEN' rtac @{thm csum_mono} THEN'
- rtac bd;
- fun gen_after _ = rtac @{thm ordIso_imp_ordLeq} THEN' rtac @{thm cprod_csum_distrib1};
- in
- unfold_thms_tac ctxt [comp_set_alt] THEN
- rtac @{thm comp_set_bd_Union_o_collect} 1 THEN
- unfold_thms_tac ctxt @{thms Union_image_insert Union_image_empty Union_Un_distrib o_apply} THEN
- (rtac ctrans THEN'
- WRAP' gen_before gen_after bds (rtac last_bd) THEN'
- rtac @{thm ordIso_imp_ordLeq} THEN'
- rtac @{thm cprod_com}) 1
- end;
-
-val comp_in_alt_thms = @{thms o_apply collect_def SUP_def image_insert image_empty Union_insert
- Union_empty Un_empty_right Union_Un_distrib Un_subset_iff conj_subset_def UN_image_subset
- conj_assoc};
-
-fun mk_comp_in_alt_tac ctxt comp_set_alts =
- unfold_thms_tac ctxt (comp_set_alts @ comp_in_alt_thms) THEN
- unfold_thms_tac ctxt @{thms set_eq_subset} THEN
- rtac conjI 1 THEN
- REPEAT_DETERM (
- rtac @{thm subsetI} 1 THEN
- unfold_thms_tac ctxt @{thms mem_Collect_eq Ball_def} THEN
- (REPEAT_DETERM (CHANGED (etac conjE 1)) THEN
- REPEAT_DETERM (CHANGED ((
- (rtac conjI THEN' (atac ORELSE' rtac subset_UNIV)) ORELSE'
- atac ORELSE'
- (rtac subset_UNIV)) 1)) ORELSE rtac subset_UNIV 1));
-
-fun mk_comp_in_bd_tac comp_in_alt Fin_bds Gin_bd Fbd_Cinfs Gbd_Card_order =
- let
- val (bds, last_bd) = split_last Fin_bds;
- val (Cinfs, _) = split_last Fbd_Cinfs;
- fun gen_before (bd, _) = rtac ctrans THEN' rtac @{thm csum_mono} THEN' rtac bd;
- fun gen_after (_, (bd_Cinf, next_bd_Cinf)) =
- TRY o (rtac @{thm csum_cexp} THEN'
- rtac bd_Cinf THEN'
- (TRY o (rtac @{thm Cinfinite_csum} THEN' rtac disjI1) THEN' rtac next_bd_Cinf ORELSE'
- rtac next_bd_Cinf) THEN'
- ((rtac Card_order_csum THEN' rtac ordLeq_csum2) ORELSE'
- (rtac Card_order_ctwo THEN' rtac @{thm ordLeq_refl})) THEN'
- rtac Card_order_ctwo);
- in
- (rtac @{thm ordIso_ordLeq_trans} THEN'
- rtac @{thm card_of_ordIso_subst} THEN'
- rtac comp_in_alt THEN'
- rtac ctrans THEN'
- rtac Gin_bd THEN'
- rtac @{thm ordLeq_ordIso_trans} THEN'
- rtac @{thm cexp_mono1} THEN'
- rtac @{thm ordLeq_ordIso_trans} THEN'
- rtac @{thm csum_mono1} THEN'
- WRAP' gen_before gen_after (bds ~~ (Cinfs ~~ tl Fbd_Cinfs)) (rtac last_bd) THEN'
- rtac @{thm csum_absorb1} THEN'
- rtac @{thm Cinfinite_cexp} THEN'
- (rtac ordLeq_csum2 ORELSE' rtac @{thm ordLeq_refl}) THEN'
- rtac Card_order_ctwo THEN'
- (TRY o (rtac @{thm Cinfinite_csum} THEN' rtac disjI1) THEN' rtac (hd Fbd_Cinfs) ORELSE'
- rtac (hd Fbd_Cinfs)) THEN'
- rtac @{thm ctwo_ordLeq_Cinfinite} THEN'
- rtac @{thm Cinfinite_cexp} THEN'
- (rtac ordLeq_csum2 ORELSE' rtac @{thm ordLeq_refl}) THEN'
- rtac Card_order_ctwo THEN'
- (TRY o (rtac @{thm Cinfinite_csum} THEN' rtac disjI1) THEN' rtac (hd Fbd_Cinfs) ORELSE'
- rtac (hd Fbd_Cinfs)) THEN'
- rtac disjI1 THEN'
- TRY o rtac csum_Cnotzero2 THEN'
- rtac ctwo_Cnotzero THEN'
- rtac Gbd_Card_order THEN'
- rtac @{thm cexp_cprod} THEN'
- TRY o rtac csum_Cnotzero2 THEN'
- rtac ctwo_Cnotzero) 1
- end;
-
-val comp_wit_thms = @{thms Union_empty_conv o_apply collect_def SUP_def
- Union_image_insert Union_image_empty};
-
-fun mk_comp_wit_tac ctxt Gwit_thms collect_set_natural Fwit_thms =
- ALLGOALS (dtac @{thm in_Union_o_assoc}) THEN
- unfold_thms_tac ctxt (collect_set_natural :: comp_wit_thms) THEN
- REPEAT_DETERM (
- atac 1 ORELSE
- REPEAT_DETERM (eresolve_tac @{thms UnionE UnE imageE} 1) THEN
- (TRY o dresolve_tac Gwit_thms THEN'
- (etac FalseE ORELSE'
- hyp_subst_tac THEN'
- dresolve_tac Fwit_thms THEN'
- (etac FalseE ORELSE' atac))) 1);
-
-
-
-(* Kill operation *)
-
-fun mk_kill_map_cong_tac ctxt n m map_cong =
- (rtac map_cong THEN' EVERY' (replicate n (rtac refl)) THEN'
- EVERY' (replicate m (Goal.assume_rule_tac ctxt))) 1;
-
-fun mk_kill_bd_card_order_tac n bd_card_order =
- (rtac @{thm card_order_cprod} THEN'
- K (REPEAT_DETERM_N (n - 1)
- ((rtac @{thm card_order_csum} THEN'
- rtac @{thm card_of_card_order_on}) 1)) THEN'
- rtac @{thm card_of_card_order_on} THEN'
- rtac bd_card_order) 1;
-
-fun mk_kill_bd_cinfinite_tac bd_Cinfinite =
- (rtac @{thm cinfinite_cprod2} THEN'
- TRY o rtac csum_Cnotzero1 THEN'
- rtac Cnotzero_UNIV THEN'
- rtac bd_Cinfinite) 1;
-
-fun mk_kill_set_bd_tac bd_Card_order set_bd =
- (rtac ctrans THEN'
- rtac set_bd THEN'
- rtac @{thm ordLeq_cprod2} THEN'
- TRY o rtac csum_Cnotzero1 THEN'
- rtac Cnotzero_UNIV THEN'
- rtac bd_Card_order) 1
-
-val kill_in_alt_tac =
- ((rtac @{thm Collect_cong} THEN' rtac iffI) 1 THEN
- REPEAT_DETERM (CHANGED (etac conjE 1)) THEN
- REPEAT_DETERM (CHANGED ((etac conjI ORELSE'
- rtac conjI THEN' rtac subset_UNIV) 1)) THEN
- (rtac subset_UNIV ORELSE' atac) 1 THEN
- REPEAT_DETERM (CHANGED (etac conjE 1)) THEN
- REPEAT_DETERM (CHANGED ((etac conjI ORELSE' atac) 1))) ORELSE
- ((rtac @{thm UNIV_eq_I} THEN' rtac CollectI) 1 THEN
- REPEAT_DETERM (TRY (rtac conjI 1) THEN rtac subset_UNIV 1));
-
-fun mk_kill_in_bd_tac n nontrivial_kill_in in_alt in_bd bd_Card_order bd_Cinfinite bd_Cnotzero =
- (rtac @{thm ordIso_ordLeq_trans} THEN'
- rtac @{thm card_of_ordIso_subst} THEN'
- rtac in_alt THEN'
- rtac ctrans THEN'
- rtac in_bd THEN'
- rtac @{thm ordIso_ordLeq_trans} THEN'
- rtac @{thm cexp_cong1}) 1 THEN
- (if nontrivial_kill_in then
- rtac ordIso_transitive 1 THEN
- REPEAT_DETERM_N (n - 1)
- ((rtac @{thm csum_cong1} THEN'
- rtac @{thm ordIso_symmetric} THEN'
- rtac @{thm csum_assoc} THEN'
- rtac ordIso_transitive) 1) THEN
- (rtac @{thm ordIso_refl} THEN'
- rtac Card_order_csum THEN'
- rtac ordIso_transitive THEN'
- rtac @{thm csum_assoc} THEN'
- rtac ordIso_transitive THEN'
- rtac @{thm csum_cong1} THEN'
- K (mk_flatten_assoc_tac
- (rtac @{thm ordIso_refl} THEN'
- FIRST' [rtac card_of_Card_order, rtac Card_order_csum])
- ordIso_transitive @{thm csum_assoc} @{thm csum_cong}) THEN'
- rtac @{thm ordIso_refl} THEN'
- (rtac card_of_Card_order ORELSE' rtac Card_order_csum)) 1
- else all_tac) THEN
- (rtac @{thm csum_com} THEN'
- rtac bd_Card_order THEN'
- rtac disjI1 THEN'
- rtac csum_Cnotzero2 THEN'
- rtac ctwo_Cnotzero THEN'
- rtac disjI1 THEN'
- rtac csum_Cnotzero2 THEN'
- TRY o rtac csum_Cnotzero1 THEN'
- rtac Cnotzero_UNIV THEN'
- rtac @{thm ordLeq_ordIso_trans} THEN'
- rtac @{thm cexp_mono1} THEN'
- rtac ctrans THEN'
- rtac @{thm csum_mono2} THEN'
- rtac @{thm ordLeq_cprod1} THEN'
- (rtac card_of_Card_order ORELSE' rtac Card_order_csum) THEN'
- rtac bd_Cnotzero THEN'
- rtac @{thm csum_cexp'} THEN'
- rtac @{thm Cinfinite_cprod2} THEN'
- TRY o rtac csum_Cnotzero1 THEN'
- rtac Cnotzero_UNIV THEN'
- rtac bd_Cinfinite THEN'
- ((rtac Card_order_ctwo THEN' rtac @{thm ordLeq_refl}) ORELSE'
- (rtac Card_order_csum THEN' rtac ordLeq_csum2)) THEN'
- rtac Card_order_ctwo THEN'
- rtac disjI1 THEN'
- rtac csum_Cnotzero2 THEN'
- TRY o rtac csum_Cnotzero1 THEN'
- rtac Cnotzero_UNIV THEN'
- rtac bd_Card_order THEN'
- rtac @{thm cexp_cprod_ordLeq} THEN'
- TRY o rtac csum_Cnotzero2 THEN'
- rtac ctwo_Cnotzero THEN'
- rtac @{thm Cinfinite_cprod2} THEN'
- TRY o rtac csum_Cnotzero1 THEN'
- rtac Cnotzero_UNIV THEN'
- rtac bd_Cinfinite THEN'
- rtac bd_Cnotzero THEN'
- rtac @{thm ordLeq_cprod2} THEN'
- TRY o rtac csum_Cnotzero1 THEN'
- rtac Cnotzero_UNIV THEN'
- rtac bd_Card_order) 1;
-
-
-
-(* Lift operation *)
-
-val empty_natural_tac = rtac @{thm empty_natural} 1;
-
-fun mk_lift_set_bd_tac bd_Card_order = (rtac @{thm Card_order_empty} THEN' rtac bd_Card_order) 1;
-
-val lift_in_alt_tac =
- ((rtac @{thm Collect_cong} THEN' rtac iffI) 1 THEN
- REPEAT_DETERM (CHANGED (etac conjE 1)) THEN
- REPEAT_DETERM (CHANGED ((etac conjI ORELSE' atac) 1)) THEN
- REPEAT_DETERM (CHANGED (etac conjE 1)) THEN
- REPEAT_DETERM (CHANGED ((etac conjI ORELSE'
- rtac conjI THEN' rtac @{thm empty_subsetI}) 1)) THEN
- (rtac @{thm empty_subsetI} ORELSE' atac) 1) ORELSE
- ((rtac sym THEN' rtac @{thm UNIV_eq_I} THEN' rtac CollectI) 1 THEN
- REPEAT_DETERM (TRY (rtac conjI 1) THEN rtac @{thm empty_subsetI} 1));
-
-fun mk_lift_in_bd_tac n in_alt in_bd bd_Card_order =
- (rtac @{thm ordIso_ordLeq_trans} THEN'
- rtac @{thm card_of_ordIso_subst} THEN'
- rtac in_alt THEN'
- rtac ctrans THEN'
- rtac in_bd THEN'
- rtac @{thm cexp_mono1}) 1 THEN
- ((rtac @{thm csum_mono1} 1 THEN
- REPEAT_DETERM_N (n - 1)
- ((rtac ctrans THEN'
- rtac ordLeq_csum2 THEN'
- (rtac Card_order_csum ORELSE' rtac card_of_Card_order)) 1) THEN
- (rtac ordLeq_csum2 THEN'
- (rtac Card_order_csum ORELSE' rtac card_of_Card_order)) 1) ORELSE
- (rtac ordLeq_csum2 THEN' rtac Card_order_ctwo) 1) THEN
- (rtac disjI1 THEN' TRY o rtac csum_Cnotzero2 THEN' rtac ctwo_Cnotzero
- THEN' rtac bd_Card_order) 1;
-
-
-
-(* Permute operation *)
-
-fun mk_permute_in_alt_tac src dest =
- (rtac @{thm Collect_cong} THEN'
- mk_rotate_eq_tac (rtac refl) trans @{thm conj_assoc} @{thm conj_commute} @{thm conj_cong}
- dest src) 1;
-
-fun mk_permute_in_bd_tac src dest in_alt in_bd bd_Card_order =
- (rtac @{thm ordIso_ordLeq_trans} THEN'
- rtac @{thm card_of_ordIso_subst} THEN'
- rtac in_alt THEN'
- rtac @{thm ordLeq_ordIso_trans} THEN'
- rtac in_bd THEN'
- rtac @{thm cexp_cong1} THEN'
- rtac @{thm csum_cong1} THEN'
- mk_rotate_eq_tac
- (rtac @{thm ordIso_refl} THEN'
- FIRST' [rtac card_of_Card_order, rtac Card_order_csum])
- ordIso_transitive @{thm csum_assoc} @{thm csum_com} @{thm csum_cong}
- src dest THEN'
- rtac bd_Card_order THEN'
- rtac disjI1 THEN'
- TRY o rtac csum_Cnotzero2 THEN'
- rtac ctwo_Cnotzero THEN'
- rtac disjI1 THEN'
- TRY o rtac csum_Cnotzero2 THEN'
- rtac ctwo_Cnotzero) 1;
-
-fun mk_map_wpull_tac comp_in_alt inner_map_wpulls outer_map_wpull =
- (rtac (@{thm wpull_cong} OF (replicate 3 comp_in_alt)) THEN' rtac outer_map_wpull) 1 THEN
- WRAP (fn thm => rtac thm 1 THEN REPEAT_DETERM (atac 1)) (K all_tac) inner_map_wpulls all_tac THEN
- TRY (REPEAT_DETERM (atac 1 ORELSE rtac @{thm wpull_id} 1));
-
-fun mk_simple_srel_O_Gr_tac ctxt srel_def srel_O_Gr in_alt_thm =
- rtac (unfold_thms ctxt [srel_def]
- (trans OF [srel_O_Gr, in_alt_thm RS @{thm subst_rel_def} RS sym])) 1;
-
-fun mk_simple_wit_tac wit_thms = ALLGOALS (atac ORELSE' eresolve_tac (@{thm emptyE} :: wit_thms));
-
-end;
--- a/src/HOL/Codatatype/Tools/bnf_def.ML Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1238 +0,0 @@
-(* Title: HOL/BNF/Tools/bnf_def.ML
- Author: Dmitriy Traytel, TU Muenchen
- Author: Jasmin Blanchette, TU Muenchen
- Copyright 2012
-
-Definition of bounded natural functors.
-*)
-
-signature BNF_DEF =
-sig
- type BNF
- type nonemptiness_witness = {I: int list, wit: term, prop: thm list}
-
- val bnf_of: Proof.context -> string -> BNF option
- val register_bnf: string -> (BNF * local_theory) -> (BNF * local_theory)
-
- val name_of_bnf: BNF -> binding
- val T_of_bnf: BNF -> typ
- val live_of_bnf: BNF -> int
- val lives_of_bnf: BNF -> typ list
- val dead_of_bnf: BNF -> int
- val deads_of_bnf: BNF -> typ list
- val nwits_of_bnf: BNF -> int
-
- val mapN: string
- val relN: string
- val setN: string
- val mk_setN: int -> string
- val srelN: string
-
- val map_of_bnf: BNF -> term
-
- val mk_T_of_bnf: typ list -> typ list -> BNF -> typ
- val mk_bd_of_bnf: typ list -> typ list -> BNF -> term
- val mk_map_of_bnf: typ list -> typ list -> typ list -> BNF -> term
- val mk_rel_of_bnf: typ list -> typ list -> typ list -> BNF -> term
- val mk_sets_of_bnf: typ list list -> typ list list -> BNF -> term list
- val mk_srel_of_bnf: typ list -> typ list -> typ list -> BNF -> term
- val mk_wits_of_bnf: typ list list -> typ list list -> BNF -> (int list * term) list
-
- val bd_Card_order_of_bnf: BNF -> thm
- val bd_Cinfinite_of_bnf: BNF -> thm
- val bd_Cnotzero_of_bnf: BNF -> thm
- val bd_card_order_of_bnf: BNF -> thm
- val bd_cinfinite_of_bnf: BNF -> thm
- val collect_set_natural_of_bnf: BNF -> thm
- val in_bd_of_bnf: BNF -> thm
- val in_cong_of_bnf: BNF -> thm
- val in_mono_of_bnf: BNF -> thm
- val in_srel_of_bnf: BNF -> thm
- val map_comp'_of_bnf: BNF -> thm
- val map_comp_of_bnf: BNF -> thm
- val map_cong_of_bnf: BNF -> thm
- val map_def_of_bnf: BNF -> thm
- val map_id'_of_bnf: BNF -> thm
- val map_id_of_bnf: BNF -> thm
- val map_wppull_of_bnf: BNF -> thm
- val map_wpull_of_bnf: BNF -> thm
- val rel_def_of_bnf: BNF -> thm
- val set_bd_of_bnf: BNF -> thm list
- val set_defs_of_bnf: BNF -> thm list
- val set_natural'_of_bnf: BNF -> thm list
- val set_natural_of_bnf: BNF -> thm list
- val sets_of_bnf: BNF -> term list
- val srel_def_of_bnf: BNF -> thm
- val srel_Gr_of_bnf: BNF -> thm
- val srel_Id_of_bnf: BNF -> thm
- val srel_O_of_bnf: BNF -> thm
- val srel_O_Gr_of_bnf: BNF -> thm
- val srel_cong_of_bnf: BNF -> thm
- val srel_converse_of_bnf: BNF -> thm
- val srel_mono_of_bnf: BNF -> thm
- val wit_thms_of_bnf: BNF -> thm list
- val wit_thmss_of_bnf: BNF -> thm list list
-
- val mk_witness: int list * term -> thm list -> nonemptiness_witness
- val minimize_wits: (''a list * 'b) list -> (''a list * 'b) list
- val wits_of_bnf: BNF -> nonemptiness_witness list
-
- val zip_axioms: 'a -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a -> 'a list
-
- datatype const_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline
- datatype fact_policy =
- Derive_Few_Facts | Derive_All_Facts | Derive_All_Facts_Note_Most | Note_All_Facts_and_Axioms
- val bnf_note_all: bool Config.T
- val user_policy: fact_policy -> Proof.context -> fact_policy
-
- val print_bnfs: Proof.context -> unit
- val bnf_def: const_policy -> (Proof.context -> fact_policy) -> (binding -> binding) ->
- ({prems: thm list, context: Proof.context} -> tactic) list ->
- ({prems: thm list, context: Proof.context} -> tactic) -> typ list option ->
- ((((binding * term) * term list) * term) * term list) * term option -> local_theory ->
- BNF * local_theory
-end;
-
-structure BNF_Def : BNF_DEF =
-struct
-
-open BNF_Util
-open BNF_Tactics
-open BNF_Def_Tactics
-
-type axioms = {
- map_id: thm,
- map_comp: thm,
- map_cong: thm,
- set_natural: thm list,
- bd_card_order: thm,
- bd_cinfinite: thm,
- set_bd: thm list,
- in_bd: thm,
- map_wpull: thm,
- srel_O_Gr: thm
-};
-
-fun mk_axioms' (((((((((id, comp), cong), nat), c_o), cinf), set_bd), in_bd), wpull), srel) =
- {map_id = id, map_comp = comp, map_cong = cong, set_natural = nat, bd_card_order = c_o,
- bd_cinfinite = cinf, set_bd = set_bd, in_bd = in_bd, map_wpull = wpull, srel_O_Gr = srel};
-
-fun dest_cons [] = raise Empty
- | dest_cons (x :: xs) = (x, xs);
-
-fun mk_axioms n thms = thms
- |> map the_single
- |> dest_cons
- ||>> dest_cons
- ||>> dest_cons
- ||>> chop n
- ||>> dest_cons
- ||>> dest_cons
- ||>> chop n
- ||>> dest_cons
- ||>> dest_cons
- ||> the_single
- |> mk_axioms';
-
-fun zip_axioms mid mcomp mcong snat bdco bdinf sbd inbd wpull srel =
- [mid, mcomp, mcong] @ snat @ [bdco, bdinf] @ sbd @ [inbd, wpull, srel];
-
-fun dest_axioms {map_id, map_comp, map_cong, set_natural, bd_card_order, bd_cinfinite, set_bd,
- in_bd, map_wpull, srel_O_Gr} =
- zip_axioms map_id map_comp map_cong set_natural bd_card_order bd_cinfinite set_bd in_bd map_wpull
- srel_O_Gr;
-
-fun map_axioms f {map_id, map_comp, map_cong, set_natural, bd_card_order, bd_cinfinite, set_bd,
- in_bd, map_wpull, srel_O_Gr} =
- {map_id = f map_id,
- map_comp = f map_comp,
- map_cong = f map_cong,
- set_natural = map f set_natural,
- bd_card_order = f bd_card_order,
- bd_cinfinite = f bd_cinfinite,
- set_bd = map f set_bd,
- in_bd = f in_bd,
- map_wpull = f map_wpull,
- srel_O_Gr = f srel_O_Gr};
-
-val morph_axioms = map_axioms o Morphism.thm;
-
-type defs = {
- map_def: thm,
- set_defs: thm list,
- rel_def: thm,
- srel_def: thm
-}
-
-fun mk_defs map sets rel srel = {map_def = map, set_defs = sets, rel_def = rel, srel_def = srel};
-
-fun map_defs f {map_def, set_defs, rel_def, srel_def} =
- {map_def = f map_def, set_defs = map f set_defs, rel_def = f rel_def, srel_def = f srel_def};
-
-val morph_defs = map_defs o Morphism.thm;
-
-type facts = {
- bd_Card_order: thm,
- bd_Cinfinite: thm,
- bd_Cnotzero: thm,
- collect_set_natural: thm lazy,
- in_cong: thm lazy,
- in_mono: thm lazy,
- in_srel: thm lazy,
- map_comp': thm lazy,
- map_id': thm lazy,
- map_wppull: thm lazy,
- set_natural': thm lazy list,
- srel_cong: thm lazy,
- srel_mono: thm lazy,
- srel_Id: thm lazy,
- srel_Gr: thm lazy,
- srel_converse: thm lazy,
- srel_O: thm lazy
-};
-
-fun mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_natural in_cong in_mono in_srel
- map_comp' map_id' map_wppull set_natural' srel_cong srel_mono srel_Id srel_Gr srel_converse
- srel_O = {
- bd_Card_order = bd_Card_order,
- bd_Cinfinite = bd_Cinfinite,
- bd_Cnotzero = bd_Cnotzero,
- collect_set_natural = collect_set_natural,
- in_cong = in_cong,
- in_mono = in_mono,
- in_srel = in_srel,
- map_comp' = map_comp',
- map_id' = map_id',
- map_wppull = map_wppull,
- set_natural' = set_natural',
- srel_cong = srel_cong,
- srel_mono = srel_mono,
- srel_Id = srel_Id,
- srel_Gr = srel_Gr,
- srel_converse = srel_converse,
- srel_O = srel_O};
-
-fun map_facts f {
- bd_Card_order,
- bd_Cinfinite,
- bd_Cnotzero,
- collect_set_natural,
- in_cong,
- in_mono,
- in_srel,
- map_comp',
- map_id',
- map_wppull,
- set_natural',
- srel_cong,
- srel_mono,
- srel_Id,
- srel_Gr,
- srel_converse,
- srel_O} =
- {bd_Card_order = f bd_Card_order,
- bd_Cinfinite = f bd_Cinfinite,
- bd_Cnotzero = f bd_Cnotzero,
- collect_set_natural = Lazy.map f collect_set_natural,
- in_cong = Lazy.map f in_cong,
- in_mono = Lazy.map f in_mono,
- in_srel = Lazy.map f in_srel,
- map_comp' = Lazy.map f map_comp',
- map_id' = Lazy.map f map_id',
- map_wppull = Lazy.map f map_wppull,
- set_natural' = map (Lazy.map f) set_natural',
- srel_cong = Lazy.map f srel_cong,
- srel_mono = Lazy.map f srel_mono,
- srel_Id = Lazy.map f srel_Id,
- srel_Gr = Lazy.map f srel_Gr,
- srel_converse = Lazy.map f srel_converse,
- srel_O = Lazy.map f srel_O};
-
-val morph_facts = map_facts o Morphism.thm;
-
-type nonemptiness_witness = {
- I: int list,
- wit: term,
- prop: thm list
-};
-
-fun mk_witness (I, wit) prop = {I = I, wit = wit, prop = prop};
-fun map_witness f g {I, wit, prop} = {I = I, wit = f wit, prop = map g prop};
-fun morph_witness phi = map_witness (Morphism.term phi) (Morphism.thm phi);
-
-datatype BNF = BNF of {
- name: binding,
- T: typ,
- live: int,
- lives: typ list, (*source type variables of map, only for composition*)
- lives': typ list, (*target type variables of map, only for composition*)
- dead: int,
- deads: typ list, (*only for composition*)
- map: term,
- sets: term list,
- bd: term,
- axioms: axioms,
- defs: defs,
- facts: facts,
- nwits: int,
- wits: nonemptiness_witness list,
- rel: term,
- srel: term
-};
-
-(* getters *)
-
-fun rep_bnf (BNF bnf) = bnf;
-val name_of_bnf = #name o rep_bnf;
-val T_of_bnf = #T o rep_bnf;
-fun mk_T_of_bnf Ds Ts bnf =
- let val bnf_rep = rep_bnf bnf
- in Term.typ_subst_atomic ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#T bnf_rep) end;
-val live_of_bnf = #live o rep_bnf;
-val lives_of_bnf = #lives o rep_bnf;
-val dead_of_bnf = #dead o rep_bnf;
-val deads_of_bnf = #deads o rep_bnf;
-val axioms_of_bnf = #axioms o rep_bnf;
-val facts_of_bnf = #facts o rep_bnf;
-val nwits_of_bnf = #nwits o rep_bnf;
-val wits_of_bnf = #wits o rep_bnf;
-
-(*terms*)
-val map_of_bnf = #map o rep_bnf;
-val sets_of_bnf = #sets o rep_bnf;
-fun mk_map_of_bnf Ds Ts Us bnf =
- let val bnf_rep = rep_bnf bnf;
- in
- Term.subst_atomic_types
- ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#map bnf_rep)
- end;
-fun mk_sets_of_bnf Dss Tss bnf =
- let val bnf_rep = rep_bnf bnf;
- in
- map2 (fn (Ds, Ts) => Term.subst_atomic_types
- ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts))) (Dss ~~ Tss) (#sets bnf_rep)
- end;
-val bd_of_bnf = #bd o rep_bnf;
-fun mk_bd_of_bnf Ds Ts bnf =
- let val bnf_rep = rep_bnf bnf;
- in Term.subst_atomic_types ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#bd bnf_rep) end;
-fun mk_wits_of_bnf Dss Tss bnf =
- let
- val bnf_rep = rep_bnf bnf;
- val wits = map (fn x => (#I x, #wit x)) (#wits bnf_rep);
- in
- map2 (fn (Ds, Ts) => apsnd (Term.subst_atomic_types
- ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)))) (Dss ~~ Tss) wits
- end;
-val rel_of_bnf = #rel o rep_bnf;
-fun mk_rel_of_bnf Ds Ts Us bnf =
- let val bnf_rep = rep_bnf bnf;
- in
- Term.subst_atomic_types
- ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#rel bnf_rep)
- end;
-val srel_of_bnf = #srel o rep_bnf;
-fun mk_srel_of_bnf Ds Ts Us bnf =
- let val bnf_rep = rep_bnf bnf;
- in
- Term.subst_atomic_types
- ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#srel bnf_rep)
- end;
-
-(*thms*)
-val bd_card_order_of_bnf = #bd_card_order o #axioms o rep_bnf;
-val bd_cinfinite_of_bnf = #bd_cinfinite o #axioms o rep_bnf;
-val bd_Card_order_of_bnf = #bd_Card_order o #facts o rep_bnf;
-val bd_Cinfinite_of_bnf = #bd_Cinfinite o #facts o rep_bnf;
-val bd_Cnotzero_of_bnf = #bd_Cnotzero o #facts o rep_bnf;
-val collect_set_natural_of_bnf = Lazy.force o #collect_set_natural o #facts o rep_bnf;
-val in_bd_of_bnf = #in_bd o #axioms o rep_bnf;
-val in_cong_of_bnf = Lazy.force o #in_cong o #facts o rep_bnf;
-val in_mono_of_bnf = Lazy.force o #in_mono o #facts o rep_bnf;
-val in_srel_of_bnf = Lazy.force o #in_srel o #facts o rep_bnf;
-val map_def_of_bnf = #map_def o #defs o rep_bnf;
-val map_id_of_bnf = #map_id o #axioms o rep_bnf;
-val map_id'_of_bnf = Lazy.force o #map_id' o #facts o rep_bnf;
-val map_comp_of_bnf = #map_comp o #axioms o rep_bnf;
-val map_comp'_of_bnf = Lazy.force o #map_comp' o #facts o rep_bnf;
-val map_cong_of_bnf = #map_cong o #axioms o rep_bnf;
-val map_wppull_of_bnf = Lazy.force o #map_wppull o #facts o rep_bnf;
-val map_wpull_of_bnf = #map_wpull o #axioms o rep_bnf;
-val rel_def_of_bnf = #rel_def o #defs o rep_bnf;
-val set_bd_of_bnf = #set_bd o #axioms o rep_bnf;
-val set_defs_of_bnf = #set_defs o #defs o rep_bnf;
-val set_natural_of_bnf = #set_natural o #axioms o rep_bnf;
-val set_natural'_of_bnf = map Lazy.force o #set_natural' o #facts o rep_bnf;
-val srel_cong_of_bnf = Lazy.force o #srel_cong o #facts o rep_bnf;
-val srel_mono_of_bnf = Lazy.force o #srel_mono o #facts o rep_bnf;
-val srel_def_of_bnf = #srel_def o #defs o rep_bnf;
-val srel_Id_of_bnf = Lazy.force o #srel_Id o #facts o rep_bnf;
-val srel_Gr_of_bnf = Lazy.force o #srel_Gr o #facts o rep_bnf;
-val srel_converse_of_bnf = Lazy.force o #srel_converse o #facts o rep_bnf;
-val srel_O_of_bnf = Lazy.force o #srel_O o #facts o rep_bnf;
-val srel_O_Gr_of_bnf = #srel_O_Gr o #axioms o rep_bnf;
-val wit_thms_of_bnf = maps #prop o wits_of_bnf;
-val wit_thmss_of_bnf = map #prop o wits_of_bnf;
-
-fun mk_bnf name T live lives lives' dead deads map sets bd axioms defs facts wits rel srel =
- BNF {name = name, T = T,
- live = live, lives = lives, lives' = lives', dead = dead, deads = deads,
- map = map, sets = sets, bd = bd,
- axioms = axioms, defs = defs, facts = facts,
- nwits = length wits, wits = wits, rel = rel, srel = srel};
-
-fun morph_bnf phi (BNF {name = name, T = T, live = live, lives = lives, lives' = lives',
- dead = dead, deads = deads, map = map, sets = sets, bd = bd,
- axioms = axioms, defs = defs, facts = facts,
- nwits = nwits, wits = wits, rel = rel, srel = srel}) =
- BNF {name = Morphism.binding phi name, T = Morphism.typ phi T,
- live = live, lives = List.map (Morphism.typ phi) lives,
- lives' = List.map (Morphism.typ phi) lives',
- dead = dead, deads = List.map (Morphism.typ phi) deads,
- map = Morphism.term phi map, sets = List.map (Morphism.term phi) sets,
- bd = Morphism.term phi bd,
- axioms = morph_axioms phi axioms,
- defs = morph_defs phi defs,
- facts = morph_facts phi facts,
- nwits = nwits,
- wits = List.map (morph_witness phi) wits,
- rel = Morphism.term phi rel, srel = Morphism.term phi srel};
-
-fun eq_bnf (BNF {T = T1, live = live1, dead = dead1, ...},
- BNF {T = T2, live = live2, dead = dead2, ...}) =
- Type.could_unify (T1, T2) andalso live1 = live2 andalso dead1 = dead2;
-
-structure Data = Generic_Data
-(
- type T = BNF Symtab.table;
- val empty = Symtab.empty;
- val extend = I;
- val merge = Symtab.merge eq_bnf;
-);
-
-val bnf_of = Symtab.lookup o Data.get o Context.Proof;
-
-
-
-(* Utilities *)
-
-fun normalize_set insts instA set =
- let
- val (T, T') = dest_funT (fastype_of set);
- val A = fst (Term.dest_TVar (HOLogic.dest_setT T'));
- val params = Term.add_tvar_namesT T [];
- in Term.subst_TVars ((A :: params) ~~ (instA :: insts)) set end;
-
-fun normalize_rel ctxt instTs instA instB rel =
- let
- val thy = Proof_Context.theory_of ctxt;
- val tyenv =
- Sign.typ_match thy (fastype_of rel, Library.foldr (op -->) (instTs, mk_pred2T instA instB))
- Vartab.empty;
- in Envir.subst_term (tyenv, Vartab.empty) rel end
- handle Type.TYPE_MATCH => error "Bad predicator";
-
-fun normalize_srel ctxt instTs instA instB srel =
- let
- val thy = Proof_Context.theory_of ctxt;
- val tyenv =
- Sign.typ_match thy (fastype_of srel, Library.foldr (op -->) (instTs, mk_relT (instA, instB)))
- Vartab.empty;
- in Envir.subst_term (tyenv, Vartab.empty) srel end
- handle Type.TYPE_MATCH => error "Bad relator";
-
-fun normalize_wit insts CA As wit =
- let
- fun strip_param (Ts, T as Type (@{type_name fun}, [T1, T2])) =
- if Type.raw_instance (CA, T) then (Ts, T) else strip_param (T1 :: Ts, T2)
- | strip_param x = x;
- val (Ts, T) = strip_param ([], fastype_of wit);
- val subst = Term.add_tvar_namesT T [] ~~ insts;
- fun find y = find_index (fn x => x = y) As;
- in
- (map (find o Term.typ_subst_TVars subst) (rev Ts), Term.subst_TVars subst wit)
- end;
-
-fun minimize_wits wits =
- let
- fun minimize done [] = done
- | minimize done ((I, wit) :: todo) =
- if exists (fn (J, _) => subset (op =) (J, I)) (done @ todo)
- then minimize done todo
- else minimize ((I, wit) :: done) todo;
- in minimize [] wits end;
-
-
-
-(* Names *)
-
-val mapN = "map";
-val setN = "set";
-fun mk_setN i = setN ^ nonzero_string_of_int i;
-val bdN = "bd";
-val witN = "wit";
-fun mk_witN i = witN ^ nonzero_string_of_int i;
-val relN = "rel";
-val srelN = "srel";
-
-val bd_card_orderN = "bd_card_order";
-val bd_cinfiniteN = "bd_cinfinite";
-val bd_Card_orderN = "bd_Card_order";
-val bd_CinfiniteN = "bd_Cinfinite";
-val bd_CnotzeroN = "bd_Cnotzero";
-val collect_set_naturalN = "collect_set_natural";
-val in_bdN = "in_bd";
-val in_monoN = "in_mono";
-val in_srelN = "in_srel";
-val map_idN = "map_id";
-val map_id'N = "map_id'";
-val map_compN = "map_comp";
-val map_comp'N = "map_comp'";
-val map_congN = "map_cong";
-val map_wpullN = "map_wpull";
-val srel_IdN = "srel_Id";
-val srel_GrN = "srel_Gr";
-val srel_converseN = "srel_converse";
-val srel_monoN = "srel_mono"
-val srel_ON = "srel_comp";
-val srel_O_GrN = "srel_comp_Gr";
-val set_naturalN = "set_natural";
-val set_natural'N = "set_natural'";
-val set_bdN = "set_bd";
-
-datatype const_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline;
-
-datatype fact_policy =
- Derive_Few_Facts | Derive_All_Facts | Derive_All_Facts_Note_Most | Note_All_Facts_and_Axioms;
-
-val bnf_note_all = Attrib.setup_config_bool @{binding bnf_note_all} (K false);
-
-fun user_policy policy ctxt =
- if Config.get ctxt bnf_note_all then Note_All_Facts_and_Axioms else policy;
-
-val smart_max_inline_size = 25; (*FUDGE*)
-
-
-(* Define new BNFs *)
-
-fun prepare_def const_policy mk_fact_policy qualify prep_term Ds_opt
- (((((raw_b, raw_map), raw_sets), raw_bd_Abs), raw_wits), raw_rel_opt) no_defs_lthy =
- let
- val fact_policy = mk_fact_policy no_defs_lthy;
- val b = qualify raw_b;
- val live = length raw_sets;
- val nwits = length raw_wits;
-
- val map_rhs = prep_term no_defs_lthy raw_map;
- val set_rhss = map (prep_term no_defs_lthy) raw_sets;
- val (bd_rhsT, bd_rhs) = (case prep_term no_defs_lthy raw_bd_Abs of
- Abs (_, T, t) => (T, t)
- | _ => error "Bad bound constant");
- val wit_rhss = map (prep_term no_defs_lthy) raw_wits;
-
- fun err T =
- error ("Trying to register the type " ^ quote (Syntax.string_of_typ no_defs_lthy T) ^
- " as unnamed BNF");
-
- val (b, key) =
- if Binding.eq_name (b, Binding.empty) then
- (case bd_rhsT of
- Type (C, Ts) => if forall (is_some o try dest_TFree) Ts
- then (Binding.qualified_name C, C) else err bd_rhsT
- | T => err T)
- else (b, Local_Theory.full_name no_defs_lthy b);
-
- fun maybe_define user_specified (b, rhs) lthy =
- let
- val inline =
- (user_specified orelse fact_policy = Derive_Few_Facts) andalso
- (case const_policy of
- Dont_Inline => false
- | Hardly_Inline => Term.is_Free rhs orelse Term.is_Const rhs
- | Smart_Inline => Term.size_of_term rhs <= smart_max_inline_size
- | Do_Inline => true)
- in
- if inline then
- ((rhs, Drule.reflexive_thm), lthy)
- else
- let val b = b () in
- apfst (apsnd snd) (Local_Theory.define ((b, NoSyn), ((Thm.def_binding b, []), rhs))
- lthy)
- end
- end;
-
- fun maybe_restore lthy_old lthy =
- lthy |> not (pointer_eq (lthy_old, lthy)) ? Local_Theory.restore;
-
- val map_bind_def = (fn () => Binding.suffix_name ("_" ^ mapN) b, map_rhs);
- val set_binds_defs =
- let
- val bs = if live = 1 then [fn () => Binding.suffix_name ("_" ^ setN) b]
- else map (fn i => fn () => Binding.suffix_name ("_" ^ mk_setN i) b) (1 upto live)
- in map2 pair bs set_rhss end;
- val bd_bind_def = (fn () => Binding.suffix_name ("_" ^ bdN) b, bd_rhs);
- val wit_binds_defs =
- let
- val bs = if nwits = 1 then [fn () => Binding.suffix_name ("_" ^ witN) b]
- else map (fn i => fn () => Binding.suffix_name ("_" ^ mk_witN i) b) (1 upto nwits);
- in map2 pair bs wit_rhss end;
-
- val (((((bnf_map_term, raw_map_def),
- (bnf_set_terms, raw_set_defs)),
- (bnf_bd_term, raw_bd_def)),
- (bnf_wit_terms, raw_wit_defs)), (lthy, lthy_old)) =
- no_defs_lthy
- |> maybe_define true map_bind_def
- ||>> apfst split_list o fold_map (maybe_define true) set_binds_defs
- ||>> maybe_define true bd_bind_def
- ||>> apfst split_list o fold_map (maybe_define true) wit_binds_defs
- ||> `(maybe_restore no_defs_lthy);
-
- val phi = Proof_Context.export_morphism lthy_old lthy;
-
- val bnf_map_def = Morphism.thm phi raw_map_def;
- val bnf_set_defs = map (Morphism.thm phi) raw_set_defs;
- val bnf_bd_def = Morphism.thm phi raw_bd_def;
- val bnf_wit_defs = map (Morphism.thm phi) raw_wit_defs;
-
- val bnf_map = Morphism.term phi bnf_map_term;
-
- (*TODO: handle errors*)
- (*simple shape analysis of a map function*)
- val ((alphas, betas), (CA, _)) =
- fastype_of bnf_map
- |> strip_typeN live
- |>> map_split dest_funT
- ||> dest_funT
- handle TYPE _ => error "Bad map function";
-
- val CA_params = map TVar (Term.add_tvarsT CA []);
-
- val bnf_sets = map2 (normalize_set CA_params) alphas (map (Morphism.term phi) bnf_set_terms);
- val bdT = Morphism.typ phi bd_rhsT;
- val bnf_bd =
- Term.subst_TVars (Term.add_tvar_namesT bdT [] ~~ CA_params) (Morphism.term phi bnf_bd_term);
- val bnf_wits = map (normalize_wit CA_params CA alphas o Morphism.term phi) bnf_wit_terms;
-
- (*TODO: assert Ds = (TVars of bnf_map) \ (alphas @ betas) as sets*)
- val deads = (case Ds_opt of
- NONE => subtract (op =) (alphas @ betas) (map TVar (Term.add_tvars bnf_map []))
- | SOME Ds => map (Morphism.typ phi) Ds);
- val dead = length deads;
-
- (*TODO: further checks of type of bnf_map*)
- (*TODO: check types of bnf_sets*)
- (*TODO: check type of bnf_bd*)
- (*TODO: check type of bnf_rel*)
-
- val ((((((((((As', Bs'), Cs), Ds), B1Ts), B2Ts), domTs), ranTs), ranTs'), ranTs''),
- (Ts, T)) = lthy
- |> mk_TFrees live
- ||>> mk_TFrees live
- ||>> mk_TFrees live
- ||>> mk_TFrees dead
- ||>> mk_TFrees live
- ||>> mk_TFrees live
- ||>> mk_TFrees live
- ||>> mk_TFrees live
- ||>> mk_TFrees live
- ||>> mk_TFrees live
- ||> fst o mk_TFrees 1
- ||> the_single
- ||> `(replicate live);
-
- fun mk_bnf_map As' Bs' =
- Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As') @ (betas ~~ Bs')) bnf_map;
- fun mk_bnf_t As' = Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As'));
- fun mk_bnf_T As' = Term.typ_subst_atomic ((deads ~~ Ds) @ (alphas ~~ As'));
-
- val (setRTs, RTs) = map_split (`HOLogic.mk_setT o HOLogic.mk_prodT) (As' ~~ Bs');
- val setRTsAsCs = map (HOLogic.mk_setT o HOLogic.mk_prodT) (As' ~~ Cs);
- val setRTsBsCs = map (HOLogic.mk_setT o HOLogic.mk_prodT) (Bs' ~~ Cs);
- val setRT's = map (HOLogic.mk_setT o HOLogic.mk_prodT) (Bs' ~~ As');
- val self_setRTs = map (HOLogic.mk_setT o HOLogic.mk_prodT) (As' ~~ As');
- val QTs = map2 mk_pred2T As' Bs';
-
- val CA' = mk_bnf_T As' CA;
- val CB' = mk_bnf_T Bs' CA;
- val CC' = mk_bnf_T Cs CA;
- val CRs' = mk_bnf_T RTs CA;
- val CA'CB' = HOLogic.mk_prodT (CA', CB');
-
- val bnf_map_AsAs = mk_bnf_map As' As';
- val bnf_map_AsBs = mk_bnf_map As' Bs';
- val bnf_map_AsCs = mk_bnf_map As' Cs;
- val bnf_map_BsCs = mk_bnf_map Bs' Cs;
- val bnf_sets_As = map (mk_bnf_t As') bnf_sets;
- val bnf_sets_Bs = map (mk_bnf_t Bs') bnf_sets;
- val bnf_bd_As = mk_bnf_t As' bnf_bd;
- val bnf_wit_As = map (apsnd (mk_bnf_t As')) bnf_wits;
-
- val (((((((((((((((((((((((((fs, fs_copy), gs), hs), p), (x, x')), (y, y')), (z, z')), zs), As),
- As_copy), Xs), B1s), B2s), f1s), f2s), e1s), e2s), p1s), p2s), bs), (Rs, Rs')), Rs_copy), Ss),
- (Qs, Qs')), _) = lthy
- |> mk_Frees "f" (map2 (curry (op -->)) As' Bs')
- ||>> mk_Frees "f" (map2 (curry (op -->)) As' Bs')
- ||>> mk_Frees "g" (map2 (curry (op -->)) Bs' Cs)
- ||>> mk_Frees "h" (map2 (curry (op -->)) As' Ts)
- ||>> yield_singleton (mk_Frees "p") CA'CB'
- ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "x") CA'
- ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "y") CB'
- ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "z") CRs'
- ||>> mk_Frees "z" As'
- ||>> mk_Frees "A" (map HOLogic.mk_setT As')
- ||>> mk_Frees "A" (map HOLogic.mk_setT As')
- ||>> mk_Frees "A" (map HOLogic.mk_setT domTs)
- ||>> mk_Frees "B1" (map HOLogic.mk_setT B1Ts)
- ||>> mk_Frees "B2" (map HOLogic.mk_setT B2Ts)
- ||>> mk_Frees "f1" (map2 (curry (op -->)) B1Ts ranTs)
- ||>> mk_Frees "f2" (map2 (curry (op -->)) B2Ts ranTs)
- ||>> mk_Frees "e1" (map2 (curry (op -->)) B1Ts ranTs')
- ||>> mk_Frees "e2" (map2 (curry (op -->)) B2Ts ranTs'')
- ||>> mk_Frees "p1" (map2 (curry (op -->)) domTs B1Ts)
- ||>> mk_Frees "p2" (map2 (curry (op -->)) domTs B2Ts)
- ||>> mk_Frees "b" As'
- ||>> mk_Frees' "R" setRTs
- ||>> mk_Frees "R" setRTs
- ||>> mk_Frees "S" setRTsBsCs
- ||>> mk_Frees' "Q" QTs;
-
- (*Gr (in R1 .. Rn) (map fst .. fst)^-1 O Gr (in R1 .. Rn) (map snd .. snd)*)
- val O_Gr =
- let
- val map1 = Term.list_comb (mk_bnf_map RTs As', map fst_const RTs);
- val map2 = Term.list_comb (mk_bnf_map RTs Bs', map snd_const RTs);
- val bnf_in = mk_in (map Free Rs') (map (mk_bnf_t RTs) bnf_sets) CRs';
- in
- mk_rel_comp (mk_converse (mk_Gr bnf_in map1), mk_Gr bnf_in map2)
- end;
-
- fun mk_predicate_of_set x_name y_name t =
- let
- val (T, U) = HOLogic.dest_prodT (HOLogic.dest_setT (fastype_of t));
- val x = Free (x_name, T);
- val y = Free (y_name, U);
- in fold_rev Term.lambda [x, y] (HOLogic.mk_mem (HOLogic.mk_prod (x, y), t)) end;
-
- val rel_rhs = (case raw_rel_opt of
- NONE =>
- fold_rev absfree Qs' (mk_predicate_of_set (fst x') (fst y')
- (Term.list_comb (fold_rev Term.absfree Rs' O_Gr, map3 (fn Q => fn T => fn U =>
- HOLogic.Collect_const (HOLogic.mk_prodT (T, U)) $ HOLogic.mk_split Q) Qs As' Bs')))
- | SOME raw_rel => prep_term no_defs_lthy raw_rel);
-
- val rel_bind_def = (fn () => Binding.suffix_name ("_" ^ relN) b, rel_rhs);
-
- val ((bnf_rel_term, raw_rel_def), (lthy, lthy_old)) =
- lthy
- |> maybe_define (is_some raw_rel_opt) rel_bind_def
- ||> `(maybe_restore lthy);
-
- val phi = Proof_Context.export_morphism lthy_old lthy;
- val bnf_rel_def = Morphism.thm phi raw_rel_def;
- val bnf_rel = Morphism.term phi bnf_rel_term;
-
- fun mk_bnf_rel QTs CA' CB' = normalize_rel lthy QTs CA' CB' bnf_rel;
-
- val rel = mk_bnf_rel QTs CA' CB';
-
- val srel_rhs =
- fold_rev Term.absfree Rs' (HOLogic.Collect_const CA'CB' $
- Term.lambda p (Term.list_comb (rel, map (mk_predicate_of_set (fst x') (fst y')) Rs) $
- HOLogic.mk_fst p $ HOLogic.mk_snd p));
-
- val srel_bind_def = (fn () => Binding.suffix_name ("_" ^ srelN) b, srel_rhs);
-
- val ((bnf_srel_term, raw_srel_def), (lthy, lthy_old)) =
- lthy
- |> maybe_define false srel_bind_def
- ||> `(maybe_restore lthy);
-
- val phi = Proof_Context.export_morphism lthy_old lthy;
- val bnf_srel_def = Morphism.thm phi raw_srel_def;
- val bnf_srel = Morphism.term phi bnf_srel_term;
-
- fun mk_bnf_srel setRTs CA' CB' = normalize_srel lthy setRTs CA' CB' bnf_srel;
-
- val srel = mk_bnf_srel setRTs CA' CB';
-
- val _ = case no_reflexive (raw_map_def :: raw_set_defs @ [raw_bd_def] @
- raw_wit_defs @ [raw_rel_def, raw_srel_def]) of
- [] => ()
- | defs => Proof_Display.print_consts true lthy_old (K false)
- (map (dest_Free o fst o Logic.dest_equals o prop_of) defs);
-
- val map_id_goal =
- let
- val bnf_map_app_id = Term.list_comb (bnf_map_AsAs, map HOLogic.id_const As');
- in
- HOLogic.mk_Trueprop
- (HOLogic.mk_eq (bnf_map_app_id, HOLogic.id_const CA'))
- end;
-
- val map_comp_goal =
- let
- val bnf_map_app_comp = Term.list_comb (bnf_map_AsCs, map2 (curry HOLogic.mk_comp) gs fs);
- val comp_bnf_map_app = HOLogic.mk_comp
- (Term.list_comb (bnf_map_BsCs, gs),
- Term.list_comb (bnf_map_AsBs, fs));
- in
- fold_rev Logic.all (fs @ gs) (mk_Trueprop_eq (bnf_map_app_comp, comp_bnf_map_app))
- end;
-
- val map_cong_goal =
- let
- fun mk_prem z set f f_copy =
- Logic.all z (Logic.mk_implies
- (HOLogic.mk_Trueprop (HOLogic.mk_mem (z, set $ x)),
- mk_Trueprop_eq (f $ z, f_copy $ z)));
- val prems = map4 mk_prem zs bnf_sets_As fs fs_copy;
- val eq = HOLogic.mk_eq (Term.list_comb (bnf_map_AsBs, fs) $ x,
- Term.list_comb (bnf_map_AsBs, fs_copy) $ x);
- in
- fold_rev Logic.all (x :: fs @ fs_copy)
- (Logic.list_implies (prems, HOLogic.mk_Trueprop eq))
- end;
-
- val set_naturals_goal =
- let
- fun mk_goal setA setB f =
- let
- val set_comp_map =
- HOLogic.mk_comp (setB, Term.list_comb (bnf_map_AsBs, fs));
- val image_comp_set = HOLogic.mk_comp (mk_image f, setA);
- in
- fold_rev Logic.all fs (mk_Trueprop_eq (set_comp_map, image_comp_set))
- end;
- in
- map3 mk_goal bnf_sets_As bnf_sets_Bs fs
- end;
-
- val card_order_bd_goal = HOLogic.mk_Trueprop (mk_card_order bnf_bd_As);
-
- val cinfinite_bd_goal = HOLogic.mk_Trueprop (mk_cinfinite bnf_bd_As);
-
- val set_bds_goal =
- let
- fun mk_goal set =
- Logic.all x (HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of (set $ x)) bnf_bd_As));
- in
- map mk_goal bnf_sets_As
- end;
-
- val in_bd_goal =
- let
- val bd = mk_cexp
- (if live = 0 then ctwo
- else mk_csum (Library.foldr1 (uncurry mk_csum) (map mk_card_of As)) ctwo)
- bnf_bd_As;
- in
- fold_rev Logic.all As
- (HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of (mk_in As bnf_sets_As CA')) bd))
- end;
-
- val map_wpull_goal =
- let
- val prems = map HOLogic.mk_Trueprop
- (map8 mk_wpull Xs B1s B2s f1s f2s (replicate live NONE) p1s p2s);
- val CX = mk_bnf_T domTs CA;
- val CB1 = mk_bnf_T B1Ts CA;
- val CB2 = mk_bnf_T B2Ts CA;
- val bnf_sets_CX = map2 (normalize_set (map (mk_bnf_T domTs) CA_params)) domTs bnf_sets;
- val bnf_sets_CB1 = map2 (normalize_set (map (mk_bnf_T B1Ts) CA_params)) B1Ts bnf_sets;
- val bnf_sets_CB2 = map2 (normalize_set (map (mk_bnf_T B2Ts) CA_params)) B2Ts bnf_sets;
- val bnf_map_app_f1 = Term.list_comb (mk_bnf_map B1Ts ranTs, f1s);
- val bnf_map_app_f2 = Term.list_comb (mk_bnf_map B2Ts ranTs, f2s);
- val bnf_map_app_p1 = Term.list_comb (mk_bnf_map domTs B1Ts, p1s);
- val bnf_map_app_p2 = Term.list_comb (mk_bnf_map domTs B2Ts, p2s);
-
- val map_wpull = mk_wpull (mk_in Xs bnf_sets_CX CX)
- (mk_in B1s bnf_sets_CB1 CB1) (mk_in B2s bnf_sets_CB2 CB2)
- bnf_map_app_f1 bnf_map_app_f2 NONE bnf_map_app_p1 bnf_map_app_p2;
- in
- fold_rev Logic.all (Xs @ B1s @ B2s @ f1s @ f2s @ p1s @ p2s)
- (Logic.list_implies (prems, HOLogic.mk_Trueprop map_wpull))
- end;
-
- val srel_O_Gr_goal = fold_rev Logic.all Rs (mk_Trueprop_eq (Term.list_comb (srel, Rs), O_Gr));
-
- val goals = zip_axioms map_id_goal map_comp_goal map_cong_goal set_naturals_goal
- card_order_bd_goal cinfinite_bd_goal set_bds_goal in_bd_goal map_wpull_goal srel_O_Gr_goal;
-
- fun mk_wit_goals (I, wit) =
- let
- val xs = map (nth bs) I;
- fun wit_goal i =
- let
- val z = nth zs i;
- val set_wit = nth bnf_sets_As i $ Term.list_comb (wit, xs);
- val concl = HOLogic.mk_Trueprop
- (if member (op =) I i then HOLogic.mk_eq (z, nth bs i)
- else @{term False});
- in
- fold_rev Logic.all (z :: xs)
- (Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (z, set_wit)), concl))
- end;
- in
- map wit_goal (0 upto live - 1)
- end;
-
- val wit_goalss = map mk_wit_goals bnf_wit_As;
-
- fun after_qed thms lthy =
- let
- val (axioms, wit_thms) = apfst (mk_axioms live) (chop (length goals) thms);
-
- val bd_Card_order = #bd_card_order axioms RS @{thm conjunct2[OF card_order_on_Card_order]};
- val bd_Cinfinite = @{thm conjI} OF [#bd_cinfinite axioms, bd_Card_order];
- val bd_Cnotzero = bd_Cinfinite RS @{thm Cinfinite_Cnotzero};
-
- fun mk_lazy f = if fact_policy <> Derive_Few_Facts then Lazy.value (f ()) else Lazy.lazy f;
-
- fun mk_collect_set_natural () =
- let
- val defT = mk_bnf_T Ts CA --> HOLogic.mk_setT T;
- val collect_map = HOLogic.mk_comp
- (mk_collect (map (mk_bnf_t Ts) bnf_sets) defT,
- Term.list_comb (mk_bnf_map As' Ts, hs));
- val image_collect = mk_collect
- (map2 (fn h => fn set => HOLogic.mk_comp (mk_image h, set)) hs bnf_sets_As)
- defT;
- (*collect {set1 ... setm} o map f1 ... fm = collect {f1` o set1 ... fm` o setm}*)
- val goal = fold_rev Logic.all hs (mk_Trueprop_eq (collect_map, image_collect));
- in
- Skip_Proof.prove lthy [] [] goal
- (fn {context = ctxt, ...} => mk_collect_set_natural_tac ctxt (#set_natural axioms))
- |> Thm.close_derivation
- end;
-
- val collect_set_natural = mk_lazy mk_collect_set_natural;
-
- fun mk_in_mono () =
- let
- val prems_mono = map2 (HOLogic.mk_Trueprop oo mk_subset) As As_copy;
- val in_mono_goal =
- fold_rev Logic.all (As @ As_copy)
- (Logic.list_implies (prems_mono, HOLogic.mk_Trueprop
- (mk_subset (mk_in As bnf_sets_As CA') (mk_in As_copy bnf_sets_As CA'))));
- in
- Skip_Proof.prove lthy [] [] in_mono_goal (K (mk_in_mono_tac live))
- |> Thm.close_derivation
- end;
-
- val in_mono = mk_lazy mk_in_mono;
-
- fun mk_in_cong () =
- let
- val prems_cong = map2 (HOLogic.mk_Trueprop oo curry HOLogic.mk_eq) As As_copy;
- val in_cong_goal =
- fold_rev Logic.all (As @ As_copy)
- (Logic.list_implies (prems_cong, HOLogic.mk_Trueprop
- (HOLogic.mk_eq (mk_in As bnf_sets_As CA', mk_in As_copy bnf_sets_As CA'))));
- in
- Skip_Proof.prove lthy [] [] in_cong_goal (K ((TRY o hyp_subst_tac THEN' rtac refl) 1))
- |> Thm.close_derivation
- end;
-
- val in_cong = mk_lazy mk_in_cong;
-
- val map_id' = mk_lazy (fn () => mk_id' (#map_id axioms));
- val map_comp' = mk_lazy (fn () => mk_comp' (#map_comp axioms));
-
- val set_natural' =
- map (fn thm => mk_lazy (fn () => mk_set_natural' thm)) (#set_natural axioms);
-
- fun mk_map_wppull () =
- let
- val prems = if live = 0 then [] else
- [HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
- (map8 mk_wpull Xs B1s B2s f1s f2s (map SOME (e1s ~~ e2s)) p1s p2s))];
- val CX = mk_bnf_T domTs CA;
- val CB1 = mk_bnf_T B1Ts CA;
- val CB2 = mk_bnf_T B2Ts CA;
- val bnf_sets_CX =
- map2 (normalize_set (map (mk_bnf_T domTs) CA_params)) domTs bnf_sets;
- val bnf_sets_CB1 =
- map2 (normalize_set (map (mk_bnf_T B1Ts) CA_params)) B1Ts bnf_sets;
- val bnf_sets_CB2 =
- map2 (normalize_set (map (mk_bnf_T B2Ts) CA_params)) B2Ts bnf_sets;
- val bnf_map_app_f1 = Term.list_comb (mk_bnf_map B1Ts ranTs, f1s);
- val bnf_map_app_f2 = Term.list_comb (mk_bnf_map B2Ts ranTs, f2s);
- val bnf_map_app_e1 = Term.list_comb (mk_bnf_map B1Ts ranTs', e1s);
- val bnf_map_app_e2 = Term.list_comb (mk_bnf_map B2Ts ranTs'', e2s);
- val bnf_map_app_p1 = Term.list_comb (mk_bnf_map domTs B1Ts, p1s);
- val bnf_map_app_p2 = Term.list_comb (mk_bnf_map domTs B2Ts, p2s);
-
- val concl = mk_wpull (mk_in Xs bnf_sets_CX CX)
- (mk_in B1s bnf_sets_CB1 CB1) (mk_in B2s bnf_sets_CB2 CB2)
- bnf_map_app_f1 bnf_map_app_f2 (SOME (bnf_map_app_e1, bnf_map_app_e2))
- bnf_map_app_p1 bnf_map_app_p2;
-
- val goal =
- fold_rev Logic.all (Xs @ B1s @ B2s @ f1s @ f2s @ e1s @ e2s @ p1s @ p2s)
- (Logic.list_implies (prems, HOLogic.mk_Trueprop concl))
- in
- Skip_Proof.prove lthy [] [] goal
- (fn _ => mk_map_wppull_tac (#map_id axioms) (#map_cong axioms)
- (#map_wpull axioms) (Lazy.force map_comp') (map Lazy.force set_natural'))
- |> Thm.close_derivation
- end;
-
- val srel_O_Grs = no_refl [#srel_O_Gr axioms];
-
- val map_wppull = mk_lazy mk_map_wppull;
-
- fun mk_srel_Gr () =
- let
- val lhs = Term.list_comb (srel, map2 mk_Gr As fs);
- val rhs = mk_Gr (mk_in As bnf_sets_As CA') (Term.list_comb (bnf_map_AsBs, fs));
- val goal = fold_rev Logic.all (As @ fs) (mk_Trueprop_eq (lhs, rhs));
- in
- Skip_Proof.prove lthy [] [] goal
- (mk_srel_Gr_tac srel_O_Grs (#map_id axioms) (#map_cong axioms) (Lazy.force map_id')
- (Lazy.force map_comp') (map Lazy.force set_natural'))
- |> Thm.close_derivation
- end;
-
- val srel_Gr = mk_lazy mk_srel_Gr;
-
- fun mk_srel_prems f = map2 (HOLogic.mk_Trueprop oo f) Rs Rs_copy
- fun mk_srel_concl f = HOLogic.mk_Trueprop
- (f (Term.list_comb (srel, Rs), Term.list_comb (srel, Rs_copy)));
-
- fun mk_srel_mono () =
- let
- val mono_prems = mk_srel_prems mk_subset;
- val mono_concl = mk_srel_concl (uncurry mk_subset);
- in
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (mono_prems, mono_concl)))
- (mk_srel_mono_tac srel_O_Grs (Lazy.force in_mono))
- |> Thm.close_derivation
- end;
-
- fun mk_srel_cong () =
- let
- val cong_prems = mk_srel_prems (curry HOLogic.mk_eq);
- val cong_concl = mk_srel_concl HOLogic.mk_eq;
- in
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (cong_prems, cong_concl)))
- (fn _ => (TRY o hyp_subst_tac THEN' rtac refl) 1)
- |> Thm.close_derivation
- end;
-
- val srel_mono = mk_lazy mk_srel_mono;
- val srel_cong = mk_lazy mk_srel_cong;
-
- fun mk_srel_Id () =
- let val relAsAs = mk_bnf_srel self_setRTs CA' CA' in
- Skip_Proof.prove lthy [] []
- (HOLogic.mk_Trueprop
- (HOLogic.mk_eq (Term.list_comb (relAsAs, map Id_const As'), Id_const CA')))
- (mk_srel_Id_tac live (Lazy.force srel_Gr) (#map_id axioms))
- |> Thm.close_derivation
- end;
-
- val srel_Id = mk_lazy mk_srel_Id;
-
- fun mk_srel_converse () =
- let
- val relBsAs = mk_bnf_srel setRT's CB' CA';
- val lhs = Term.list_comb (relBsAs, map mk_converse Rs);
- val rhs = mk_converse (Term.list_comb (srel, Rs));
- val le_goal = fold_rev Logic.all Rs (HOLogic.mk_Trueprop (mk_subset lhs rhs));
- val le_thm = Skip_Proof.prove lthy [] [] le_goal
- (mk_srel_converse_le_tac srel_O_Grs (Lazy.force srel_Id) (#map_cong axioms)
- (Lazy.force map_comp') (map Lazy.force set_natural'))
- |> Thm.close_derivation
- val goal = fold_rev Logic.all Rs (mk_Trueprop_eq (lhs, rhs));
- in
- Skip_Proof.prove lthy [] [] goal (fn _ => mk_srel_converse_tac le_thm)
- |> Thm.close_derivation
- end;
-
- val srel_converse = mk_lazy mk_srel_converse;
-
- fun mk_srel_O () =
- let
- val relAsCs = mk_bnf_srel setRTsAsCs CA' CC';
- val relBsCs = mk_bnf_srel setRTsBsCs CB' CC';
- val lhs = Term.list_comb (relAsCs, map2 (curry mk_rel_comp) Rs Ss);
- val rhs = mk_rel_comp (Term.list_comb (srel, Rs), Term.list_comb (relBsCs, Ss));
- val goal = fold_rev Logic.all (Rs @ Ss) (mk_Trueprop_eq (lhs, rhs));
- in
- Skip_Proof.prove lthy [] [] goal
- (mk_srel_O_tac srel_O_Grs (Lazy.force srel_Id) (#map_cong axioms)
- (Lazy.force map_wppull) (Lazy.force map_comp') (map Lazy.force set_natural'))
- |> Thm.close_derivation
- end;
-
- val srel_O = mk_lazy mk_srel_O;
-
- fun mk_in_srel () =
- let
- val bnf_in = mk_in Rs (map (mk_bnf_t RTs) bnf_sets) CRs';
- val map1 = Term.list_comb (mk_bnf_map RTs As', map fst_const RTs);
- val map2 = Term.list_comb (mk_bnf_map RTs Bs', map snd_const RTs);
- val map_fst_eq = HOLogic.mk_eq (map1 $ z, x);
- val map_snd_eq = HOLogic.mk_eq (map2 $ z, y);
- val lhs = HOLogic.mk_mem (HOLogic.mk_prod (x, y), Term.list_comb (srel, Rs));
- val rhs =
- HOLogic.mk_exists (fst z', snd z', HOLogic.mk_conj (HOLogic.mk_mem (z, bnf_in),
- HOLogic.mk_conj (map_fst_eq, map_snd_eq)));
- val goal =
- fold_rev Logic.all (x :: y :: Rs) (mk_Trueprop_eq (lhs, rhs));
- in
- Skip_Proof.prove lthy [] [] goal (mk_in_srel_tac srel_O_Grs (length bnf_sets))
- |> Thm.close_derivation
- end;
-
- val in_srel = mk_lazy mk_in_srel;
-
- val defs = mk_defs bnf_map_def bnf_set_defs bnf_rel_def bnf_srel_def;
-
- val facts = mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_natural in_cong
- in_mono in_srel map_comp' map_id' map_wppull set_natural' srel_cong srel_mono srel_Id
- srel_Gr srel_converse srel_O;
-
- val wits = map2 mk_witness bnf_wits wit_thms;
-
- val bnf_rel =
- Term.subst_atomic_types ((Ds ~~ deads) @ (As' ~~ alphas) @ (Bs' ~~ betas)) rel;
- val bnf_srel =
- Term.subst_atomic_types ((Ds ~~ deads) @ (As' ~~ alphas) @ (Bs' ~~ betas)) srel;
-
- val bnf = mk_bnf b CA live alphas betas dead deads bnf_map bnf_sets bnf_bd axioms defs facts
- wits bnf_rel bnf_srel;
- in
- (bnf, lthy
- |> (if fact_policy = Note_All_Facts_and_Axioms then
- let
- val witNs = if length wits = 1 then [witN] else map mk_witN (1 upto length wits);
- val notes =
- [(bd_card_orderN, [#bd_card_order axioms]),
- (bd_cinfiniteN, [#bd_cinfinite axioms]),
- (bd_Card_orderN, [#bd_Card_order facts]),
- (bd_CinfiniteN, [#bd_Cinfinite facts]),
- (bd_CnotzeroN, [#bd_Cnotzero facts]),
- (collect_set_naturalN, [Lazy.force (#collect_set_natural facts)]),
- (in_bdN, [#in_bd axioms]),
- (in_monoN, [Lazy.force (#in_mono facts)]),
- (in_srelN, [Lazy.force (#in_srel facts)]),
- (map_compN, [#map_comp axioms]),
- (map_idN, [#map_id axioms]),
- (map_wpullN, [#map_wpull axioms]),
- (set_naturalN, #set_natural axioms),
- (set_bdN, #set_bd axioms)] @
- map2 pair witNs wit_thms
- |> map (fn (thmN, thms) =>
- ((qualify (Binding.qualify true (Binding.name_of b) (Binding.name thmN)), []),
- [(thms, [])]));
- in
- Local_Theory.notes notes #> snd
- end
- else
- I)
- |> (if fact_policy = Note_All_Facts_and_Axioms orelse
- fact_policy = Derive_All_Facts_Note_Most then
- let
- val notes =
- [(map_comp'N, [Lazy.force (#map_comp' facts)]),
- (map_congN, [#map_cong axioms]),
- (map_id'N, [Lazy.force (#map_id' facts)]),
- (set_natural'N, map Lazy.force (#set_natural' facts)),
- (srel_O_GrN, srel_O_Grs),
- (srel_IdN, [Lazy.force (#srel_Id facts)]),
- (srel_GrN, [Lazy.force (#srel_Gr facts)]),
- (srel_converseN, [Lazy.force (#srel_converse facts)]),
- (srel_monoN, [Lazy.force (#srel_mono facts)]),
- (srel_ON, [Lazy.force (#srel_O facts)])]
- |> filter_out (null o #2)
- |> map (fn (thmN, thms) =>
- ((qualify (Binding.qualify true (Binding.name_of b) (Binding.name thmN)), []),
- [(thms, [])]));
- in
- Local_Theory.notes notes #> snd
- end
- else
- I))
- end;
-
- val one_step_defs =
- no_reflexive (bnf_map_def :: bnf_bd_def :: bnf_set_defs @ bnf_wit_defs @ [bnf_rel_def,
- bnf_srel_def]);
- in
- (key, goals, wit_goalss, after_qed, lthy, one_step_defs)
- end;
-
-fun register_bnf key (bnf, lthy) =
- (bnf, Local_Theory.declaration {syntax = false, pervasive = true}
- (fn phi => Data.map (Symtab.update_new (key, morph_bnf phi bnf))) lthy);
-
-(* TODO: Once the invariant "nwits > 0" holds, remove "mk_conjunction_balanced'" and "rtac TrueI"
- below *)
-fun mk_conjunction_balanced' [] = @{prop True}
- | mk_conjunction_balanced' ts = Logic.mk_conjunction_balanced ts;
-
-fun bnf_def const_policy fact_policy qualify tacs wit_tac Ds =
- (fn (_, goals, wit_goalss, after_qed, lthy, one_step_defs) =>
- let
- val wits_tac =
- K (TRYALL Goal.conjunction_tac) THEN' K (TRYALL (rtac TrueI)) THEN'
- mk_unfold_thms_then_tac lthy one_step_defs wit_tac;
- val wit_goals = map mk_conjunction_balanced' wit_goalss;
- val wit_thms =
- Skip_Proof.prove lthy [] [] (mk_conjunction_balanced' wit_goals) wits_tac
- |> Conjunction.elim_balanced (length wit_goals)
- |> map2 (Conjunction.elim_balanced o length) wit_goalss
- |> map (map (Thm.close_derivation o Thm.forall_elim_vars 0));
- in
- map2 (Thm.close_derivation oo Skip_Proof.prove lthy [] [])
- goals (map (mk_unfold_thms_then_tac lthy one_step_defs) tacs)
- |> (fn thms => after_qed (map single thms @ wit_thms) lthy)
- end) oo prepare_def const_policy fact_policy qualify (K I) Ds;
-
-val bnf_def_cmd = (fn (key, goals, wit_goals, after_qed, lthy, defs) =>
- Proof.unfolding ([[(defs, [])]])
- (Proof.theorem NONE (snd o register_bnf key oo after_qed)
- (map (single o rpair []) goals @ map (map (rpair [])) wit_goals) lthy)) oo
- prepare_def Do_Inline (user_policy Derive_All_Facts_Note_Most) I Syntax.read_term NONE;
-
-fun print_bnfs ctxt =
- let
- fun pretty_set sets i = Pretty.block
- [Pretty.str (mk_setN (i + 1) ^ ":"), Pretty.brk 1,
- Pretty.quote (Syntax.pretty_term ctxt (nth sets i))];
-
- fun pretty_bnf (key, BNF {T = T, map = map, sets = sets, bd = bd,
- live = live, lives = lives, dead = dead, deads = deads, ...}) =
- Pretty.big_list
- (Pretty.string_of (Pretty.block [Pretty.str key, Pretty.str ":", Pretty.brk 1,
- Pretty.quote (Syntax.pretty_typ ctxt T)]))
- ([Pretty.block [Pretty.str "live:", Pretty.brk 1, Pretty.str (string_of_int live),
- Pretty.brk 3, Pretty.list "[" "]" (List.map (Syntax.pretty_typ ctxt) lives)],
- Pretty.block [Pretty.str "dead:", Pretty.brk 1, Pretty.str (string_of_int dead),
- Pretty.brk 3, Pretty.list "[" "]" (List.map (Syntax.pretty_typ ctxt) deads)],
- Pretty.block [Pretty.str (mapN ^ ":"), Pretty.brk 1,
- Pretty.quote (Syntax.pretty_term ctxt map)]] @
- List.map (pretty_set sets) (0 upto length sets - 1) @
- [Pretty.block [Pretty.str (bdN ^ ":"), Pretty.brk 1,
- Pretty.quote (Syntax.pretty_term ctxt bd)]]);
- in
- Pretty.big_list "BNFs:" (map pretty_bnf (Symtab.dest (Data.get (Context.Proof ctxt))))
- |> Pretty.writeln
- end;
-
-val _ =
- Outer_Syntax.improper_command @{command_spec "print_bnfs"} "print all BNFs"
- (Scan.succeed (Toplevel.keep (print_bnfs o Toplevel.context_of)));
-
-val _ =
- Outer_Syntax.local_theory_to_proof @{command_spec "bnf_def"} "define a BNF for an existing type"
- ((parse_opt_binding_colon -- Parse.term --
- (@{keyword "["} |-- Parse.list Parse.term --| @{keyword "]"}) -- Parse.term --
- (@{keyword "["} |-- Parse.list Parse.term --| @{keyword "]"}) -- Scan.option Parse.term)
- >> bnf_def_cmd);
-
-end;
--- a/src/HOL/Codatatype/Tools/bnf_def_tactics.ML Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,209 +0,0 @@
-(* Title: HOL/BNF/Tools/bnf_def_tactics.ML
- Author: Dmitriy Traytel, TU Muenchen
- Author: Jasmin Blanchette, TU Muenchen
- Copyright 2012
-
-Tactics for definition of bounded natural functors.
-*)
-
-signature BNF_DEF_TACTICS =
-sig
- val mk_collect_set_natural_tac: Proof.context -> thm list -> tactic
- val mk_id': thm -> thm
- val mk_comp': thm -> thm
- val mk_in_mono_tac: int -> tactic
- val mk_map_wppull_tac: thm -> thm -> thm -> thm -> thm list -> tactic
- val mk_set_natural': thm -> thm
-
- val mk_srel_Gr_tac: thm list -> thm -> thm -> thm -> thm -> thm list ->
- {prems: thm list, context: Proof.context} -> tactic
- val mk_srel_Id_tac: int -> thm -> thm -> {prems: 'a, context: Proof.context} -> tactic
- val mk_srel_O_tac: thm list -> thm -> thm -> thm -> thm -> thm list ->
- {prems: thm list, context: Proof.context} -> tactic
- val mk_in_srel_tac: thm list -> int -> {prems: 'b, context: Proof.context} -> tactic
- val mk_srel_converse_tac: thm -> tactic
- val mk_srel_converse_le_tac: thm list -> thm -> thm -> thm -> thm list ->
- {prems: thm list, context: Proof.context} -> tactic
- val mk_srel_mono_tac: thm list -> thm -> {prems: 'a, context: Proof.context} -> tactic
-end;
-
-structure BNF_Def_Tactics : BNF_DEF_TACTICS =
-struct
-
-open BNF_Util
-open BNF_Tactics
-
-fun mk_id' id = mk_trans (fun_cong OF [id]) @{thm id_apply};
-fun mk_comp' comp = @{thm o_eq_dest_lhs} OF [mk_sym comp];
-fun mk_set_natural' set_natural = set_natural RS @{thm pointfreeE};
-fun mk_in_mono_tac n = if n = 0 then rtac subset_UNIV 1
- else (rtac subsetI THEN'
- rtac CollectI) 1 THEN
- REPEAT_DETERM (eresolve_tac [CollectE, conjE] 1) THEN
- REPEAT_DETERM_N (n - 1)
- ((rtac conjI THEN' etac subset_trans THEN' atac) 1) THEN
- (etac subset_trans THEN' atac) 1;
-
-fun mk_collect_set_natural_tac ctxt set_naturals =
- substs_tac ctxt (@{thms collect_o image_insert image_empty} @ set_naturals) 1 THEN rtac refl 1;
-
-fun mk_map_wppull_tac map_id map_cong map_wpull map_comp set_naturals =
- if null set_naturals then
- EVERY' [rtac @{thm wppull_id}, rtac map_wpull, rtac map_id, rtac map_id] 1
- else EVERY' [REPEAT_DETERM o etac conjE, REPEAT_DETERM o dtac @{thm wppull_thePull},
- REPEAT_DETERM o etac exE, rtac @{thm wpull_wppull}, rtac map_wpull,
- REPEAT_DETERM o rtac @{thm wpull_thePull}, rtac ballI,
- REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac conjI, rtac CollectI,
- CONJ_WRAP' (fn thm => EVERY' [rtac (thm RS @{thm ord_eq_le_trans}),
- rtac @{thm image_subsetI}, rtac conjunct1, etac bspec, etac set_mp, atac])
- set_naturals,
- CONJ_WRAP' (fn thm => EVERY' [rtac (map_comp RS trans), rtac (map_comp RS trans),
- rtac (map_comp RS trans RS sym), rtac map_cong,
- REPEAT_DETERM_N (length set_naturals) o EVERY' [rtac (o_apply RS trans),
- rtac (o_apply RS trans RS sym), rtac (o_apply RS trans), rtac thm,
- rtac conjunct2, etac bspec, etac set_mp, atac]]) [conjunct1, conjunct2]] 1;
-
-fun mk_srel_Gr_tac srel_O_Grs map_id map_cong map_id' map_comp set_naturals
- {context = ctxt, prems = _} =
- let
- val n = length set_naturals;
- in
- if null set_naturals then
- unfold_thms_tac ctxt srel_O_Grs THEN EVERY' [rtac @{thm Gr_UNIV_id}, rtac map_id] 1
- else unfold_thms_tac ctxt (@{thm Gr_def} :: srel_O_Grs) THEN
- EVERY' [rtac equalityI, rtac subsetI,
- REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE, @{thm relcompE}, @{thm converseE}],
- REPEAT_DETERM o dtac Pair_eqD,
- REPEAT_DETERM o etac conjE, hyp_subst_tac,
- rtac CollectI, rtac exI, rtac conjI, rtac Pair_eqI, rtac conjI, rtac refl,
- rtac sym, rtac trans, rtac map_comp, rtac map_cong,
- REPEAT_DETERM_N n o EVERY' [dtac @{thm set_rev_mp}, atac,
- REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac,
- rtac (o_apply RS trans), rtac (@{thm fst_conv} RS arg_cong RS trans),
- rtac (@{thm snd_conv} RS sym)],
- rtac CollectI,
- CONJ_WRAP' (fn thm => EVERY' [rtac (thm RS @{thm ord_eq_le_trans}),
- rtac @{thm image_subsetI}, dtac @{thm set_rev_mp}, atac,
- REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac,
- stac @{thm fst_conv}, atac]) set_naturals,
- rtac @{thm subrelI}, etac CollectE, REPEAT_DETERM o eresolve_tac [exE, conjE],
- REPEAT_DETERM o dtac Pair_eqD,
- REPEAT_DETERM o etac conjE, hyp_subst_tac,
- rtac @{thm relcompI}, rtac @{thm converseI},
- EVERY' (map2 (fn convol => fn map_id =>
- EVERY' [rtac CollectI, rtac exI, rtac conjI,
- rtac Pair_eqI, rtac conjI, rtac refl, rtac sym,
- rtac (box_equals OF [map_cong, map_comp RS sym, map_id]),
- REPEAT_DETERM_N n o rtac (convol RS fun_cong),
- REPEAT_DETERM o eresolve_tac [CollectE, conjE],
- rtac CollectI,
- CONJ_WRAP' (fn thm =>
- EVERY' [rtac @{thm ord_eq_le_trans}, rtac thm, rtac @{thm image_subsetI},
- rtac @{thm convol_memI[of id _ "%x. x", OF id_apply refl]}, etac set_mp, atac])
- set_naturals])
- @{thms fst_convol snd_convol} [map_id', refl])] 1
- end;
-
-fun mk_srel_Id_tac n srel_Gr map_id {context = ctxt, prems = _} =
- unfold_thms_tac ctxt [srel_Gr, @{thm Id_alt}] THEN
- subst_tac ctxt [map_id] 1 THEN
- (if n = 0 then rtac refl 1
- else EVERY' [rtac @{thm arg_cong2[of _ _ _ _ Gr]},
- rtac equalityI, rtac subset_UNIV, rtac subsetI, rtac CollectI,
- CONJ_WRAP' (K (rtac subset_UNIV)) (1 upto n), rtac refl] 1);
-
-fun mk_srel_mono_tac srel_O_Grs in_mono {context = ctxt, prems = _} =
- unfold_thms_tac ctxt srel_O_Grs THEN
- EVERY' [rtac @{thm relcomp_mono}, rtac @{thm iffD2[OF converse_mono]},
- rtac @{thm Gr_mono}, rtac in_mono, REPEAT_DETERM o atac,
- rtac @{thm Gr_mono}, rtac in_mono, REPEAT_DETERM o atac] 1;
-
-fun mk_srel_converse_le_tac srel_O_Grs srel_Id map_cong map_comp set_naturals
- {context = ctxt, prems = _} =
- let
- val n = length set_naturals;
- in
- if null set_naturals then
- unfold_thms_tac ctxt [srel_Id] THEN rtac equalityD2 1 THEN rtac @{thm converse_Id} 1
- else unfold_thms_tac ctxt (@{thm Gr_def} :: srel_O_Grs) THEN
- EVERY' [rtac @{thm subrelI},
- REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE, @{thm relcompE}, @{thm converseE}],
- REPEAT_DETERM o dtac Pair_eqD,
- REPEAT_DETERM o etac conjE, hyp_subst_tac, rtac @{thm converseI},
- rtac @{thm relcompI}, rtac @{thm converseI},
- EVERY' (map (fn thm => EVERY' [rtac CollectI, rtac exI,
- rtac conjI, rtac Pair_eqI, rtac conjI, rtac refl, rtac trans,
- rtac map_cong, REPEAT_DETERM_N n o rtac thm,
- rtac (map_comp RS sym), rtac CollectI,
- CONJ_WRAP' (fn thm => EVERY' [rtac (thm RS @{thm ord_eq_le_trans}),
- etac @{thm flip_rel}]) set_naturals]) [@{thm snd_fst_flip}, @{thm fst_snd_flip}])] 1
- end;
-
-fun mk_srel_converse_tac le_converse =
- EVERY' [rtac equalityI, rtac le_converse, rtac @{thm xt1(6)}, rtac @{thm converse_shift},
- rtac le_converse, REPEAT_DETERM o stac @{thm converse_converse}, rtac subset_refl] 1;
-
-fun mk_srel_O_tac srel_O_Grs srel_Id map_cong map_wppull map_comp set_naturals
- {context = ctxt, prems = _} =
- let
- val n = length set_naturals;
- fun in_tac nthO_in = rtac CollectI THEN'
- CONJ_WRAP' (fn thm => EVERY' [rtac (thm RS @{thm ord_eq_le_trans}),
- rtac @{thm image_subsetI}, rtac nthO_in, etac set_mp, atac]) set_naturals;
- in
- if null set_naturals then unfold_thms_tac ctxt [srel_Id] THEN rtac (@{thm Id_O_R} RS sym) 1
- else unfold_thms_tac ctxt (@{thm Gr_def} :: srel_O_Grs) THEN
- EVERY' [rtac equalityI, rtac @{thm subrelI},
- REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE, @{thm relcompE}, @{thm converseE}],
- REPEAT_DETERM o dtac Pair_eqD,
- REPEAT_DETERM o etac conjE, hyp_subst_tac,
- rtac @{thm relcompI}, rtac @{thm relcompI}, rtac @{thm converseI},
- rtac CollectI, rtac exI, rtac conjI, rtac Pair_eqI, rtac conjI, rtac refl,
- rtac sym, rtac trans, rtac map_comp, rtac sym, rtac map_cong,
- REPEAT_DETERM_N n o rtac @{thm fst_fstO},
- in_tac @{thm fstO_in},
- rtac CollectI, rtac exI, rtac conjI, rtac Pair_eqI, rtac conjI, rtac refl,
- rtac sym, rtac trans, rtac map_comp, rtac map_cong,
- REPEAT_DETERM_N n o EVERY' [rtac trans, rtac o_apply, rtac ballE, rtac subst,
- rtac @{thm csquare_def}, rtac @{thm csquare_fstO_sndO}, atac, etac notE,
- etac set_mp, atac],
- in_tac @{thm fstO_in},
- rtac @{thm relcompI}, rtac @{thm converseI},
- rtac CollectI, rtac exI, rtac conjI, rtac Pair_eqI, rtac conjI, rtac refl,
- rtac sym, rtac trans, rtac map_comp, rtac map_cong,
- REPEAT_DETERM_N n o rtac o_apply,
- in_tac @{thm sndO_in},
- rtac CollectI, rtac exI, rtac conjI, rtac Pair_eqI, rtac conjI, rtac refl,
- rtac sym, rtac trans, rtac map_comp, rtac sym, rtac map_cong,
- REPEAT_DETERM_N n o rtac @{thm snd_sndO},
- in_tac @{thm sndO_in},
- rtac @{thm subrelI},
- REPEAT_DETERM o eresolve_tac [CollectE, @{thm relcompE}, @{thm converseE}],
- REPEAT_DETERM o eresolve_tac [exE, conjE],
- REPEAT_DETERM o dtac Pair_eqD,
- REPEAT_DETERM o etac conjE, hyp_subst_tac,
- rtac allE, rtac subst, rtac @{thm wppull_def}, rtac map_wppull,
- CONJ_WRAP' (K (rtac @{thm wppull_fstO_sndO})) set_naturals,
- etac allE, etac impE, etac conjI, etac conjI, atac,
- REPEAT_DETERM o eresolve_tac [bexE, conjE],
- rtac @{thm relcompI}, rtac @{thm converseI},
- EVERY' (map (fn thm => EVERY' [rtac CollectI, rtac exI,
- rtac conjI, rtac Pair_eqI, rtac conjI, rtac refl, rtac sym, rtac trans,
- rtac trans, rtac map_cong, REPEAT_DETERM_N n o rtac thm,
- rtac (map_comp RS sym), atac, atac]) [@{thm fst_fstO}, @{thm snd_sndO}])] 1
- end;
-
-fun mk_in_srel_tac srel_O_Grs m {context = ctxt, prems = _} =
- let
- val ls' = replicate (Int.max (1, m)) ();
- in
- unfold_thms_tac ctxt (srel_O_Grs @
- @{thms Gr_def converse_unfold relcomp_unfold mem_Collect_eq prod.cases Pair_eq}) THEN
- EVERY' [rtac iffI, REPEAT_DETERM o eresolve_tac [exE, conjE], hyp_subst_tac, rtac exI,
- rtac conjI, CONJ_WRAP' (K atac) ls', rtac conjI, rtac refl, rtac refl,
- REPEAT_DETERM o eresolve_tac [exE, conjE], rtac exI, rtac conjI,
- REPEAT_DETERM_N 2 o EVERY' [rtac exI, rtac conjI, etac @{thm conjI[OF refl sym]},
- CONJ_WRAP' (K atac) ls']] 1
- end;
-
-end;
--- a/src/HOL/Codatatype/Tools/bnf_fp.ML Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,442 +0,0 @@
-(* Title: HOL/BNF/Tools/bnf_fp.ML
- Author: Dmitriy Traytel, TU Muenchen
- Copyright 2012
-
-Shared library for the datatype and codatatype constructions.
-*)
-
-signature BNF_FP =
-sig
- val time: Timer.real_timer -> string -> Timer.real_timer
-
- 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 corecN: string
- val corecsN: string
- val ctorN: string
- val ctor_dtorN: string
- val ctor_dtor_unfoldsN: string
- val ctor_dtor_corecsN: string
- val ctor_exhaustN: string
- val ctor_induct2N: string
- val ctor_inductN: string
- val ctor_injectN: string
- val ctor_foldN: string
- val ctor_fold_uniqueN: string
- val ctor_foldsN: string
- val ctor_recN: string
- val ctor_recsN: string
- val disc_unfold_iffN: string
- val disc_unfoldsN: string
- val disc_corec_iffN: string
- val disc_corecsN: string
- val dtorN: string
- val dtor_coinductN: string
- val dtor_unfoldN: string
- val dtor_unfold_uniqueN: string
- val dtor_unfoldsN: string
- val dtor_corecN: string
- val dtor_corecsN: string
- val dtor_exhaustN: string
- val dtor_ctorN: string
- val dtor_injectN: string
- val dtor_strong_coinductN: string
- val exhaustN: string
- val foldN: string
- val foldsN: string
- val hsetN: string
- val hset_recN: string
- val inductN: string
- val injectN: string
- val isNodeN: string
- val lsbisN: string
- val map_simpsN: string
- val map_uniqueN: string
- val min_algN: string
- val morN: string
- val nchotomyN: string
- val recN: string
- val recsN: string
- val rel_coinductN: string
- val rel_simpN: string
- val rel_strong_coinductN: string
- val rvN: string
- val sel_unfoldsN: string
- val sel_corecsN: string
- val set_inclN: string
- val set_set_inclN: string
- val simpsN: string
- val srel_coinductN: string
- val srel_simpN: string
- val srel_strong_coinductN: string
- val strTN: string
- val str_initN: string
- val strongN: string
- val sum_bdN: string
- val sum_bdTN: string
- val unfoldN: string
- val unfoldsN: string
- val uniqueN: string
-
- val mk_exhaustN: string -> string
- val mk_injectN: string -> string
- val mk_nchotomyN: string -> string
- val mk_set_simpsN: int -> string
- val mk_set_minimalN: int -> string
- val mk_set_inductN: int -> string
-
- val mk_common_name: string list -> string
-
- val split_conj_thm: thm -> thm list
- val split_conj_prems: int -> thm -> thm
-
- val retype_free: typ -> term -> term
-
- val mk_sumTN: typ list -> typ
- val mk_sumTN_balanced: typ list -> typ
-
- val id_const: typ -> term
- val id_abs: typ -> term
-
- val Inl_const: typ -> typ -> term
- val Inr_const: typ -> typ -> term
-
- val mk_Inl: typ -> term -> term
- val mk_Inr: typ -> term -> term
- val mk_InN: typ list -> term -> int -> term
- val mk_InN_balanced: typ -> int -> term -> int -> term
- val mk_sum_case: term * term -> term
- val mk_sum_caseN: term list -> term
- val mk_sum_caseN_balanced: term list -> term
-
- val dest_sumT: typ -> typ * typ
- val dest_sumTN: int -> typ -> typ list
- val dest_sumTN_balanced: int -> typ -> typ list
- val dest_tupleT: int -> typ -> typ list
-
- val mk_Field: term -> term
- val mk_If: term -> term -> term -> term
- val mk_union: term * term -> term
-
- val mk_sumEN: int -> thm
- val mk_sumEN_balanced: int -> thm
- val mk_sumEN_tupled_balanced: int list -> thm
- val mk_sum_casesN: int -> int -> thm
- val mk_sum_casesN_balanced: int -> int -> thm
-
- val fixpoint: ('a * 'a -> bool) -> ('a list -> 'a list) -> 'a list -> 'a list
-
- val fp_bnf: (mixfix list -> (string * sort) list option -> binding list ->
- typ list * typ list list -> BNF_Def.BNF list -> local_theory -> 'a) ->
- binding list -> mixfix list -> (string * sort) list -> ((string * sort) * typ) list ->
- local_theory -> BNF_Def.BNF list * 'a
- val fp_bnf_cmd: (mixfix list -> (string * sort) list option -> binding list ->
- typ list * typ list list -> BNF_Def.BNF list -> local_theory -> 'a) ->
- binding list * (string list * string list) -> local_theory -> 'a
-end;
-
-structure BNF_FP : BNF_FP =
-struct
-
-open BNF_Comp
-open BNF_Def
-open BNF_Util
-
-val timing = true;
-fun time timer msg = (if timing
- then warning (msg ^ ": " ^ ATP_Util.string_from_time (Timer.checkRealTimer timer))
- 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 foldsN = foldN ^ "s"
-val unfoldN = unN ^ foldN
-val unfoldsN = unfoldN ^ "s"
-val uniqueN = "_unique"
-val simpsN = "simps"
-val ctorN = "ctor"
-val dtorN = "dtor"
-val ctor_foldN = ctorN ^ "_" ^ foldN
-val ctor_foldsN = ctor_foldN ^ "s"
-val dtor_unfoldN = dtorN ^ "_" ^ unfoldN
-val dtor_unfoldsN = dtor_unfoldN ^ "s"
-val ctor_fold_uniqueN = ctor_foldN ^ uniqueN
-val dtor_unfold_uniqueN = dtor_unfoldN ^ uniqueN
-val ctor_dtor_unfoldsN = ctorN ^ "_" ^ dtor_unfoldN ^ "s"
-val map_simpsN = mapN ^ "_" ^ simpsN
-val map_uniqueN = mapN ^ 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"
-fun mk_set_simpsN i = mk_setN i ^ "_" ^ simpsN
-fun mk_set_minimalN i = mk_setN i ^ "_minimal"
-fun mk_set_inductN i = mk_setN i ^ "_induct"
-
-val str_initN = "str_init"
-val recN = "rec"
-val recsN = recN ^ "s"
-val corecN = coN ^ recN
-val corecsN = corecN ^ "s"
-val ctor_recN = ctorN ^ "_" ^ recN
-val ctor_recsN = ctor_recN ^ "s"
-val dtor_corecN = dtorN ^ "_" ^ corecN
-val dtor_corecsN = dtor_corecN ^ "s"
-val ctor_dtor_corecsN = ctorN ^ "_" ^ dtor_corecN ^ "s"
-
-val ctor_dtorN = ctorN ^ "_" ^ dtorN
-val dtor_ctorN = dtorN ^ "_" ^ ctorN
-val nchotomyN = "nchotomy"
-fun mk_nchotomyN s = s ^ "_" ^ nchotomyN
-val injectN = "inject"
-fun mk_injectN s = s ^ "_" ^ injectN
-val exhaustN = "exhaust"
-fun mk_exhaustN s = s ^ "_" ^ exhaustN
-val ctor_injectN = mk_injectN ctorN
-val ctor_exhaustN = mk_exhaustN ctorN
-val dtor_injectN = mk_injectN dtorN
-val dtor_exhaustN = mk_exhaustN dtorN
-val inductN = "induct"
-val coinductN = coN ^ inductN
-val ctor_inductN = ctorN ^ "_" ^ inductN
-val ctor_induct2N = ctor_inductN ^ "2"
-val dtor_coinductN = dtorN ^ "_" ^ coinductN
-val rel_coinductN = relN ^ "_" ^ coinductN
-val srel_coinductN = srelN ^ "_" ^ coinductN
-val simpN = "_simp";
-val srel_simpN = srelN ^ simpN;
-val rel_simpN = relN ^ simpN;
-val strongN = "strong_"
-val dtor_strong_coinductN = dtorN ^ "_" ^ strongN ^ coinductN
-val rel_strong_coinductN = relN ^ "_" ^ strongN ^ coinductN
-val srel_strong_coinductN = srelN ^ "_" ^ strongN ^ coinductN
-val hsetN = "Hset"
-val hset_recN = hsetN ^ "_rec"
-val set_inclN = "set_incl"
-val set_set_inclN = "set_set_incl"
-
-val caseN = "case"
-val discN = "disc"
-val disc_unfoldsN = discN ^ "_" ^ unfoldsN
-val disc_corecsN = discN ^ "_" ^ corecsN
-val iffN = "_iff"
-val disc_unfold_iffN = discN ^ "_" ^ unfoldN ^ iffN
-val disc_corec_iffN = discN ^ "_" ^ corecN ^ iffN
-val selN = "sel"
-val sel_unfoldsN = selN ^ "_" ^ unfoldsN
-val sel_corecsN = selN ^ "_" ^ corecsN
-
-val mk_common_name = space_implode "_";
-
-fun retype_free T (Free (s, _)) = Free (s, T);
-
-fun dest_sumT (Type (@{type_name sum}, [T, T'])) = (T, T');
-
-fun dest_sumTN 1 T = [T]
- | dest_sumTN n (Type (@{type_name sum}, [T, T'])) = T :: dest_sumTN (n - 1) T';
-
-val dest_sumTN_balanced = Balanced_Tree.dest dest_sumT;
-
-(* TODO: move something like this to "HOLogic"? *)
-fun dest_tupleT 0 @{typ unit} = []
- | dest_tupleT 1 T = [T]
- | dest_tupleT n (Type (@{type_name prod}, [T, T'])) = T :: dest_tupleT (n - 1) T';
-
-val mk_sumTN = Library.foldr1 mk_sumT;
-val mk_sumTN_balanced = Balanced_Tree.make mk_sumT;
-
-fun id_const T = Const (@{const_name id}, T --> T);
-fun id_abs T = Abs (Name.uu, T, Bound 0);
-
-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_InN [_] t 1 = t
- | mk_InN (_ :: Ts) t 1 = mk_Inl (mk_sumTN Ts) t
- | mk_InN (LT :: Ts) t m = mk_Inr LT (mk_InN Ts t (m - 1))
- | mk_InN Ts t _ = raise (TYPE ("mk_InN", Ts, [t]));
-
-fun mk_InN_balanced sum_T n t k =
- let
- fun repair_types T (Const (s as @{const_name Inl}, _) $ t) = repair_inj_types T s fst t
- | repair_types T (Const (s as @{const_name Inr}, _) $ t) = repair_inj_types T s snd t
- | repair_types _ t = t
- and repair_inj_types T s get t =
- let val T' = get (dest_sumT T) in
- Const (s, T' --> T) $ repair_types T' t
- end;
- in
- Balanced_Tree.access {left = mk_Inl dummyT, right = mk_Inr dummyT, init = t} n k
- |> repair_types sum_T
- end;
-
-fun mk_sum_case (f, g) =
- let
- val fT = fastype_of f;
- val gT = fastype_of g;
- in
- Const (@{const_name sum_case},
- fT --> gT --> mk_sumT (domain_type fT, domain_type gT) --> range_type fT) $ f $ g
- end;
-
-val mk_sum_caseN = Library.foldr1 mk_sum_case;
-val mk_sum_caseN_balanced = Balanced_Tree.make mk_sum_case;
-
-fun mk_If p t f =
- let val T = fastype_of t;
- in Const (@{const_name If}, HOLogic.boolT --> T --> T --> 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;
-
-val mk_union = HOLogic.mk_binop @{const_name sup};
-
-(*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_sumEN 1 = @{thm one_pointE}
- | mk_sumEN 2 = @{thm sumE}
- | mk_sumEN n =
- (fold (fn i => fn thm => @{thm obj_sum_step} RSN (i, thm)) (2 upto n - 1) @{thm obj_sumE}) OF
- replicate n (impI RS allI);
-
-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_sumEN_balanced' n all_impIs = mk_obj_sumEN_balanced n OF all_impIs RS @{thm obj_one_pointE};
-
-fun mk_sumEN_balanced 1 = @{thm one_pointE} (*optimization*)
- | mk_sumEN_balanced 2 = @{thm sumE} (*optimization*)
- | mk_sumEN_balanced n = mk_sumEN_balanced' n (replicate n (impI RS allI));
-
-fun mk_tupled_allIN 0 = @{thm unit_all_impI}
- | mk_tupled_allIN 1 = @{thm impI[THEN allI]}
- | mk_tupled_allIN 2 = @{thm prod_all_impI} (*optimization*)
- | mk_tupled_allIN n = mk_tupled_allIN (n - 1) RS @{thm prod_all_impI_step};
-
-fun mk_sumEN_tupled_balanced ms =
- let val n = length ms in
- if forall (curry (op =) 1) ms then mk_sumEN_balanced n
- else mk_sumEN_balanced' n (map mk_tupled_allIN ms)
- end;
-
-fun mk_sum_casesN 1 1 = refl
- | mk_sum_casesN _ 1 = @{thm sum.cases(1)}
- | mk_sum_casesN 2 2 = @{thm sum.cases(2)}
- | mk_sum_casesN n k = trans OF [@{thm sum_case_step(2)}, mk_sum_casesN (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_casesN_balanced 1 1 = refl
- | mk_sum_casesN_balanced n k =
- Balanced_Tree.access {left = mk_sum_step @{thm sum.cases(1)} @{thm sum_case_step(1)},
- right = mk_sum_step @{thm sum.cases(2)} @{thm sum_case_step(2)}, init = refl} n k;
-
-(* FIXME: because of "@ lhss", the output could contain type variables that are not in the input;
- also, "fp_sort" should put the "resBs" first and in the order in which they appear *)
-fun fp_sort lhss NONE Ass = Library.sort (Term_Ord.typ_ord o pairself TFree)
- (subtract (op =) lhss (fold (fold (insert (op =))) Ass [])) @ lhss
- | fp_sort lhss (SOME resBs) Ass =
- (subtract (op =) lhss (filter (fn T => exists (fn Ts => member (op =) Ts T) Ass) resBs)) @ lhss;
-
-fun mk_fp_bnf timer construct resBs bs sort lhss bnfs deadss livess unfold_set lthy =
- let
- val name = mk_common_name (map Binding.name_of bs);
- fun qualify i =
- let val namei = name ^ nonzero_string_of_int i;
- in Binding.qualify true namei end;
-
- val Ass = map (map dest_TFree) livess;
- val resDs = (case resBs of NONE => [] | SOME Ts => fold (subtract (op =)) Ass Ts);
- val Ds = fold (fold Term.add_tfreesT) deadss [];
-
- val _ = (case Library.inter (op =) Ds lhss of [] => ()
- | A :: _ => error ("Nonadmissible type recursion (cannot take fixed point of dead type \
- \variable " ^ quote (Syntax.string_of_typ lthy (TFree A)) ^ ")"));
-
- val timer = time (timer "Construction of BNFs");
-
- val ((kill_poss, _), (bnfs', (unfold_set', lthy'))) =
- normalize_bnfs qualify Ass Ds sort bnfs unfold_set lthy;
-
- val Dss = map3 (append oo map o nth) livess kill_poss deadss;
-
- val ((bnfs'', deadss), lthy'') =
- fold_map3 (seal_bnf unfold_set') (map (Binding.prefix_name preN) bs) Dss bnfs' lthy'
- |>> split_list;
-
- val timer = time (timer "Normalization & sealing of BNFs");
-
- val res = construct resBs bs (map TFree resDs, deadss) bnfs'' lthy'';
-
- val timer = time (timer "FP construction in total");
- in
- timer; (bnfs'', res)
- end;
-
-fun fp_bnf construct bs mixfixes resBs eqs lthy =
- let
- val timer = time (Timer.startRealTimer ());
- val (lhss, rhss) = split_list eqs;
- val sort = fp_sort lhss (SOME resBs);
- fun qualify b = Binding.qualify true (Binding.name_of (Binding.prefix_name rawN b));
- val ((bnfs, (Dss, Ass)), (unfold_set, lthy')) = apfst (apsnd split_list o split_list)
- (fold_map2 (fn b => bnf_of_typ Smart_Inline (qualify b) sort) bs rhss
- (empty_unfolds, lthy));
- in
- mk_fp_bnf timer (construct mixfixes) (SOME resBs) bs sort lhss bnfs Dss Ass unfold_set lthy'
- end;
-
-fun fp_bnf_cmd construct (bs, (raw_lhss, raw_bnfs)) lthy =
- let
- val timer = time (Timer.startRealTimer ());
- val lhss = map (dest_TFree o Syntax.read_typ lthy) raw_lhss;
- val sort = fp_sort lhss NONE;
- fun qualify b = Binding.qualify true (Binding.name_of (Binding.prefix_name rawN b));
- val ((bnfs, (Dss, Ass)), (unfold_set, lthy')) = apfst (apsnd split_list o split_list)
- (fold_map2 (fn b => fn rawT =>
- (bnf_of_typ Smart_Inline (qualify b) sort (Syntax.read_typ lthy rawT)))
- bs raw_bnfs (empty_unfolds, lthy));
- in
- snd (mk_fp_bnf timer
- (construct (map (K NoSyn) bs)) NONE bs sort lhss bnfs Dss Ass unfold_set lthy')
- end;
-
-end;
--- a/src/HOL/Codatatype/Tools/bnf_fp_sugar.ML Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,911 +0,0 @@
-(* Title: HOL/BNF/Tools/bnf_fp_sugar.ML
- Author: Jasmin Blanchette, TU Muenchen
- Copyright 2012
-
-Sugared datatype and codatatype constructions.
-*)
-
-signature BNF_FP_SUGAR =
-sig
- val datatyp: bool ->
- (mixfix list -> (string * sort) list option -> binding list -> typ list * typ list list ->
- BNF_Def.BNF list -> local_theory ->
- (term list * term list * term list * term list * thm * thm list * thm list * thm list *
- thm list * thm list) * local_theory) ->
- bool * ((((typ * sort) list * binding) * mixfix) * ((((binding * binding) *
- (binding * typ) list) * (binding * term) list) * mixfix) list) list ->
- local_theory -> local_theory
- val parse_datatype_cmd: bool ->
- (mixfix list -> (string * sort) list option -> binding list -> typ list * typ list list ->
- BNF_Def.BNF list -> local_theory ->
- (term list * term list * term list * term list * thm * thm list * thm list * thm list *
- thm list * thm list) * local_theory) ->
- (local_theory -> local_theory) parser
-end;
-
-structure BNF_FP_Sugar : BNF_FP_SUGAR =
-struct
-
-open BNF_Util
-open BNF_Wrap
-open BNF_Def
-open BNF_FP
-open BNF_FP_Sugar_Tactics
-
-val simp_attrs = @{attributes [simp]};
-
-fun split_list8 xs =
- (map #1 xs, map #2 xs, map #3 xs, map #4 xs, map #5 xs, map #6 xs, map #7 xs, map #8 xs);
-
-fun resort_tfree S (TFree (s, _)) = TFree (s, S);
-
-fun typ_subst inst (T as Type (s, Ts)) =
- (case AList.lookup (op =) inst T of
- NONE => Type (s, map (typ_subst inst) Ts)
- | SOME T' => T')
- | typ_subst inst T = the_default T (AList.lookup (op =) inst T);
-
-val lists_bmoc = fold (fn xs => fn t => Term.list_comb (t, xs));
-
-fun mk_tupled_fun x f xs = HOLogic.tupled_lambda x (Term.list_comb (f, xs));
-fun mk_uncurried_fun f xs = mk_tupled_fun (HOLogic.mk_tuple xs) f xs;
-fun mk_uncurried2_fun f xss =
- mk_tupled_fun (HOLogic.mk_tuple (map HOLogic.mk_tuple xss)) f (flat xss);
-
-fun tick u f = Term.lambda u (HOLogic.mk_prod (u, f $ u));
-
-fun tack z_name (c, u) f =
- let val z = Free (z_name, mk_sumT (fastype_of u, fastype_of c)) in
- Term.lambda z (mk_sum_case (Term.lambda u u, Term.lambda c (f $ c)) $ z)
- end;
-
-fun cannot_merge_types () = error "Mutually recursive types must have the same type parameters";
-
-fun merge_type_arg T T' = if T = T' then T else cannot_merge_types ();
-
-fun merge_type_args (As, As') =
- if length As = length As' then map2 merge_type_arg As As' else cannot_merge_types ();
-
-fun is_triv_implies thm =
- op aconv (Logic.dest_implies (Thm.prop_of thm))
- handle TERM _ => false;
-
-fun type_args_constrained_of (((cAs, _), _), _) = cAs;
-fun type_binding_of (((_, b), _), _) = b;
-fun mixfix_of ((_, mx), _) = mx;
-fun ctr_specs_of (_, ctr_specs) = ctr_specs;
-
-fun disc_of ((((disc, _), _), _), _) = disc;
-fun ctr_of ((((_, ctr), _), _), _) = ctr;
-fun args_of (((_, args), _), _) = args;
-fun defaults_of ((_, ds), _) = ds;
-fun ctr_mixfix_of (_, mx) = mx;
-
-fun define_datatype prepare_constraint prepare_typ prepare_term lfp construct (no_dests, specs)
- no_defs_lthy0 =
- let
- (* TODO: sanity checks on arguments *)
- (* TODO: integration with function package ("size") *)
-
- val _ = if not lfp andalso no_dests then error "Cannot define destructor-less codatatypes"
- else ();
-
- val nn = length specs;
- val fp_bs = map type_binding_of specs;
- val fp_b_names = map Binding.name_of fp_bs;
- val fp_common_name = mk_common_name fp_b_names;
-
- fun prepare_type_arg (ty, c) =
- let val TFree (s, _) = prepare_typ no_defs_lthy0 ty in
- TFree (s, prepare_constraint no_defs_lthy0 c)
- end;
-
- val Ass0 = map (map prepare_type_arg o type_args_constrained_of) specs;
- val unsorted_Ass0 = map (map (resort_tfree HOLogic.typeS)) Ass0;
- val unsorted_As = Library.foldr1 merge_type_args unsorted_Ass0;
-
- val ((Bs, Cs), no_defs_lthy) =
- no_defs_lthy0
- |> fold (Variable.declare_typ o resort_tfree dummyS) unsorted_As
- |> mk_TFrees nn
- ||>> mk_TFrees nn;
-
- (* TODO: cleaner handling of fake contexts, without "background_theory" *)
- (*the "perhaps o try" below helps gracefully handles the case where the new type is defined in a
- locale and shadows an existing global type*)
- val fake_thy =
- Theory.copy #> fold (fn spec => perhaps (try (Sign.add_type no_defs_lthy
- (type_binding_of spec, length (type_args_constrained_of spec), mixfix_of spec)))) specs;
- val fake_lthy = Proof_Context.background_theory fake_thy no_defs_lthy;
-
- fun mk_fake_T b =
- Type (fst (Term.dest_Type (Proof_Context.read_type_name fake_lthy true (Binding.name_of b))),
- unsorted_As);
-
- val fake_Ts = map mk_fake_T fp_bs;
-
- val mixfixes = map mixfix_of specs;
-
- val _ = (case duplicates Binding.eq_name fp_bs of [] => ()
- | b :: _ => error ("Duplicate type name declaration " ^ quote (Binding.name_of b)));
-
- val ctr_specss = map ctr_specs_of specs;
-
- val disc_bindingss = map (map disc_of) ctr_specss;
- val ctr_bindingss =
- map2 (fn fp_b_name => map (Binding.qualify false fp_b_name o ctr_of)) fp_b_names ctr_specss;
- val ctr_argsss = map (map args_of) ctr_specss;
- val ctr_mixfixess = map (map ctr_mixfix_of) ctr_specss;
-
- val sel_bindingsss = map (map (map fst)) ctr_argsss;
- val fake_ctr_Tsss0 = map (map (map (prepare_typ fake_lthy o snd))) ctr_argsss;
- val raw_sel_defaultsss = map (map defaults_of) ctr_specss;
-
- val (As :: _) :: fake_ctr_Tsss =
- burrow (burrow (Syntax.check_typs fake_lthy)) (Ass0 :: fake_ctr_Tsss0);
-
- val _ = (case duplicates (op =) unsorted_As of [] => ()
- | A :: _ => error ("Duplicate type parameter " ^
- quote (Syntax.string_of_typ no_defs_lthy A)));
-
- val rhs_As' = fold (fold (fold Term.add_tfreesT)) fake_ctr_Tsss [];
- val _ = (case subtract (op =) (map dest_TFree As) rhs_As' of
- [] => ()
- | A' :: _ => error ("Extra type variable on right-hand side: " ^
- quote (Syntax.string_of_typ no_defs_lthy (TFree A'))));
-
- fun eq_fpT (T as Type (s, Us)) (Type (s', Us')) =
- s = s' andalso (Us = Us' orelse error ("Illegal occurrence of recursive type " ^
- quote (Syntax.string_of_typ fake_lthy T)))
- | eq_fpT _ _ = false;
-
- fun freeze_fp (T as Type (s, Us)) =
- (case find_index (eq_fpT T) fake_Ts of ~1 => Type (s, map freeze_fp Us) | j => nth Bs j)
- | freeze_fp T = T;
-
- val ctr_TsssBs = map (map (map freeze_fp)) fake_ctr_Tsss;
- val ctr_sum_prod_TsBs = map (mk_sumTN_balanced o map HOLogic.mk_tupleT) ctr_TsssBs;
-
- val fp_eqs =
- map dest_TFree Bs ~~ map (Term.typ_subst_atomic (As ~~ unsorted_As)) ctr_sum_prod_TsBs;
-
- val (pre_bnfs, ((dtors0, ctors0, fp_folds0, fp_recs0, fp_induct, dtor_ctors, ctor_dtors,
- ctor_injects, fp_fold_thms, fp_rec_thms), lthy)) =
- fp_bnf construct fp_bs mixfixes (map dest_TFree unsorted_As) fp_eqs no_defs_lthy0;
-
- fun add_nesty_bnf_names Us =
- let
- fun add (Type (s, Ts)) ss =
- let val (needs, ss') = fold_map add Ts ss in
- if exists I needs then (true, insert (op =) s ss') else (false, ss')
- end
- | add T ss = (member (op =) Us T, ss);
- in snd oo add end;
-
- fun nesty_bnfs Us =
- map_filter (bnf_of lthy) (fold (fold (fold (add_nesty_bnf_names Us))) ctr_TsssBs []);
-
- val nesting_bnfs = nesty_bnfs As;
- val nested_bnfs = nesty_bnfs Bs;
-
- val timer = time (Timer.startRealTimer ());
-
- fun mk_ctor_or_dtor get_T Ts t =
- let val Type (_, Ts0) = get_T (fastype_of t) in
- Term.subst_atomic_types (Ts0 ~~ Ts) t
- end;
-
- val mk_ctor = mk_ctor_or_dtor range_type;
- val mk_dtor = mk_ctor_or_dtor domain_type;
-
- val ctors = map (mk_ctor As) ctors0;
- val dtors = map (mk_dtor As) dtors0;
-
- val fpTs = map (domain_type o fastype_of) dtors;
-
- val exists_fp_subtype = exists_subtype (member (op =) fpTs);
-
- val ctr_Tsss = map (map (map (Term.typ_subst_atomic (Bs ~~ fpTs)))) ctr_TsssBs;
- val ns = map length ctr_Tsss;
- val kss = map (fn n => 1 upto n) ns;
- val mss = map (map length) ctr_Tsss;
- val Css = map2 replicate ns Cs;
-
- fun mk_rec_like Ts Us t =
- let
- val (bindings, body) = strip_type (fastype_of t);
- val (f_Us, prebody) = split_last bindings;
- val Type (_, Ts0) = if lfp then prebody else body;
- val Us0 = distinct (op =) (map (if lfp then body_type else domain_type) f_Us);
- in
- Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t
- end;
-
- val fp_folds as fp_fold1 :: _ = map (mk_rec_like As Cs) fp_folds0;
- val fp_recs as fp_rec1 :: _ = map (mk_rec_like As Cs) fp_recs0;
-
- val fp_fold_fun_Ts = fst (split_last (binder_types (fastype_of fp_fold1)));
- val fp_rec_fun_Ts = fst (split_last (binder_types (fastype_of fp_rec1)));
-
- val (((fold_only as (gss, _, _), rec_only as (hss, _, _)),
- (zs, cs, cpss, unfold_only as ((pgss, crgsss), _), corec_only as ((phss, cshsss), _))),
- names_lthy) =
- if lfp then
- let
- val y_Tsss =
- map3 (fn n => fn ms => map2 dest_tupleT ms o dest_sumTN_balanced n o domain_type)
- ns mss fp_fold_fun_Ts;
- val g_Tss = map2 (map2 (curry (op --->))) y_Tsss Css;
-
- val ((gss, ysss), lthy) =
- lthy
- |> mk_Freess "f" g_Tss
- ||>> mk_Freesss "x" y_Tsss;
- val yssss = map (map (map single)) ysss;
-
- fun dest_rec_prodT (T as Type (@{type_name prod}, Us as [_, U])) =
- if member (op =) Cs U then Us else [T]
- | dest_rec_prodT T = [T];
-
- val z_Tssss =
- map3 (fn n => fn ms => map2 (map dest_rec_prodT oo dest_tupleT) ms o
- dest_sumTN_balanced n o domain_type) ns mss fp_rec_fun_Ts;
- val h_Tss = map2 (map2 (fold_rev (curry (op --->)))) z_Tssss Css;
-
- val hss = map2 (map2 retype_free) h_Tss gss;
- val zssss_hd = map2 (map2 (map2 (retype_free o hd))) z_Tssss ysss;
- val (zssss_tl, lthy) =
- lthy
- |> mk_Freessss "y" (map (map (map tl)) z_Tssss);
- val zssss = map2 (map2 (map2 cons)) zssss_hd zssss_tl;
- in
- ((((gss, g_Tss, yssss), (hss, h_Tss, zssss)),
- ([], [], [], (([], []), ([], [])), (([], []), ([], [])))), lthy)
- end
- else
- let
- (*avoid "'a itself" arguments in coiterators and corecursors*)
- val mss' = map (fn [0] => [1] | ms => ms) mss;
-
- val p_Tss = map2 (fn n => replicate (Int.max (0, n - 1)) o mk_pred1T) ns Cs;
-
- fun zip_predss_getterss qss fss = maps (op @) (qss ~~ fss);
-
- fun zip_preds_predsss_gettersss [] [qss] [fss] = zip_predss_getterss qss fss
- | zip_preds_predsss_gettersss (p :: ps) (qss :: qsss) (fss :: fsss) =
- p :: zip_predss_getterss qss fss @ zip_preds_predsss_gettersss ps qsss fsss;
-
- fun mk_types maybe_dest_sumT fun_Ts =
- let
- val f_sum_prod_Ts = map range_type fun_Ts;
- val f_prod_Tss = map2 dest_sumTN_balanced ns f_sum_prod_Ts;
- val f_Tssss =
- map3 (fn C => map2 (map (map (curry (op -->) C) o maybe_dest_sumT) oo dest_tupleT))
- Cs mss' f_prod_Tss;
- val q_Tssss =
- map (map (map (fn [_] => [] | [_, C] => [mk_pred1T (domain_type C)]))) f_Tssss;
- val pf_Tss = map3 zip_preds_predsss_gettersss p_Tss q_Tssss f_Tssss;
- in (q_Tssss, f_sum_prod_Ts, f_Tssss, pf_Tss) end;
-
- val (r_Tssss, g_sum_prod_Ts, g_Tssss, pg_Tss) = mk_types single fp_fold_fun_Ts;
-
- val ((((Free (z, _), cs), pss), gssss), lthy) =
- lthy
- |> yield_singleton (mk_Frees "z") dummyT
- ||>> mk_Frees "a" Cs
- ||>> mk_Freess "p" p_Tss
- ||>> mk_Freessss "g" g_Tssss;
- val rssss = map (map (map (fn [] => []))) r_Tssss;
-
- fun dest_corec_sumT (T as Type (@{type_name sum}, Us as [_, U])) =
- if member (op =) Cs U then Us else [T]
- | dest_corec_sumT T = [T];
-
- val (s_Tssss, h_sum_prod_Ts, h_Tssss, ph_Tss) = mk_types dest_corec_sumT fp_rec_fun_Ts;
-
- val hssss_hd = map2 (map2 (map2 (fn T :: _ => fn [g] => retype_free T g))) h_Tssss gssss;
- val ((sssss, hssss_tl), lthy) =
- lthy
- |> mk_Freessss "q" s_Tssss
- ||>> mk_Freessss "h" (map (map (map tl)) h_Tssss);
- val hssss = map2 (map2 (map2 cons)) hssss_hd hssss_tl;
-
- val cpss = map2 (fn c => map (fn p => p $ c)) cs pss;
-
- fun mk_preds_getters_join [] [cf] = cf
- | mk_preds_getters_join [cq] [cf, cf'] =
- mk_If cq (mk_Inl (fastype_of cf') cf) (mk_Inr (fastype_of cf) cf');
-
- fun mk_terms qssss fssss =
- let
- val pfss = map3 zip_preds_predsss_gettersss pss qssss fssss;
- val cqssss = map2 (fn c => map (map (map (fn f => f $ c)))) cs qssss;
- val cfssss = map2 (fn c => map (map (map (fn f => f $ c)))) cs fssss;
- val cqfsss = map2 (map2 (map2 mk_preds_getters_join)) cqssss cfssss;
- in (pfss, cqfsss) end;
- in
- (((([], [], []), ([], [], [])),
- ([z], cs, cpss, (mk_terms rssss gssss, (g_sum_prod_Ts, pg_Tss)),
- (mk_terms sssss hssss, (h_sum_prod_Ts, ph_Tss)))), lthy)
- end;
-
- fun define_ctrs_case_for_type ((((((((((((((((((fp_b, fpT), C), ctor), dtor), fp_fold), fp_rec),
- ctor_dtor), dtor_ctor), ctor_inject), n), ks), ms), ctr_bindings), ctr_mixfixes), ctr_Tss),
- disc_bindings), sel_bindingss), raw_sel_defaultss) no_defs_lthy =
- let
- val fp_b_name = Binding.name_of fp_b;
-
- val dtorT = domain_type (fastype_of ctor);
- val ctr_prod_Ts = map HOLogic.mk_tupleT ctr_Tss;
- val ctr_sum_prod_T = mk_sumTN_balanced ctr_prod_Ts;
- val case_Ts = map (fn Ts => Ts ---> C) ctr_Tss;
-
- val ((((w, fs), xss), u'), _) =
- no_defs_lthy
- |> yield_singleton (mk_Frees "w") dtorT
- ||>> mk_Frees "f" case_Ts
- ||>> mk_Freess "x" ctr_Tss
- ||>> yield_singleton Variable.variant_fixes fp_b_name;
-
- val u = Free (u', fpT);
-
- val ctr_rhss =
- map2 (fn k => fn xs => fold_rev Term.lambda xs (ctor $
- mk_InN_balanced ctr_sum_prod_T n (HOLogic.mk_tuple xs) k)) ks xss;
-
- val case_binding = Binding.suffix_name ("_" ^ caseN) fp_b;
-
- val case_rhs =
- fold_rev Term.lambda (fs @ [u])
- (mk_sum_caseN_balanced (map2 mk_uncurried_fun fs xss) $ (dtor $ u));
-
- val ((raw_case :: raw_ctrs, raw_case_def :: raw_ctr_defs), (lthy', lthy)) = no_defs_lthy
- |> apfst split_list o fold_map3 (fn b => fn mx => fn rhs =>
- Local_Theory.define ((b, mx), ((Thm.def_binding b, []), rhs)) #>> apsnd snd)
- (case_binding :: ctr_bindings) (NoSyn :: ctr_mixfixes) (case_rhs :: ctr_rhss)
- ||> `Local_Theory.restore;
-
- val phi = Proof_Context.export_morphism lthy lthy';
-
- val ctr_defs = map (Morphism.thm phi) raw_ctr_defs;
- val case_def = Morphism.thm phi raw_case_def;
-
- val ctrs0 = map (Morphism.term phi) raw_ctrs;
- val casex0 = Morphism.term phi raw_case;
-
- val ctrs = map (mk_ctr As) ctrs0;
-
- fun exhaust_tac {context = ctxt, ...} =
- let
- val ctor_iff_dtor_thm =
- let
- val goal =
- fold_rev Logic.all [w, u]
- (mk_Trueprop_eq (HOLogic.mk_eq (u, ctor $ w), HOLogic.mk_eq (dtor $ u, w)));
- in
- Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
- mk_ctor_iff_dtor_tac ctxt (map (SOME o certifyT lthy) [dtorT, fpT])
- (certify lthy ctor) (certify lthy dtor) ctor_dtor dtor_ctor)
- |> Thm.close_derivation
- |> Morphism.thm phi
- end;
-
- val sumEN_thm' =
- unfold_thms lthy @{thms all_unit_eq}
- (Drule.instantiate' (map (SOME o certifyT lthy) ctr_prod_Ts) []
- (mk_sumEN_balanced n))
- |> Morphism.thm phi;
- in
- mk_exhaust_tac ctxt n ctr_defs ctor_iff_dtor_thm sumEN_thm'
- end;
-
- val inject_tacss =
- map2 (fn 0 => K [] | _ => fn ctr_def => [fn {context = ctxt, ...} =>
- mk_inject_tac ctxt ctr_def ctor_inject]) ms ctr_defs;
-
- val half_distinct_tacss =
- map (map (fn (def, def') => fn {context = ctxt, ...} =>
- mk_half_distinct_tac ctxt ctor_inject [def, def'])) (mk_half_pairss ctr_defs);
-
- val case_tacs =
- map3 (fn k => fn m => fn ctr_def => fn {context = ctxt, ...} =>
- mk_case_tac ctxt n k m case_def ctr_def dtor_ctor) ks ms ctr_defs;
-
- val tacss = [exhaust_tac] :: inject_tacss @ half_distinct_tacss @ [case_tacs];
-
- fun define_fold_rec (wrap_res, no_defs_lthy) =
- let
- val fpT_to_C = fpT --> C;
-
- fun generate_rec_like (suf, fp_rec_like, (fss, f_Tss, xssss)) =
- let
- val res_T = fold_rev (curry (op --->)) f_Tss fpT_to_C;
- val binding = Binding.suffix_name ("_" ^ suf) fp_b;
- val spec =
- mk_Trueprop_eq (lists_bmoc fss (Free (Binding.name_of binding, res_T)),
- Term.list_comb (fp_rec_like,
- map2 (mk_sum_caseN_balanced oo map2 mk_uncurried2_fun) fss xssss));
- in (binding, spec) end;
-
- val rec_like_infos =
- [(foldN, fp_fold, fold_only),
- (recN, fp_rec, rec_only)];
-
- val (bindings, specs) = map generate_rec_like rec_like_infos |> split_list;
-
- val ((csts, defs), (lthy', lthy)) = no_defs_lthy
- |> apfst split_list o fold_map2 (fn b => fn spec =>
- Specification.definition (SOME (b, NONE, NoSyn), ((Thm.def_binding b, []), spec))
- #>> apsnd snd) bindings specs
- ||> `Local_Theory.restore;
-
- val phi = Proof_Context.export_morphism lthy lthy';
-
- val [fold_def, rec_def] = map (Morphism.thm phi) defs;
-
- val [foldx, recx] = map (mk_rec_like As Cs o Morphism.term phi) csts;
- in
- ((wrap_res, ctrs, foldx, recx, xss, ctr_defs, fold_def, rec_def), lthy)
- end;
-
- fun define_unfold_corec (wrap_res, no_defs_lthy) =
- let
- val B_to_fpT = C --> fpT;
-
- fun mk_preds_getterss_join c n cps sum_prod_T cqfss =
- Term.lambda c (mk_IfN sum_prod_T cps
- (map2 (mk_InN_balanced sum_prod_T n) (map HOLogic.mk_tuple cqfss) (1 upto n)));
-
- fun generate_corec_like (suf, fp_rec_like, ((pfss, cqfsss), (f_sum_prod_Ts,
- pf_Tss))) =
- let
- val res_T = fold_rev (curry (op --->)) pf_Tss B_to_fpT;
- val binding = Binding.suffix_name ("_" ^ suf) fp_b;
- val spec =
- mk_Trueprop_eq (lists_bmoc pfss (Free (Binding.name_of binding, res_T)),
- Term.list_comb (fp_rec_like,
- map5 mk_preds_getterss_join cs ns cpss f_sum_prod_Ts cqfsss));
- in (binding, spec) end;
-
- val corec_like_infos =
- [(unfoldN, fp_fold, unfold_only),
- (corecN, fp_rec, corec_only)];
-
- val (bindings, specs) = map generate_corec_like corec_like_infos |> split_list;
-
- val ((csts, defs), (lthy', lthy)) = no_defs_lthy
- |> apfst split_list o fold_map2 (fn b => fn spec =>
- Specification.definition (SOME (b, NONE, NoSyn), ((Thm.def_binding b, []), spec))
- #>> apsnd snd) bindings specs
- ||> `Local_Theory.restore;
-
- val phi = Proof_Context.export_morphism lthy lthy';
-
- val [unfold_def, corec_def] = map (Morphism.thm phi) defs;
-
- val [unfold, corec] = map (mk_rec_like As Cs o Morphism.term phi) csts;
- in
- ((wrap_res, ctrs, unfold, corec, xss, ctr_defs, unfold_def, corec_def), lthy)
- end;
-
- fun wrap lthy =
- let val sel_defaultss = map (map (apsnd (prepare_term lthy))) raw_sel_defaultss in
- wrap_datatype tacss (((no_dests, ctrs0), casex0), (disc_bindings, (sel_bindingss,
- sel_defaultss))) lthy
- end;
-
- val define_rec_likes = if lfp then define_fold_rec else define_unfold_corec;
- in
- ((wrap, define_rec_likes), lthy')
- end;
-
- val pre_map_defs = map map_def_of_bnf pre_bnfs;
- val pre_set_defss = map set_defs_of_bnf pre_bnfs;
- val nested_set_natural's = maps set_natural'_of_bnf nested_bnfs;
- val nesting_map_ids = map map_id_of_bnf nesting_bnfs;
-
- fun mk_map live Ts Us t =
- let
- val (Type (_, Ts0), Type (_, Us0)) = strip_typeN (live + 1) (fastype_of t) |>> List.last
- in
- Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t
- end;
-
- fun build_map build_arg (Type (s, Ts)) (Type (_, Us)) =
- let
- val bnf = the (bnf_of lthy s);
- val live = live_of_bnf bnf;
- val mapx = mk_map live Ts Us (map_of_bnf bnf);
- val TUs = map dest_funT (fst (strip_typeN live (fastype_of mapx)));
- val args = map build_arg TUs;
- in Term.list_comb (mapx, args) end;
-
- val mk_simp_thmss =
- map3 (fn (_, _, injects, distincts, cases, _, _, _) => fn rec_likes => fn fold_likes =>
- injects @ distincts @ cases @ rec_likes @ fold_likes);
-
- fun derive_induct_fold_rec_thms_for_types ((wrap_ress, ctrss, folds, recs, xsss, ctr_defss,
- fold_defs, rec_defs), lthy) =
- let
- val (((phis, phis'), us'), names_lthy) =
- lthy
- |> mk_Frees' "P" (map mk_pred1T fpTs)
- ||>> Variable.variant_fixes fp_b_names;
-
- val us = map2 (curry Free) us' fpTs;
-
- fun mk_sets_nested bnf =
- let
- val Type (T_name, Us) = T_of_bnf bnf;
- val lives = lives_of_bnf bnf;
- val sets = sets_of_bnf bnf;
- fun mk_set U =
- (case find_index (curry (op =) U) lives of
- ~1 => Term.dummy
- | i => nth sets i);
- in
- (T_name, map mk_set Us)
- end;
-
- val setss_nested = map mk_sets_nested nested_bnfs;
-
- val (induct_thms, induct_thm) =
- let
- fun mk_set Ts t =
- let val Type (_, Ts0) = domain_type (fastype_of t) in
- Term.subst_atomic_types (Ts0 ~~ Ts) t
- end;
-
- fun mk_raw_prem_prems names_lthy (x as Free (s, T as Type (T_name, Ts0))) =
- (case find_index (curry (op =) T) fpTs of
- ~1 =>
- (case AList.lookup (op =) setss_nested T_name of
- NONE => []
- | SOME raw_sets0 =>
- let
- val (Ts, raw_sets) =
- split_list (filter (exists_fp_subtype o fst) (Ts0 ~~ raw_sets0));
- val sets = map (mk_set Ts0) raw_sets;
- val (ys, names_lthy') = names_lthy |> mk_Frees s Ts;
- val xysets = map (pair x) (ys ~~ sets);
- val ppremss = map (mk_raw_prem_prems names_lthy') ys;
- in
- flat (map2 (map o apfst o cons) xysets ppremss)
- end)
- | i => [([], (i + 1, x))])
- | mk_raw_prem_prems _ _ = [];
-
- fun close_prem_prem xs t =
- fold_rev Logic.all (map Free (drop (nn + length xs)
- (rev (Term.add_frees t (map dest_Free xs @ phis'))))) t;
-
- fun mk_prem_prem xs (xysets, (j, x)) =
- close_prem_prem xs (Logic.list_implies (map (fn (x', (y, set)) =>
- HOLogic.mk_Trueprop (HOLogic.mk_mem (y, set $ x'))) xysets,
- HOLogic.mk_Trueprop (nth phis (j - 1) $ x)));
-
- fun mk_raw_prem phi ctr ctr_Ts =
- let
- val (xs, names_lthy') = names_lthy |> mk_Frees "x" ctr_Ts;
- val pprems = maps (mk_raw_prem_prems names_lthy') xs;
- in (xs, pprems, HOLogic.mk_Trueprop (phi $ Term.list_comb (ctr, xs))) end;
-
- fun mk_prem (xs, raw_pprems, concl) =
- fold_rev Logic.all xs (Logic.list_implies (map (mk_prem_prem xs) raw_pprems, concl));
-
- val raw_premss = map3 (map2 o mk_raw_prem) phis ctrss ctr_Tsss;
-
- val goal =
- Library.foldr (Logic.list_implies o apfst (map mk_prem)) (raw_premss,
- HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (curry (op $)) phis us)));
-
- val kksss = map (map (map (fst o snd) o #2)) raw_premss;
-
- val ctor_induct' = fp_induct OF (map mk_sumEN_tupled_balanced mss);
-
- val induct_thm =
- Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
- mk_induct_tac ctxt ns mss kksss (flat ctr_defss) ctor_induct'
- nested_set_natural's pre_set_defss)
- |> singleton (Proof_Context.export names_lthy lthy)
- in
- `(conj_dests nn) induct_thm
- end;
-
- (* TODO: Generate nicer names in case of clashes *)
- val induct_cases = Datatype_Prop.indexify_names (maps (map base_name_of_ctr) ctrss);
-
- val (fold_thmss, rec_thmss) =
- let
- val xctrss = map2 (map2 (curry Term.list_comb)) ctrss xsss;
- val gfolds = map (lists_bmoc gss) folds;
- val hrecs = map (lists_bmoc hss) recs;
-
- fun mk_goal fss frec_like xctr f xs fxs =
- fold_rev (fold_rev Logic.all) (xs :: fss)
- (mk_Trueprop_eq (frec_like $ xctr, Term.list_comb (f, fxs)));
-
- fun build_call frec_likes maybe_tick (T, U) =
- if T = U then
- id_const T
- else
- (case find_index (curry (op =) T) fpTs of
- ~1 => build_map (build_call frec_likes maybe_tick) T U
- | j => maybe_tick (nth us j) (nth frec_likes j));
-
- fun mk_U maybe_mk_prodT =
- typ_subst (map2 (fn fpT => fn C => (fpT, maybe_mk_prodT fpT C)) fpTs Cs);
-
- fun intr_calls frec_likes maybe_cons maybe_tick maybe_mk_prodT (x as Free (_, T)) =
- if member (op =) fpTs T then
- maybe_cons x [build_call frec_likes (K I) (T, mk_U (K I) T) $ x]
- else if exists_fp_subtype T then
- [build_call frec_likes maybe_tick (T, mk_U maybe_mk_prodT T) $ x]
- else
- [x];
-
- val gxsss = map (map (maps (intr_calls gfolds (K I) (K I) (K I)))) xsss;
- val hxsss = map (map (maps (intr_calls hrecs cons tick (curry HOLogic.mk_prodT)))) xsss;
-
- val fold_goalss = map5 (map4 o mk_goal gss) gfolds xctrss gss xsss gxsss;
- val rec_goalss = map5 (map4 o mk_goal hss) hrecs xctrss hss xsss hxsss;
-
- val fold_tacss =
- map2 (map o mk_rec_like_tac pre_map_defs nesting_map_ids fold_defs) fp_fold_thms
- ctr_defss;
- val rec_tacss =
- map2 (map o mk_rec_like_tac pre_map_defs nesting_map_ids rec_defs) fp_rec_thms
- ctr_defss;
-
- fun prove goal tac = Skip_Proof.prove lthy [] [] goal (tac o #context);
- in
- (map2 (map2 prove) fold_goalss fold_tacss,
- map2 (map2 prove) rec_goalss rec_tacss)
- end;
-
- val simp_thmss = mk_simp_thmss wrap_ress rec_thmss fold_thmss;
-
- val induct_case_names_attr = Attrib.internal (K (Rule_Cases.case_names induct_cases));
- fun induct_type_attr T_name = Attrib.internal (K (Induct.induct_type T_name));
-
- (* TODO: Also note "recs", "simps", and "splits" if "nn > 1" (for compatibility with the
- old package)? And for codatatypes as well? *)
- val common_notes =
- (if nn > 1 then [(inductN, [induct_thm], [induct_case_names_attr])] else [])
- |> map (fn (thmN, thms, attrs) =>
- ((Binding.qualify true fp_common_name (Binding.name thmN), attrs), [(thms, [])]));
-
- val notes =
- [(inductN, map single induct_thms,
- fn T_name => [induct_case_names_attr, induct_type_attr T_name]),
- (foldsN, fold_thmss, K (Code.add_default_eqn_attrib :: simp_attrs)),
- (recsN, rec_thmss, K (Code.add_default_eqn_attrib :: simp_attrs)),
- (simpsN, simp_thmss, K [])]
- |> maps (fn (thmN, thmss, attrs) =>
- map3 (fn b => fn Type (T_name, _) => fn thms =>
- ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), attrs T_name),
- [(thms, [])])) fp_bs fpTs thmss);
- in
- lthy |> Local_Theory.notes (common_notes @ notes) |> snd
- end;
-
- fun derive_coinduct_unfold_corec_thms_for_types ((wrap_ress, ctrss, unfolds, corecs, _,
- ctr_defss, unfold_defs, corec_defs), lthy) =
- let
- val discss = map (map (mk_disc_or_sel As) o #1) wrap_ress;
- val selsss = map #2 wrap_ress;
- val disc_thmsss = map #6 wrap_ress;
- val discIss = map #7 wrap_ress;
- val sel_thmsss = map #8 wrap_ress;
-
- val (us', _) =
- lthy
- |> Variable.variant_fixes fp_b_names;
-
- val us = map2 (curry Free) us' fpTs;
-
- val (coinduct_thms, coinduct_thm) =
- let
- val coinduct_thm = fp_induct;
- in
- `(conj_dests nn) coinduct_thm
- end;
-
- fun mk_maybe_not pos = not pos ? HOLogic.mk_not;
-
- val z = the_single zs;
- val gunfolds = map (lists_bmoc pgss) unfolds;
- val hcorecs = map (lists_bmoc phss) corecs;
-
- val (unfold_thmss, corec_thmss, safe_unfold_thmss, safe_corec_thmss) =
- let
- fun mk_goal pfss c cps fcorec_like n k ctr m cfs' =
- fold_rev (fold_rev Logic.all) ([c] :: pfss)
- (Logic.list_implies (seq_conds (HOLogic.mk_Trueprop oo mk_maybe_not) n k cps,
- mk_Trueprop_eq (fcorec_like $ c, Term.list_comb (ctr, take m cfs'))));
-
- fun build_call frec_likes maybe_tack (T, U) =
- if T = U then
- id_const T
- else
- (case find_index (curry (op =) U) fpTs of
- ~1 => build_map (build_call frec_likes maybe_tack) T U
- | j => maybe_tack (nth cs j, nth us j) (nth frec_likes j));
-
- fun mk_U maybe_mk_sumT =
- typ_subst (map2 (fn C => fn fpT => (maybe_mk_sumT fpT C, fpT)) Cs fpTs);
-
- fun intr_calls frec_likes maybe_mk_sumT maybe_tack cqf =
- let val T = fastype_of cqf in
- if exists_subtype (member (op =) Cs) T then
- build_call frec_likes maybe_tack (T, mk_U maybe_mk_sumT T) $ cqf
- else
- cqf
- end;
-
- val crgsss' = map (map (map (intr_calls gunfolds (K I) (K I)))) crgsss;
- val cshsss' = map (map (map (intr_calls hcorecs (curry mk_sumT) (tack z)))) cshsss;
-
- val unfold_goalss =
- map8 (map4 oooo mk_goal pgss) cs cpss gunfolds ns kss ctrss mss crgsss';
- val corec_goalss =
- map8 (map4 oooo mk_goal phss) cs cpss hcorecs ns kss ctrss mss cshsss';
-
- val unfold_tacss =
- map3 (map oo mk_corec_like_tac unfold_defs nesting_map_ids) fp_fold_thms pre_map_defs
- ctr_defss;
- val corec_tacss =
- map3 (map oo mk_corec_like_tac corec_defs nesting_map_ids) fp_rec_thms pre_map_defs
- ctr_defss;
-
- fun prove goal tac =
- Skip_Proof.prove lthy [] [] goal (tac o #context) |> Thm.close_derivation;
-
- val unfold_thmss = map2 (map2 prove) unfold_goalss unfold_tacss;
- val corec_thmss =
- map2 (map2 prove) corec_goalss corec_tacss
- |> map (map (unfold_thms lthy @{thms sum_case_if}));
-
- val unfold_safesss = map2 (map2 (map2 (curry (op =)))) crgsss' crgsss;
- val corec_safesss = map2 (map2 (map2 (curry (op =)))) cshsss' cshsss;
-
- val filter_safesss =
- map2 (map_filter (fn (safes, thm) => if forall I safes then SOME thm else NONE) oo
- curry (op ~~));
-
- val safe_unfold_thmss = filter_safesss unfold_safesss unfold_thmss;
- val safe_corec_thmss = filter_safesss corec_safesss corec_thmss;
- in
- (unfold_thmss, corec_thmss, safe_unfold_thmss, safe_corec_thmss)
- end;
-
- val (disc_unfold_iff_thmss, disc_corec_iff_thmss) =
- let
- fun mk_goal c cps fcorec_like n k disc =
- mk_Trueprop_eq (disc $ (fcorec_like $ c),
- if n = 1 then @{const True}
- else Library.foldr1 HOLogic.mk_conj (seq_conds mk_maybe_not n k cps));
-
- val unfold_goalss = map6 (map2 oooo mk_goal) cs cpss gunfolds ns kss discss;
- val corec_goalss = map6 (map2 oooo mk_goal) cs cpss hcorecs ns kss discss;
-
- fun mk_case_split' cp =
- Drule.instantiate' [] [SOME (certify lthy cp)] @{thm case_split};
-
- val case_splitss' = map (map mk_case_split') cpss;
-
- val unfold_tacss =
- map3 (map oo mk_disc_corec_like_iff_tac) case_splitss' unfold_thmss disc_thmsss;
- val corec_tacss =
- map3 (map oo mk_disc_corec_like_iff_tac) case_splitss' corec_thmss disc_thmsss;
-
- fun prove goal tac =
- Skip_Proof.prove lthy [] [] goal (tac o #context)
- |> Thm.close_derivation
- |> singleton (Proof_Context.export names_lthy no_defs_lthy);
-
- fun proves [_] [_] = []
- | proves goals tacs = map2 prove goals tacs;
- in
- (map2 proves unfold_goalss unfold_tacss,
- map2 proves corec_goalss corec_tacss)
- end;
-
- fun mk_disc_corec_like_thms corec_likes discIs =
- map (op RS) (filter_out (is_triv_implies o snd) (corec_likes ~~ discIs));
-
- val disc_unfold_thmss = map2 mk_disc_corec_like_thms unfold_thmss discIss;
- val disc_corec_thmss = map2 mk_disc_corec_like_thms corec_thmss discIss;
-
- fun mk_sel_corec_like_thm corec_like_thm sel sel_thm =
- let
- val (domT, ranT) = dest_funT (fastype_of sel);
- val arg_cong' =
- Drule.instantiate' (map (SOME o certifyT lthy) [domT, ranT])
- [NONE, NONE, SOME (certify lthy sel)] arg_cong
- |> Thm.varifyT_global;
- val sel_thm' = sel_thm RSN (2, trans);
- in
- corec_like_thm RS arg_cong' RS sel_thm'
- end;
-
- fun mk_sel_corec_like_thms corec_likess =
- map3 (map3 (map2 o mk_sel_corec_like_thm)) corec_likess selsss sel_thmsss |> map flat;
-
- val sel_unfold_thmss = mk_sel_corec_like_thms unfold_thmss;
- val sel_corec_thmss = mk_sel_corec_like_thms corec_thmss;
-
- fun zip_corec_like_thms corec_likes disc_corec_likes sel_corec_likes =
- corec_likes @ disc_corec_likes @ sel_corec_likes;
-
- val simp_thmss =
- mk_simp_thmss wrap_ress
- (map3 zip_corec_like_thms safe_corec_thmss disc_corec_thmss sel_corec_thmss)
- (map3 zip_corec_like_thms safe_unfold_thmss disc_unfold_thmss sel_unfold_thmss);
-
- val anonymous_notes =
- [(flat safe_unfold_thmss @ flat safe_corec_thmss, simp_attrs)]
- |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])]));
-
- val common_notes =
- (if nn > 1 then [(coinductN, [coinduct_thm], [])] (* FIXME: attribs *) else [])
- |> map (fn (thmN, thms, attrs) =>
- ((Binding.qualify true fp_common_name (Binding.name thmN), attrs), [(thms, [])]));
-
- val notes =
- [(coinductN, map single coinduct_thms, []), (* FIXME: attribs *)
- (unfoldsN, unfold_thmss, []),
- (corecsN, corec_thmss, []),
- (disc_unfold_iffN, disc_unfold_iff_thmss, simp_attrs),
- (disc_unfoldsN, disc_unfold_thmss, simp_attrs),
- (disc_corec_iffN, disc_corec_iff_thmss, simp_attrs),
- (disc_corecsN, disc_corec_thmss, simp_attrs),
- (sel_unfoldsN, sel_unfold_thmss, simp_attrs),
- (sel_corecsN, sel_corec_thmss, simp_attrs),
- (simpsN, simp_thmss, [])]
- |> maps (fn (thmN, thmss, attrs) =>
- map_filter (fn (_, []) => NONE | (b, thms) =>
- SOME ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), attrs),
- [(thms, [])])) (fp_bs ~~ thmss));
- in
- lthy |> Local_Theory.notes (anonymous_notes @ common_notes @ notes) |> snd
- end;
-
- fun wrap_types_and_define_rec_likes ((wraps, define_rec_likess), lthy) =
- fold_map2 (curry (op o)) define_rec_likess wraps lthy |>> split_list8
-
- val lthy' = lthy
- |> fold_map define_ctrs_case_for_type (fp_bs ~~ fpTs ~~ Cs ~~ ctors ~~ dtors ~~ fp_folds ~~
- fp_recs ~~ ctor_dtors ~~ dtor_ctors ~~ ctor_injects ~~ ns ~~ kss ~~ mss ~~ ctr_bindingss ~~
- ctr_mixfixess ~~ ctr_Tsss ~~ disc_bindingss ~~ sel_bindingsss ~~ raw_sel_defaultsss)
- |>> split_list |> wrap_types_and_define_rec_likes
- |> (if lfp then derive_induct_fold_rec_thms_for_types
- else derive_coinduct_unfold_corec_thms_for_types);
-
- val timer = time (timer ("Constructors, discriminators, selectors, etc., for the new " ^
- (if lfp then "" else "co") ^ "datatype"));
- in
- timer; lthy'
- end;
-
-val datatyp = define_datatype (K I) (K I) (K I);
-
-val datatype_cmd = define_datatype Typedecl.read_constraint Syntax.parse_typ Syntax.read_term;
-
-val parse_ctr_arg =
- @{keyword "("} |-- parse_binding_colon -- Parse.typ --| @{keyword ")"} ||
- (Parse.typ >> pair Binding.empty);
-
-val parse_defaults =
- @{keyword "("} |-- @{keyword "defaults"} |-- Scan.repeat parse_bound_term --| @{keyword ")"};
-
-val parse_single_spec =
- Parse.type_args_constrained -- Parse.binding -- Parse.opt_mixfix --
- (@{keyword "="} |-- Parse.enum1 "|" (parse_opt_binding_colon -- Parse.binding --
- Scan.repeat parse_ctr_arg -- Scan.optional parse_defaults [] -- Parse.opt_mixfix));
-
-val parse_datatype = parse_wrap_options -- Parse.and_list1 parse_single_spec;
-
-fun parse_datatype_cmd lfp construct = parse_datatype >> datatype_cmd lfp construct;
-
-end;
--- a/src/HOL/Codatatype/Tools/bnf_fp_sugar_tactics.ML Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,133 +0,0 @@
-(* Title: HOL/BNF/Tools/bnf_fp_sugar_tactics.ML
- Author: Jasmin Blanchette, TU Muenchen
- Copyright 2012
-
-Tactics for datatype and codatatype sugar.
-*)
-
-signature BNF_FP_SUGAR_TACTICS =
-sig
- val mk_case_tac: Proof.context -> int -> int -> int -> thm -> thm -> thm -> tactic
- val mk_corec_like_tac: thm list -> thm list -> thm -> thm -> thm -> Proof.context -> tactic
- val mk_ctor_iff_dtor_tac: Proof.context -> ctyp option list -> cterm -> cterm -> thm -> thm ->
- tactic
- val mk_disc_corec_like_iff_tac: thm list -> thm list -> thm list -> Proof.context -> tactic
- val mk_exhaust_tac: Proof.context -> int -> thm list -> thm -> thm -> tactic
- val mk_half_distinct_tac: Proof.context -> thm -> thm list -> tactic
- val mk_induct_tac: Proof.context -> int list -> int list list -> int list list list -> thm list ->
- thm -> thm list -> thm list list -> tactic
- val mk_inject_tac: Proof.context -> thm -> thm -> tactic
- val mk_rec_like_tac: thm list -> thm list -> thm list -> thm -> thm -> Proof.context -> tactic
-end;
-
-structure BNF_FP_Sugar_Tactics : BNF_FP_SUGAR_TACTICS =
-struct
-
-open BNF_Tactics
-open BNF_Util
-open BNF_FP
-
-val meta_mp = @{thm meta_mp};
-val meta_spec = @{thm meta_spec};
-
-fun inst_spurious_fs lthy thm =
- let
- val fs =
- Term.add_vars (prop_of thm) []
- |> filter (fn (_, Type (@{type_name fun}, [_, T'])) => T' <> HOLogic.boolT | _ => false);
- val cfs =
- map (fn f as (_, T) => (certify lthy (Var f), certify lthy (id_abs (domain_type T)))) fs;
- in
- Drule.cterm_instantiate cfs thm
- end;
-
-val inst_spurious_fs_tac = PRIMITIVE o inst_spurious_fs;
-
-fun mk_case_tac ctxt n k m case_def ctr_def dtor_ctor =
- unfold_thms_tac ctxt [case_def, ctr_def, dtor_ctor] THEN
- (rtac (mk_sum_casesN_balanced n k RS ssubst) THEN'
- REPEAT_DETERM_N (Int.max (0, m - 1)) o rtac (@{thm split} RS ssubst) THEN'
- rtac refl) 1;
-
-fun mk_exhaust_tac ctxt n ctr_defs ctor_iff_dtor sumEN' =
- unfold_thms_tac ctxt (ctor_iff_dtor :: ctr_defs) THEN rtac sumEN' 1 THEN
- unfold_thms_tac ctxt @{thms all_prod_eq} THEN
- EVERY' (maps (fn k => [select_prem_tac n (rotate_tac 1) k, REPEAT_DETERM o dtac meta_spec,
- etac meta_mp, atac]) (1 upto n)) 1;
-
-fun mk_ctor_iff_dtor_tac ctxt cTs cctor cdtor ctor_dtor dtor_ctor =
- (rtac iffI THEN'
- EVERY' (map3 (fn cTs => fn cx => fn th =>
- dtac (Drule.instantiate' cTs [NONE, NONE, SOME cx] arg_cong) THEN'
- SELECT_GOAL (unfold_thms_tac ctxt [th]) THEN'
- atac) [rev cTs, cTs] [cdtor, cctor] [dtor_ctor, ctor_dtor])) 1;
-
-fun mk_half_distinct_tac ctxt ctor_inject ctr_defs =
- unfold_thms_tac ctxt (ctor_inject :: @{thms sum.inject} @ ctr_defs) THEN
- rtac @{thm sum.distinct(1)} 1;
-
-fun mk_inject_tac ctxt ctr_def ctor_inject =
- unfold_thms_tac ctxt [ctr_def] THEN rtac (ctor_inject RS ssubst) 1 THEN
- unfold_thms_tac ctxt @{thms sum.inject Pair_eq conj_assoc} THEN rtac refl 1;
-
-val rec_like_unfold_thms =
- @{thms case_unit comp_def convol_def id_apply map_pair_def sum.simps(5,6) sum_map.simps
- split_conv};
-
-fun mk_rec_like_tac pre_map_defs map_ids rec_like_defs ctor_rec_like ctr_def ctxt =
- unfold_thms_tac ctxt (ctr_def :: ctor_rec_like :: rec_like_defs @ pre_map_defs @ map_ids @
- rec_like_unfold_thms) THEN unfold_thms_tac ctxt @{thms id_def} THEN rtac refl 1;
-
-val corec_like_ss = ss_only @{thms if_True if_False};
-val corec_like_unfold_thms = @{thms id_apply map_pair_def sum_map.simps prod.cases};
-
-fun mk_corec_like_tac corec_like_defs map_ids ctor_dtor_corec_like pre_map_def ctr_def ctxt =
- unfold_thms_tac ctxt (ctr_def :: corec_like_defs) THEN
- subst_tac ctxt [ctor_dtor_corec_like] 1 THEN asm_simp_tac corec_like_ss 1 THEN
- unfold_thms_tac ctxt (pre_map_def :: corec_like_unfold_thms @ map_ids) THEN
- unfold_thms_tac ctxt @{thms id_def} THEN
- TRY ((rtac refl ORELSE' subst_tac ctxt @{thms unit_eq} THEN' rtac refl) 1);
-
-fun mk_disc_corec_like_iff_tac case_splits' corec_likes discs ctxt =
- EVERY (map3 (fn case_split_tac => fn corec_like_thm => fn disc =>
- case_split_tac 1 THEN unfold_thms_tac ctxt [corec_like_thm] THEN
- asm_simp_tac (ss_only @{thms simp_thms(7,8,12,14,22,24)}) 1 THEN
- (if is_refl disc then all_tac else rtac disc 1))
- (map rtac case_splits' @ [K all_tac]) corec_likes discs);
-
-val solve_prem_prem_tac =
- REPEAT o (eresolve_tac @{thms bexE rev_bexI} ORELSE' rtac @{thm rev_bexI[OF UNIV_I]} ORELSE'
- hyp_subst_tac ORELSE' resolve_tac @{thms disjI1 disjI2}) THEN'
- (rtac refl ORELSE' atac ORELSE' rtac @{thm singletonI});
-
-val induct_prem_prem_thms =
- @{thms SUP_empty Sup_empty Sup_insert UN_insert Un_empty_left Un_empty_right Un_iff
- Union_Un_distrib collect_def[abs_def] image_def o_apply map_pair_simp
- mem_Collect_eq mem_UN_compreh_eq prod_set_simps sum_map.simps sum_set_simps};
-
-fun mk_induct_leverage_prem_prems_tac ctxt nn kks set_natural's pre_set_defs =
- EVERY' (maps (fn kk => [select_prem_tac nn (dtac meta_spec) kk, etac meta_mp,
- SELECT_GOAL (unfold_thms_tac ctxt (pre_set_defs @ set_natural's @ induct_prem_prem_thms)),
- solve_prem_prem_tac]) (rev kks)) 1;
-
-fun mk_induct_discharge_prem_tac ctxt nn n set_natural's pre_set_defs m k kks =
- let val r = length kks in
- EVERY' [select_prem_tac n (rotate_tac 1) k, rotate_tac ~1, hyp_subst_tac,
- REPEAT_DETERM_N m o (dtac meta_spec THEN' rotate_tac ~1)] 1 THEN
- EVERY [REPEAT_DETERM_N r
- (rotate_tac ~1 1 THEN dtac meta_mp 1 THEN rotate_tac 1 1 THEN prefer_tac 2),
- if r > 0 then PRIMITIVE Raw_Simplifier.norm_hhf else all_tac, atac 1,
- mk_induct_leverage_prem_prems_tac ctxt nn kks set_natural's pre_set_defs]
- end;
-
-fun mk_induct_tac ctxt ns mss kkss ctr_defs ctor_induct' set_natural's pre_set_defss =
- let
- val nn = length ns;
- val n = Integer.sum ns;
- in
- unfold_thms_tac ctxt ctr_defs THEN rtac ctor_induct' 1 THEN inst_spurious_fs_tac ctxt THEN
- EVERY (map4 (EVERY oooo map3 o mk_induct_discharge_prem_tac ctxt nn n set_natural's)
- pre_set_defss mss (unflat mss (1 upto n)) kkss)
- end;
-
-end;
--- a/src/HOL/Codatatype/Tools/bnf_gfp.ML Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,3002 +0,0 @@
-(* Title: HOL/BNF/Tools/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 bnf_gfp: mixfix list -> (string * sort) list option -> binding list ->
- typ list * typ list list -> BNF_Def.BNF list -> local_theory ->
- (term list * term list * term list * term list * thm * thm list * thm list * thm list *
- thm list * thm list) * local_theory
-end;
-
-structure BNF_GFP : BNF_GFP =
-struct
-
-open BNF_Def
-open BNF_Util
-open BNF_Tactics
-open BNF_FP
-open BNF_FP_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 bnf_gfp mixfixes resBs bs (resDs, Dss) bnfs lthy =
- let
- 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 b = Binding.name (mk_common_name (map Binding.name_of bs));
-
- (* TODO: check if m, n, etc., are sane *)
-
- val deads = fold (union (op =)) Dss resDs;
- val names_lthy = fold Variable.declare_typ deads lthy;
-
- (* tvars *)
- val ((((((((passiveAs, activeAs), allAs)), (passiveBs, activeBs)),
- (passiveCs, activeCs)), passiveXs), passiveYs), idxT) = names_lthy
- |> mk_TFrees live
- |> apfst (`(chop m))
- ||> mk_TFrees live
- ||>> apfst (chop m)
- ||> mk_TFrees live
- ||>> apfst (chop m)
- ||>> mk_TFrees m
- ||>> mk_TFrees m
- ||> fst o mk_TFrees 1
- ||> the_single;
-
- 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';
-
- (* typs *)
- val dead_poss =
- (case resBs of
- NONE => map SOME deads @ replicate m NONE
- | SOME Ts => map (fn T => if member (op =) deads (TFree T) then SOME (TFree T) else NONE) Ts);
- 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 all_fTs = map2 (fn T => fn U => T --> U) allAs allBs;
- 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);
- val sumBsAs = map2 (curry mk_sumT) activeBs activeAs;
- val sumFTs = mk_FTs (passiveAs @ sumBsAs);
- val sum_sTs = map2 (fn T => fn U => T --> U) activeAs sumFTs;
-
- (* terms *)
- val mapsAsAs = map4 mk_map_of_bnf Dss Ass Ass bnfs;
- val mapsAsBs = map4 mk_map_of_bnf Dss Ass Bss bnfs;
- val mapsBsCs' = map4 mk_map_of_bnf Dss Bss Css' bnfs;
- val mapsAsCs' = map4 mk_map_of_bnf Dss Ass Css' bnfs;
- val map_Inls = map4 mk_map_of_bnf Dss Bss (replicate n (passiveAs @ sumBsAs)) bnfs;
- val map_Inls_rev = map4 mk_map_of_bnf Dss (replicate n (passiveAs @ sumBsAs)) Bss bnfs;
- val map_fsts = map4 mk_map_of_bnf Dss (replicate n (passiveAs @ RTs)) Ass bnfs;
- val map_snds = map4 mk_map_of_bnf Dss (replicate n (passiveAs @ RTs)) Bss bnfs;
- fun mk_setss Ts = map3 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 = map4 mk_srel_of_bnf Dss Ass Bss bnfs;
- val bds = map3 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 emptys = map (fn T => HOLogic.mk_set T []) passiveAs;
- val Zeros = map (fn empty =>
- HOLogic.mk_tuple (map (fn U => absdummy U empty) activeAs)) emptys;
- val hrecTs = map fastype_of Zeros;
- val hsetTs = map (fn hrecT => Library.foldr (op -->) (sTs, HOLogic.natT --> hrecT)) hrecTs;
-
- val (((((((((((((((((((((((((((((((((((zs, zs'), zs_copy), zs_copy2),
- z's), As), As_copy), Bs), Bs_copy), B's), B''s), ss), sum_ss), s's), s''s), fs), fs_copy),
- self_fs), all_fs), gs), all_gs), xFs), xFs_copy), RFs), (Rtuple, Rtuple')), (hrecs, hrecs')),
- (nat, nat')), Rs), Rs_copy), R's), sRs), (idx, idx')), Idx), Ris), Kss),
- names_lthy) = lthy
- |> mk_Frees' "b" activeAs
- ||>> mk_Frees "b" activeAs
- ||>> mk_Frees "b" activeAs
- ||>> mk_Frees "b" activeBs
- ||>> mk_Frees "A" ATs
- ||>> mk_Frees "A" ATs
- ||>> mk_Frees "B" BTs
- ||>> mk_Frees "B" BTs
- ||>> mk_Frees "B'" B'Ts
- ||>> mk_Frees "B''" B''Ts
- ||>> mk_Frees "s" sTs
- ||>> mk_Frees "sums" sum_sTs
- ||>> mk_Frees "s'" s'Ts
- ||>> mk_Frees "s''" s''Ts
- ||>> mk_Frees "f" fTs
- ||>> mk_Frees "f" fTs
- ||>> mk_Frees "f" self_fTs
- ||>> mk_Frees "f" all_fTs
- ||>> mk_Frees "g" gTs
- ||>> mk_Frees "g" all_gTs
- ||>> mk_Frees "x" FTsAs
- ||>> mk_Frees "x" FTsAs
- ||>> mk_Frees "x" FRTs
- ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "Rtuple") all_sbisT
- ||>> mk_Frees' "rec" hrecTs
- ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "n") HOLogic.natT
- ||>> 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)
- ||>> mk_Freess "K" (map (fn AT => map (fn T => T --> AT) activeAs) ATs);
-
- val passive_UNIVs = map HOLogic.mk_UNIV passiveAs;
- val passive_diags = map mk_diag As;
- val active_UNIVs = map HOLogic.mk_UNIV activeAs;
- val sum_UNIVs = map HOLogic.mk_UNIV sumBsAs;
- val passive_ids = map HOLogic.id_const passiveAs;
- val active_ids = map HOLogic.id_const activeAs;
- val Inls = map2 Inl_const activeBs 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 bd_Cnotzeros = map bd_Cnotzero_of_bnf bnfs;
- val bd_Cnotzero = hd bd_Cnotzeros;
- val in_bds = map in_bd_of_bnf bnfs;
- val in_monos = map in_mono_of_bnf bnfs;
- val map_comps = map map_comp_of_bnf bnfs;
- val map_comp's = map map_comp'_of_bnf bnfs;
- val map_congs = map map_cong_of_bnf bnfs;
- val map_id's = map map_id'_of_bnf bnfs;
- val map_wpulls = map map_wpull_of_bnf bnfs;
- val set_bdss = map set_bd_of_bnf bnfs;
- val set_natural'ss = map set_natural'_of_bnf bnfs;
- val srel_congs = map srel_cong_of_bnf bnfs;
- val srel_converses = map srel_converse_of_bnf bnfs;
- val srel_defs = map srel_def_of_bnf bnfs;
- val srel_Grs = map srel_Gr_of_bnf bnfs;
- val srel_Ids = map srel_Id_of_bnf bnfs;
- val srel_monos = map srel_mono_of_bnf bnfs;
- val srel_Os = map srel_O_of_bnf bnfs;
- val srel_O_Grs = map srel_O_Gr_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_comp =
- 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;
- in
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (x :: fs @ all_gs) (mk_Trueprop_eq (lhs, rhs)))
- (K (mk_map_comp_id_tac map_comp))
- |> Thm.close_derivation
- end;
-
- val map_comp_id_thms = map5 mk_map_comp_id xFs mapsAsBs mapsBsCs' mapsAsCs' map_comp's;
-
- (*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_congL x mapAsAs sets map_cong 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 = map4 mk_prem (drop m sets) self_fs zs zs';
- val goal = mk_Trueprop_eq (Term.list_comb (mapAsAs, passive_ids @ self_fs) $ x, x);
- in
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (x :: self_fs) (Logic.list_implies (prems, goal)))
- (K (mk_map_congL_tac m map_cong map_id'))
- |> Thm.close_derivation
- end;
-
- val map_congL_thms = map5 mk_map_congL xFs mapsAsAs setssAs map_congs map_id's;
- 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) xFs xFs_copy;
- val maps = map (fn mapx => Term.list_comb (mapx, all_fs)) mapsAsBs;
- val concls =
- map3 (fn x => fn y => fn mapx => mk_Trueprop_eq (mapx $ x, mapx $ y)) xFs xFs_copy maps;
- val goals =
- map4 (fn prem => fn concl => fn x => fn y =>
- fold_rev Logic.all (x :: y :: all_fs) (Logic.mk_implies (prem, concl)))
- prems concls xFs xFs_copy;
- in
- map (fn goal => Skip_Proof.prove lthy [] [] goal
- (K ((hyp_subst_tac THEN' rtac refl) 1)) |> Thm.close_derivation) goals
- end;
-
- val timer = time (timer "Derived simple theorems");
-
- (* coalgebra *)
-
- val coalg_bind = Binding.suffix_name ("_" ^ coN ^ algN) b;
- val coalg_name = Binding.name_of coalg_bind;
- val coalg_def_bind = (Thm.def_binding coalg_bind, []);
-
- (*forall i = 1 ... n: (\<forall>x \<in> Bi. si \<in> Fi_in A1 .. Am B1 ... Bn)*)
- val coalg_spec =
- let
- val coalgT = Library.foldr (op -->) (ATs @ BTs @ sTs, HOLogic.boolT);
-
- val ins = map3 mk_in (replicate n (As @ 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 lhs = Term.list_comb (Free (coalg_name, coalgT), As @ Bs @ ss);
- val rhs = Library.foldr1 HOLogic.mk_conj (map5 mk_coalg_conjunct Bs ss ins zs zs')
- in
- mk_Trueprop_eq (lhs, rhs)
- end;
-
- val ((coalg_free, (_, coalg_def_free)), (lthy, lthy_old)) =
- lthy
- |> Specification.definition (SOME (coalg_bind, NONE, NoSyn), (coalg_def_bind, coalg_spec))
- ||> `Local_Theory.restore;
-
- val phi = Proof_Context.export_morphism lthy_old lthy;
- val coalg = fst (Term.dest_Const (Morphism.term phi coalg_free));
- val coalg_def = Morphism.thm phi coalg_def_free;
-
- fun mk_coalg As Bs ss =
- let
- val args = As @ 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 coalg_prem = HOLogic.mk_Trueprop (mk_coalg As Bs ss);
-
- val coalg_in_thms = map (fn i =>
- coalg_def RS @{thm subst[of _ _ "%x. x"]} RS mk_conjunctN n i RS bspec) ks
-
- val coalg_set_thmss =
- let
- val coalg_prem = HOLogic.mk_Trueprop (mk_coalg As Bs ss);
- fun mk_prem x B = HOLogic.mk_Trueprop (HOLogic.mk_mem (x, B));
- fun mk_concl s x B set = HOLogic.mk_Trueprop (mk_subset (set $ (s $ x)) B);
- val prems = map2 mk_prem zs Bs;
- val conclss = map3 (fn s => fn x => fn sets => map2 (mk_concl s x) (As @ Bs) sets)
- ss zs setssAs;
- val goalss = map3 (fn x => fn prem => fn concls => map (fn concl =>
- fold_rev Logic.all (x :: As @ Bs @ ss)
- (Logic.list_implies (coalg_prem :: [prem], concl))) concls) zs prems conclss;
- in
- map (fn goals => map (fn goal => Skip_Proof.prove lthy [] [] goal
- (K (mk_coalg_set_tac coalg_def)) |> Thm.close_derivation) goals) goalss
- end;
-
- val coalg_set_thmss' = transpose coalg_set_thmss;
-
- fun mk_tcoalg ATs BTs = mk_coalg (map HOLogic.mk_UNIV ATs) (map HOLogic.mk_UNIV BTs);
-
- val tcoalg_thm =
- let
- val goal = fold_rev Logic.all ss
- (HOLogic.mk_Trueprop (mk_tcoalg passiveAs activeAs ss))
- in
- Skip_Proof.prove lthy [] [] goal
- (K (stac coalg_def 1 THEN CONJ_WRAP
- (K (EVERY' [rtac ballI, rtac CollectI,
- CONJ_WRAP' (K (EVERY' [rtac @{thm subset_UNIV}])) allAs] 1)) ss))
- |> Thm.close_derivation
- end;
-
- val timer = time (timer "Coalgebra definition & thms");
-
- (* morphism *)
-
- val mor_bind = Binding.suffix_name ("_" ^ morN) b;
- val mor_name = Binding.name_of mor_bind;
- val mor_def_bind = (Thm.def_binding mor_bind, []);
-
- (*fbetw) forall i = 1 ... n: (\<forall>x \<in> Bi. fi x \<in> B'i)*)
- (*mor) forall i = 1 ... n: (\<forall>x \<in> Bi.
- Fi_map id ... id f1 ... fn (si x) = si' (fi x)*)
- val mor_spec =
- let
- val morT = Library.foldr (op -->) (BTs @ sTs @ B'Ts @ s'Ts @ fTs, HOLogic.boolT);
-
- 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 lhs = Term.list_comb (Free (mor_name, morT), Bs @ ss @ B's @ s's @ fs);
- val rhs = HOLogic.mk_conj
- (Library.foldr1 HOLogic.mk_conj (map5 mk_fbetw fs Bs B's zs zs'),
- Library.foldr1 HOLogic.mk_conj (map7 mk_mor Bs mapsAsBs fs ss s's zs zs'))
- in
- mk_Trueprop_eq (lhs, rhs)
- end;
-
- val ((mor_free, (_, mor_def_free)), (lthy, lthy_old)) =
- lthy
- |> Specification.definition (SOME (mor_bind, NONE, NoSyn), (mor_def_bind, mor_spec))
- ||> `Local_Theory.restore;
-
- val phi = Proof_Context.export_morphism lthy_old lthy;
- val mor = fst (Term.dest_Const (Morphism.term phi mor_free));
- val mor_def = 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 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 = fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs)
- (Logic.mk_implies (prem, HOLogic.mk_Trueprop (mk_subset (mk_image f $ B1) B2)));
- val image_goals = map3 mk_image_goal fs Bs B's;
- fun mk_elim_goal B mapAsBs f s s' x =
- fold_rev Logic.all (x :: Bs @ ss @ B's @ s's @ fs)
- (Logic.list_implies ([prem, HOLogic.mk_Trueprop (HOLogic.mk_mem (x, B))],
- mk_Trueprop_eq (Term.list_comb (mapAsBs, passive_ids @ fs @ [s $ x]), s' $ (f $ x))));
- val elim_goals = map6 mk_elim_goal Bs mapsAsBs fs ss s's zs;
- fun prove goal =
- Skip_Proof.prove lthy [] [] goal (K (mk_mor_elim_tac mor_def))
- |> Thm.close_derivation;
- 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_subset) Bs Bs_copy;
- val concl = HOLogic.mk_Trueprop (mk_mor Bs ss Bs_copy ss active_ids);
- in
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (Bs @ ss @ Bs_copy) (Logic.list_implies (prems, concl)))
- (K (mk_mor_incl_tac mor_def map_id's))
- |> Thm.close_derivation
- 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));
- in
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (Bs @ ss @ B's @ s's @ B''s @ s''s @ fs @ gs)
- (Logic.list_implies (prems, concl)))
- (K (mk_mor_comp_tac mor_def mor_image'_thms morE_thms map_comp_id_thms))
- |> Thm.close_derivation
- 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);
- in
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs @ fs_copy)
- (Logic.list_implies (prems, concl)))
- (K ((hyp_subst_tac THEN' atac) 1))
- |> Thm.close_derivation
- 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 (map4 mk_conjunct mapsAsBs fs ss s's);
- in
- Skip_Proof.prove lthy [] [] (fold_rev Logic.all (ss @ s's @ fs) (mk_Trueprop_eq (lhs, rhs)))
- (K (mk_mor_UNIV_tac morE_thms mor_def))
- |> Thm.close_derivation
- 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;
- in
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all ss (HOLogic.mk_Trueprop
- (mk_mor active_UNIVs ss (map HOLogic.mk_UNIV FTsAs) maps ss)))
- (K (mk_mor_str_tac ks mor_UNIV_thm))
- |> Thm.close_derivation
- end;
-
- val mor_sum_case_thm =
- let
- val maps = map3 (fn s => fn sum_s => fn mapx =>
- mk_sum_case (HOLogic.mk_comp (Term.list_comb (mapx, passive_ids @ Inls), s), sum_s))
- s's sum_ss map_Inls;
- in
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (s's @ sum_ss) (HOLogic.mk_Trueprop
- (mk_mor (map HOLogic.mk_UNIV activeBs) s's sum_UNIVs maps Inls)))
- (K (mk_mor_sum_case_tac ks mor_UNIV_thm))
- |> Thm.close_derivation
- end;
-
- val timer = time (timer "Morphism definition & thms");
-
- fun hset_rec_bind j = Binding.suffix_name ("_" ^ hset_recN ^ (if m = 1 then "" else
- string_of_int j)) b;
- val hset_rec_name = Binding.name_of o hset_rec_bind;
- val hset_rec_def_bind = rpair [] o Thm.def_binding o hset_rec_bind;
-
- fun hset_rec_spec j Zero hsetT hrec hrec' =
- let
- fun mk_Suc s setsAs z z' =
- let
- val (set, sets) = apfst (fn xs => nth xs (j - 1)) (chop m setsAs);
- fun mk_UN set k = mk_UNION (set $ (s $ z)) (mk_nthN n hrec k);
- in
- Term.absfree z'
- (mk_union (set $ (s $ z), Library.foldl1 mk_union (map2 mk_UN sets ks)))
- end;
-
- val Suc = Term.absdummy HOLogic.natT (Term.absfree hrec'
- (HOLogic.mk_tuple (map4 mk_Suc ss setssAs zs zs')));
-
- val lhs = Term.list_comb (Free (hset_rec_name j, hsetT), ss);
- val rhs = mk_nat_rec Zero Suc;
- in
- mk_Trueprop_eq (lhs, rhs)
- end;
-
- val ((hset_rec_frees, (_, hset_rec_def_frees)), (lthy, lthy_old)) =
- lthy
- |> fold_map5 (fn j => fn Zero => fn hsetT => fn hrec => fn hrec' => Specification.definition
- (SOME (hset_rec_bind j, NONE, NoSyn),
- (hset_rec_def_bind j, hset_rec_spec j Zero hsetT hrec hrec')))
- ls Zeros hsetTs hrecs hrecs'
- |>> apsnd split_list o split_list
- ||> `Local_Theory.restore;
-
- val phi = Proof_Context.export_morphism lthy_old lthy;
-
- val hset_rec_defs = map (Morphism.thm phi) hset_rec_def_frees;
- val hset_recs = map (fst o Term.dest_Const o Morphism.term phi) hset_rec_frees;
-
- fun mk_hset_rec ss nat i j T =
- let
- val args = ss @ [nat];
- val Ts = map fastype_of ss;
- val bTs = map domain_type Ts;
- val hrecT = HOLogic.mk_tupleT (map (fn U => U --> HOLogic.mk_setT T) bTs)
- val hset_recT = Library.foldr (op -->) (Ts, HOLogic.natT --> hrecT);
- in
- mk_nthN n (Term.list_comb (Const (nth hset_recs (j - 1), hset_recT), args)) i
- end;
-
- val hset_rec_0ss = mk_rec_simps n @{thm nat_rec_0} hset_rec_defs;
- val hset_rec_Sucss = mk_rec_simps n @{thm nat_rec_Suc} hset_rec_defs;
- val hset_rec_0ss' = transpose hset_rec_0ss;
- val hset_rec_Sucss' = transpose hset_rec_Sucss;
-
- fun hset_bind i j = Binding.suffix_name ("_" ^ hsetN ^
- (if m = 1 then "" else string_of_int j)) (nth bs (i - 1));
- val hset_name = Binding.name_of oo hset_bind;
- val hset_def_bind = rpair [] o Thm.def_binding oo hset_bind;
-
- fun hset_spec i j =
- let
- val U = nth activeAs (i - 1);
- val z = nth zs (i - 1);
- val T = nth passiveAs (j - 1);
- val setT = HOLogic.mk_setT T;
- val hsetT = Library.foldr (op -->) (sTs, U --> setT);
-
- val lhs = Term.list_comb (Free (hset_name i j, hsetT), ss @ [z]);
- val rhs = mk_UNION (HOLogic.mk_UNIV HOLogic.natT)
- (Term.absfree nat' (mk_hset_rec ss nat i j T $ z));
- in
- mk_Trueprop_eq (lhs, rhs)
- end;
-
- val ((hset_frees, (_, hset_def_frees)), (lthy, lthy_old)) =
- lthy
- |> fold_map (fn i => fold_map (fn j => Specification.definition
- (SOME (hset_bind i j, NONE, NoSyn), (hset_def_bind i j, hset_spec i j))) ls) ks
- |>> map (apsnd split_list o split_list)
- |>> apsnd split_list o split_list
- ||> `Local_Theory.restore;
-
- val phi = Proof_Context.export_morphism lthy_old lthy;
-
- val hset_defss = map (map (Morphism.thm phi)) hset_def_frees;
- val hset_defss' = transpose hset_defss;
- val hset_namess = map (map (fst o Term.dest_Const o Morphism.term phi)) hset_frees;
-
- fun mk_hset ss i j T =
- let
- val Ts = map fastype_of ss;
- val bTs = map domain_type Ts;
- val hsetT = Library.foldr (op -->) (Ts, nth bTs (i - 1) --> HOLogic.mk_setT T);
- in
- Term.list_comb (Const (nth (nth hset_namess (i - 1)) (j - 1), hsetT), ss)
- end;
-
- val hsetssAs = map (fn i => map2 (mk_hset ss i) ls passiveAs) ks;
-
- val (set_incl_hset_thmss, set_hset_incl_hset_thmsss) =
- let
- fun mk_set_incl_hset s x set hset = fold_rev Logic.all (x :: ss)
- (HOLogic.mk_Trueprop (mk_subset (set $ (s $ x)) (hset $ x)));
-
- fun mk_set_hset_incl_hset s x y set hset1 hset2 =
- fold_rev Logic.all (x :: y :: ss)
- (Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (x, set $ (s $ y))),
- HOLogic.mk_Trueprop (mk_subset (hset1 $ x) (hset2 $ y))));
-
- val set_incl_hset_goalss =
- map4 (fn s => fn x => fn sets => fn hsets =>
- map2 (mk_set_incl_hset s x) (take m sets) hsets)
- ss zs setssAs hsetssAs;
-
- (*xk : F(i)set(m+k) (si yi) ==> F(k)_hset(j) s1 ... sn xk <= F(i)_hset(j) s1 ... sn yi*)
- val set_hset_incl_hset_goalsss =
- map4 (fn si => fn yi => fn sets => fn hsetsi =>
- map3 (fn xk => fn set => fn hsetsk =>
- map2 (mk_set_hset_incl_hset si xk yi set) hsetsk hsetsi)
- zs_copy (drop m sets) hsetssAs)
- ss zs setssAs hsetssAs;
- in
- (map3 (fn goals => fn defs => fn rec_Sucs =>
- map3 (fn goal => fn def => fn rec_Suc =>
- Skip_Proof.prove lthy [] [] goal (K (mk_set_incl_hset_tac def rec_Suc))
- |> Thm.close_derivation)
- goals defs rec_Sucs)
- set_incl_hset_goalss hset_defss hset_rec_Sucss,
- map3 (fn goalss => fn defsi => fn rec_Sucs =>
- map3 (fn k => fn goals => fn defsk =>
- map4 (fn goal => fn defk => fn defi => fn rec_Suc =>
- Skip_Proof.prove lthy [] [] goal
- (K (mk_set_hset_incl_hset_tac n [defk, defi] rec_Suc k))
- |> Thm.close_derivation)
- goals defsk defsi rec_Sucs)
- ks goalss hset_defss)
- set_hset_incl_hset_goalsss hset_defss hset_rec_Sucss)
- end;
-
- val set_incl_hset_thmss' = transpose set_incl_hset_thmss;
- val set_hset_incl_hset_thmsss' = transpose (map transpose set_hset_incl_hset_thmsss);
- val set_hset_incl_hset_thmsss'' = map transpose set_hset_incl_hset_thmsss';
- val set_hset_thmss = map (map (fn thm => thm RS @{thm set_mp})) set_incl_hset_thmss;
- val set_hset_hset_thmsss = map (map (map (fn thm => thm RS @{thm set_mp})))
- set_hset_incl_hset_thmsss;
- val set_hset_thmss' = transpose set_hset_thmss;
- val set_hset_hset_thmsss' = transpose (map transpose set_hset_hset_thmsss);
-
- val set_incl_hin_thmss =
- let
- fun mk_set_incl_hin s x hsets1 set hsets2 T =
- fold_rev Logic.all (x :: ss @ As)
- (Logic.list_implies
- (map2 (fn hset => fn A => HOLogic.mk_Trueprop (mk_subset (hset $ x) A)) hsets1 As,
- HOLogic.mk_Trueprop (mk_subset (set $ (s $ x)) (mk_in As hsets2 T))));
-
- val set_incl_hin_goalss =
- map4 (fn s => fn x => fn sets => fn hsets =>
- map3 (mk_set_incl_hin s x hsets) (drop m sets) hsetssAs activeAs)
- ss zs setssAs hsetssAs;
- in
- map2 (map2 (fn goal => fn thms =>
- Skip_Proof.prove lthy [] [] goal (K (mk_set_incl_hin_tac thms))
- |> Thm.close_derivation))
- set_incl_hin_goalss set_hset_incl_hset_thmsss
- end;
-
- val hset_minimal_thms =
- let
- fun mk_passive_prem set s x K =
- Logic.all x (HOLogic.mk_Trueprop (mk_subset (set $ (s $ x)) (K $ x)));
-
- fun mk_active_prem s x1 K1 set x2 K2 =
- fold_rev Logic.all [x1, x2]
- (Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (x2, set $ (s $ x1))),
- HOLogic.mk_Trueprop (mk_subset (K2 $ x2) (K1 $ x1))));
-
- val premss = map2 (fn j => fn Ks =>
- map4 mk_passive_prem (map (fn xs => nth xs (j - 1)) setssAs) ss zs Ks @
- flat (map4 (fn sets => fn s => fn x1 => fn K1 =>
- map3 (mk_active_prem s x1 K1) (drop m sets) zs_copy Ks) setssAs ss zs Ks))
- ls Kss;
-
- val hset_rec_minimal_thms =
- let
- fun mk_conjunct j T i K x = mk_subset (mk_hset_rec ss nat i j T $ x) (K $ x);
- fun mk_concl j T Ks = list_all_free zs
- (Library.foldr1 HOLogic.mk_conj (map3 (mk_conjunct j T) ks Ks zs));
- val concls = map3 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 certify lthy) [Term.absfree nat' phi, nat]) concls;
- in
- map4 (fn goal => fn cts => fn hset_rec_0s => fn hset_rec_Sucs =>
- singleton (Proof_Context.export names_lthy lthy)
- (Skip_Proof.prove lthy [] [] goal
- (mk_hset_rec_minimal_tac m cts hset_rec_0s hset_rec_Sucs))
- |> Thm.close_derivation)
- goals ctss hset_rec_0ss' hset_rec_Sucss'
- end;
-
- fun mk_conjunct j T i K x = mk_subset (mk_hset ss i j T $ x) (K $ x);
- fun mk_concl j T Ks = Library.foldr1 HOLogic.mk_conj (map3 (mk_conjunct j T) ks Ks zs);
- val concls = map3 mk_concl ls passiveAs Kss;
-
- val goals = map3 (fn Ks => fn prems => fn concl =>
- fold_rev Logic.all (Ks @ ss @ zs)
- (Logic.list_implies (prems, HOLogic.mk_Trueprop concl))) Kss premss concls;
- in
- map3 (fn goal => fn hset_defs => fn hset_rec_minimal =>
- Skip_Proof.prove lthy [] [] goal
- (mk_hset_minimal_tac n hset_defs hset_rec_minimal)
- |> Thm.close_derivation)
- goals hset_defss' hset_rec_minimal_thms
- end;
-
- val mor_hset_thmss =
- let
- val mor_hset_rec_thms =
- let
- fun mk_conjunct j T i f x B =
- HOLogic.mk_imp (HOLogic.mk_mem (x, B), HOLogic.mk_eq
- (mk_hset_rec s's nat i j T $ (f $ x), mk_hset_rec ss nat i j T $ x));
-
- fun mk_concl j T = list_all_free zs
- (Library.foldr1 HOLogic.mk_conj (map4 (mk_conjunct j T) ks fs zs Bs));
- val concls = map2 mk_concl ls passiveAs;
-
- val ctss =
- map (fn phi => map (SOME o certify lthy) [Term.absfree nat' phi, nat]) concls;
-
- val goals = map (fn concl =>
- Logic.list_implies ([coalg_prem, mor_prem], HOLogic.mk_Trueprop concl)) concls;
- in
- map5 (fn j => fn goal => fn cts => fn hset_rec_0s => fn hset_rec_Sucs =>
- singleton (Proof_Context.export names_lthy lthy)
- (Skip_Proof.prove lthy [] [] goal
- (K (mk_mor_hset_rec_tac m n cts j hset_rec_0s hset_rec_Sucs
- morE_thms set_natural'ss coalg_set_thmss)))
- |> Thm.close_derivation)
- ls goals ctss hset_rec_0ss' hset_rec_Sucss'
- end;
-
- val mor_hset_rec_thmss = map (fn thm => map (fn i =>
- mk_specN n thm RS mk_conjunctN n i RS mp) ks) mor_hset_rec_thms;
-
- fun mk_prem x B = HOLogic.mk_Trueprop (HOLogic.mk_mem (x, B));
-
- fun mk_concl j T i f x =
- mk_Trueprop_eq (mk_hset s's i j T $ (f $ x), mk_hset ss i j T $ x);
-
- val goalss = map2 (fn j => fn T => map4 (fn i => fn f => fn x => fn B =>
- fold_rev Logic.all (x :: As @ Bs @ ss @ B's @ s's @ fs)
- (Logic.list_implies ([coalg_prem, mor_prem,
- mk_prem x B], mk_concl j T i f x))) ks fs zs Bs) ls passiveAs;
- in
- map3 (map3 (fn goal => fn hset_def => fn mor_hset_rec =>
- Skip_Proof.prove lthy [] [] goal
- (K (mk_mor_hset_tac hset_def mor_hset_rec))
- |> Thm.close_derivation))
- goalss hset_defss' mor_hset_rec_thmss
- end;
-
- val timer = time (timer "Hereditary sets");
-
- (* bisimulation *)
-
- val bis_bind = Binding.suffix_name ("_" ^ bisN) b;
- val bis_name = Binding.name_of bis_bind;
- val bis_def_bind = (Thm.def_binding bis_bind, []);
-
- fun mk_bis_le_conjunct R B1 B2 = mk_subset R (mk_Times (B1, B2));
- val bis_le = Library.foldr1 HOLogic.mk_conj (map3 mk_bis_le_conjunct Rs Bs B's)
-
- val bis_spec =
- let
- val bisT = Library.foldr (op -->) (ATs @ BTs @ sTs @ B'Ts @ s'Ts @ setRTs, HOLogic.boolT);
-
- 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 (As @ 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 lhs = Term.list_comb (Free (bis_name, bisT), As @ Bs @ ss @ B's @ s's @ Rs);
- val rhs = HOLogic.mk_conj
- (bis_le, Library.foldr1 HOLogic.mk_conj
- (map9 mk_bis Rs ss s's zs z's RFs map_fsts map_snds bis_setss))
- in
- mk_Trueprop_eq (lhs, rhs)
- end;
-
- val ((bis_free, (_, bis_def_free)), (lthy, lthy_old)) =
- lthy
- |> Specification.definition (SOME (bis_bind, NONE, NoSyn), (bis_def_bind, bis_spec))
- ||> `Local_Theory.restore;
-
- val phi = Proof_Context.export_morphism lthy_old lthy;
- val bis = fst (Term.dest_Const (Morphism.term phi bis_free));
- val bis_def = Morphism.thm phi bis_def_free;
-
- fun mk_bis As Bs1 ss1 Bs2 ss2 Rs =
- let
- val args = As @ 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 bis_cong_thm =
- let
- val prems = map HOLogic.mk_Trueprop
- (mk_bis As Bs ss B's s's Rs :: map2 (curry HOLogic.mk_eq) Rs_copy Rs)
- val concl = HOLogic.mk_Trueprop (mk_bis As Bs ss B's s's Rs_copy);
- in
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (As @ Bs @ ss @ B's @ s's @ Rs @ Rs_copy)
- (Logic.list_implies (prems, concl)))
- (K ((hyp_subst_tac THEN' atac) 1))
- |> Thm.close_derivation
- end;
-
- val bis_srel_thm =
- let
- fun mk_conjunct R s s' b1 b2 srel =
- list_all_free [b1, b2] (HOLogic.mk_imp
- (HOLogic.mk_mem (HOLogic.mk_prod (b1, b2), R),
- HOLogic.mk_mem (HOLogic.mk_prod (s $ b1, s' $ b2),
- Term.list_comb (srel, passive_diags @ Rs))));
-
- val rhs = HOLogic.mk_conj
- (bis_le, Library.foldr1 HOLogic.mk_conj
- (map6 mk_conjunct Rs ss s's zs z's relsAsBs))
- in
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (As @ Bs @ ss @ B's @ s's @ Rs)
- (mk_Trueprop_eq (mk_bis As Bs ss B's s's Rs, rhs)))
- (K (mk_bis_srel_tac m bis_def srel_O_Grs map_comp's map_congs set_natural'ss))
- |> Thm.close_derivation
- end;
-
- val bis_converse_thm =
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (As @ Bs @ ss @ B's @ s's @ Rs)
- (Logic.mk_implies
- (HOLogic.mk_Trueprop (mk_bis As Bs ss B's s's Rs),
- HOLogic.mk_Trueprop (mk_bis As B's s's Bs ss (map mk_converse Rs)))))
- (K (mk_bis_converse_tac m bis_srel_thm srel_congs srel_converses))
- |> Thm.close_derivation;
-
- val bis_O_thm =
- let
- val prems =
- [HOLogic.mk_Trueprop (mk_bis As Bs ss B's s's Rs),
- HOLogic.mk_Trueprop (mk_bis As B's s's B''s s''s R's)];
- val concl =
- HOLogic.mk_Trueprop (mk_bis As Bs ss B''s s''s (map2 (curry mk_rel_comp) Rs R's));
- in
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (As @ Bs @ ss @ B's @ s's @ B''s @ s''s @ Rs @ R's)
- (Logic.list_implies (prems, concl)))
- (K (mk_bis_O_tac m bis_srel_thm srel_congs srel_Os))
- |> Thm.close_derivation
- end;
-
- val bis_Gr_thm =
- let
- val concl =
- HOLogic.mk_Trueprop (mk_bis As Bs ss B's s's (map2 mk_Gr Bs fs));
- in
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (As @ Bs @ ss @ B's @ s's @ fs)
- (Logic.list_implies ([coalg_prem, mor_prem], concl)))
- (mk_bis_Gr_tac bis_srel_thm srel_Grs mor_image_thms morE_thms coalg_in_thms)
- |> Thm.close_derivation
- 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_diag_thm = bis_cong_thm OF ((mor_id_thm RSN (2, bis_Gr_thm)) ::
- replicate n @{thm diag_Gr});
-
- val bis_Union_thm =
- let
- val prem =
- HOLogic.mk_Trueprop (mk_Ball Idx
- (Term.absfree idx' (mk_bis As Bs ss B's s's (map (fn R => R $ idx) Ris))));
- val concl =
- HOLogic.mk_Trueprop (mk_bis As Bs ss B's s's (map (mk_UNION Idx) Ris));
- in
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (Idx :: As @ Bs @ ss @ B's @ s's @ Ris)
- (Logic.mk_implies (prem, concl)))
- (mk_bis_Union_tac bis_def in_mono'_thms)
- |> Thm.close_derivation
- end;
-
- (* self-bisimulation *)
-
- fun mk_sbis As Bs ss Rs = mk_bis As Bs ss Bs ss Rs;
-
- val sbis_prem = HOLogic.mk_Trueprop (mk_sbis As Bs ss sRs);
-
- (* largest self-bisimulation *)
-
- fun lsbis_bind i = Binding.suffix_name ("_" ^ lsbisN ^ (if n = 1 then "" else
- string_of_int i)) b;
- val lsbis_name = Binding.name_of o lsbis_bind;
- 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 As Bs ss sRs)));
-
- fun lsbis_spec i RT =
- let
- fun mk_lsbisT RT =
- Library.foldr (op -->) (map fastype_of (As @ Bs @ ss), RT);
- val lhs = Term.list_comb (Free (lsbis_name i, mk_lsbisT RT), As @ Bs @ ss);
- val rhs = mk_UNION all_sbis (Term.absfree Rtuple' (mk_nthN n Rtuple i));
- in
- mk_Trueprop_eq (lhs, rhs)
- end;
-
- val ((lsbis_frees, (_, lsbis_def_frees)), (lthy, lthy_old)) =
- lthy
- |> fold_map2 (fn i => fn RT => Specification.definition
- (SOME (lsbis_bind i, NONE, NoSyn), (lsbis_def_bind i, lsbis_spec i RT))) ks setsRTs
- |>> apsnd split_list o split_list
- ||> `Local_Theory.restore;
-
- val phi = Proof_Context.export_morphism lthy_old lthy;
-
- val lsbis_defs = map (Morphism.thm phi) lsbis_def_frees;
- val lsbiss = map (fst o Term.dest_Const o Morphism.term phi) lsbis_frees;
-
- fun mk_lsbis As Bs ss i =
- let
- val args = As @ 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 sbis_lsbis_thm =
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (As @ Bs @ ss)
- (HOLogic.mk_Trueprop (mk_sbis As Bs ss (map (mk_lsbis As Bs ss) ks))))
- (K (mk_sbis_lsbis_tac lsbis_defs bis_Union_thm bis_cong_thm))
- |> Thm.close_derivation;
-
- val lsbis_incl_thms = map (fn i => sbis_lsbis_thm RS
- (bis_def RS @{thm subst[of _ _ "%x. x"]} RS conjunct1 RS mk_conjunctN n i)) ks;
- val lsbisE_thms = map (fn i => (mk_specN 2 (sbis_lsbis_thm RS
- (bis_def RS @{thm subst[of _ _ "%x. x"]} RS conjunct2 RS mk_conjunctN n i))) RS mp) ks;
-
- val incl_lsbis_thms =
- let
- fun mk_concl i R = HOLogic.mk_Trueprop (mk_subset R (mk_lsbis As Bs ss i));
- val goals = map2 (fn i => fn R => fold_rev Logic.all (As @ Bs @ ss @ sRs)
- (Logic.mk_implies (sbis_prem, mk_concl i R))) ks sRs;
- in
- map3 (fn goal => fn i => fn def => Skip_Proof.prove lthy [] [] goal
- (K (mk_incl_lsbis_tac n i def)) |> Thm.close_derivation) goals ks lsbis_defs
- end;
-
- val equiv_lsbis_thms =
- let
- fun mk_concl i B = HOLogic.mk_Trueprop (mk_equiv B (mk_lsbis As Bs ss i));
- val goals = map2 (fn i => fn B => fold_rev Logic.all (As @ Bs @ ss)
- (Logic.mk_implies (coalg_prem, mk_concl i B))) ks Bs;
- in
- map3 (fn goal => fn l_incl => fn incl_l =>
- Skip_Proof.prove lthy [] [] goal
- (K (mk_equiv_lsbis_tac sbis_lsbis_thm l_incl incl_l
- bis_diag_thm bis_converse_thm bis_O_thm))
- |> Thm.close_derivation)
- goals lsbis_incl_thms incl_lsbis_thms
- end;
-
- val timer = time (timer "Bisimulations");
-
- (* bounds *)
-
- val (lthy, sbd, sbdT,
- sbd_card_order, sbd_Cinfinite, sbd_Cnotzero, sbd_Card_order, set_sbdss, in_sbds) =
- if n = 1
- then (lthy, sum_bd, sum_bdT,
- bd_card_order, bd_Cinfinite, bd_Cnotzero, bd_Card_order, set_bdss, in_bds)
- else
- let
- val sbdT_bind = Binding.suffix_name ("_" ^ sum_bdTN) b;
-
- val ((sbdT_name, (sbdT_glob_info, sbdT_loc_info)), lthy) =
- typedef false NONE (sbdT_bind, params, NoSyn)
- (HOLogic.mk_UNIV sum_bdT) NONE (EVERY' [rtac exI, rtac UNIV_I] 1) lthy;
-
- val sbdT = Type (sbdT_name, params');
- val Abs_sbdT = Const (#Abs_name sbdT_glob_info, sum_bdT --> sbdT);
-
- val sbd_bind = Binding.suffix_name ("_" ^ sum_bdN) b;
- val sbd_name = Binding.name_of sbd_bind;
- val sbd_def_bind = (Thm.def_binding sbd_bind, []);
-
- val sbd_spec = HOLogic.mk_Trueprop
- (HOLogic.mk_eq (Free (sbd_name, mk_relT (`I sbdT)), mk_dir_image sum_bd Abs_sbdT));
-
- val ((sbd_free, (_, sbd_def_free)), (lthy, lthy_old)) =
- lthy
- |> Specification.definition (SOME (sbd_bind, NONE, NoSyn), (sbd_def_bind, sbd_spec))
- ||> `Local_Theory.restore;
-
- val phi = Proof_Context.export_morphism lthy_old lthy;
-
- val sbd_def = 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);
-
- fun mk_sum_Cinfinite [thm] = thm
- | mk_sum_Cinfinite (thm :: thms) =
- @{thm Cinfinite_csum_strong} OF [thm, mk_sum_Cinfinite thms];
-
- val sum_Cinfinite = mk_sum_Cinfinite bd_Cinfinites;
- val sum_Card_order = sum_Cinfinite RS conjunct2;
-
- fun mk_sum_card_order [thm] = thm
- | mk_sum_card_order (thm :: thms) =
- @{thm card_order_csum} OF [thm, mk_sum_card_order thms];
-
- val sum_card_order = mk_sum_card_order bd_card_orders;
-
- val sbd_ordIso = fold_thms lthy [sbd_def]
- (@{thm dir_image} OF [Abs_sbdT_inj, sum_Card_order]);
- val sbd_card_order = fold_thms lthy [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_Cnotzero = sbd_Cinfinite RS @{thm Cinfinite_Cnotzero};
- 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 = map3 mk_set_sbd ks bd_Card_orders set_bdss;
-
- fun mk_in_sbd 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 _ csum_Cnotzero2[OF ctwo_Cnotzero]]]}));
- val in_sbds = map4 mk_in_sbd ks bd_Card_orders bd_Cnotzeros in_bds;
- in
- (lthy, sbd, sbdT,
- sbd_card_order, sbd_Cinfinite, sbd_Cnotzero, sbd_Card_order, set_sbdss, in_sbds)
- end;
-
- fun mk_sbd_sbd 1 = sbd_Card_order RS @{thm ordIso_refl}
- | mk_sbd_sbd n = @{thm csum_absorb1} OF
- [sbd_Cinfinite, mk_sbd_sbd (n - 1) RS @{thm ordIso_imp_ordLeq}];
-
- val sbd_sbd_thm = mk_sbd_sbd n;
-
- val sbdTs = replicate n sbdT;
- val sum_sbd = Library.foldr1 (uncurry mk_csum) (replicate n sbd);
- 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 = map4 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 = map4 mk_map_of_bnf Dss (replicate n bdTs) (replicate n treeTs) bnfs;
- val final_maps = map4 mk_map_of_bnf Dss (replicate n treeTs) (replicate n treeQTs) bnfs;
- val tree_setss = mk_setss treeTs;
- 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 LevT = Library.foldr (op -->) (sTs, HOLogic.natT --> Lev_recT);
-
- val Nil = HOLogic.mk_tuple (map3 (fn i => fn z => fn z'=>
- Term.absfree z' (mk_InN activeAs z i)) ks zs zs');
- val rv_recT = fastype_of Nil;
- val rvT = Library.foldr (op -->) (sTs, sum_sbd_listT --> rv_recT);
-
- val (((((((((((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')),
- names_lthy) = names_lthy
- |> 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 *)
-
- fun isNode_bind i = Binding.suffix_name ("_" ^ isNodeN ^ (if n = 1 then "" else
- string_of_int i)) b;
- val isNode_name = Binding.name_of o isNode_bind;
- val isNode_def_bind = rpair [] o Thm.def_binding o isNode_bind;
-
- val isNodeT =
- Library.foldr (op -->) (map fastype_of (As @ [Kl, lab, kl]), HOLogic.boolT);
-
- val Succs = map3 (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 (passive_sets, active_sets) = chop m (map (fn set => set $ x) sets);
- val lhs = Term.list_comb (Free (isNode_name i, isNodeT), As @ [Kl, lab, kl]);
- val rhs = list_exists_free [x]
- (Library.foldr1 HOLogic.mk_conj (HOLogic.mk_eq (lab $ kl, mk_InN bdFTs x i) ::
- map2 mk_subset passive_sets As @ map2 (curry HOLogic.mk_eq) active_sets Succs));
- in
- mk_Trueprop_eq (lhs, rhs)
- end;
-
- val ((isNode_frees, (_, isNode_def_frees)), (lthy, lthy_old)) =
- lthy
- |> fold_map3 (fn i => fn x => fn sets => Specification.definition
- (SOME (isNode_bind i, NONE, NoSyn), (isNode_def_bind i, isNode_spec sets x i)))
- ks xs isNode_setss
- |>> apsnd split_list o split_list
- ||> `Local_Theory.restore;
-
- val phi = Proof_Context.export_morphism lthy_old lthy;
-
- val isNode_defs = map (Morphism.thm phi) isNode_def_frees;
- val isNodes = map (fst o Term.dest_Const o Morphism.term phi) isNode_frees;
-
- fun mk_isNode As kl i =
- Term.list_comb (Const (nth isNodes (i - 1), isNodeT), As @ [Kl, lab, kl]);
-
- val isTree =
- let
- val empty = HOLogic.mk_mem (HOLogic.mk_list sum_sbdT [], Kl);
- val Field = mk_subset Kl (mk_Field (mk_clists sum_sbd));
- val prefCl = mk_prefCl Kl;
-
- val tree = mk_Ball Kl (Term.absfree kl'
- (HOLogic.mk_conj
- (Library.foldr1 HOLogic.mk_disj (map (mk_isNode As kl) ks),
- Library.foldr1 HOLogic.mk_conj (map4 (fn Succ => fn i => fn k => fn k' =>
- mk_Ball Succ (Term.absfree k' (mk_isNode As
- (mk_append (kl, HOLogic.mk_list sum_sbdT [mk_InN sbdTs k i])) i)))
- Succs ks kks kks'))));
-
- val undef = list_all_free [kl] (HOLogic.mk_imp
- (HOLogic.mk_not (HOLogic.mk_mem (kl, Kl)),
- HOLogic.mk_eq (lab $ kl, mk_undefined sbdFT)));
- in
- Library.foldr1 HOLogic.mk_conj [empty, Field, prefCl, tree, undef]
- end;
-
- fun carT_bind i = Binding.suffix_name ("_" ^ carTN ^ (if n = 1 then "" else
- string_of_int i)) b;
- val carT_name = Binding.name_of o carT_bind;
- val carT_def_bind = rpair [] o Thm.def_binding o carT_bind;
-
- fun carT_spec i =
- let
- val carTT = Library.foldr (op -->) (ATs, HOLogic.mk_setT treeT);
-
- val lhs = Term.list_comb (Free (carT_name i, carTT), As);
- val rhs = 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 As (HOLogic.mk_list sum_sbdT []) i))));
- in
- mk_Trueprop_eq (lhs, rhs)
- end;
-
- val ((carT_frees, (_, carT_def_frees)), (lthy, lthy_old)) =
- lthy
- |> fold_map (fn i => Specification.definition
- (SOME (carT_bind i, NONE, NoSyn), (carT_def_bind i, carT_spec i))) ks
- |>> apsnd split_list o split_list
- ||> `Local_Theory.restore;
-
- val phi = Proof_Context.export_morphism lthy_old lthy;
-
- val carT_defs = map (Morphism.thm phi) carT_def_frees;
- val carTs = map (fst o Term.dest_Const o Morphism.term phi) carT_frees;
-
- fun mk_carT As i = Term.list_comb
- (Const (nth carTs (i - 1),
- Library.foldr (op -->) (map fastype_of As, HOLogic.mk_setT treeT)), As);
-
- fun strT_bind i = Binding.suffix_name ("_" ^ strTN ^ (if n = 1 then "" else
- string_of_int i)) b;
- val strT_name = Binding.name_of o strT_bind;
- val strT_def_bind = rpair [] o Thm.def_binding o strT_bind;
-
- fun strT_spec mapFT FT i =
- let
- val strTT = treeT --> FT;
-
- 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 @ map3 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);
- val lhs = Free (strT_name i, strTT);
- val rhs = HOLogic.mk_split (Term.absfree Kl' (Term.absfree lab'
- (mk_sum_caseN fs $ (lab $ HOLogic.mk_list sum_sbdT []))));
- in
- mk_Trueprop_eq (lhs, rhs)
- end;
-
- val ((strT_frees, (_, strT_def_frees)), (lthy, lthy_old)) =
- lthy
- |> fold_map3 (fn i => fn mapFT => fn FT => Specification.definition
- (SOME (strT_bind i, NONE, NoSyn), (strT_def_bind i, strT_spec mapFT FT i)))
- ks tree_maps treeFTs
- |>> apsnd split_list o split_list
- ||> `Local_Theory.restore;
-
- val phi = Proof_Context.export_morphism lthy_old lthy;
-
- val strT_defs = map ((fn def => trans OF [def RS fun_cong, @{thm prod.cases}]) o
- Morphism.thm phi) 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 As) ks;
- val carTAs_copy = map (mk_carT As_copy) ks;
- val strTAs = map2 mk_strT treeFTs ks;
- val hset_strTss = map (fn i => map2 (mk_hset strTAs i) ls passiveAs) ks;
-
- val coalgT_thm =
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all As (HOLogic.mk_Trueprop (mk_coalg As carTAs strTAs)))
- (mk_coalgT_tac m (coalg_def :: isNode_defs @ carT_defs) strT_defs set_natural'ss)
- |> Thm.close_derivation;
-
- val card_of_carT_thms =
- let
- val lhs = mk_card_of
- (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)), isTree))));
- val rhs = mk_cexp
- (if m = 0 then ctwo else
- (mk_csum (Library.foldr1 (uncurry mk_csum) (map mk_card_of As)) ctwo))
- (mk_cexp sbd sbd);
- val card_of_carT =
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all As (HOLogic.mk_Trueprop (mk_ordLeq lhs rhs)))
- (K (mk_card_of_carT_tac m isNode_defs sbd_sbd_thm
- sbd_card_order sbd_Card_order sbd_Cinfinite sbd_Cnotzero in_sbds))
- |> Thm.close_derivation
- in
- map (fn def => @{thm ordLeq_transitive[OF
- card_of_mono1[OF ord_eq_le_trans[OF _ Collect_restrict']]]} OF [def, card_of_carT])
- carT_defs
- end;
-
- val carT_set_thmss =
- let
- val Kl_lab = HOLogic.mk_prod (Kl, lab);
- fun mk_goal carT strT set k i =
- fold_rev Logic.all (sumx :: Kl :: lab :: k :: kl :: As)
- (Logic.list_implies (map HOLogic.mk_Trueprop
- [HOLogic.mk_mem (Kl_lab, carT), HOLogic.mk_mem (mk_Cons sumx kl, Kl),
- HOLogic.mk_eq (sumx, mk_InN sbdTs k i)],
- HOLogic.mk_Trueprop (HOLogic.mk_mem
- (HOLogic.mk_prod (mk_Shift Kl sumx, mk_shift lab sumx),
- set $ (strT $ Kl_lab)))));
-
- val goalss = map3 (fn carT => fn strT => fn sets =>
- map3 (mk_goal carT strT) (drop m sets) kks ks) carTAs strTAs tree_setss;
- in
- map6 (fn i => fn goals =>
- fn carT_def => fn strT_def => fn isNode_def => fn set_naturals =>
- map2 (fn goal => fn set_natural =>
- Skip_Proof.prove lthy [] [] goal
- (mk_carT_set_tac n i carT_def strT_def isNode_def set_natural)
- |> Thm.close_derivation)
- goals (drop m set_naturals))
- ks goalss carT_defs strT_defs isNode_defs set_natural'ss
- end;
-
- val carT_set_thmss' = transpose carT_set_thmss;
-
- val isNode_hset_thmss =
- let
- val Kl_lab = HOLogic.mk_prod (Kl, lab);
- fun mk_Kl_lab carT = HOLogic.mk_mem (Kl_lab, carT);
-
- val strT_hset_thmsss =
- let
- val strT_hset_thms =
- let
- fun mk_lab_kl i x = HOLogic.mk_eq (lab $ kl, mk_InN bdFTs x i);
-
- fun mk_inner_conjunct j T i x set i' carT =
- HOLogic.mk_imp (HOLogic.mk_conj (mk_Kl_lab carT, mk_lab_kl i x),
- mk_subset (set $ x) (mk_hset strTAs i' j T $ Kl_lab));
-
- fun mk_conjunct j T i x set =
- Library.foldr1 HOLogic.mk_conj (map2 (mk_inner_conjunct j T i x set) ks carTAs);
-
- fun mk_concl j T = list_all_free (Kl :: lab :: xs @ As)
- (HOLogic.mk_imp (HOLogic.mk_mem (kl, Kl),
- Library.foldr1 HOLogic.mk_conj (map3 (mk_conjunct j T)
- ks xs (map (fn xs => nth xs (j - 1)) isNode_setss))));
- val concls = map2 mk_concl ls passiveAs;
-
- val cTs = [SOME (certifyT lthy sum_sbdT)];
- val arg_cong_cTs = map (SOME o certifyT lthy) treeFTs;
- val ctss =
- map (fn phi => map (SOME o certify lthy) [Term.absfree kl' phi, kl]) concls;
-
- val goals = map HOLogic.mk_Trueprop concls;
- in
- map5 (fn j => fn goal => fn cts => fn set_incl_hsets => fn set_hset_incl_hsetss =>
- singleton (Proof_Context.export names_lthy lthy)
- (Skip_Proof.prove lthy [] [] goal
- (K (mk_strT_hset_tac n m j arg_cong_cTs cTs cts
- carT_defs strT_defs isNode_defs
- set_incl_hsets set_hset_incl_hsetss coalg_set_thmss' carT_set_thmss'
- coalgT_thm set_natural'ss)))
- |> Thm.close_derivation)
- ls goals ctss set_incl_hset_thmss' set_hset_incl_hset_thmsss''
- end;
-
- val strT_hset'_thms = map (fn thm => mk_specN (2 + n + m) thm RS mp) strT_hset_thms;
- in
- map (fn thm => map (fn i => map (fn i' =>
- thm RS mk_conjunctN n i RS mk_conjunctN n i' RS mp) ks) ks) strT_hset'_thms
- end;
-
- val carT_prems = map (fn carT =>
- HOLogic.mk_Trueprop (HOLogic.mk_mem (Kl_lab, carT))) carTAs_copy;
- val prem = HOLogic.mk_Trueprop (HOLogic.mk_mem (kl, Kl));
- val in_prems = map (fn hsets =>
- HOLogic.mk_Trueprop (HOLogic.mk_mem (Kl_lab, mk_in As hsets treeT))) hset_strTss;
- val isNode_premss = replicate n (map (HOLogic.mk_Trueprop o mk_isNode As_copy kl) ks);
- val conclss = replicate n (map (HOLogic.mk_Trueprop o mk_isNode As kl) ks);
- in
- map5 (fn carT_prem => fn isNode_prems => fn in_prem => fn concls => fn strT_hset_thmss =>
- map4 (fn isNode_prem => fn concl => fn isNode_def => fn strT_hset_thms =>
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (Kl :: lab :: kl :: As @ As_copy)
- (Logic.list_implies ([carT_prem, prem, isNode_prem, in_prem], concl)))
- (mk_isNode_hset_tac n isNode_def strT_hset_thms)
- |> Thm.close_derivation)
- isNode_prems concls isNode_defs
- (if m = 0 then replicate n [] else transpose strT_hset_thmss))
- carT_prems isNode_premss in_prems conclss
- (if m = 0 then replicate n [] else transpose (map transpose strT_hset_thmsss))
- end;
-
- 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 to_sbd_thmss = mk_to_sbd_thmss @{thm toCard};
- val from_to_sbd_thmss = mk_to_sbd_thmss @{thm fromCard_toCard};
-
- val Lev_bind = Binding.suffix_name ("_" ^ LevN) b;
- val Lev_name = Binding.name_of Lev_bind;
- 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 (map3 mk_set ks sets zs_copy))
- end;
-
- val Suc = Term.absdummy HOLogic.natT (Term.absfree Lev_rec'
- (HOLogic.mk_tuple (map5 mk_Suc ks ss setssAs zs zs')));
-
- val lhs = Term.list_comb (Free (Lev_name, LevT), ss);
- val rhs = mk_nat_rec Zero Suc;
- in
- mk_Trueprop_eq (lhs, rhs)
- end;
-
- val ((Lev_free, (_, Lev_def_free)), (lthy, lthy_old)) =
- lthy
- |> Specification.definition (SOME (Lev_bind, NONE, NoSyn), (Lev_def_bind, Lev_spec))
- ||> `Local_Theory.restore;
-
- val phi = Proof_Context.export_morphism lthy_old lthy;
-
- val Lev_def = 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 nat_rec_0} [Lev_def]);
- val Lev_Sucs = flat (mk_rec_simps n @{thm nat_rec_Suc} [Lev_def]);
-
- val rv_bind = Binding.suffix_name ("_" ^ rvN) b;
- val rv_name = Binding.name_of rv_bind;
- 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_sum_caseN (map mk_case ks) $ sumx)
- end;
-
- val Cons = Term.absfree sumx' (Term.absdummy sum_sbd_listT (Term.absfree rv_rec'
- (HOLogic.mk_tuple (map4 mk_Cons ks ss zs zs'))));
-
- val lhs = Term.list_comb (Free (rv_name, rvT), ss);
- val rhs = mk_list_rec Nil Cons;
- in
- mk_Trueprop_eq (lhs, rhs)
- end;
-
- val ((rv_free, (_, rv_def_free)), (lthy, lthy_old)) =
- lthy
- |> Specification.definition (SOME (rv_bind, NONE, NoSyn), (rv_def_bind, rv_spec))
- ||> `Local_Theory.restore;
-
- val phi = Proof_Context.export_morphism lthy_old lthy;
-
- val rv_def = 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 list_rec_Nil} [rv_def]);
- val rv_Conss = flat (mk_rec_simps n @{thm list_rec_Cons} [rv_def]);
-
- fun beh_bind i = Binding.suffix_name ("_" ^ behN ^ (if n = 1 then "" else
- string_of_int i)) b;
- val beh_name = Binding.name_of o beh_bind;
- val beh_def_bind = rpair [] o Thm.def_binding o beh_bind;
-
- fun beh_spec i z =
- let
- val mk_behT = Library.foldr (op -->) (map fastype_of (ss @ [z]), treeT);
-
- 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_If
- (HOLogic.mk_mem (kl, mk_Lev ss (mk_size kl) i $ z))
- (mk_sum_caseN (map5 mk_case ks to_sbd_maps ss zs zs') $ (mk_rv ss kl i $ z))
- (mk_undefined sbdFT));
-
- val lhs = Term.list_comb (Free (beh_name i, mk_behT), ss) $ z;
- val rhs = HOLogic.mk_prod (mk_UNION (HOLogic.mk_UNIV HOLogic.natT)
- (Term.absfree nat' (mk_Lev ss nat i $ z)), Lab);
- in
- mk_Trueprop_eq (lhs, rhs)
- end;
-
- val ((beh_frees, (_, beh_def_frees)), (lthy, lthy_old)) =
- lthy
- |> fold_map2 (fn i => fn z => Specification.definition
- (SOME (beh_bind i, NONE, NoSyn), (beh_def_bind i, beh_spec i z))) ks zs
- |>> apsnd split_list o split_list
- ||> `Local_Theory.restore;
-
- val phi = Proof_Context.export_morphism lthy_old lthy;
-
- val beh_defs = map (Morphism.thm phi) 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 Lev_sbd_thms =
- let
- fun mk_conjunct i z = mk_subset (mk_Lev ss nat i $ z) (mk_Field (mk_clists sum_sbd));
- val goal = list_all_free zs
- (Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct ks zs));
-
- val cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
-
- val Lev_sbd = singleton (Proof_Context.export names_lthy lthy)
- (Skip_Proof.prove lthy [] [] (HOLogic.mk_Trueprop goal)
- (K (mk_Lev_sbd_tac cts Lev_0s Lev_Sucs to_sbd_thmss))
- |> Thm.close_derivation);
-
- val Lev_sbd' = mk_specN n Lev_sbd;
- in
- map (fn i => Lev_sbd' RS mk_conjunctN n i) ks
- end;
-
- 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 cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
-
- val length_Lev = singleton (Proof_Context.export names_lthy lthy)
- (Skip_Proof.prove lthy [] [] (HOLogic.mk_Trueprop goal)
- (K (mk_length_Lev_tac cts Lev_0s Lev_Sucs))
- |> Thm.close_derivation);
-
- 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 = fold_rev Logic.all (z :: kl :: nat :: ss) (Logic.mk_implies
- (HOLogic.mk_Trueprop (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z)),
- HOLogic.mk_Trueprop (HOLogic.mk_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 =>
- Skip_Proof.prove lthy [] [] goal (K (mk_length_Lev'_tac length_Lev))
- |> Thm.close_derivation) goals length_Levs;
- in
- (length_Levs, length_Levs')
- end;
-
- val prefCl_Lev_thms =
- let
- fun mk_conjunct i z = HOLogic.mk_imp
- (HOLogic.mk_conj (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z), mk_subset kl_copy kl),
- HOLogic.mk_mem (kl_copy, mk_Lev ss (mk_size kl_copy) i $ z));
- val goal = list_all_free (kl :: kl_copy :: zs)
- (Library.foldr1 HOLogic.mk_conj (map2 mk_conjunct ks zs));
-
- val cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
-
- val prefCl_Lev = singleton (Proof_Context.export names_lthy lthy)
- (Skip_Proof.prove lthy [] [] (HOLogic.mk_Trueprop goal)
- (K (mk_prefCl_Lev_tac cts Lev_0s Lev_Sucs)))
- |> Thm.close_derivation;
-
- val prefCl_Lev' = mk_specN (n + 2) prefCl_Lev;
- in
- map (fn i => prefCl_Lev' RS mk_conjunctN n i RS mp) ks
- 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 cTs = [SOME (certifyT lthy sum_sbdT)];
- val cts = map (SOME o certify lthy) [Term.absfree kl' goal, kl];
-
- val rv_last = singleton (Proof_Context.export names_lthy lthy)
- (Skip_Proof.prove lthy [] [] (HOLogic.mk_Trueprop goal)
- (K (mk_rv_last_tac cTs cts rv_Nils rv_Conss)))
- |> Thm.close_derivation;
-
- 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_rv_Lev_thmsss = if m = 0 then replicate n (replicate n []) else
- let
- fun mk_case s sets z z_free = Term.absfree z_free (Library.foldr1 HOLogic.mk_conj
- (map2 (fn set => fn A => mk_subset (set $ (s $ z)) A) (take m sets) As));
-
- fun mk_conjunct i z B = HOLogic.mk_imp
- (HOLogic.mk_conj (HOLogic.mk_mem (kl, mk_Lev ss nat i $ z), HOLogic.mk_mem (z, B)),
- mk_sum_caseN (map4 mk_case ss setssAs zs zs') $ (mk_rv ss kl i $ z));
-
- val goal = list_all_free (kl :: zs)
- (Library.foldr1 HOLogic.mk_conj (map3 mk_conjunct ks zs Bs));
-
- val cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
-
- val set_rv_Lev = singleton (Proof_Context.export names_lthy lthy)
- (Skip_Proof.prove lthy [] []
- (Logic.mk_implies (coalg_prem, HOLogic.mk_Trueprop goal))
- (K (mk_set_rv_Lev_tac m cts Lev_0s Lev_Sucs rv_Nils rv_Conss
- coalg_set_thmss from_to_sbd_thmss)))
- |> Thm.close_derivation;
-
- val set_rv_Lev' = mk_specN (n + 1) set_rv_Lev;
- in
- map (fn i => map (fn i' =>
- split_conj_thm (if n = 1 then set_rv_Lev' RS mk_conjunctN n i RS mp
- else set_rv_Lev' RS mk_conjunctN n i RS mp RSN
- (2, @{thm sum_case_weak_cong} RS @{thm subst[of _ _ "%x. x"]}) RS
- (mk_sum_casesN n i' RS @{thm subst[of _ _ "%x. x"]}))) 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 (map3 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 (map4 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 cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
-
- val set_Lev = singleton (Proof_Context.export names_lthy lthy)
- (Skip_Proof.prove lthy [] [] (HOLogic.mk_Trueprop goal)
- (K (mk_set_Lev_tac cts Lev_0s Lev_Sucs rv_Nils rv_Conss from_to_sbd_thmss)))
- |> Thm.close_derivation;
-
- 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 (map4 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 cts = map (SOME o certify lthy) [Term.absfree nat' goal, nat];
-
- val set_image_Lev = singleton (Proof_Context.export names_lthy lthy)
- (Skip_Proof.prove lthy [] [] (HOLogic.mk_Trueprop goal)
- (K (mk_set_image_Lev_tac cts Lev_0s Lev_Sucs rv_Nils rv_Conss
- from_to_sbd_thmss to_sbd_inj_thmss)))
- |> Thm.close_derivation;
-
- 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 =
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (As @ Bs @ ss) (Logic.mk_implies (coalg_prem,
- HOLogic.mk_Trueprop (mk_mor Bs ss carTAs strTAs (map (mk_beh ss) ks)))))
- (mk_mor_beh_tac 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 Lev_sbd_thms
- length_Lev_thms length_Lev'_thms prefCl_Lev_thms rv_last_thmss
- set_rv_Lev_thmsss set_Lev_thmsss set_image_Lev_thmsss
- set_natural'ss coalg_set_thmss map_comp_id_thms map_congs map_arg_cong_thms)
- |> Thm.close_derivation;
-
- val timer = time (timer "Behavioral morphism");
-
- fun mk_LSBIS As i = mk_lsbis As (map (mk_carT As) ks) strTAs i;
- fun mk_car_final As i =
- mk_quotient (mk_carT As i) (mk_LSBIS As i);
- fun mk_str_final As i =
- mk_univ (HOLogic.mk_comp (Term.list_comb (nth final_maps (i - 1),
- passive_ids @ map (mk_proj o mk_LSBIS As) ks), nth strTAs (i - 1)));
-
- val car_finalAs = map (mk_car_final As) ks;
- val str_finalAs = map (mk_str_final As) ks;
- val car_finals = map (mk_car_final passive_UNIVs) ks;
- val str_finals = map (mk_str_final passive_UNIVs) 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 =
- fold_rev Logic.all As (HOLogic.mk_Trueprop
- (mk_congruent R (HOLogic.mk_comp
- (Term.list_comb (final_map, passive_ids @ map (mk_proj o mk_LSBIS As) ks), strT))));
-
- val goals = map3 mk_goal (map (mk_LSBIS As) ks) final_maps strTAs;
- in
- map4 (fn goal => fn lsbisE => fn map_comp_id => fn map_cong =>
- Skip_Proof.prove lthy [] [] goal
- (K (mk_congruent_str_final_tac m lsbisE map_comp_id map_cong equiv_LSBIS_thms))
- |> Thm.close_derivation)
- goals lsbisE_thms map_comp_id_thms map_congs
- end;
-
- val coalg_final_thm = Skip_Proof.prove lthy [] [] (fold_rev Logic.all As
- (HOLogic.mk_Trueprop (mk_coalg As car_finalAs str_finalAs)))
- (K (mk_coalg_final_tac m coalg_def congruent_str_final_thms equiv_LSBIS_thms
- set_natural'ss coalgT_set_thmss))
- |> Thm.close_derivation;
-
- val mor_T_final_thm = Skip_Proof.prove lthy [] [] (fold_rev Logic.all As
- (HOLogic.mk_Trueprop (mk_mor carTAs strTAs car_finalAs str_finalAs
- (map (mk_proj o mk_LSBIS As) ks))))
- (K (mk_mor_T_final_tac mor_def congruent_str_final_thms equiv_LSBIS_thms))
- |> Thm.close_derivation;
-
- val mor_final_thm = mor_comp_thm OF [mor_beh_thm, mor_T_final_thm];
- val in_car_final_thms = map (fn mor_image' => mor_image' OF
- [tcoalg_thm RS 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_map4 (fn b => fn mx => fn car_final => fn in_car_final =>
- typedef false NONE (b, params, mx) car_final NONE
- (EVERY' [rtac exI, rtac 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 Ts'' = mk_Ts passiveCs;
- 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 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);
- val prodTs = map (HOLogic.mk_prodT o `I) Ts;
- val prodFTs = mk_FTs (passiveAs @ prodTs);
- val FTs_setss = mk_setss (passiveAs @ Ts);
- val prodFT_setss = mk_setss (passiveAs @ prodTs);
- val map_FTs = map2 (fn Ds => mk_map_of_bnf Ds treeQTs (passiveAs @ Ts)) Dss bnfs;
- val map_FT_nths = map2 (fn Ds =>
- mk_map_of_bnf Ds (passiveAs @ prodTs) (passiveAs @ Ts)) Dss bnfs;
- val fstsTs = map fst_const prodTs;
- val sndsTs = map snd_const prodTs;
- val dtorTs = map2 (curry (op -->)) Ts FTs;
- val ctorTs = map2 (curry (op -->)) FTs Ts;
- val unfold_fTs = map2 (curry op -->) activeAs Ts;
- val corec_sTs = map (Term.typ_subst_atomic (activeBs ~~ Ts)) sum_sTs;
- val corec_maps = map (Term.subst_atomic_types (activeBs ~~ Ts)) map_Inls;
- val corec_maps_rev = map (Term.subst_atomic_types (activeBs ~~ Ts)) map_Inls_rev;
- val corec_Inls = map (Term.subst_atomic_types (activeBs ~~ Ts)) Inls;
-
- val (((((((((((((Jzs, Jzs'), (Jz's, Jz's')), Jzs_copy), Jzs1), Jzs2), Jpairs),
- FJzs), TRs), unfold_fs), unfold_fs_copy), corec_ss), phis), names_lthy) = names_lthy
- |> mk_Frees' "z" Ts
- ||>> mk_Frees' "z" Ts'
- ||>> mk_Frees "z" Ts
- ||>> mk_Frees "z1" Ts
- ||>> mk_Frees "z2" Ts
- ||>> mk_Frees "j" (map2 (curry HOLogic.mk_prodT) Ts Ts')
- ||>> mk_Frees "x" prodFTs
- ||>> mk_Frees "R" (map (mk_relT o `I) Ts)
- ||>> mk_Frees "f" unfold_fTs
- ||>> mk_Frees "g" unfold_fTs
- ||>> mk_Frees "s" corec_sTs
- ||>> mk_Frees "P" (map2 mk_pred2T Ts Ts);
-
- fun dtor_bind i = Binding.suffix_name ("_" ^ dtorN) (nth bs (i - 1));
- val dtor_name = Binding.name_of o dtor_bind;
- val dtor_def_bind = rpair [] o Thm.def_binding o dtor_bind;
-
- fun dtor_spec i rep str map_FT dtorT Jz Jz' =
- let
- val lhs = Free (dtor_name i, dtorT);
- val rhs = Term.absfree Jz'
- (Term.list_comb (map_FT, map HOLogic.id_const passiveAs @ Abs_Ts) $
- (str $ (rep $ Jz)));
- in
- mk_Trueprop_eq (lhs, rhs)
- end;
-
- val ((dtor_frees, (_, dtor_def_frees)), (lthy, lthy_old)) =
- lthy
- |> fold_map7 (fn i => fn rep => fn str => fn mapx => fn dtorT => fn Jz => fn Jz' =>
- Specification.definition (SOME (dtor_bind i, NONE, NoSyn),
- (dtor_def_bind i, dtor_spec i rep str mapx dtorT Jz Jz')))
- ks Rep_Ts str_finals map_FTs dtorTs Jzs Jzs'
- |>> apsnd split_list o split_list
- ||> `Local_Theory.restore;
-
- 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 thm => thm RS fun_cong) o Morphism.thm phi) 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 =
- Skip_Proof.prove lthy [] []
- (HOLogic.mk_Trueprop (mk_mor UNIVs dtors car_finals str_finals Rep_Ts))
- (mk_mor_Rep_tac m (mor_def :: dtor_defs) Reps Abs_inverses coalg_final_set_thmss
- map_comp_id_thms map_congL_thms)
- |> Thm.close_derivation;
-
- val mor_Abs =
- Skip_Proof.prove lthy [] []
- (HOLogic.mk_Trueprop (mk_mor car_finals str_finals UNIVs dtors Abs_Ts))
- (mk_mor_Abs_tac (mor_def :: dtor_defs) Abs_inverses)
- |> Thm.close_derivation;
- in
- (mor_Rep, mor_Abs)
- end;
-
- val timer = time (timer "dtor definitions & thms");
-
- fun unfold_bind i = Binding.suffix_name ("_" ^ dtor_unfoldN) (nth bs (i - 1));
- val unfold_name = Binding.name_of o unfold_bind;
- val unfold_def_bind = rpair [] o Thm.def_binding o unfold_bind;
-
- fun unfold_spec i T AT abs f z z' =
- let
- val unfoldT = Library.foldr (op -->) (sTs, AT --> T);
-
- val lhs = Term.list_comb (Free (unfold_name i, unfoldT), ss);
- val rhs = Term.absfree z' (abs $ (f $ z));
- in
- mk_Trueprop_eq (lhs, rhs)
- end;
-
- val ((unfold_frees, (_, unfold_def_frees)), (lthy, lthy_old)) =
- lthy
- |> fold_map7 (fn i => fn T => fn AT => fn abs => fn f => fn z => fn z' =>
- Specification.definition
- (SOME (unfold_bind i, NONE, NoSyn), (unfold_def_bind i, unfold_spec i T AT abs f z z')))
- ks Ts activeAs Abs_Ts (map (fn i => HOLogic.mk_comp
- (mk_proj (mk_LSBIS passive_UNIVs i), mk_beh ss i)) ks) zs zs'
- |>> apsnd split_list o split_list
- ||> `Local_Theory.restore;
-
- 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_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 thm => thm RS fun_cong) o Morphism.thm phi) unfold_def_frees;
-
- 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 [tcoalg_thm RS mor_final_thm, UNIV_I]) RS sym) morE_thms;
- in
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all ss
- (HOLogic.mk_Trueprop (mk_mor active_UNIVs ss UNIVs dtors (map (mk_unfold Ts ss) ks))))
- (K (mk_mor_unfold_tac m mor_UNIV_thm dtor_defs unfold_defs Abs_inverses' morEs'
- map_comp_id_thms map_congs))
- |> Thm.close_derivation
- 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 passive_UNIVs UNIVs dtors TRs);
- val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
- (map2 (fn R => fn T => mk_subset R (Id_const T)) TRs Ts));
- val goal = fold_rev Logic.all TRs (Logic.mk_implies (prem, concl));
- in
- `split_conj_thm (Skip_Proof.prove lthy [] [] goal
- (K (mk_raw_coind_tac 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)
- end;
-
- val unique_mor_thms =
- let
- val prems = [HOLogic.mk_Trueprop (mk_coalg passive_UNIVs Bs ss), HOLogic.mk_Trueprop
- (HOLogic.mk_conj (mk_mor Bs ss UNIVs dtors unfold_fs,
- mk_mor Bs ss UNIVs dtors unfold_fs_copy))];
- fun mk_fun_eq B f g z = HOLogic.mk_imp
- (HOLogic.mk_mem (z, B), HOLogic.mk_eq (f $ z, g $ z));
- val unique = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
- (map4 mk_fun_eq Bs unfold_fs unfold_fs_copy zs));
-
- val unique_mor = Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (Bs @ ss @ unfold_fs @ unfold_fs_copy @ zs)
- (Logic.list_implies (prems, unique)))
- (K (mk_unique_mor_tac raw_coind_thms bis_image2_thm))
- |> Thm.close_derivation;
- in
- map (fn thm => conjI RSN (2, thm RS mp)) (split_conj_thm unique_mor)
- 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 bis_thm = tcoalg_thm RSN (2, tcoalg_thm RS bis_image2_thm);
- val mor_thm = mor_comp_thm OF [tcoalg_thm RS mor_final_thm, mor_Abs_thm];
-
- val unique_mor = Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (ss @ unfold_fs) (Logic.mk_implies (prem, unique)))
- (K (mk_unfold_unique_mor_tac raw_coind_thms bis_thm mor_thm unfold_defs))
- |> Thm.close_derivation;
- 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 @{thm ssubst[of _ _ "%x. x"]} 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 = Binding.suffix_name ("_" ^ ctorN) (nth bs (i - 1));
- val ctor_name = Binding.name_of o ctor_bind;
- val ctor_def_bind = rpair [] o Thm.def_binding o ctor_bind;
-
- fun ctor_spec i ctorT =
- let
- val lhs = Free (ctor_name i, ctorT);
- val rhs = mk_unfold Ts map_dtors i;
- in
- mk_Trueprop_eq (lhs, rhs)
- end;
-
- val ((ctor_frees, (_, ctor_def_frees)), (lthy, lthy_old)) =
- lthy
- |> fold_map2 (fn i => fn ctorT =>
- Specification.definition
- (SOME (ctor_bind i, NONE, NoSyn), (ctor_def_bind i, ctor_spec i ctorT))) ks ctorTs
- |>> apsnd split_list o split_list
- ||> `Local_Theory.restore;
-
- 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 (Morphism.thm phi) ctor_def_frees;
-
- val ctor_o_dtor_thms = map2 (fold_thms 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 = map3 mk_goal dtors ctors FTs;
- in
- map5 (fn goal => fn ctor_def => fn unfold => fn map_comp_id => fn map_congL =>
- Skip_Proof.prove lthy [] [] goal
- (mk_dtor_o_ctor_tac ctor_def unfold map_comp_id map_congL unfold_o_dtor_thms)
- |> Thm.close_derivation)
- goals ctor_defs dtor_unfold_thms map_comp_id_thms map_congL_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;
-
- fun mk_ctor_dtor_unfold_like_thm dtor_inject dtor_ctor unfold =
- iffD1 OF [dtor_inject, trans OF [unfold, dtor_ctor RS sym]];
-
- val ctor_dtor_unfold_thms =
- map3 mk_ctor_dtor_unfold_like_thm dtor_inject_thms dtor_ctor_thms dtor_unfold_thms;
-
- val timer = time (timer "ctor definitions & thms");
-
- val corec_Inl_sum_thms =
- let
- val mor = mor_comp_thm OF [mor_sum_case_thm, mor_unfold_thm];
- in
- map2 (fn unique => fn unfold_dtor =>
- trans OF [mor RS unique, unfold_dtor]) unfold_unique_mor_thms unfold_dtor_thms
- end;
-
- fun corec_bind i = Binding.suffix_name ("_" ^ dtor_corecN) (nth bs (i - 1));
- val corec_name = Binding.name_of o corec_bind;
- val corec_def_bind = rpair [] o Thm.def_binding o corec_bind;
-
- fun corec_spec i T AT =
- let
- val corecT = Library.foldr (op -->) (corec_sTs, AT --> T);
- val maps = map3 (fn dtor => fn sum_s => fn mapx => mk_sum_case
- (HOLogic.mk_comp (Term.list_comb (mapx, passive_ids @ corec_Inls), dtor), sum_s))
- dtors corec_ss corec_maps;
-
- val lhs = Term.list_comb (Free (corec_name i, corecT), corec_ss);
- val rhs = HOLogic.mk_comp (mk_unfold Ts maps i, Inr_const T AT);
- in
- mk_Trueprop_eq (lhs, rhs)
- end;
-
- val ((corec_frees, (_, corec_def_frees)), (lthy, lthy_old)) =
- lthy
- |> fold_map3 (fn i => fn T => fn AT =>
- Specification.definition
- (SOME (corec_bind i, NONE, NoSyn), (corec_def_bind i, corec_spec i T AT)))
- ks Ts activeAs
- |>> apsnd split_list o split_list
- ||> `Local_Theory.restore;
-
- val phi = Proof_Context.export_morphism lthy_old lthy;
- val corecs = map (Morphism.term phi) corec_frees;
- val corec_names = map (fst o dest_Const) corecs;
- fun mk_corec ss i = Term.list_comb (Const (nth corec_names (i - 1), Library.foldr (op -->)
- (map fastype_of ss, domain_type (fastype_of (nth ss (i - 1))) --> nth Ts (i - 1))), ss);
- val corec_defs = map (Morphism.thm phi) corec_def_frees;
-
- val sum_cases =
- map2 (fn T => fn i => mk_sum_case (HOLogic.id_const T, mk_corec corec_ss i)) Ts ks;
- val dtor_corec_thms =
- let
- fun mk_goal i corec_s corec_map dtor z =
- let
- val lhs = dtor $ (mk_corec corec_ss i $ z);
- val rhs = Term.list_comb (corec_map, passive_ids @ sum_cases) $ (corec_s $ z);
- in
- fold_rev Logic.all (z :: corec_ss) (mk_Trueprop_eq (lhs, rhs))
- end;
- val goals = map5 mk_goal ks corec_ss corec_maps_rev dtors zs;
- in
- map3 (fn goal => fn unfold => fn map_cong =>
- Skip_Proof.prove lthy [] [] goal
- (mk_corec_tac m corec_defs unfold map_cong corec_Inl_sum_thms)
- |> Thm.close_derivation)
- goals dtor_unfold_thms map_congs
- end;
-
- val ctor_dtor_corec_thms =
- map3 mk_ctor_dtor_unfold_like_thm dtor_inject_thms dtor_ctor_thms dtor_corec_thms;
-
- val timer = time (timer "corec definitions & thms");
-
- val (dtor_coinduct_thm, coinduct_params, srel_coinduct_thm, rel_coinduct_thm,
- dtor_strong_coinduct_thm, srel_strong_coinduct_thm, rel_strong_coinduct_thm) =
- let
- val zs = Jzs1 @ Jzs2;
- val frees = phis @ zs;
-
- fun mk_Ids Id = if Id then map Id_const passiveAs else map mk_diag passive_UNIVs;
-
- fun mk_phi upto_eq phi z1 z2 = if upto_eq
- then Term.absfree (dest_Free z1) (Term.absfree (dest_Free z2)
- (HOLogic.mk_disj (phi $ z1 $ z2, HOLogic.mk_eq (z1, z2))))
- else phi;
-
- fun phi_srels upto_eq = map4 (fn phi => fn T => fn z1 => fn z2 =>
- HOLogic.Collect_const (HOLogic.mk_prodT (T, T)) $
- HOLogic.mk_split (mk_phi upto_eq phi z1 z2)) phis Ts Jzs1 Jzs2;
-
- val srels = 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
- (map3 mk_concl phis Jzs1 Jzs2));
-
- fun mk_srel_prem upto_eq phi dtor srel Jz Jz_copy =
- let
- val concl = HOLogic.mk_mem (HOLogic.mk_tuple [dtor $ Jz, dtor $ Jz_copy],
- Term.list_comb (srel, mk_Ids upto_eq @ phi_srels upto_eq));
- in
- HOLogic.mk_Trueprop
- (list_all_free [Jz, Jz_copy] (HOLogic.mk_imp (phi $ Jz $ Jz_copy, concl)))
- end;
-
- val srel_prems = map5 (mk_srel_prem false) phis dtors srels Jzs Jzs_copy;
- val srel_upto_prems = map5 (mk_srel_prem true) phis dtors srels Jzs Jzs_copy;
-
- val srel_coinduct_goal = fold_rev Logic.all frees (Logic.list_implies (srel_prems, concl));
- val coinduct_params = rev (Term.add_tfrees srel_coinduct_goal []);
-
- val srel_coinduct = unfold_thms lthy @{thms diag_UNIV}
- (Skip_Proof.prove lthy [] [] srel_coinduct_goal
- (K (mk_srel_coinduct_tac ks raw_coind_thm bis_srel_thm))
- |> Thm.close_derivation);
-
- fun mk_dtor_prem upto_eq phi dtor map_nth sets Jz Jz_copy FJz =
- let
- val xs = [Jz, Jz_copy];
-
- fun mk_map_conjunct nths x =
- HOLogic.mk_eq (Term.list_comb (map_nth, passive_ids @ nths) $ FJz, dtor $ x);
-
- fun mk_set_conjunct set phi z1 z2 =
- list_all_free [z1, z2]
- (HOLogic.mk_imp (HOLogic.mk_mem (HOLogic.mk_prod (z1, z2), set $ FJz),
- mk_phi upto_eq phi z1 z2 $ z1 $ z2));
-
- val concl = list_exists_free [FJz] (HOLogic.mk_conj
- (Library.foldr1 HOLogic.mk_conj (map2 mk_map_conjunct [fstsTs, sndsTs] xs),
- Library.foldr1 HOLogic.mk_conj
- (map4 mk_set_conjunct (drop m sets) phis Jzs1 Jzs2)));
- in
- fold_rev Logic.all xs (Logic.mk_implies
- (HOLogic.mk_Trueprop (Term.list_comb (phi, xs)), HOLogic.mk_Trueprop concl))
- end;
-
- fun mk_dtor_prems upto_eq =
- map7 (mk_dtor_prem upto_eq) phis dtors map_FT_nths prodFT_setss Jzs Jzs_copy FJzs;
-
- val dtor_prems = mk_dtor_prems false;
- val dtor_upto_prems = mk_dtor_prems true;
-
- val dtor_coinduct_goal = fold_rev Logic.all frees (Logic.list_implies (dtor_prems, concl));
- val dtor_coinduct = Skip_Proof.prove lthy [] [] dtor_coinduct_goal
- (K (mk_dtor_coinduct_tac m ks raw_coind_thm bis_def))
- |> Thm.close_derivation;
-
- val cTs = map (SOME o certifyT lthy o TFree) coinduct_params;
- val cts = map3 (SOME o certify lthy ooo mk_phi true) phis Jzs1 Jzs2;
-
- val srel_strong_coinduct = singleton (Proof_Context.export names_lthy lthy)
- (Skip_Proof.prove lthy [] []
- (fold_rev Logic.all zs (Logic.list_implies (srel_upto_prems, concl)))
- (K (mk_srel_strong_coinduct_tac m cTs cts srel_coinduct srel_monos srel_Ids)))
- |> Thm.close_derivation;
-
- val dtor_strong_coinduct = singleton (Proof_Context.export names_lthy lthy)
- (Skip_Proof.prove lthy [] []
- (fold_rev Logic.all zs (Logic.list_implies (dtor_upto_prems, concl)))
- (K (mk_dtor_strong_coinduct_tac ks cTs cts dtor_coinduct bis_def
- (tcoalg_thm RS bis_diag_thm))))
- |> Thm.close_derivation;
-
- val rel_of_srel_thms =
- srel_defs @ @{thms Id_def' mem_Collect_eq fst_conv snd_conv split_conv};
-
- val rel_coinduct = unfold_thms lthy rel_of_srel_thms srel_coinduct;
- val rel_strong_coinduct = unfold_thms lthy rel_of_srel_thms srel_strong_coinduct;
- in
- (dtor_coinduct, rev (Term.add_tfrees dtor_coinduct_goal []), srel_coinduct, rel_coinduct,
- dtor_strong_coinduct, srel_strong_coinduct, rel_strong_coinduct)
- end;
-
- val timer = time (timer "coinduction");
-
- (*register new codatatypes as BNFs*)
- val lthy = if m = 0 then lthy else
- let
- val fTs = map2 (curry op -->) passiveAs passiveBs;
- val gTs = map2 (curry op -->) passiveBs passiveCs;
- val f1Ts = map2 (curry op -->) passiveAs passiveYs;
- val f2Ts = map2 (curry op -->) passiveBs passiveYs;
- val p1Ts = map2 (curry op -->) passiveXs passiveAs;
- val p2Ts = map2 (curry op -->) passiveXs passiveBs;
- val pTs = map2 (curry op -->) passiveXs passiveCs;
- val uTs = map2 (curry op -->) Ts Ts';
- val JRTs = map2 (curry mk_relT) passiveAs passiveBs;
- val JphiTs = map2 mk_pred2T passiveAs passiveBs;
- val prodTs = map2 (curry HOLogic.mk_prodT) Ts Ts';
- val B1Ts = map HOLogic.mk_setT passiveAs;
- val B2Ts = map HOLogic.mk_setT passiveBs;
- val AXTs = map HOLogic.mk_setT passiveXs;
- val XTs = mk_Ts passiveXs;
- val YTs = mk_Ts passiveYs;
-
- val ((((((((((((((((((((fs, fs'), fs_copy), gs), us),
- (Jys, Jys')), (Jys_copy, Jys'_copy)), set_induct_phiss), JRs), Jphis),
- B1s), B2s), AXs), f1s), f2s), p1s), p2s), ps), (ys, ys')), (ys_copy, ys'_copy)),
- names_lthy) = names_lthy
- |> 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_Freess "P" (map (fn A => map (mk_pred2T A) Ts) passiveAs)
- ||>> mk_Frees "R" JRTs
- ||>> mk_Frees "P" JphiTs
- ||>> mk_Frees "B1" B1Ts
- ||>> mk_Frees "B2" B2Ts
- ||>> mk_Frees "A" AXTs
- ||>> mk_Frees "f1" f1Ts
- ||>> mk_Frees "f2" f2Ts
- ||>> mk_Frees "p1" p1Ts
- ||>> mk_Frees "p2" p2Ts
- ||>> mk_Frees "p" pTs
- ||>> mk_Frees' "y" passiveAs
- ||>> mk_Frees' "y" passiveAs;
-
- 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 mk_mapsBC = mk_maps passiveBs passiveCs;
- val mk_mapsAC = mk_maps passiveAs passiveCs;
- val mk_mapsAY = mk_maps passiveAs passiveYs;
- val mk_mapsBY = mk_maps passiveBs passiveYs;
- val mk_mapsXA = mk_maps passiveXs passiveAs;
- val mk_mapsXB = mk_maps passiveXs passiveBs;
- val mk_mapsXC = mk_maps passiveXs passiveCs;
- val fs_maps = map (mk_map_id Ts fs Ts' dtors mk_mapsAB) ks;
- val fs_copy_maps = map (mk_map_id Ts fs_copy Ts' dtors mk_mapsAB) ks;
- val gs_maps = map (mk_map_id Ts' gs Ts'' dtor's mk_mapsBC) ks;
- val fgs_maps =
- map (mk_map_id Ts (map2 (curry HOLogic.mk_comp) gs fs) Ts'' dtors mk_mapsAC) ks;
- val Xdtors = mk_dtors passiveXs;
- val UNIV's = map HOLogic.mk_UNIV Ts';
- val CUNIVs = map HOLogic.mk_UNIV passiveCs;
- val UNIV''s = map HOLogic.mk_UNIV Ts'';
- val fstsTsTs' = map fst_const prodTs;
- val sndsTsTs' = map snd_const prodTs;
- val dtor''s = mk_dtors passiveCs;
- val f1s_maps = map (mk_map_id Ts f1s YTs dtors mk_mapsAY) ks;
- val f2s_maps = map (mk_map_id Ts' f2s YTs dtor's mk_mapsBY) ks;
- val pid_maps = map (mk_map_id XTs ps Ts'' Xdtors mk_mapsXC) ks;
- val pfst_Fmaps =
- map (mk_Fmap fst_const p1s prodTs) (mk_mapsXA prodTs (fst o HOLogic.dest_prodT));
- val psnd_Fmaps =
- map (mk_Fmap snd_const p2s prodTs) (mk_mapsXB prodTs (snd o HOLogic.dest_prodT));
- val p1id_Fmaps = map (mk_Fmap HOLogic.id_const p1s prodTs) (mk_mapsXA prodTs I);
- val p2id_Fmaps = map (mk_Fmap HOLogic.id_const p2s prodTs) (mk_mapsXB prodTs I);
- val pid_Fmaps = map (mk_Fmap HOLogic.id_const ps prodTs) (mk_mapsXC prodTs I);
-
- val (map_simp_thms, map_thms) =
- let
- fun mk_goal fs_map map dtor dtor' = fold_rev Logic.all fs
- (mk_Trueprop_eq (HOLogic.mk_comp (dtor', fs_map),
- HOLogic.mk_comp (Term.list_comb (map, fs @ fs_maps), dtor)));
- val goals = map4 mk_goal fs_maps map_FTFT's dtors dtor's;
- val cTs = map (SOME o certifyT lthy) FTs';
- val maps =
- map5 (fn goal => fn cT => fn unfold => fn map_comp' => fn map_cong =>
- Skip_Proof.prove lthy [] [] goal
- (K (mk_map_tac m n cT unfold map_comp' map_cong))
- |> Thm.close_derivation)
- goals cTs dtor_unfold_thms map_comp's map_congs;
- in
- map_split (fn thm => (thm RS @{thm pointfreeE}, thm)) maps
- end;
-
- val map_comp_thms =
- let
- val goal = fold_rev Logic.all (fs @ gs)
- (HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
- (map3 (fn fmap => fn gmap => fn fgmap =>
- HOLogic.mk_eq (HOLogic.mk_comp (gmap, fmap), fgmap))
- fs_maps gs_maps fgs_maps)))
- in
- split_conj_thm (Skip_Proof.prove lthy [] [] goal
- (K (mk_map_comp_tac m n map_thms map_comps map_congs dtor_unfold_unique_thm))
- |> Thm.close_derivation)
- end;
-
- val map_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 = map4 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_maps));
- in
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (us @ fs) (Logic.list_implies (prems, goal)))
- (mk_map_unique_tac dtor_unfold_unique_thm map_comps)
- |> Thm.close_derivation
- end;
-
- val timer = time (timer "map functions for the new codatatypes");
-
- val bd = mk_ccexp sbd sbd;
-
- val timer = time (timer "bounds for the new codatatypes");
-
- val setss_by_bnf = map (fn i => map2 (mk_hset dtors i) ls passiveAs) ks;
- val setss_by_bnf' = map (fn i => map2 (mk_hset dtor's i) ls passiveBs) ks;
- val setss_by_range = transpose setss_by_bnf;
-
- val set_simp_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 =>
- map4 (fn Fsets =>
- mk_simp_goal eq (nth Fsets (i - 1)) (drop m Fsets) sets)
- FTs_setss dtors Jzs sets)
- ls setss_by_range;
-
- val le_goals = map
- (fold_rev Logic.all Jzs o HOLogic.mk_Trueprop o Library.foldr1 HOLogic.mk_conj)
- (mk_goals (uncurry mk_subset));
- val set_le_thmss = map split_conj_thm
- (map4 (fn goal => fn hset_minimal => fn set_hsets => fn set_hset_hsetss =>
- Skip_Proof.prove lthy [] [] goal
- (K (mk_set_le_tac n hset_minimal set_hsets set_hset_hsetss))
- |> Thm.close_derivation)
- le_goals hset_minimal_thms set_hset_thmss' set_hset_hset_thmsss');
-
- val simp_goalss = map (map2 (fn z => fn goal =>
- Logic.all z (HOLogic.mk_Trueprop goal)) Jzs)
- (mk_goals HOLogic.mk_eq);
- in
- map4 (map4 (fn goal => fn set_le => fn set_incl_hset => fn set_hset_incl_hsets =>
- Skip_Proof.prove lthy [] [] goal
- (K (mk_set_simp_tac n set_le set_incl_hset set_hset_incl_hsets))
- |> Thm.close_derivation))
- simp_goalss set_le_thmss set_incl_hset_thmss' set_hset_incl_hset_thmsss'
- end;
-
- val timer = time (timer "set functions for the new codatatypes");
-
- val colss = map2 (fn j => fn T =>
- map (fn i => mk_hset_rec dtors nat i j T) ks) ls passiveAs;
- val colss' = map2 (fn j => fn T =>
- map (fn i => mk_hset_rec dtor's nat i j T) ks) ls passiveBs;
- val Xcolss = map2 (fn j => fn T =>
- map (fn i => mk_hset_rec Xdtors nat i j T) ks) ls passiveXs;
-
- 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
- (map4 (mk_col_natural f) fs_maps Jzs cols cols'));
-
- val goals = map3 mk_goal fs colss colss';
-
- val ctss =
- map (fn phi => map (SOME o certify lthy) [Term.absfree nat' phi, nat]) goals;
-
- val thms =
- map4 (fn goal => fn cts => fn rec_0s => fn rec_Sucs =>
- singleton (Proof_Context.export names_lthy lthy)
- (Skip_Proof.prove lthy [] [] (HOLogic.mk_Trueprop goal)
- (mk_col_natural_tac cts rec_0s rec_Sucs map_simp_thms set_natural'ss))
- |> Thm.close_derivation)
- goals ctss hset_rec_0ss' hset_rec_Sucss';
- in
- map (split_conj_thm o mk_specN n) thms
- end;
-
- val col_bd_thmss =
- let
- fun mk_col_bd z col = mk_ordLeq (mk_card_of (col $ z)) sbd;
-
- fun mk_goal cols = list_all_free Jzs (Library.foldr1 HOLogic.mk_conj
- (map2 mk_col_bd Jzs cols));
-
- val goals = map mk_goal colss;
-
- val ctss =
- map (fn phi => map (SOME o certify lthy) [Term.absfree nat' phi, nat]) goals;
-
- val thms =
- map5 (fn j => fn goal => fn cts => fn rec_0s => fn rec_Sucs =>
- singleton (Proof_Context.export names_lthy lthy)
- (Skip_Proof.prove lthy [] [] (HOLogic.mk_Trueprop goal)
- (K (mk_col_bd_tac m j cts rec_0s rec_Sucs
- sbd_Card_order sbd_Cinfinite set_sbdss)))
- |> Thm.close_derivation)
- ls goals ctss hset_rec_0ss' hset_rec_Sucss';
- in
- map (split_conj_thm o mk_specN n) thms
- end;
-
- val map_cong_thms =
- let
- val cTs = map (SOME o certifyT 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 (map5 (mk_prem z) sets fs fs_copy ys ys')
-
- fun mk_map_cong 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'
- |> certify lthy;
-
- val cphis =
- map9 mk_cphi setss_by_bnf Jzs' Jzs fs_maps fs_copy_maps Jys' Jys Jys'_copy Jys_copy;
-
- val coinduct = Drule.instantiate' cTs (map SOME cphis) dtor_coinduct_thm;
-
- val goal =
- HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
- (map4 mk_map_cong setss_by_bnf Jzs fs_maps fs_copy_maps));
-
- val thm = singleton (Proof_Context.export names_lthy lthy)
- (Skip_Proof.prove lthy [] [] goal
- (K (mk_mcong_tac m (rtac coinduct) map_comp's map_simp_thms map_congs set_natural'ss
- set_hset_thmss set_hset_hset_thmsss)))
- |> Thm.close_derivation
- in
- split_conj_thm thm
- end;
-
- val B1_ins = map2 (mk_in B1s) setss_by_bnf Ts;
- val B2_ins = map2 (mk_in B2s) setss_by_bnf' Ts';
- val thePulls = map4 mk_thePull B1_ins B2_ins f1s_maps f2s_maps;
- val thePullTs = passiveXs @ map2 (curry HOLogic.mk_prodT) Ts Ts';
- val thePull_ins = map2 (mk_in (AXs @ thePulls)) (mk_setss thePullTs) (mk_FTs thePullTs);
- val pickFs = map5 mk_pickWP thePull_ins pfst_Fmaps psnd_Fmaps
- (map2 (curry (op $)) dtors Jzs) (map2 (curry (op $)) dtor's Jz's);
- val pickF_ss = map3 (fn pickF => fn z => fn z' =>
- HOLogic.mk_split (Term.absfree z (Term.absfree z' pickF))) pickFs Jzs' Jz's';
- val picks = map (mk_unfold XTs pickF_ss) ks;
-
- val wpull_prem = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
- (map8 mk_wpull AXs B1s B2s f1s f2s (replicate m NONE) p1s p2s));
-
- val map_eq_thms = map2 (fn simp => fn diff => box_equals OF [diff RS iffD2, simp, simp])
- map_simp_thms dtor_inject_thms;
- val map_wpull_thms = map (fn thm => thm OF
- (replicate m asm_rl @ replicate n @{thm wpull_thePull})) map_wpulls;
- val pickWP_assms_tacs =
- map3 mk_pickWP_assms_tac set_incl_hset_thmss set_incl_hin_thmss map_eq_thms;
-
- val coalg_thePull_thm =
- let
- val coalg = HOLogic.mk_Trueprop
- (mk_coalg CUNIVs thePulls (map2 (curry HOLogic.mk_comp) pid_Fmaps pickF_ss));
- val goal = fold_rev Logic.all (AXs @ B1s @ B2s @ f1s @ f2s @ p1s @ p2s @ ps)
- (Logic.mk_implies (wpull_prem, coalg));
- in
- Skip_Proof.prove lthy [] [] goal (mk_coalg_thePull_tac m coalg_def map_wpull_thms
- set_natural'ss pickWP_assms_tacs)
- |> Thm.close_derivation
- end;
-
- val (mor_thePull_fst_thm, mor_thePull_snd_thm, mor_thePull_pick_thm) =
- let
- val mor_fst = HOLogic.mk_Trueprop
- (mk_mor thePulls (map2 (curry HOLogic.mk_comp) p1id_Fmaps pickF_ss)
- UNIVs dtors fstsTsTs');
- val mor_snd = HOLogic.mk_Trueprop
- (mk_mor thePulls (map2 (curry HOLogic.mk_comp) p2id_Fmaps pickF_ss)
- UNIV's dtor's sndsTsTs');
- val mor_pick = HOLogic.mk_Trueprop
- (mk_mor thePulls (map2 (curry HOLogic.mk_comp) pid_Fmaps pickF_ss)
- UNIV''s dtor''s (map2 (curry HOLogic.mk_comp) pid_maps picks));
-
- val fst_goal = fold_rev Logic.all (AXs @ B1s @ B2s @ f1s @ f2s @ p1s @ p2s)
- (Logic.mk_implies (wpull_prem, mor_fst));
- val snd_goal = fold_rev Logic.all (AXs @ B1s @ B2s @ f1s @ f2s @ p1s @ p2s)
- (Logic.mk_implies (wpull_prem, mor_snd));
- val pick_goal = fold_rev Logic.all (AXs @ B1s @ B2s @ f1s @ f2s @ p1s @ p2s @ ps)
- (Logic.mk_implies (wpull_prem, mor_pick));
- in
- (Skip_Proof.prove lthy [] [] fst_goal (mk_mor_thePull_fst_tac m mor_def map_wpull_thms
- map_comp's pickWP_assms_tacs) |> Thm.close_derivation,
- Skip_Proof.prove lthy [] [] snd_goal (mk_mor_thePull_snd_tac m mor_def map_wpull_thms
- map_comp's pickWP_assms_tacs) |> Thm.close_derivation,
- Skip_Proof.prove lthy [] [] pick_goal (mk_mor_thePull_pick_tac mor_def dtor_unfold_thms
- map_comp's) |> Thm.close_derivation)
- end;
-
- val pick_col_thmss =
- let
- fun mk_conjunct AX Jpair pick thePull col =
- HOLogic.mk_imp (HOLogic.mk_mem (Jpair, thePull), mk_subset (col $ (pick $ Jpair)) AX);
-
- fun mk_concl AX cols =
- list_all_free Jpairs (Library.foldr1 HOLogic.mk_conj
- (map4 (mk_conjunct AX) Jpairs picks thePulls cols));
-
- val concls = map2 mk_concl AXs Xcolss;
-
- val ctss =
- map (fn phi => map (SOME o certify lthy) [Term.absfree nat' phi, nat]) concls;
-
- val goals =
- map (fn concl => Logic.mk_implies (wpull_prem, HOLogic.mk_Trueprop concl)) concls;
-
- val thms =
- map5 (fn j => fn goal => fn cts => fn rec_0s => fn rec_Sucs =>
- singleton (Proof_Context.export names_lthy lthy) (Skip_Proof.prove lthy [] [] goal
- (mk_pick_col_tac m j cts rec_0s rec_Sucs dtor_unfold_thms set_natural'ss
- map_wpull_thms pickWP_assms_tacs))
- |> Thm.close_derivation)
- ls goals ctss hset_rec_0ss' hset_rec_Sucss';
- in
- map (map (fn thm => thm RS mp) o split_conj_thm o mk_specN n) thms
- end;
-
- val timer = time (timer "helpers for BNF properties");
-
- val map_id_tacs =
- map2 (K oo mk_map_id_tac map_thms) dtor_unfold_unique_thms unfold_dtor_thms;
- val map_comp_tacs = map (fn thm => K (rtac (thm RS sym) 1)) map_comp_thms;
- val map_cong_tacs = map (mk_map_cong_tac m) map_cong_thms;
- val set_nat_tacss =
- map2 (map2 (K oo mk_set_natural_tac)) hset_defss (transpose col_natural_thmss);
-
- val bd_co_tacs = replicate n (K (mk_bd_card_order_tac sbd_card_order));
- val bd_cinf_tacs = replicate n (K (mk_bd_cinfinite_tac sbd_Cinfinite));
-
- val set_bd_tacss =
- map2 (map2 (K oo mk_set_bd_tac sbd_Cinfinite)) hset_defss (transpose col_bd_thmss);
-
- val in_bd_tacs = map7 (fn i => fn isNode_hsets => fn carT_def =>
- fn card_of_carT => fn mor_image => fn Rep_inverse => fn mor_hsets =>
- K (mk_in_bd_tac (nth isNode_hsets (i - 1)) isNode_hsets carT_def
- card_of_carT mor_image Rep_inverse mor_hsets
- sbd_Cnotzero sbd_Card_order mor_Rep_thm coalgT_thm mor_T_final_thm tcoalg_thm))
- ks isNode_hset_thmss carT_defs card_of_carT_thms
- mor_image'_thms Rep_inverses (transpose mor_hset_thmss);
-
- val map_wpull_tacs =
- map3 (K ooo mk_wpull_tac m coalg_thePull_thm mor_thePull_fst_thm mor_thePull_snd_thm
- mor_thePull_pick_thm) unique_mor_thms (transpose pick_col_thmss) hset_defss;
-
- val srel_O_Gr_tacs = replicate n (simple_srel_O_Gr_tac o #context);
-
- val tacss = map10 zip_axioms map_id_tacs map_comp_tacs map_cong_tacs set_nat_tacss
- bd_co_tacs bd_cinf_tacs set_bd_tacss in_bd_tacs map_wpull_tacs srel_O_Gr_tacs;
-
- val (hset_dtor_incl_thmss, hset_hset_dtor_incl_thmsss, hset_induct_thms) =
- let
- fun tinst_of dtor =
- map (SOME o certify lthy) (dtor :: remove (op =) dtor dtors);
- fun tinst_of' dtor = case tinst_of dtor of t :: ts => t :: NONE :: ts;
- val Tinst = map (pairself (certifyT lthy))
- (map Logic.varifyT_global (deads @ allAs) ~~ (deads @ passiveAs @ Ts));
- val set_incl_thmss =
- map2 (fn dtor => map (singleton (Proof_Context.export names_lthy lthy) o
- Drule.instantiate' [] (tinst_of' dtor) o
- Thm.instantiate (Tinst, []) o Drule.zero_var_indexes))
- dtors set_incl_hset_thmss;
-
- val tinst = interleave (map (SOME o certify lthy) dtors) (replicate n NONE)
- val set_minimal_thms =
- map (Drule.instantiate' [] tinst o Thm.instantiate (Tinst, []) o
- Drule.zero_var_indexes)
- hset_minimal_thms;
-
- val set_set_incl_thmsss =
- map2 (fn dtor => map (map (singleton (Proof_Context.export names_lthy lthy) o
- Drule.instantiate' [] (NONE :: tinst_of' dtor) o
- Thm.instantiate (Tinst, []) o Drule.zero_var_indexes)))
- dtors set_hset_incl_hset_thmsss;
-
- val set_set_incl_thmsss' = transpose (map transpose set_set_incl_thmsss);
-
- val incls =
- maps (map (fn thm => thm RS @{thm subset_Collect_iff})) set_incl_thmss @
- @{thms subset_Collect_iff[OF subset_refl]};
-
- fun mk_induct_tinst phis jsets y y' =
- map4 (fn phi => fn jset => fn Jz => fn Jz' =>
- SOME (certify 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';
- val set_induct_thms =
- map6 (fn set_minimal => fn set_set_inclss => fn jsets => fn y => fn y' => fn phis =>
- ((set_minimal
- |> Drule.instantiate' [] (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 asm_rl)))
- set_minimal_thms set_set_incl_thmsss' setss_by_range ys ys' set_induct_phiss
- in
- (set_incl_thmss, set_set_incl_thmsss, set_induct_thms)
- end;
-
- 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;
-
- 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
- fold_rev Logic.all (map (nth ys) I @ Jzs) (HOLogic.mk_Trueprop
- (Library.foldr1 HOLogic.mk_conj (map4 mk_conjunct sets Jzs dummys wits)))
- end;
- val goals = map5 mk_goal setss_by_range ys ys_copy ys'_copy ls;
- in
- map2 (fn goal => fn induct =>
- Skip_Proof.prove lthy [] [] goal
- (mk_coind_wit_tac induct dtor_unfold_thms (flat set_natural'ss) wit_thms)
- |> Thm.close_derivation)
- goals hset_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 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;
-
- val 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));
-
- val wit_tac = mk_wit_tac n dtor_ctor_thms (flat set_simp_thmss) (maps wit_thms_of_bnf bnfs);
-
- val policy = user_policy Derive_All_Facts_Note_Most;
-
- val (Jbnfs, lthy) =
- fold_map6 (fn tacs => fn b => fn mapx => fn sets => fn T => fn (thms, wits) => fn lthy =>
- bnf_def Dont_Inline policy I tacs (wit_tac thms) (SOME deads)
- (((((b, fold_rev Term.absfree fs' mapx), sets), absdummy T bd), wits), NONE) lthy
- |> register_bnf (Local_Theory.full_name lthy b))
- tacss bs fs_maps setss_by_bnf Ts all_witss lthy;
-
- val fold_maps = fold_thms lthy (map (fn bnf =>
- mk_unabs_def m (map_def_of_bnf bnf RS @{thm meta_eq_to_obj_eq})) Jbnfs);
-
- val fold_sets = fold_thms lthy (maps (fn bnf =>
- map (fn thm => thm RS @{thm meta_eq_to_obj_eq}) (set_defs_of_bnf bnf)) Jbnfs);
-
- val timer = time (timer "registered new codatatypes as BNFs");
-
- val set_incl_thmss = map (map fold_sets) hset_dtor_incl_thmss;
- val set_set_incl_thmsss = map (map (map fold_sets)) hset_hset_dtor_incl_thmsss;
- val set_induct_thms = map fold_sets hset_induct_thms;
-
- val srels = map2 (fn Ds => mk_srel_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs;
- val Jsrels = map (mk_srel_of_bnf deads passiveAs passiveBs) Jbnfs;
- val rels = map2 (fn Ds => mk_rel_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs;
- val Jrels = map (mk_rel_of_bnf deads passiveAs passiveBs) Jbnfs;
-
- val JrelRs = map (fn Jsrel => Term.list_comb (Jsrel, JRs)) Jsrels;
- val relRs = map (fn srel => Term.list_comb (srel, JRs @ JrelRs)) srels;
- val Jpredphis = map (fn Jsrel => Term.list_comb (Jsrel, Jphis)) Jrels;
- val predphis = map (fn srel => Term.list_comb (srel, Jphis @ Jpredphis)) rels;
-
- val in_srels = map in_srel_of_bnf bnfs;
- val in_Jsrels = map in_srel_of_bnf Jbnfs;
- val Jsrel_defs = map srel_def_of_bnf Jbnfs;
- val Jrel_defs = map rel_def_of_bnf Jbnfs;
-
- val folded_map_simp_thms = map fold_maps map_simp_thms;
- val folded_set_simp_thmss = map (map fold_sets) set_simp_thmss;
- val folded_set_simp_thmss' = transpose folded_set_simp_thmss;
-
- val Jsrel_simp_thms =
- let
- fun mk_goal Jz Jz' dtor dtor' JrelR relR = fold_rev Logic.all (Jz :: Jz' :: JRs)
- (mk_Trueprop_eq (HOLogic.mk_mem (HOLogic.mk_prod (Jz, Jz'), JrelR),
- HOLogic.mk_mem (HOLogic.mk_prod (dtor $ Jz, dtor' $ Jz'), relR)));
- val goals = map6 mk_goal Jzs Jz's dtors dtor's JrelRs relRs;
- in
- map12 (fn i => fn goal => fn in_srel => fn map_comp => fn map_cong =>
- fn map_simp => fn set_simps => fn dtor_inject => fn dtor_ctor =>
- fn set_naturals => fn set_incls => fn set_set_inclss =>
- Skip_Proof.prove lthy [] [] goal
- (K (mk_srel_simp_tac in_Jsrels i in_srel map_comp map_cong map_simp set_simps
- dtor_inject dtor_ctor set_naturals set_incls set_set_inclss))
- |> Thm.close_derivation)
- ks goals in_srels map_comp's map_congs folded_map_simp_thms folded_set_simp_thmss'
- dtor_inject_thms dtor_ctor_thms set_natural'ss set_incl_thmss set_set_incl_thmsss
- end;
-
- val Jrel_simp_thms =
- let
- fun mk_goal Jz Jz' dtor dtor' Jpredphi predphi = fold_rev Logic.all (Jz :: Jz' :: Jphis)
- (mk_Trueprop_eq (Jpredphi $ Jz $ Jz', predphi $ (dtor $ Jz) $ (dtor' $ Jz')));
- val goals = map6 mk_goal Jzs Jz's dtors dtor's Jpredphis predphis;
- in
- map3 (fn goal => fn srel_def => fn Jsrel_simp =>
- Skip_Proof.prove lthy [] [] goal
- (mk_rel_simp_tac srel_def Jrel_defs Jsrel_defs Jsrel_simp)
- |> Thm.close_derivation)
- goals srel_defs Jsrel_simp_thms
- end;
-
- val timer = time (timer "additional properties");
-
- val ls' = if m = 1 then [0] else ls;
-
- val Jbnf_common_notes =
- [(map_uniqueN, [fold_maps map_unique_thm])] @
- map2 (fn i => fn thm => (mk_set_inductN i, [thm])) ls' set_induct_thms
- |> map (fn (thmN, thms) =>
- ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]));
-
- val Jbnf_notes =
- [(map_simpsN, map single folded_map_simp_thms),
- (rel_simpN, map single Jrel_simp_thms),
- (set_inclN, set_incl_thmss),
- (set_set_inclN, map flat set_set_incl_thmsss),
- (srel_simpN, map single Jsrel_simp_thms)] @
- map2 (fn i => fn thms => (mk_set_simpsN i, map single thms)) ls' folded_set_simp_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; lthy |> Local_Theory.notes (Jbnf_common_notes @ Jbnf_notes) |> snd
- end;
-
- val common_notes =
- [(dtor_coinductN, [dtor_coinduct_thm]),
- (dtor_strong_coinductN, [dtor_strong_coinduct_thm]),
- (rel_coinductN, [rel_coinduct_thm]),
- (rel_strong_coinductN, [rel_strong_coinduct_thm]),
- (srel_coinductN, [srel_coinduct_thm]),
- (srel_strong_coinductN, [srel_strong_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_dtor_unfoldsN, ctor_dtor_unfold_thms),
- (ctor_dtor_corecsN, ctor_dtor_corec_thms),
- (ctor_exhaustN, ctor_exhaust_thms),
- (ctor_injectN, ctor_inject_thms),
- (dtor_corecsN, dtor_corec_thms),
- (dtor_ctorN, dtor_ctor_thms),
- (dtor_exhaustN, dtor_exhaust_thms),
- (dtor_injectN, dtor_inject_thms),
- (dtor_unfold_uniqueN, dtor_unfold_unique_thms),
- (dtor_unfoldsN, dtor_unfold_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)
- in
- ((dtors, ctors, unfolds, corecs, dtor_coinduct_thm, dtor_ctor_thms, ctor_dtor_thms,
- ctor_inject_thms, ctor_dtor_unfold_thms, ctor_dtor_corec_thms),
- lthy |> Local_Theory.notes (common_notes @ notes) |> snd)
- end;
-
-val _ =
- Outer_Syntax.local_theory @{command_spec "codata_raw"}
- "define BNF-based coinductive datatypes (low-level)"
- (Parse.and_list1
- ((Parse.binding --| @{keyword ":"}) -- (Parse.typ --| @{keyword "="} -- Parse.typ)) >>
- (snd oo fp_bnf_cmd bnf_gfp o apsnd split_list o split_list));
-
-val _ =
- Outer_Syntax.local_theory @{command_spec "codata"} "define BNF-based coinductive datatypes"
- (parse_datatype_cmd false bnf_gfp);
-
-end;
--- a/src/HOL/Codatatype/Tools/bnf_gfp_tactics.ML Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1554 +0,0 @@
-(* Title: HOL/BNF/Tools/bnf_gfp_tactics.ML
- Author: Dmitriy Traytel, TU Muenchen
- Author: Andrei Popescu, TU Muenchen
- Author: Jasmin Blanchette, TU Muenchen
- Copyright 2012
-
-Tactics for the codatatype construction.
-*)
-
-signature BNF_GFP_TACTICS =
-sig
- val mk_Lev_sbd_tac: cterm option list -> thm list -> thm list -> thm list list -> tactic
- val mk_bd_card_order_tac: thm -> tactic
- val mk_bd_cinfinite_tac: thm -> tactic
- val mk_bis_Gr_tac: thm -> thm list -> thm list -> thm list -> thm list ->
- {prems: 'a, context: Proof.context} -> tactic
- val mk_bis_O_tac: int -> thm -> thm list -> thm list -> tactic
- val mk_bis_Union_tac: thm -> thm list -> {prems: 'a, context: Proof.context} -> tactic
- val mk_bis_converse_tac: int -> thm -> thm list -> thm list -> tactic
- val mk_bis_srel_tac: int -> thm -> thm list -> thm list -> thm list -> thm list list -> tactic
- val mk_carT_set_tac: int -> int -> thm -> thm -> thm -> thm ->
- {prems: 'a, context: Proof.context} -> tactic
- val mk_card_of_carT_tac: int -> thm list -> thm -> thm -> thm -> thm -> thm -> thm list -> tactic
- val mk_coalgT_tac: int -> thm list -> thm list -> thm list list ->
- {prems: 'a, context: Proof.context} -> tactic
- val mk_coalg_final_tac: int -> thm -> thm list -> thm list -> thm list list -> thm list list ->
- tactic
- val mk_coalg_set_tac: thm -> tactic
- val mk_coalg_thePull_tac: int -> thm -> thm list -> thm list list -> (int -> tactic) list ->
- {prems: 'a, context: Proof.context} -> tactic
- val mk_coind_wit_tac: thm -> thm list -> thm list -> thm list ->
- {prems: 'a, context: Proof.context} -> tactic
- val mk_col_bd_tac: int -> int -> cterm option list -> thm list -> thm list -> thm -> thm ->
- thm list list -> tactic
- val mk_col_natural_tac: cterm option list -> thm list -> thm list -> thm list -> thm list list ->
- {prems: 'a, context: Proof.context} -> tactic
- val mk_congruent_str_final_tac: int -> thm -> thm -> thm -> thm list -> tactic
- val mk_corec_tac: int -> thm list -> thm -> thm -> thm list ->
- {prems: 'a, context: Proof.context} -> tactic
- val mk_dtor_coinduct_tac: int -> int list -> thm -> thm -> tactic
- val mk_dtor_strong_coinduct_tac: int list -> ctyp option list -> cterm option list -> thm ->
- thm -> thm -> tactic
- val mk_dtor_o_ctor_tac: thm -> thm -> thm -> thm -> thm list ->
- {prems: 'a, context: Proof.context} -> tactic
- val mk_equiv_lsbis_tac: thm -> thm -> thm -> thm -> thm -> thm -> tactic
- val mk_hset_minimal_tac: int -> thm list -> thm -> {prems: 'a, context: Proof.context} -> tactic
- val mk_hset_rec_minimal_tac: int -> cterm option list -> thm list -> thm list ->
- {prems: 'a, context: Proof.context} -> tactic
- val mk_in_bd_tac: thm -> thm list -> thm -> thm -> thm -> thm -> thm list -> thm -> thm -> thm ->
- thm -> thm -> thm -> tactic
- val mk_incl_lsbis_tac: int -> int -> thm -> tactic
- val mk_isNode_hset_tac: int -> thm -> thm list -> {prems: 'a, context: Proof.context} -> tactic
- val mk_length_Lev'_tac: thm -> tactic
- val mk_length_Lev_tac: cterm option list -> thm list -> thm list -> tactic
- val mk_map_comp_tac: int -> int -> thm list -> thm list -> thm list -> thm -> tactic
- val mk_mcong_tac: int -> (int -> tactic) -> thm list -> thm list -> thm list -> thm list list ->
- thm list list -> thm list list list -> tactic
- val mk_map_id_tac: thm list -> thm -> thm -> tactic
- val mk_map_tac: int -> int -> ctyp option -> thm -> thm -> thm -> tactic
- val mk_map_unique_tac: thm -> thm list -> {prems: 'a, context: Proof.context} -> tactic
- val mk_mor_Abs_tac: thm list -> thm list -> {prems: 'a, context: Proof.context} -> tactic
- val mk_mor_Rep_tac: int -> thm list -> thm list -> thm list -> thm list list -> thm list ->
- thm list -> {prems: 'a, context: Proof.context} -> tactic
- val mk_mor_T_final_tac: thm -> thm list -> thm list -> tactic
- val mk_mor_UNIV_tac: thm list -> thm -> tactic
- val mk_mor_beh_tac: int -> thm -> thm -> thm list -> thm list -> thm list -> thm list ->
- thm list list -> thm list list -> thm list -> thm list -> thm list -> thm list -> thm list ->
- thm list -> thm list -> thm list -> thm list list -> thm list list list -> thm list list list ->
- thm list list list -> thm list list -> thm list list -> thm list -> thm list -> thm list ->
- {prems: 'a, context: Proof.context} -> tactic
- val mk_mor_comp_tac: thm -> thm list -> thm list -> thm list -> tactic
- val mk_mor_elim_tac: thm -> tactic
- val mk_mor_hset_rec_tac: int -> int -> cterm option list -> int -> thm list -> thm list ->
- thm list -> thm list list -> thm list list -> tactic
- val mk_mor_hset_tac: thm -> thm -> tactic
- val mk_mor_incl_tac: thm -> thm list -> tactic
- val mk_mor_str_tac: 'a list -> thm -> tactic
- val mk_mor_sum_case_tac: 'a list -> thm -> tactic
- val mk_mor_thePull_fst_tac: int -> thm -> thm list -> thm list -> (int -> tactic) list ->
- {prems: thm list, context: Proof.context} -> tactic
- val mk_mor_thePull_snd_tac: int -> thm -> thm list -> thm list -> (int -> tactic) list ->
- {prems: thm list, context: Proof.context} -> tactic
- val mk_mor_thePull_pick_tac: thm -> thm list -> thm list ->
- {prems: 'a, context: Proof.context} -> tactic
- val mk_mor_unfold_tac: int -> thm -> thm list -> thm list -> thm list -> thm list -> thm list ->
- thm list -> tactic
- val mk_prefCl_Lev_tac: cterm option list -> thm list -> thm list -> tactic
- val mk_pickWP_assms_tac: thm list -> thm list -> thm -> (int -> tactic)
- val mk_pick_col_tac: int -> int -> cterm option list -> thm list -> thm list -> thm list ->
- thm list list -> thm list -> (int -> tactic) list -> {prems: 'a, context: Proof.context} ->
- tactic
- val mk_raw_coind_tac: thm -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> thm -> thm list ->
- thm list -> thm list -> thm -> thm list -> tactic
- val mk_rv_last_tac: ctyp option list -> cterm option list -> thm list -> thm list -> tactic
- val mk_sbis_lsbis_tac: thm list -> thm -> thm -> tactic
- val mk_set_Lev_tac: cterm option list -> thm list -> thm list -> thm list -> thm list ->
- thm list list -> tactic
- val mk_set_bd_tac: thm -> thm -> thm -> tactic
- val mk_set_hset_incl_hset_tac: int -> thm list -> thm -> int -> tactic
- val mk_set_image_Lev_tac: cterm option list -> thm list -> thm list -> thm list -> thm list ->
- thm list list -> thm list list -> tactic
- val mk_set_incl_hin_tac: thm list -> tactic
- val mk_set_incl_hset_tac: thm -> thm -> tactic
- val mk_set_le_tac: int -> thm -> thm list -> thm list list -> tactic
- val mk_set_natural_tac: thm -> thm -> tactic
- val mk_set_rv_Lev_tac: int -> cterm option list -> thm list -> thm list -> thm list -> thm list ->
- thm list list -> thm list list -> tactic
- val mk_set_simp_tac: int -> thm -> thm -> thm list -> tactic
- val mk_srel_coinduct_tac: 'a list -> thm -> thm -> tactic
- val mk_srel_strong_coinduct_tac: int -> ctyp option list -> cterm option list -> thm ->
- thm list -> thm list -> tactic
- val mk_srel_simp_tac: thm list -> int -> thm -> thm -> thm -> thm -> thm list -> thm -> thm ->
- thm list -> thm list -> thm list list -> tactic
- val mk_strT_hset_tac: int -> int -> int -> ctyp option list -> ctyp option list ->
- cterm option list -> thm list -> thm list -> thm list -> thm list -> thm list list ->
- thm list list -> thm list list -> thm -> thm list list -> tactic
- val mk_unfold_unique_mor_tac: thm list -> thm -> thm -> thm list -> tactic
- val mk_unique_mor_tac: thm list -> thm -> tactic
- val mk_wit_tac: int -> thm list -> thm list -> thm list -> thm list ->
- {prems: 'a, context: Proof.context} -> tactic
- val mk_wpull_tac: int -> thm -> thm -> thm -> thm -> thm -> thm list -> thm list -> tactic
-end;
-
-structure BNF_GFP_Tactics : BNF_GFP_TACTICS =
-struct
-
-open BNF_Tactics
-open BNF_Util
-open BNF_FP
-open BNF_GFP_Util
-
-val fst_convol_fun_cong_sym = @{thm fst_convol} RS fun_cong RS sym;
-val list_inject_iffD1 = @{thm list.inject[THEN iffD1]};
-val nat_induct = @{thm nat_induct};
-val o_apply_trans_sym = o_apply RS trans RS sym;
-val ord_eq_le_trans = @{thm ord_eq_le_trans};
-val ord_eq_le_trans_trans_fun_cong_image_id_id_apply =
- @{thm ord_eq_le_trans[OF trans[OF fun_cong[OF image_id] id_apply]]};
-val ordIso_ordLeq_trans = @{thm ordIso_ordLeq_trans};
-val snd_convol_fun_cong_sym = @{thm snd_convol} RS fun_cong RS sym;
-val sum_case_weak_cong = @{thm sum_case_weak_cong};
-val trans_fun_cong_image_id_id_apply = @{thm trans[OF fun_cong[OF image_id] id_apply]};
-
-fun mk_coalg_set_tac coalg_def =
- dtac (coalg_def RS iffD1) 1 THEN
- REPEAT_DETERM (etac conjE 1) THEN
- EVERY' [dtac @{thm rev_bspec}, atac] 1 THEN
- REPEAT_DETERM (eresolve_tac [CollectE, conjE] 1) THEN atac 1;
-
-fun mk_mor_elim_tac mor_def =
- (dtac (subst OF [mor_def]) THEN'
- REPEAT o etac conjE THEN'
- TRY o rtac @{thm image_subsetI} THEN'
- etac bspec THEN'
- atac) 1;
-
-fun mk_mor_incl_tac mor_def map_id's =
- (stac mor_def THEN'
- rtac conjI THEN'
- CONJ_WRAP' (K (EVERY' [rtac ballI, etac set_mp, stac @{thm id_apply}, atac]))
- map_id's THEN'
- CONJ_WRAP' (fn thm =>
- (EVERY' [rtac ballI, rtac (thm RS trans), rtac sym, rtac (@{thm id_apply} RS arg_cong)]))
- map_id's) 1;
-
-fun mk_mor_comp_tac mor_def mor_images morEs map_comp_ids =
- let
- fun fbetw_tac image = EVERY' [rtac ballI, stac o_apply, etac image, etac image, atac];
- fun mor_tac ((mor_image, morE), map_comp_id) =
- EVERY' [rtac ballI, stac o_apply, rtac trans, rtac (map_comp_id RS sym), rtac trans,
- etac (morE RS arg_cong), atac, etac morE, etac mor_image, atac];
- in
- (stac mor_def THEN' rtac conjI THEN'
- CONJ_WRAP' fbetw_tac mor_images THEN'
- CONJ_WRAP' mor_tac ((mor_images ~~ morEs) ~~ map_comp_ids)) 1
- end;
-
-fun mk_mor_UNIV_tac morEs mor_def =
- let
- val n = length morEs;
- fun mor_tac morE = EVERY' [rtac ext, rtac trans, rtac o_apply, rtac trans, etac morE,
- rtac UNIV_I, rtac sym, rtac o_apply];
- in
- EVERY' [rtac iffI, CONJ_WRAP' mor_tac morEs,
- stac mor_def, rtac conjI, CONJ_WRAP' (K (rtac ballI THEN' rtac UNIV_I)) morEs,
- CONJ_WRAP' (fn i =>
- EVERY' [dtac (mk_conjunctN n i), rtac ballI, etac @{thm pointfreeE}]) (1 upto n)] 1
- end;
-
-fun mk_mor_str_tac ks mor_UNIV =
- (stac mor_UNIV THEN' CONJ_WRAP' (K (rtac refl)) ks) 1;
-
-fun mk_mor_sum_case_tac ks mor_UNIV =
- (stac mor_UNIV THEN' CONJ_WRAP' (K (rtac @{thm sum_case_comp_Inl[symmetric]})) ks) 1;
-
-fun mk_set_incl_hset_tac def rec_Suc =
- EVERY' (stac def ::
- map rtac [@{thm incl_UNION_I}, UNIV_I, @{thm ord_le_eq_trans}, @{thm Un_upper1},
- sym, rec_Suc]) 1;
-
-fun mk_set_hset_incl_hset_tac n defs rec_Suc i =
- EVERY' (map (TRY oo stac) defs @
- map rtac [@{thm UN_least}, subsetI, @{thm UN_I}, UNIV_I, set_mp, equalityD2, rec_Suc, UnI2,
- mk_UnIN n i] @
- [etac @{thm UN_I}, atac]) 1;
-
-fun mk_set_incl_hin_tac incls =
- if null incls then rtac subset_UNIV 1
- else EVERY' [rtac subsetI, rtac CollectI,
- CONJ_WRAP' (fn incl => EVERY' [rtac subset_trans, etac incl, atac]) incls] 1;
-
-fun mk_hset_rec_minimal_tac m cts rec_0s rec_Sucs {context = ctxt, prems = _} =
- EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
- REPEAT_DETERM o rtac allI,
- CONJ_WRAP' (fn thm => EVERY'
- [rtac ord_eq_le_trans, rtac thm, rtac @{thm empty_subsetI}]) rec_0s,
- REPEAT_DETERM o rtac allI,
- CONJ_WRAP' (fn rec_Suc => EVERY'
- [rtac ord_eq_le_trans, rtac rec_Suc,
- if m = 0 then K all_tac
- else (rtac @{thm Un_least} THEN' Goal.assume_rule_tac ctxt),
- CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_least}))
- (K (EVERY' [rtac @{thm UN_least}, REPEAT_DETERM o eresolve_tac [allE, conjE],
- rtac subset_trans, atac, Goal.assume_rule_tac ctxt])) rec_0s])
- rec_Sucs] 1;
-
-fun mk_hset_minimal_tac n hset_defs hset_rec_minimal {context = ctxt, prems = _} =
- (CONJ_WRAP' (fn def => (EVERY' [rtac ord_eq_le_trans, rtac def,
- rtac @{thm UN_least}, rtac rev_mp, rtac hset_rec_minimal,
- EVERY' (replicate ((n + 1) * n) (Goal.assume_rule_tac ctxt)), rtac impI,
- REPEAT_DETERM o eresolve_tac [allE, conjE], atac])) hset_defs) 1
-
-fun mk_mor_hset_rec_tac m n cts j rec_0s rec_Sucs morEs set_naturalss coalg_setss =
- EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
- REPEAT_DETERM o rtac allI,
- CONJ_WRAP' (fn thm => EVERY' (map rtac [impI, thm RS trans, thm RS sym])) rec_0s,
- REPEAT_DETERM o rtac allI,
- CONJ_WRAP'
- (fn (rec_Suc, (morE, ((passive_set_naturals, active_set_naturals), coalg_sets))) =>
- EVERY' [rtac impI, rtac (rec_Suc RS trans), rtac (rec_Suc RS trans RS sym),
- if m = 0 then K all_tac
- else EVERY' [rtac @{thm Un_cong}, rtac box_equals,
- rtac (nth passive_set_naturals (j - 1) RS sym),
- rtac trans_fun_cong_image_id_id_apply, etac (morE RS arg_cong), atac],
- CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_cong}))
- (fn (i, (set_natural, coalg_set)) =>
- EVERY' [rtac sym, rtac trans, rtac (refl RSN (2, @{thm UN_cong})),
- etac (morE RS sym RS arg_cong RS trans), atac, rtac set_natural,
- rtac (@{thm UN_simps(10)} RS trans), rtac (refl RS @{thm UN_cong}),
- ftac coalg_set, atac, dtac set_mp, atac, rtac mp, rtac (mk_conjunctN n i),
- REPEAT_DETERM o etac allE, atac, atac])
- (rev ((1 upto n) ~~ (active_set_naturals ~~ coalg_sets)))])
- (rec_Sucs ~~ (morEs ~~ (map (chop m) set_naturalss ~~ map (drop m) coalg_setss)))] 1;
-
-fun mk_mor_hset_tac hset_def mor_hset_rec =
- EVERY' [rtac (hset_def RS trans), rtac (refl RS @{thm UN_cong} RS trans), etac mor_hset_rec,
- atac, atac, rtac (hset_def RS sym)] 1
-
-fun mk_bis_srel_tac m bis_def srel_O_Grs map_comps map_congs set_naturalss =
- let
- val n = length srel_O_Grs;
- val thms = ((1 upto n) ~~ map_comps ~~ map_congs ~~ set_naturalss ~~ srel_O_Grs);
-
- fun mk_if_tac ((((i, map_comp), map_cong), set_naturals), srel_O_Gr) =
- EVERY' [rtac allI, rtac allI, rtac impI, dtac (mk_conjunctN n i),
- etac allE, etac allE, etac impE, atac, etac bexE, etac conjE,
- rtac (srel_O_Gr RS equalityD2 RS set_mp),
- rtac @{thm relcompI}, rtac @{thm converseI},
- EVERY' (map (fn thm =>
- EVERY' [rtac @{thm GrI}, REPEAT_DETERM o eresolve_tac [CollectE, conjE],
- rtac CollectI,
- CONJ_WRAP' (fn (i, thm) =>
- if i <= m
- then EVERY' [rtac ord_eq_le_trans, rtac thm, rtac subset_trans,
- etac @{thm image_mono}, rtac @{thm image_subsetI}, etac @{thm diagI}]
- else EVERY' [rtac ord_eq_le_trans, rtac trans, rtac thm,
- rtac trans_fun_cong_image_id_id_apply, atac])
- (1 upto (m + n) ~~ set_naturals),
- rtac trans, rtac trans, rtac map_comp, rtac map_cong, REPEAT_DETERM_N m o rtac thm,
- REPEAT_DETERM_N n o rtac (@{thm o_id} RS fun_cong), atac])
- @{thms fst_diag_id snd_diag_id})];
-
- fun mk_only_if_tac ((((i, map_comp), map_cong), set_naturals), srel_O_Gr) =
- EVERY' [dtac (mk_conjunctN n i), rtac allI, rtac allI, rtac impI,
- etac allE, etac allE, etac impE, atac,
- dtac (srel_O_Gr RS equalityD1 RS set_mp),
- REPEAT_DETERM o eresolve_tac [CollectE, @{thm relcompE}, @{thm converseE}],
- REPEAT_DETERM o eresolve_tac [@{thm GrE}, exE, conjE],
- REPEAT_DETERM o dtac Pair_eqD,
- REPEAT_DETERM o etac conjE,
- hyp_subst_tac,
- REPEAT_DETERM o eresolve_tac [CollectE, conjE],
- rtac bexI, rtac conjI, rtac trans, rtac map_comp,
- REPEAT_DETERM_N m o stac @{thm id_o},
- REPEAT_DETERM_N n o stac @{thm o_id},
- etac sym, rtac trans, rtac map_comp,
- REPEAT_DETERM_N m o stac @{thm id_o},
- REPEAT_DETERM_N n o stac @{thm o_id},
- rtac trans, rtac map_cong,
- REPEAT_DETERM_N m o EVERY' [rtac @{thm diagE'}, etac set_mp, atac],
- REPEAT_DETERM_N n o rtac refl,
- etac sym, rtac CollectI,
- CONJ_WRAP' (fn (i, thm) =>
- if i <= m
- then EVERY' [rtac ord_eq_le_trans, rtac thm, rtac @{thm image_subsetI},
- rtac @{thm diag_fst}, etac set_mp, atac]
- else EVERY' [rtac ord_eq_le_trans, rtac trans, rtac thm,
- rtac trans_fun_cong_image_id_id_apply, atac])
- (1 upto (m + n) ~~ set_naturals)];
- in
- EVERY' [rtac (bis_def RS trans),
- rtac iffI, etac conjE, etac conjI, CONJ_WRAP' mk_if_tac thms,
- etac conjE, etac conjI, CONJ_WRAP' mk_only_if_tac thms] 1
- end;
-
-fun mk_bis_converse_tac m bis_srel srel_congs srel_converses =
- EVERY' [stac bis_srel, dtac (bis_srel RS iffD1),
- REPEAT_DETERM o etac conjE, rtac conjI,
- CONJ_WRAP' (K (EVERY' [rtac @{thm converse_shift}, etac subset_trans,
- rtac equalityD2, rtac @{thm converse_Times}])) srel_congs,
- CONJ_WRAP' (fn (srel_cong, srel_converse) =>
- EVERY' [rtac allI, rtac allI, rtac impI, rtac @{thm set_mp[OF equalityD2]},
- rtac (srel_cong RS trans),
- REPEAT_DETERM_N m o rtac @{thm diag_converse},
- REPEAT_DETERM_N (length srel_congs) o rtac refl,
- rtac srel_converse,
- REPEAT_DETERM o etac allE,
- rtac @{thm converseI}, etac mp, etac @{thm converseD}]) (srel_congs ~~ srel_converses)] 1;
-
-fun mk_bis_O_tac m bis_srel srel_congs srel_Os =
- EVERY' [stac bis_srel, REPEAT_DETERM o dtac (bis_srel RS iffD1),
- REPEAT_DETERM o etac conjE, rtac conjI,
- CONJ_WRAP' (K (EVERY' [etac @{thm relcomp_subset_Sigma}, atac])) srel_congs,
- CONJ_WRAP' (fn (srel_cong, srel_O) =>
- EVERY' [rtac allI, rtac allI, rtac impI, rtac @{thm set_mp[OF equalityD2]},
- rtac (srel_cong RS trans),
- REPEAT_DETERM_N m o rtac @{thm diag_Comp},
- REPEAT_DETERM_N (length srel_congs) o rtac refl,
- rtac srel_O,
- etac @{thm relcompE},
- REPEAT_DETERM o dtac Pair_eqD,
- etac conjE, hyp_subst_tac,
- REPEAT_DETERM o etac allE, rtac @{thm relcompI},
- etac mp, atac, etac mp, atac]) (srel_congs ~~ srel_Os)] 1;
-
-fun mk_bis_Gr_tac bis_srel srel_Grs mor_images morEs coalg_ins
- {context = ctxt, prems = _} =
- unfold_thms_tac ctxt (bis_srel :: @{thm diag_Gr} :: srel_Grs) THEN
- EVERY' [rtac conjI,
- CONJ_WRAP' (fn thm => rtac (@{thm Gr_incl} RS ssubst) THEN' etac thm) mor_images,
- CONJ_WRAP' (fn (coalg_in, morE) =>
- EVERY' [rtac allI, rtac allI, rtac impI, rtac @{thm GrI}, etac coalg_in,
- etac @{thm GrD1}, etac (morE RS trans), etac @{thm GrD1},
- etac (@{thm GrD2} RS arg_cong)]) (coalg_ins ~~ morEs)] 1;
-
-fun mk_bis_Union_tac bis_def in_monos {context = ctxt, prems = _} =
- let
- val n = length in_monos;
- val ks = 1 upto n;
- in
- unfold_thms_tac ctxt [bis_def] THEN
- EVERY' [rtac conjI,
- CONJ_WRAP' (fn i =>
- EVERY' [rtac @{thm UN_least}, dtac bspec, atac,
- dtac conjunct1, etac (mk_conjunctN n i)]) ks,
- CONJ_WRAP' (fn (i, in_mono) =>
- EVERY' [rtac allI, rtac allI, rtac impI, etac @{thm UN_E}, dtac bspec, atac,
- dtac conjunct2, dtac (mk_conjunctN n i), etac allE, etac allE, dtac mp,
- atac, etac bexE, rtac bexI, atac, rtac in_mono,
- REPEAT_DETERM_N n o etac @{thm incl_UNION_I[OF _ subset_refl]},
- atac]) (ks ~~ in_monos)] 1
- end;
-
-fun mk_sbis_lsbis_tac lsbis_defs bis_Union bis_cong =
- let
- val n = length lsbis_defs;
- in
- EVERY' [rtac (Thm.permute_prems 0 1 bis_cong), EVERY' (map rtac lsbis_defs),
- rtac bis_Union, rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, conjE, exE],
- hyp_subst_tac, etac bis_cong, EVERY' (map (rtac o mk_nth_conv n) (1 upto n))] 1
- end;
-
-fun mk_incl_lsbis_tac n i lsbis_def =
- EVERY' [rtac @{thm xt1(3)}, rtac lsbis_def, rtac @{thm incl_UNION_I}, rtac CollectI,
- REPEAT_DETERM_N n o rtac exI, rtac conjI, rtac refl, atac, rtac equalityD2,
- rtac (mk_nth_conv n i)] 1;
-
-fun mk_equiv_lsbis_tac sbis_lsbis lsbis_incl incl_lsbis bis_diag bis_converse bis_O =
- EVERY' [rtac (@{thm equiv_def} RS iffD2),
-
- rtac conjI, rtac (@{thm refl_on_def} RS iffD2),
- rtac conjI, rtac lsbis_incl, rtac ballI, rtac set_mp,
- rtac incl_lsbis, rtac bis_diag, atac, etac @{thm diagI},
-
- rtac conjI, rtac (@{thm sym_def} RS iffD2),
- rtac allI, rtac allI, rtac impI, rtac set_mp,
- rtac incl_lsbis, rtac bis_converse, rtac sbis_lsbis, etac @{thm converseI},
-
- rtac (@{thm trans_def} RS iffD2),
- rtac allI, rtac allI, rtac allI, rtac impI, rtac impI, rtac set_mp,
- rtac incl_lsbis, rtac bis_O, rtac sbis_lsbis, rtac sbis_lsbis,
- etac @{thm relcompI}, atac] 1;
-
-fun mk_coalgT_tac m defs strT_defs set_naturalss {context = ctxt, prems = _} =
- let
- val n = length strT_defs;
- val ks = 1 upto n;
- fun coalg_tac (i, ((passive_sets, active_sets), def)) =
- EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE],
- hyp_subst_tac, rtac (def RS trans RS @{thm ssubst_mem}), etac (arg_cong RS trans),
- rtac (mk_sum_casesN n i), rtac CollectI,
- EVERY' (map (fn thm => EVERY' [rtac conjI, rtac (thm RS ord_eq_le_trans),
- etac ((trans OF [@{thm image_id} RS fun_cong, @{thm id_apply}]) RS ord_eq_le_trans)])
- passive_sets),
- CONJ_WRAP' (fn (i, thm) => EVERY' [rtac (thm RS ord_eq_le_trans),
- rtac @{thm image_subsetI}, rtac CollectI, rtac exI, rtac exI, rtac conjI, rtac refl,
- rtac conjI,
- rtac conjI, etac @{thm empty_Shift}, dtac set_rev_mp,
- etac equalityD1, etac CollectD,
- rtac conjI, etac @{thm Shift_clists},
- rtac conjI, etac @{thm Shift_prefCl},
- rtac conjI, rtac ballI,
- rtac conjI, dtac @{thm iffD1[OF ball_conj_distrib]}, dtac conjunct1,
- SELECT_GOAL (unfold_thms_tac ctxt @{thms Succ_Shift shift_def}),
- etac bspec, etac @{thm ShiftD},
- CONJ_WRAP' (fn i => EVERY' [rtac ballI, etac CollectE, dtac @{thm ShiftD},
- dtac bspec, etac thin_rl, atac, dtac conjunct2, dtac (mk_conjunctN n i),
- dtac bspec, rtac CollectI, etac @{thm set_mp[OF equalityD1[OF Succ_Shift]]},
- REPEAT_DETERM o eresolve_tac [exE, conjE], rtac exI,
- rtac conjI, rtac (@{thm shift_def} RS fun_cong RS trans),
- rtac (@{thm append_Cons} RS sym RS arg_cong RS trans), atac,
- REPEAT_DETERM_N m o (rtac conjI THEN' atac),
- CONJ_WRAP' (K (EVERY' [etac trans, rtac @{thm Collect_cong},
- rtac @{thm eqset_imp_iff}, rtac sym, rtac trans, rtac @{thm Succ_Shift},
- rtac (@{thm append_Cons} RS sym RS arg_cong)])) ks]) ks,
- rtac allI, rtac impI, REPEAT_DETERM o eresolve_tac [allE, impE],
- etac @{thm not_in_Shift}, rtac trans, rtac (@{thm shift_def} RS fun_cong), atac,
- dtac bspec, atac, dtac conjunct2, dtac (mk_conjunctN n i), dtac bspec,
- etac @{thm set_mp[OF equalityD1]}, atac,
- REPEAT_DETERM o eresolve_tac [exE, conjE], rtac exI,
- rtac conjI, rtac (@{thm shift_def} RS fun_cong RS trans),
- etac (@{thm append_Nil} RS sym RS arg_cong RS trans),
- REPEAT_DETERM_N m o (rtac conjI THEN' atac),
- CONJ_WRAP' (K (EVERY' [etac trans, rtac @{thm Collect_cong},
- rtac @{thm eqset_imp_iff}, rtac sym, rtac trans, rtac @{thm Succ_Shift},
- rtac (@{thm append_Nil} RS sym RS arg_cong)])) ks]) (ks ~~ active_sets)];
- in
- unfold_thms_tac ctxt defs THEN
- CONJ_WRAP' coalg_tac (ks ~~ (map (chop m) set_naturalss ~~ strT_defs)) 1
- end;
-
-fun mk_card_of_carT_tac m isNode_defs sbd_sbd
- sbd_card_order sbd_Card_order sbd_Cinfinite sbd_Cnotzero in_sbds =
- let
- val n = length isNode_defs;
- in
- EVERY' [rtac (Thm.permute_prems 0 1 ctrans),
- rtac @{thm card_of_Sigma_ordLeq_Cinfinite}, rtac @{thm Cinfinite_cexp},
- if m = 0 then rtac @{thm ordLeq_refl} else rtac @{thm ordLeq_csum2},
- rtac @{thm Card_order_ctwo}, rtac @{thm Cinfinite_cexp},
- rtac @{thm ctwo_ordLeq_Cinfinite}, rtac sbd_Cinfinite, rtac sbd_Cinfinite,
- rtac ctrans, rtac @{thm card_of_diff},
- rtac ordIso_ordLeq_trans, rtac @{thm card_of_Field_ordIso},
- rtac @{thm Card_order_cpow}, rtac ordIso_ordLeq_trans,
- rtac @{thm cpow_cexp_ctwo}, rtac ctrans, rtac @{thm cexp_mono1_Cnotzero},
- if m = 0 then rtac @{thm ordLeq_refl} else rtac @{thm ordLeq_csum2},
- rtac @{thm Card_order_ctwo}, rtac @{thm ctwo_Cnotzero}, rtac @{thm Card_order_clists},
- rtac @{thm cexp_mono2_Cnotzero}, rtac ordIso_ordLeq_trans,
- rtac @{thm clists_Cinfinite},
- if n = 1 then rtac sbd_Cinfinite else rtac (sbd_Cinfinite RS @{thm Cinfinite_csum1}),
- rtac ordIso_ordLeq_trans, rtac sbd_sbd, rtac @{thm infinite_ordLeq_cexp},
- rtac sbd_Cinfinite,
- if m = 0 then rtac @{thm ctwo_Cnotzero} else rtac @{thm csum_Cnotzero2[OF ctwo_Cnotzero]},
- rtac @{thm Cnotzero_clists},
- rtac ballI, rtac ordIso_ordLeq_trans, rtac @{thm card_of_Func_Ffunc},
- rtac ordIso_ordLeq_trans, rtac @{thm Func_cexp},
- rtac ctrans, rtac @{thm cexp_mono},
- rtac @{thm ordLeq_ordIso_trans},
- CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1
- (sbd_Cinfinite RS @{thm Cinfinite_cexp[OF ordLeq_csum2[OF Card_order_ctwo]]}
- RSN (3, @{thm Un_Cinfinite_bound}))))
- (fn thm => EVERY' [rtac ctrans, rtac @{thm card_of_image}, rtac thm]) (rev in_sbds),
- rtac @{thm cexp_cong1_Cnotzero}, rtac @{thm csum_cong1},
- REPEAT_DETERM_N m o rtac @{thm csum_cong2},
- CONJ_WRAP_GEN' (rtac @{thm csum_cong})
- (K (rtac (sbd_Card_order RS @{thm card_of_Field_ordIso}))) in_sbds,
- rtac sbd_Card_order,
- rtac @{thm csum_Cnotzero2}, rtac @{thm ctwo_Cnotzero},
- rtac @{thm csum_Cnotzero2}, rtac @{thm ctwo_Cnotzero},
- rtac @{thm ordLeq_ordIso_trans}, etac @{thm clists_bound},
- rtac @{thm clists_Cinfinite}, TRY o rtac @{thm Cinfinite_csum1}, rtac sbd_Cinfinite,
- rtac disjI2, rtac @{thm cone_ordLeq_cexp}, rtac @{thm cone_ordLeq_cexp},
- rtac ctrans, rtac @{thm cone_ordLeq_ctwo}, rtac @{thm ordLeq_csum2},
- rtac @{thm Card_order_ctwo}, rtac FalseE, etac @{thm cpow_clists_czero}, atac,
- rtac @{thm card_of_Card_order},
- rtac ordIso_ordLeq_trans, rtac @{thm cexp_cprod},
- rtac @{thm csum_Cnotzero2}, rtac @{thm ctwo_Cnotzero},
- rtac ordIso_ordLeq_trans, rtac @{thm cexp_cong2_Cnotzero},
- rtac @{thm ordIso_transitive}, rtac @{thm cprod_cong2}, rtac sbd_sbd,
- rtac @{thm cprod_infinite}, rtac sbd_Cinfinite,
- rtac @{thm csum_Cnotzero2}, rtac @{thm ctwo_Cnotzero}, rtac @{thm Card_order_cprod},
- rtac ctrans, rtac @{thm cexp_mono1_Cnotzero},
- rtac ordIso_ordLeq_trans, rtac @{thm ordIso_transitive}, rtac @{thm csum_cong1},
- rtac @{thm ordIso_transitive},
- REPEAT_DETERM_N m o rtac @{thm csum_cong2},
- rtac sbd_sbd,
- BNF_Tactics.mk_rotate_eq_tac (rtac @{thm ordIso_refl} THEN'
- FIRST' [rtac @{thm card_of_Card_order},
- rtac @{thm Card_order_csum}, rtac sbd_Card_order])
- @{thm ordIso_transitive} @{thm csum_assoc} @{thm csum_com} @{thm csum_cong}
- (1 upto m + 1) (m + 1 :: (1 upto m)),
- if m = 0 then K all_tac else EVERY' [rtac @{thm ordIso_transitive}, rtac @{thm csum_assoc}],
- rtac @{thm csum_com}, rtac @{thm csum_cexp'}, rtac sbd_Cinfinite,
- if m = 0 then rtac @{thm Card_order_ctwo} else rtac @{thm Card_order_csum},
- if m = 0 then rtac @{thm ordLeq_refl} else rtac @{thm ordLeq_csum2},
- rtac @{thm Card_order_ctwo},
- rtac @{thm csum_Cnotzero2}, rtac @{thm ctwo_Cnotzero}, rtac sbd_Card_order,
- rtac ordIso_ordLeq_trans, rtac @{thm cexp_cprod_ordLeq},
- if m = 0 then rtac @{thm ctwo_Cnotzero} else rtac @{thm csum_Cnotzero2[OF ctwo_Cnotzero]},
- rtac sbd_Cinfinite, rtac sbd_Cnotzero, rtac @{thm ordLeq_refl}, rtac sbd_Card_order,
- rtac @{thm cexp_mono2_Cnotzero}, rtac @{thm infinite_ordLeq_cexp},
- rtac sbd_Cinfinite,
- if m = 0 then rtac @{thm ctwo_Cnotzero} else rtac @{thm csum_Cnotzero2[OF ctwo_Cnotzero]},
- rtac sbd_Cnotzero,
- rtac @{thm card_of_mono1}, rtac subsetI,
- REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE, @{thm prod_caseE}], hyp_subst_tac,
- rtac @{thm SigmaI}, rtac @{thm DiffI}, rtac set_mp, rtac equalityD2,
- rtac (@{thm cpow_def} RS arg_cong RS trans), rtac (@{thm Pow_def} RS arg_cong RS trans),
- rtac @{thm Field_card_of}, rtac CollectI, atac, rtac notI, etac @{thm singletonE},
- hyp_subst_tac, etac @{thm emptyE}, rtac (@{thm Ffunc_def} RS equalityD2 RS set_mp),
- rtac CollectI, rtac conjI, rtac ballI, dtac bspec, etac thin_rl, atac, dtac conjunct1,
- CONJ_WRAP_GEN' (etac disjE) (fn (i, def) => EVERY'
- [rtac (mk_UnIN n i), dtac (def RS iffD1),
- REPEAT_DETERM o eresolve_tac [exE, conjE], rtac @{thm image_eqI}, atac, rtac CollectI,
- REPEAT_DETERM_N m o (rtac conjI THEN' atac),
- CONJ_WRAP' (K (EVERY' [etac ord_eq_le_trans, rtac subset_trans,
- rtac subset_UNIV, rtac equalityD2, rtac @{thm Field_card_order},
- rtac sbd_card_order])) isNode_defs]) (1 upto n ~~ isNode_defs),
- atac] 1
- end;
-
-fun mk_carT_set_tac n i carT_def strT_def isNode_def set_natural {context = ctxt, prems = _}=
- EVERY' [dtac (carT_def RS equalityD1 RS set_mp),
- REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE],
- dtac Pair_eqD,
- etac conjE, hyp_subst_tac,
- dtac (isNode_def RS iffD1),
- REPEAT_DETERM o eresolve_tac [exE, conjE],
- rtac (equalityD2 RS set_mp),
- rtac (strT_def RS arg_cong RS trans),
- etac (arg_cong RS trans),
- fo_rtac (mk_sum_casesN n i RS arg_cong RS trans) ctxt,
- rtac set_natural, rtac imageI, etac (equalityD2 RS set_mp), rtac CollectI,
- etac @{thm prefCl_Succ}, atac] 1;
-
-fun mk_strT_hset_tac n m j arg_cong_cTs cTs cts carT_defs strT_defs isNode_defs
- set_incl_hsets set_hset_incl_hsetss coalg_setss carT_setss coalgT set_naturalss =
- let
- val set_naturals = map (fn xs => nth xs (j - 1)) set_naturalss;
- val ks = 1 upto n;
- fun base_tac (i, (cT, (strT_def, (set_incl_hset, set_natural)))) =
- CONJ_WRAP' (fn (i', (carT_def, isNode_def)) => rtac impI THEN' etac conjE THEN'
- (if i = i'
- then EVERY' [rtac @{thm xt1(4)}, rtac set_incl_hset,
- rtac (strT_def RS arg_cong RS trans), etac (arg_cong RS trans),
- rtac (Thm.permute_prems 0 1 (set_natural RS box_equals)),
- rtac (trans OF [@{thm image_id} RS fun_cong, @{thm id_apply}]),
- rtac (mk_sum_casesN n i RS (Drule.instantiate' [cT] [] arg_cong) RS sym)]
- else EVERY' [dtac (carT_def RS equalityD1 RS set_mp),
- REPEAT_DETERM o eresolve_tac [CollectE, exE], etac conjE,
- dtac conjunct2, dtac Pair_eqD, etac conjE,
- hyp_subst_tac, dtac (isNode_def RS iffD1),
- REPEAT_DETERM o eresolve_tac [exE, conjE],
- rtac (mk_InN_not_InM i i' RS notE), etac (sym RS trans), atac]))
- (ks ~~ (carT_defs ~~ isNode_defs));
- fun step_tac (i, (coalg_sets, (carT_sets, set_hset_incl_hsets))) =
- dtac (mk_conjunctN n i) THEN'
- CONJ_WRAP' (fn (coalg_set, (carT_set, set_hset_incl_hset)) =>
- EVERY' [rtac impI, etac conjE, etac impE, rtac conjI,
- rtac (coalgT RS coalg_set RS set_mp), atac, etac carT_set, atac, atac,
- etac (@{thm shift_def} RS fun_cong RS trans), etac subset_trans,
- rtac set_hset_incl_hset, etac carT_set, atac, atac])
- (coalg_sets ~~ (carT_sets ~~ set_hset_incl_hsets));
- in
- EVERY' [rtac (Drule.instantiate' cTs cts @{thm list.induct}),
- REPEAT_DETERM o rtac allI, rtac impI,
- CONJ_WRAP' base_tac
- (ks ~~ (arg_cong_cTs ~~ (strT_defs ~~ (set_incl_hsets ~~ set_naturals)))),
- REPEAT_DETERM o rtac allI, rtac impI,
- REPEAT_DETERM o eresolve_tac [allE, impE], etac @{thm ShiftI},
- CONJ_WRAP' (fn i => dtac (mk_conjunctN n i) THEN' rtac (mk_sumEN n) THEN'
- CONJ_WRAP_GEN' (K all_tac) step_tac
- (ks ~~ (drop m coalg_setss ~~ (carT_setss ~~ set_hset_incl_hsetss)))) ks] 1
- end;
-
-fun mk_isNode_hset_tac n isNode_def strT_hsets {context = ctxt, prems = _} =
- let
- val m = length strT_hsets;
- in
- if m = 0 then atac 1
- else (unfold_thms_tac ctxt [isNode_def] THEN
- EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE],
- rtac exI, rtac conjI, atac,
- CONJ_WRAP' (fn (thm, i) => if i > m then atac
- else EVERY' [rtac (thm RS subset_trans), atac, rtac conjI, atac, atac, atac])
- (strT_hsets @ (replicate n mp) ~~ (1 upto (m + n)))] 1)
- end;
-
-fun mk_Lev_sbd_tac cts Lev_0s Lev_Sucs to_sbdss =
- let
- val n = length Lev_0s;
- val ks = 1 upto n;
- in
- EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
- REPEAT_DETERM o rtac allI,
- CONJ_WRAP' (fn Lev_0 =>
- EVERY' (map rtac [ord_eq_le_trans, Lev_0, @{thm Nil_clists}])) Lev_0s,
- REPEAT_DETERM o rtac allI,
- CONJ_WRAP' (fn (Lev_Suc, to_sbds) =>
- EVERY' [rtac ord_eq_le_trans, rtac Lev_Suc,
- CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_least}))
- (fn (i, to_sbd) => EVERY' [rtac subsetI,
- REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac,
- rtac @{thm Cons_clists}, rtac (mk_InN_Field n i), etac to_sbd,
- etac set_rev_mp, REPEAT_DETERM o etac allE,
- etac (mk_conjunctN n i)])
- (rev (ks ~~ to_sbds))])
- (Lev_Sucs ~~ to_sbdss)] 1
- end;
-
-fun mk_length_Lev_tac cts Lev_0s Lev_Sucs =
- let
- val n = length Lev_0s;
- val ks = n downto 1;
- in
- EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
- REPEAT_DETERM o rtac allI,
- CONJ_WRAP' (fn Lev_0 =>
- EVERY' [rtac impI, dtac (Lev_0 RS equalityD1 RS set_mp),
- etac @{thm singletonE}, etac ssubst, rtac @{thm list.size(3)}]) Lev_0s,
- REPEAT_DETERM o rtac allI,
- CONJ_WRAP' (fn Lev_Suc =>
- EVERY' [rtac impI, dtac (Lev_Suc RS equalityD1 RS set_mp),
- CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE))
- (fn i => EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac,
- rtac trans, rtac @{thm length_Cons}, rtac @{thm arg_cong[of _ _ Suc]},
- REPEAT_DETERM o etac allE, dtac (mk_conjunctN n i), etac mp, atac]) ks])
- Lev_Sucs] 1
- end;
-
-fun mk_length_Lev'_tac length_Lev =
- EVERY' [ftac length_Lev, etac ssubst, atac] 1;
-
-fun mk_prefCl_Lev_tac cts Lev_0s Lev_Sucs =
- let
- val n = length Lev_0s;
- val ks = n downto 1;
- in
- EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
- REPEAT_DETERM o rtac allI,
- CONJ_WRAP' (fn Lev_0 =>
- EVERY' [rtac impI, etac conjE, dtac (Lev_0 RS equalityD1 RS set_mp),
- etac @{thm singletonE}, hyp_subst_tac, dtac @{thm prefix_Nil[THEN subst, of "%x. x"]},
- hyp_subst_tac, rtac @{thm set_mp[OF equalityD2[OF trans[OF arg_cong[OF list.size(3)]]]]},
- rtac Lev_0, rtac @{thm singletonI}]) Lev_0s,
- REPEAT_DETERM o rtac allI,
- CONJ_WRAP' (fn (Lev_0, Lev_Suc) =>
- EVERY' [rtac impI, etac conjE, dtac (Lev_Suc RS equalityD1 RS set_mp),
- CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE))
- (fn i => EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac,
- dtac @{thm prefix_Cons[THEN subst, of "%x. x"]}, etac disjE, hyp_subst_tac,
- rtac @{thm set_mp[OF equalityD2[OF trans[OF arg_cong[OF list.size(3)]]]]},
- rtac Lev_0, rtac @{thm singletonI},
- REPEAT_DETERM o eresolve_tac [exE, conjE], hyp_subst_tac,
- rtac @{thm set_mp[OF equalityD2[OF trans[OF arg_cong[OF length_Cons]]]]},
- rtac Lev_Suc, rtac (mk_UnIN n i), rtac CollectI, REPEAT_DETERM o rtac exI, rtac conjI,
- rtac refl, etac conjI, REPEAT_DETERM o etac allE, dtac (mk_conjunctN n i),
- etac mp, etac conjI, atac]) ks])
- (Lev_0s ~~ Lev_Sucs)] 1
- end;
-
-fun mk_rv_last_tac cTs cts rv_Nils rv_Conss =
- let
- val n = length rv_Nils;
- val ks = 1 upto n;
- in
- EVERY' [rtac (Drule.instantiate' cTs cts @{thm list.induct}),
- REPEAT_DETERM o rtac allI,
- CONJ_WRAP' (fn rv_Cons =>
- CONJ_WRAP' (fn (i, rv_Nil) => (EVERY' [rtac exI,
- rtac (@{thm append_Nil} RS arg_cong RS trans),
- rtac (rv_Cons RS trans), rtac (mk_sum_casesN n i RS arg_cong RS trans), rtac rv_Nil]))
- (ks ~~ rv_Nils))
- rv_Conss,
- REPEAT_DETERM o rtac allI, rtac (mk_sumEN n),
- EVERY' (map (fn i =>
- CONJ_WRAP' (fn rv_Cons => EVERY' [REPEAT_DETERM o etac allE, dtac (mk_conjunctN n i),
- CONJ_WRAP' (fn i' => EVERY' [dtac (mk_conjunctN n i'), etac exE, rtac exI,
- rtac (@{thm append_Cons} RS arg_cong RS trans),
- rtac (rv_Cons RS trans), etac (sum_case_weak_cong RS arg_cong RS trans),
- rtac (mk_sum_casesN n i RS arg_cong RS trans), atac])
- ks])
- rv_Conss)
- ks)] 1
- end;
-
-fun mk_set_rv_Lev_tac m cts Lev_0s Lev_Sucs rv_Nils rv_Conss coalg_setss from_to_sbdss =
- let
- val n = length Lev_0s;
- val ks = 1 upto n;
- in
- EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
- REPEAT_DETERM o rtac allI,
- CONJ_WRAP' (fn (i, ((Lev_0, rv_Nil), coalg_sets)) =>
- EVERY' [rtac impI, REPEAT_DETERM o etac conjE,
- dtac (Lev_0 RS equalityD1 RS set_mp), etac @{thm singletonE}, etac ssubst,
- rtac (rv_Nil RS arg_cong RS iffD2),
- rtac (mk_sum_casesN n i RS iffD2),
- CONJ_WRAP' (fn thm => etac thm THEN' atac) (take m coalg_sets)])
- (ks ~~ ((Lev_0s ~~ rv_Nils) ~~ coalg_setss)),
- REPEAT_DETERM o rtac allI,
- CONJ_WRAP' (fn ((Lev_Suc, rv_Cons), (from_to_sbds, coalg_sets)) =>
- EVERY' [rtac impI, etac conjE, dtac (Lev_Suc RS equalityD1 RS set_mp),
- CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE))
- (fn (i, (from_to_sbd, coalg_set)) =>
- EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac,
- rtac (rv_Cons RS arg_cong RS iffD2),
- rtac (mk_sum_casesN n i RS arg_cong RS trans RS iffD2),
- etac (from_to_sbd RS arg_cong), REPEAT_DETERM o etac allE,
- dtac (mk_conjunctN n i), etac mp, etac conjI, etac set_rev_mp,
- etac coalg_set, atac])
- (rev (ks ~~ (from_to_sbds ~~ drop m coalg_sets)))])
- ((Lev_Sucs ~~ rv_Conss) ~~ (from_to_sbdss ~~ coalg_setss))] 1
- end;
-
-fun mk_set_Lev_tac cts Lev_0s Lev_Sucs rv_Nils rv_Conss from_to_sbdss =
- let
- val n = length Lev_0s;
- val ks = 1 upto n;
- in
- EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
- REPEAT_DETERM o rtac allI,
- CONJ_WRAP' (fn ((i, (Lev_0, Lev_Suc)), rv_Nil) =>
- EVERY' [rtac impI, dtac (Lev_0 RS equalityD1 RS set_mp),
- etac @{thm singletonE}, hyp_subst_tac,
- CONJ_WRAP' (fn i' => rtac impI THEN' dtac (sym RS trans) THEN' rtac rv_Nil THEN'
- (if i = i'
- then EVERY' [dtac (mk_InN_inject n i), hyp_subst_tac,
- CONJ_WRAP' (fn (i'', Lev_0'') =>
- EVERY' [rtac impI, rtac @{thm ssubst_mem[OF append_Nil]},
- rtac (Lev_Suc RS equalityD2 RS set_mp), rtac (mk_UnIN n i''),
- rtac CollectI, REPEAT_DETERM o rtac exI, rtac conjI, rtac refl,
- etac conjI, rtac (Lev_0'' RS equalityD2 RS set_mp),
- rtac @{thm singletonI}])
- (ks ~~ Lev_0s)]
- else etac (mk_InN_not_InM i' i RS notE)))
- ks])
- ((ks ~~ (Lev_0s ~~ Lev_Sucs)) ~~ rv_Nils),
- REPEAT_DETERM o rtac allI,
- CONJ_WRAP' (fn ((Lev_Suc, rv_Cons), from_to_sbds) =>
- EVERY' [rtac impI, dtac (Lev_Suc RS equalityD1 RS set_mp),
- CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE))
- (fn (i, from_to_sbd) =>
- EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac,
- CONJ_WRAP' (fn i' => rtac impI THEN'
- CONJ_WRAP' (fn i'' =>
- EVERY' [rtac impI, rtac (Lev_Suc RS equalityD2 RS set_mp),
- rtac @{thm ssubst_mem[OF append_Cons]}, rtac (mk_UnIN n i),
- rtac CollectI, REPEAT_DETERM o rtac exI, rtac conjI, rtac refl,
- rtac conjI, atac, dtac (sym RS trans RS sym),
- rtac (rv_Cons RS trans), rtac (mk_sum_casesN n i RS trans),
- etac (from_to_sbd RS arg_cong), REPEAT_DETERM o etac allE,
- dtac (mk_conjunctN n i), dtac mp, atac,
- dtac (mk_conjunctN n i'), dtac mp, atac,
- dtac (mk_conjunctN n i''), etac mp, atac])
- ks)
- ks])
- (rev (ks ~~ from_to_sbds))])
- ((Lev_Sucs ~~ rv_Conss) ~~ from_to_sbdss)] 1
- end;
-
-fun mk_set_image_Lev_tac cts Lev_0s Lev_Sucs rv_Nils rv_Conss from_to_sbdss to_sbd_injss =
- let
- val n = length Lev_0s;
- val ks = 1 upto n;
- in
- EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
- REPEAT_DETERM o rtac allI,
- CONJ_WRAP' (fn ((i, (Lev_0, Lev_Suc)), rv_Nil) =>
- EVERY' [rtac impI, dtac (Lev_0 RS equalityD1 RS set_mp),
- etac @{thm singletonE}, hyp_subst_tac,
- CONJ_WRAP' (fn i' => rtac impI THEN'
- CONJ_WRAP' (fn i'' => rtac impI THEN' dtac (sym RS trans) THEN' rtac rv_Nil THEN'
- (if i = i''
- then EVERY' [dtac @{thm ssubst_mem[OF sym[OF append_Nil]]},
- dtac (Lev_Suc RS equalityD1 RS set_mp), dtac (mk_InN_inject n i),
- hyp_subst_tac,
- CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE))
- (fn k => REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE] THEN'
- dtac list_inject_iffD1 THEN' etac conjE THEN'
- (if k = i'
- then EVERY' [dtac (mk_InN_inject n k), hyp_subst_tac, etac imageI]
- else etac (mk_InN_not_InM i' k RS notE)))
- (rev ks)]
- else etac (mk_InN_not_InM i'' i RS notE)))
- ks)
- ks])
- ((ks ~~ (Lev_0s ~~ Lev_Sucs)) ~~ rv_Nils),
- REPEAT_DETERM o rtac allI,
- CONJ_WRAP' (fn ((Lev_Suc, rv_Cons), (from_to_sbds, to_sbd_injs)) =>
- EVERY' [rtac impI, dtac (Lev_Suc RS equalityD1 RS set_mp),
- CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE))
- (fn (i, (from_to_sbd, to_sbd_inj)) =>
- REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE] THEN' hyp_subst_tac THEN'
- CONJ_WRAP' (fn i' => rtac impI THEN'
- dtac @{thm ssubst_mem[OF sym[OF append_Cons]]} THEN'
- dtac (Lev_Suc RS equalityD1 RS set_mp) THEN'
- CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE)) (fn k =>
- REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE] THEN'
- dtac list_inject_iffD1 THEN' etac conjE THEN'
- (if k = i
- then EVERY' [dtac (mk_InN_inject n i),
- dtac (Thm.permute_prems 0 2 (to_sbd_inj RS iffD1)),
- atac, atac, hyp_subst_tac] THEN'
- CONJ_WRAP' (fn i'' =>
- EVERY' [rtac impI, dtac (sym RS trans),
- rtac (rv_Cons RS trans), rtac (mk_sum_casesN n i RS arg_cong RS trans),
- etac (from_to_sbd RS arg_cong),
- REPEAT_DETERM o etac allE,
- dtac (mk_conjunctN n i), dtac mp, atac,
- dtac (mk_conjunctN n i'), dtac mp, atac,
- dtac (mk_conjunctN n i''), etac mp, etac sym])
- ks
- else etac (mk_InN_not_InM i k RS notE)))
- (rev ks))
- ks)
- (rev (ks ~~ (from_to_sbds ~~ to_sbd_injs)))])
- ((Lev_Sucs ~~ rv_Conss) ~~ (from_to_sbdss ~~ to_sbd_injss))] 1
- end;
-
-fun mk_mor_beh_tac m mor_def mor_cong beh_defs carT_defs strT_defs isNode_defs
- to_sbd_injss from_to_sbdss Lev_0s Lev_Sucs rv_Nils rv_Conss Lev_sbds length_Levs length_Lev's
- prefCl_Levs rv_lastss set_rv_Levsss set_Levsss set_image_Levsss set_naturalss coalg_setss
- map_comp_ids map_congs map_arg_congs {context = ctxt, prems = _} =
- let
- val n = length beh_defs;
- val ks = 1 upto n;
-
- fun fbetw_tac (i, (carT_def, (isNode_def, (Lev_0, (rv_Nil, (Lev_sbd,
- ((length_Lev, length_Lev'), (prefCl_Lev, (rv_lasts, (set_naturals,
- (coalg_sets, (set_rv_Levss, (set_Levss, set_image_Levss))))))))))))) =
- EVERY' [rtac ballI, rtac (carT_def RS equalityD2 RS set_mp),
- rtac CollectI, REPEAT_DETERM o rtac exI, rtac conjI, rtac refl, rtac conjI,
- rtac conjI,
- rtac @{thm UN_I}, rtac UNIV_I, rtac (Lev_0 RS equalityD2 RS set_mp),
- rtac @{thm singletonI},
- rtac conjI,
- rtac @{thm UN_least}, rtac Lev_sbd,
- rtac conjI,
- rtac @{thm prefCl_UN}, rtac ssubst, rtac @{thm PrefCl_def}, REPEAT_DETERM o rtac allI,
- rtac impI, etac conjE, rtac exI, rtac conjI, rtac @{thm ord_le_eq_trans},
- etac @{thm prefix_length_le}, etac length_Lev, rtac prefCl_Lev, etac conjI, atac,
- rtac conjI,
- rtac ballI, etac @{thm UN_E}, rtac conjI,
- if n = 1 then K all_tac else rtac (mk_sumEN n),
- EVERY' (map6 (fn i => fn isNode_def => fn set_naturals =>
- fn set_rv_Levs => fn set_Levs => fn set_image_Levs =>
- EVERY' [rtac (mk_disjIN n i), rtac (isNode_def RS ssubst),
- rtac exI, rtac conjI,
- (if n = 1 then rtac @{thm if_P} THEN' etac length_Lev'
- else rtac (@{thm if_P} RS arg_cong RS trans) THEN' etac length_Lev' THEN'
- etac (sum_case_weak_cong RS trans) THEN' rtac (mk_sum_casesN n i)),
- EVERY' (map2 (fn set_natural => fn set_rv_Lev =>
- EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac (set_natural RS trans),
- rtac trans_fun_cong_image_id_id_apply,
- etac set_rv_Lev, TRY o atac, etac conjI, atac])
- (take m set_naturals) set_rv_Levs),
- CONJ_WRAP' (fn (set_natural, (set_Lev, set_image_Lev)) =>
- EVERY' [rtac (set_natural RS trans), rtac equalityI, rtac @{thm image_subsetI},
- rtac CollectI, rtac @{thm SuccI}, rtac @{thm UN_I}, rtac UNIV_I, etac set_Lev,
- if n = 1 then rtac refl else atac, atac, rtac subsetI,
- REPEAT_DETERM o eresolve_tac [CollectE, @{thm SuccE}, @{thm UN_E}],
- rtac set_image_Lev, atac, dtac length_Lev, hyp_subst_tac, dtac length_Lev',
- etac @{thm set_mp[OF equalityD1[OF arg_cong[OF length_append_singleton]]]},
- if n = 1 then rtac refl else atac])
- (drop m set_naturals ~~ (set_Levs ~~ set_image_Levs))])
- ks isNode_defs set_naturalss set_rv_Levss set_Levss set_image_Levss),
- CONJ_WRAP' (fn (i, (rv_last, (isNode_def, (set_naturals,
- (set_rv_Levs, (set_Levs, set_image_Levs)))))) =>
- EVERY' [rtac ballI,
- REPEAT_DETERM o eresolve_tac [CollectE, @{thm SuccE}, @{thm UN_E}],
- rtac (rev_mp OF [rv_last, impI]), etac exE, rtac (isNode_def RS ssubst),
- rtac exI, rtac conjI,
- (if n = 1 then rtac @{thm if_P} THEN' etac length_Lev'
- else rtac (@{thm if_P} RS trans) THEN' etac length_Lev' THEN'
- etac (sum_case_weak_cong RS trans) THEN' rtac (mk_sum_casesN n i)),
- EVERY' (map2 (fn set_natural => fn set_rv_Lev =>
- EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac (set_natural RS trans),
- rtac trans_fun_cong_image_id_id_apply,
- etac set_rv_Lev, TRY o atac, etac conjI, atac])
- (take m set_naturals) set_rv_Levs),
- CONJ_WRAP' (fn (set_natural, (set_Lev, set_image_Lev)) =>
- EVERY' [rtac (set_natural RS trans), rtac equalityI, rtac @{thm image_subsetI},
- rtac CollectI, rtac @{thm SuccI}, rtac @{thm UN_I}, rtac UNIV_I, etac set_Lev,
- if n = 1 then rtac refl else atac, atac, rtac subsetI,
- REPEAT_DETERM o eresolve_tac [CollectE, @{thm SuccE}, @{thm UN_E}],
- REPEAT_DETERM_N 4 o etac thin_rl,
- rtac set_image_Lev,
- atac, dtac length_Lev, hyp_subst_tac, dtac length_Lev',
- etac @{thm set_mp[OF equalityD1[OF arg_cong[OF length_append_singleton]]]},
- if n = 1 then rtac refl else atac])
- (drop m set_naturals ~~ (set_Levs ~~ set_image_Levs))])
- (ks ~~ (rv_lasts ~~ (isNode_defs ~~ (set_naturalss ~~
- (set_rv_Levss ~~ (set_Levss ~~ set_image_Levss)))))),
- (**)
- rtac allI, rtac impI, rtac @{thm if_not_P}, rtac notI,
- etac notE, etac @{thm UN_I[OF UNIV_I]},
- (*root isNode*)
- rtac (isNode_def RS ssubst), rtac exI, rtac conjI, rtac (@{thm if_P} RS trans),
- rtac length_Lev', rtac (Lev_0 RS equalityD2 RS set_mp), rtac @{thm singletonI},
- CONVERSION (Conv.top_conv
- (K (Conv.try_conv (Conv.rewr_conv (rv_Nil RS eq_reflection)))) ctxt),
- if n = 1 then rtac refl else rtac (mk_sum_casesN n i),
- EVERY' (map2 (fn set_natural => fn coalg_set =>
- EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac (set_natural RS trans),
- rtac trans_fun_cong_image_id_id_apply, etac coalg_set, atac])
- (take m set_naturals) (take m coalg_sets)),
- CONJ_WRAP' (fn (set_natural, (set_Lev, set_image_Lev)) =>
- EVERY' [rtac (set_natural RS trans), rtac equalityI, rtac @{thm image_subsetI},
- rtac CollectI, rtac @{thm SuccI}, rtac @{thm UN_I}, rtac UNIV_I, rtac set_Lev,
- rtac (Lev_0 RS equalityD2 RS set_mp), rtac @{thm singletonI}, rtac rv_Nil,
- atac, rtac subsetI,
- REPEAT_DETERM o eresolve_tac [CollectE, @{thm SuccE}, @{thm UN_E}],
- rtac set_image_Lev, rtac (Lev_0 RS equalityD2 RS set_mp),
- rtac @{thm singletonI}, dtac length_Lev',
- etac @{thm set_mp[OF equalityD1[OF arg_cong[OF
- trans[OF length_append_singleton arg_cong[of _ _ Suc, OF list.size(3)]]]]]},
- rtac rv_Nil])
- (drop m set_naturals ~~ (nth set_Levss (i - 1) ~~ nth set_image_Levss (i - 1)))];
-
- fun mor_tac (i, (strT_def, (((Lev_0, Lev_Suc), (rv_Nil, rv_Cons)),
- ((map_comp_id, (map_cong, map_arg_cong)), (length_Lev', (from_to_sbds, to_sbd_injs)))))) =
- EVERY' [rtac ballI, rtac sym, rtac trans, rtac strT_def,
- rtac (@{thm if_P} RS
- (if n = 1 then map_arg_cong else sum_case_weak_cong) RS trans),
- rtac (@{thm list.size(3)} RS arg_cong RS trans RS equalityD2 RS set_mp),
- rtac Lev_0, rtac @{thm singletonI},
- CONVERSION (Conv.top_conv
- (K (Conv.try_conv (Conv.rewr_conv (rv_Nil RS eq_reflection)))) ctxt),
- if n = 1 then K all_tac
- else (rtac (sum_case_weak_cong RS trans) THEN'
- rtac (mk_sum_casesN n i) THEN' rtac (mk_sum_casesN n i RS trans)),
- rtac (map_comp_id RS trans), rtac (map_cong OF replicate m refl),
- EVERY' (map3 (fn i' => fn to_sbd_inj => fn from_to_sbd =>
- DETERM o EVERY' [rtac trans, rtac o_apply, rtac Pair_eqI, rtac conjI,
- rtac trans, rtac @{thm Shift_def},
- rtac equalityI, rtac subsetI, etac thin_rl, etac thin_rl,
- REPEAT_DETERM o eresolve_tac [CollectE, @{thm UN_E}], dtac length_Lev', dtac asm_rl,
- etac thin_rl, dtac @{thm set_rev_mp[OF _ equalityD1]},
- rtac (@{thm length_Cons} RS arg_cong RS trans), rtac Lev_Suc,
- CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE)) (fn i'' =>
- EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE],
- dtac list_inject_iffD1, etac conjE,
- if i' = i'' then EVERY' [dtac (mk_InN_inject n i'),
- dtac (Thm.permute_prems 0 2 (to_sbd_inj RS iffD1)),
- atac, atac, hyp_subst_tac, etac @{thm UN_I[OF UNIV_I]}]
- else etac (mk_InN_not_InM i' i'' RS notE)])
- (rev ks),
- rtac @{thm UN_least}, rtac subsetI, rtac CollectI, rtac @{thm UN_I[OF UNIV_I]},
- rtac (Lev_Suc RS equalityD2 RS set_mp), rtac (mk_UnIN n i'), rtac CollectI,
- REPEAT_DETERM o rtac exI, rtac conjI, rtac refl, etac conjI, atac,
- rtac trans, rtac @{thm shift_def}, rtac ssubst, rtac @{thm fun_eq_iff}, rtac allI,
- rtac @{thm if_cong}, rtac (@{thm length_Cons} RS arg_cong RS trans), rtac iffI,
- dtac asm_rl, dtac asm_rl, dtac asm_rl,
- dtac (Lev_Suc RS equalityD1 RS set_mp),
- CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE)) (fn i'' =>
- EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE],
- dtac list_inject_iffD1, etac conjE,
- if i' = i'' then EVERY' [dtac (mk_InN_inject n i'),
- dtac (Thm.permute_prems 0 2 (to_sbd_inj RS iffD1)),
- atac, atac, hyp_subst_tac, atac]
- else etac (mk_InN_not_InM i' i'' RS notE)])
- (rev ks),
- rtac (Lev_Suc RS equalityD2 RS set_mp), rtac (mk_UnIN n i'), rtac CollectI,
- REPEAT_DETERM o rtac exI, rtac conjI, rtac refl, etac conjI, atac,
- CONVERSION (Conv.top_conv
- (K (Conv.try_conv (Conv.rewr_conv (rv_Cons RS eq_reflection)))) ctxt),
- if n = 1 then K all_tac
- else rtac sum_case_weak_cong THEN' rtac (mk_sum_casesN n i' RS trans),
- SELECT_GOAL (unfold_thms_tac ctxt [from_to_sbd]), rtac refl,
- rtac refl])
- ks to_sbd_injs from_to_sbds)];
- in
- (rtac mor_cong THEN'
- EVERY' (map (fn thm => rtac (thm RS ext)) beh_defs) THEN'
- stac mor_def THEN' rtac conjI THEN'
- CONJ_WRAP' fbetw_tac
- (ks ~~ (carT_defs ~~ (isNode_defs ~~ (Lev_0s ~~ (rv_Nils ~~ (Lev_sbds ~~
- ((length_Levs ~~ length_Lev's) ~~ (prefCl_Levs ~~ (rv_lastss ~~
- (set_naturalss ~~ (coalg_setss ~~
- (set_rv_Levsss ~~ (set_Levsss ~~ set_image_Levsss))))))))))))) THEN'
- CONJ_WRAP' mor_tac
- (ks ~~ (strT_defs ~~ (((Lev_0s ~~ Lev_Sucs) ~~ (rv_Nils ~~ rv_Conss)) ~~
- ((map_comp_ids ~~ (map_congs ~~ map_arg_congs)) ~~
- (length_Lev's ~~ (from_to_sbdss ~~ to_sbd_injss))))))) 1
- end;
-
-fun mk_congruent_str_final_tac m lsbisE map_comp_id map_cong equiv_LSBISs =
- EVERY' [rtac @{thm congruentI}, dtac lsbisE,
- REPEAT_DETERM o eresolve_tac [CollectE, conjE, bexE], rtac (o_apply RS trans),
- etac (sym RS arg_cong RS trans), rtac (map_comp_id RS trans),
- rtac (map_cong RS trans), REPEAT_DETERM_N m o rtac refl,
- EVERY' (map (fn equiv_LSBIS =>
- EVERY' [rtac @{thm equiv_proj}, rtac equiv_LSBIS, etac set_mp, atac])
- equiv_LSBISs), rtac sym, rtac (o_apply RS trans),
- etac (sym RS arg_cong RS trans), rtac map_comp_id] 1;
-
-fun mk_coalg_final_tac m coalg_def congruent_str_finals equiv_LSBISs set_naturalss coalgT_setss =
- EVERY' [stac coalg_def,
- CONJ_WRAP' (fn ((set_naturals, coalgT_sets), (equiv_LSBIS, congruent_str_final)) =>
- EVERY' [rtac @{thm univ_preserves}, rtac equiv_LSBIS, rtac congruent_str_final,
- rtac ballI, rtac @{thm ssubst_mem}, rtac o_apply, rtac CollectI,
- EVERY' (map2 (fn set_natural => fn coalgT_set =>
- EVERY' [rtac conjI, rtac (set_natural RS ord_eq_le_trans),
- rtac ord_eq_le_trans_trans_fun_cong_image_id_id_apply,
- etac coalgT_set])
- (take m set_naturals) (take m coalgT_sets)),
- CONJ_WRAP' (fn (equiv_LSBIS, (set_natural, coalgT_set)) =>
- EVERY' [rtac (set_natural RS ord_eq_le_trans),
- rtac @{thm image_subsetI}, rtac ssubst, rtac @{thm proj_in_iff},
- rtac equiv_LSBIS, etac set_rev_mp, etac coalgT_set])
- (equiv_LSBISs ~~ drop m (set_naturals ~~ coalgT_sets))])
- ((set_naturalss ~~ coalgT_setss) ~~ (equiv_LSBISs ~~ congruent_str_finals))] 1;
-
-fun mk_mor_T_final_tac mor_def congruent_str_finals equiv_LSBISs =
- EVERY' [stac mor_def, rtac conjI,
- CONJ_WRAP' (fn equiv_LSBIS =>
- EVERY' [rtac ballI, rtac ssubst, rtac @{thm proj_in_iff}, rtac equiv_LSBIS, atac])
- equiv_LSBISs,
- CONJ_WRAP' (fn (equiv_LSBIS, congruent_str_final) =>
- EVERY' [rtac ballI, rtac sym, rtac trans, rtac @{thm univ_commute}, rtac equiv_LSBIS,
- rtac congruent_str_final, atac, rtac o_apply])
- (equiv_LSBISs ~~ congruent_str_finals)] 1;
-
-fun mk_mor_Rep_tac m defs Reps Abs_inverses coalg_final_setss map_comp_ids map_congLs
- {context = ctxt, prems = _} =
- unfold_thms_tac ctxt defs THEN
- EVERY' [rtac conjI,
- CONJ_WRAP' (fn thm => rtac ballI THEN' rtac thm) Reps,
- CONJ_WRAP' (fn (Rep, ((map_comp_id, map_congL), coalg_final_sets)) =>
- EVERY' [rtac ballI, rtac (map_comp_id RS trans), rtac map_congL,
- EVERY' (map2 (fn Abs_inverse => fn coalg_final_set =>
- EVERY' [rtac ballI, rtac (o_apply RS trans), rtac Abs_inverse,
- etac set_rev_mp, rtac coalg_final_set, rtac Rep])
- Abs_inverses (drop m coalg_final_sets))])
- (Reps ~~ ((map_comp_ids ~~ map_congLs) ~~ coalg_final_setss))] 1;
-
-fun mk_mor_Abs_tac defs Abs_inverses {context = ctxt, prems = _} =
- unfold_thms_tac ctxt defs THEN
- EVERY' [rtac conjI,
- CONJ_WRAP' (K (rtac ballI THEN' rtac UNIV_I)) Abs_inverses,
- CONJ_WRAP' (fn thm => rtac ballI THEN' etac (thm RS arg_cong RS sym)) Abs_inverses] 1;
-
-fun mk_mor_unfold_tac m mor_UNIV dtor_defs unfold_defs Abs_inverses morEs map_comp_ids map_congs =
- EVERY' [rtac iffD2, rtac mor_UNIV,
- CONJ_WRAP' (fn ((Abs_inverse, morE), ((dtor_def, unfold_def), (map_comp_id, map_cong))) =>
- EVERY' [rtac ext, rtac (o_apply RS trans RS sym), rtac (dtor_def RS trans),
- rtac (unfold_def RS arg_cong RS trans), rtac (Abs_inverse RS arg_cong RS trans),
- rtac (morE RS arg_cong RS trans), rtac (map_comp_id RS trans),
- rtac (o_apply RS trans RS sym), rtac map_cong,
- REPEAT_DETERM_N m o rtac refl,
- EVERY' (map (fn thm => rtac (thm RS trans) THEN' rtac (o_apply RS sym)) unfold_defs)])
- ((Abs_inverses ~~ morEs) ~~ ((dtor_defs ~~ unfold_defs) ~~ (map_comp_ids ~~ map_congs)))] 1;
-
-fun mk_raw_coind_tac bis_def bis_cong bis_O bis_converse bis_Gr tcoalg coalgT mor_T_final
- sbis_lsbis lsbis_incls incl_lsbiss equiv_LSBISs mor_Rep Rep_injects =
- let
- val n = length Rep_injects;
- in
- EVERY' [rtac rev_mp, ftac (bis_def RS iffD1),
- REPEAT_DETERM o etac conjE, rtac bis_cong, rtac bis_O, rtac bis_converse,
- rtac bis_Gr, rtac tcoalg, rtac mor_Rep, rtac bis_O, atac, rtac bis_Gr, rtac tcoalg,
- rtac mor_Rep, REPEAT_DETERM_N n o etac @{thm relImage_Gr},
- rtac impI, rtac rev_mp, rtac bis_cong, rtac bis_O, rtac bis_Gr, rtac coalgT,
- rtac mor_T_final, rtac bis_O, rtac sbis_lsbis, rtac bis_converse, rtac bis_Gr, rtac coalgT,
- rtac mor_T_final, EVERY' (map (fn thm => rtac (thm RS @{thm relInvImage_Gr})) lsbis_incls),
- rtac impI,
- CONJ_WRAP' (fn (Rep_inject, (equiv_LSBIS , (incl_lsbis, lsbis_incl))) =>
- EVERY' [rtac subset_trans, rtac @{thm relInvImage_UNIV_relImage}, rtac subset_trans,
- rtac @{thm relInvImage_mono}, rtac subset_trans, etac incl_lsbis,
- rtac ord_eq_le_trans, rtac @{thm sym[OF relImage_relInvImage]},
- rtac @{thm xt1(3)}, rtac @{thm Sigma_cong},
- rtac @{thm proj_image}, rtac @{thm proj_image}, rtac lsbis_incl,
- rtac subset_trans, rtac @{thm relImage_mono}, rtac incl_lsbis, atac,
- rtac @{thm relImage_proj}, rtac equiv_LSBIS, rtac @{thm relInvImage_diag},
- rtac Rep_inject])
- (Rep_injects ~~ (equiv_LSBISs ~~ (incl_lsbiss ~~ lsbis_incls)))] 1
- end;
-
-fun mk_unique_mor_tac raw_coinds bis =
- CONJ_WRAP' (fn raw_coind =>
- EVERY' [rtac impI, rtac (bis RS raw_coind RS set_mp RS @{thm IdD}), atac,
- etac conjunct1, atac, etac conjunct2, rtac @{thm image2_eqI}, rtac refl, rtac refl, atac])
- raw_coinds 1;
-
-fun mk_unfold_unique_mor_tac raw_coinds bis mor unfold_defs =
- CONJ_WRAP' (fn (raw_coind, unfold_def) =>
- EVERY' [rtac ext, etac (bis RS raw_coind RS set_mp RS @{thm IdD}), rtac mor,
- rtac @{thm image2_eqI}, rtac refl, rtac (unfold_def RS arg_cong RS trans),
- rtac (o_apply RS sym), rtac UNIV_I]) (raw_coinds ~~ unfold_defs) 1;
-
-fun mk_dtor_o_ctor_tac ctor_def unfold map_comp_id map_congL unfold_o_dtors
- {context = ctxt, prems = _} =
- unfold_thms_tac ctxt [ctor_def] THEN EVERY' [rtac ext, rtac trans, rtac o_apply,
- rtac trans, rtac unfold, rtac trans, rtac map_comp_id, rtac trans, rtac map_congL,
- EVERY' (map (fn thm =>
- rtac ballI THEN' rtac (trans OF [thm RS fun_cong, @{thm id_apply}])) unfold_o_dtors),
- rtac sym, rtac @{thm id_apply}] 1;
-
-fun mk_corec_tac m corec_defs unfold map_cong corec_Inls {context = ctxt, prems = _} =
- unfold_thms_tac ctxt corec_defs THEN EVERY' [rtac trans, rtac (o_apply RS arg_cong),
- rtac trans, rtac unfold, fo_rtac (@{thm sum.cases(2)} RS arg_cong RS trans) ctxt, rtac map_cong,
- REPEAT_DETERM_N m o rtac refl,
- EVERY' (map (fn thm => rtac @{thm sum_case_expand_Inr} THEN' rtac thm) corec_Inls)] 1;
-
-fun mk_srel_coinduct_tac ks raw_coind bis_srel =
- EVERY' [rtac rev_mp, rtac raw_coind, rtac ssubst, rtac bis_srel, rtac conjI,
- CONJ_WRAP' (K (rtac @{thm ord_le_eq_trans[OF subset_UNIV UNIV_Times_UNIV[THEN sym]]})) ks,
- CONJ_WRAP' (K (EVERY' [rtac allI, rtac allI, rtac impI,
- REPEAT_DETERM o etac allE, etac mp, etac CollectE, etac @{thm splitD}])) ks,
- rtac impI, REPEAT_DETERM o etac conjE,
- CONJ_WRAP' (K (EVERY' [rtac impI, rtac @{thm IdD}, etac set_mp,
- rtac CollectI, etac @{thm prod_caseI}])) ks] 1;
-
-fun mk_srel_strong_coinduct_tac m cTs cts srel_coinduct srel_monos srel_Ids =
- EVERY' [rtac rev_mp, rtac (Drule.instantiate' cTs cts srel_coinduct),
- EVERY' (map2 (fn srel_mono => fn srel_Id =>
- EVERY' [REPEAT_DETERM o resolve_tac [allI, impI], REPEAT_DETERM o etac allE,
- etac disjE, etac mp, atac, hyp_subst_tac, rtac (srel_mono RS set_mp),
- REPEAT_DETERM_N m o rtac @{thm subset_refl},
- REPEAT_DETERM_N (length srel_monos) o rtac @{thm Id_subset},
- rtac (srel_Id RS equalityD2 RS set_mp), rtac @{thm IdI}])
- srel_monos srel_Ids),
- rtac impI, REPEAT_DETERM o etac conjE,
- CONJ_WRAP' (K (rtac impI THEN' etac mp THEN' etac disjI1)) srel_Ids] 1;
-
-fun mk_dtor_coinduct_tac m ks raw_coind bis_def =
- let
- val n = length ks;
- in
- EVERY' [rtac rev_mp, rtac raw_coind, rtac ssubst, rtac bis_def, rtac conjI,
- CONJ_WRAP' (K (rtac @{thm ord_le_eq_trans[OF subset_UNIV UNIV_Times_UNIV[THEN sym]]})) ks,
- CONJ_WRAP' (fn i => EVERY' [select_prem_tac n (dtac asm_rl) i, REPEAT_DETERM o rtac allI,
- rtac impI, REPEAT_DETERM o dtac @{thm meta_spec}, etac CollectE, etac @{thm meta_impE},
- atac, etac exE, etac conjE, etac conjE, rtac bexI, rtac conjI,
- etac @{thm fst_conv[THEN subst]}, etac @{thm snd_conv[THEN subst]},
- rtac CollectI, REPEAT_DETERM_N m o (rtac conjI THEN' rtac subset_UNIV),
- CONJ_WRAP' (fn i' => EVERY' [rtac subsetI, rtac CollectI, dtac (mk_conjunctN n i'),
- REPEAT_DETERM o etac allE, etac mp, rtac @{thm ssubst_mem[OF pair_collapse]}, atac])
- ks])
- ks,
- rtac impI,
- CONJ_WRAP' (fn i => EVERY' [rtac impI, dtac (mk_conjunctN n i),
- rtac @{thm subst[OF pair_in_Id_conv]}, etac set_mp,
- rtac CollectI, etac (refl RSN (2, @{thm subst_Pair}))]) ks] 1
- end;
-
-fun mk_dtor_strong_coinduct_tac ks cTs cts dtor_coinduct bis_def bis_diag =
- EVERY' [rtac rev_mp, rtac (Drule.instantiate' cTs cts dtor_coinduct),
- EVERY' (map (fn i =>
- EVERY' [etac disjE, REPEAT_DETERM o dtac @{thm meta_spec}, etac @{thm meta_mp},
- atac, rtac rev_mp, rtac subst, rtac bis_def, rtac bis_diag,
- rtac impI, dtac conjunct2, dtac (mk_conjunctN (length ks) i), REPEAT_DETERM o etac allE,
- etac impE, etac @{thm diag_UNIV_I}, REPEAT_DETERM o eresolve_tac [bexE, conjE, CollectE],
- rtac exI, rtac conjI, etac conjI, atac,
- CONJ_WRAP' (K (EVERY' [REPEAT_DETERM o resolve_tac [allI, impI],
- rtac disjI2, rtac @{thm diagE}, etac set_mp, atac])) ks])
- ks),
- rtac impI, REPEAT_DETERM o etac conjE,
- CONJ_WRAP' (K (rtac impI THEN' etac mp THEN' etac disjI1)) ks] 1;
-
-fun mk_map_tac m n cT unfold map_comp' map_cong =
- EVERY' [rtac ext, rtac (o_apply RS trans RS sym), rtac (o_apply RS trans RS sym),
- rtac (unfold RS trans), rtac (Thm.permute_prems 0 1 (map_comp' RS box_equals)), rtac map_cong,
- REPEAT_DETERM_N m o rtac (@{thm id_o} RS fun_cong),
- REPEAT_DETERM_N n o rtac (@{thm o_id} RS fun_cong),
- rtac (o_apply RS (Drule.instantiate' [cT] [] arg_cong) RS sym)] 1;
-
-fun mk_set_le_tac n hset_minimal set_hsets set_hset_hsetss =
- EVERY' [rtac hset_minimal,
- REPEAT_DETERM_N n o rtac @{thm Un_upper1},
- REPEAT_DETERM_N n o
- EVERY' (map3 (fn i => fn set_hset => fn set_hset_hsets =>
- EVERY' [rtac subsetI, rtac @{thm UnI2}, rtac (mk_UnIN n i), etac @{thm UN_I},
- etac UnE, etac set_hset, REPEAT_DETERM_N (n - 1) o etac UnE,
- EVERY' (map (fn thm => EVERY' [etac @{thm UN_E}, etac thm, atac]) set_hset_hsets)])
- (1 upto n) set_hsets set_hset_hsetss)] 1;
-
-fun mk_set_simp_tac n set_le set_incl_hset set_hset_incl_hsets =
- EVERY' [rtac equalityI, rtac set_le, rtac @{thm Un_least}, rtac set_incl_hset,
- REPEAT_DETERM_N (n - 1) o rtac @{thm Un_least},
- EVERY' (map (fn thm => rtac @{thm UN_least} THEN' etac thm) set_hset_incl_hsets)] 1;
-
-fun mk_map_id_tac maps unfold_unique unfold_dtor =
- EVERY' [rtac (unfold_unique RS trans), EVERY' (map (fn thm => rtac (thm RS sym)) maps),
- rtac unfold_dtor] 1;
-
-fun mk_map_comp_tac m n maps map_comps map_congs unfold_unique =
- EVERY' [rtac unfold_unique,
- EVERY' (map3 (fn map_thm => fn map_comp => fn map_cong =>
- EVERY' (map rtac
- ([@{thm o_assoc} RS trans,
- @{thm arg_cong2[of _ _ _ _ "op o"]} OF [map_comp RS sym, refl] RS trans,
- @{thm o_assoc} RS trans RS sym,
- @{thm arg_cong2[of _ _ _ _ "op o"]} OF [map_thm, refl] RS trans,
- @{thm o_assoc} RS sym RS trans, map_thm RS arg_cong RS trans, @{thm o_assoc} RS trans,
- @{thm arg_cong2[of _ _ _ _ "op o"]} OF [map_comp RS sym, refl] RS trans,
- ext, o_apply RS trans, o_apply RS trans RS sym, map_cong] @
- replicate m (@{thm id_o} RS fun_cong) @
- replicate n (@{thm o_id} RS fun_cong))))
- maps map_comps map_congs)] 1;
-
-fun mk_mcong_tac m coinduct_tac map_comp's map_simps map_congs set_naturalss set_hsetss
- set_hset_hsetsss =
- let
- val n = length map_comp's;
- val ks = 1 upto n;
- in
- EVERY' ([rtac rev_mp,
- coinduct_tac] @
- maps (fn (((((map_comp'_trans, map_simps_trans), map_cong), set_naturals), set_hsets),
- set_hset_hsetss) =>
- [REPEAT_DETERM o eresolve_tac [exE, conjE], hyp_subst_tac, rtac exI, rtac conjI, rtac conjI,
- rtac map_comp'_trans, rtac sym, rtac map_simps_trans, rtac map_cong,
- REPEAT_DETERM_N m o (rtac o_apply_trans_sym THEN' rtac @{thm id_apply}),
- REPEAT_DETERM_N n o rtac fst_convol_fun_cong_sym,
- rtac map_comp'_trans, rtac sym, rtac map_simps_trans, rtac map_cong,
- EVERY' (maps (fn set_hset =>
- [rtac o_apply_trans_sym, rtac (@{thm id_apply} RS trans), etac CollectE,
- REPEAT_DETERM o etac conjE, etac bspec, etac set_hset]) set_hsets),
- REPEAT_DETERM_N n o rtac snd_convol_fun_cong_sym,
- CONJ_WRAP' (fn (set_natural, set_hset_hsets) =>
- EVERY' [REPEAT_DETERM o rtac allI, rtac impI, rtac @{thm image_convolD},
- etac set_rev_mp, rtac ord_eq_le_trans, rtac set_natural,
- rtac @{thm image_mono}, rtac subsetI, rtac CollectI, etac CollectE,
- REPEAT_DETERM o etac conjE,
- CONJ_WRAP' (fn set_hset_hset =>
- EVERY' [rtac ballI, etac bspec, etac set_hset_hset, atac]) set_hset_hsets])
- (drop m set_naturals ~~ set_hset_hsetss)])
- (map (fn th => th RS trans) map_comp's ~~ map (fn th => th RS trans) map_simps ~~
- map_congs ~~ set_naturalss ~~ set_hsetss ~~ set_hset_hsetsss) @
- [rtac impI,
- CONJ_WRAP' (fn k =>
- EVERY' [rtac impI, dtac (mk_conjunctN n k), etac mp, rtac exI, rtac conjI, etac CollectI,
- rtac conjI, rtac refl, rtac refl]) ks]) 1
- end
-
-fun mk_map_unique_tac unfold_unique map_comps {context = ctxt, prems = _} =
- rtac unfold_unique 1 THEN
- unfold_thms_tac ctxt (map (fn thm => thm RS sym) map_comps @ @{thms o_assoc id_o o_id}) THEN
- ALLGOALS (etac sym);
-
-fun mk_col_natural_tac cts rec_0s rec_Sucs map_simps set_naturalss
- {context = ctxt, prems = _} =
- let
- val n = length map_simps;
- in
- EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
- REPEAT_DETERM o rtac allI, SELECT_GOAL (unfold_thms_tac ctxt rec_0s),
- CONJ_WRAP' (K (rtac @{thm image_empty})) rec_0s,
- REPEAT_DETERM o rtac allI,
- CONJ_WRAP' (fn (rec_Suc, (map_simp, set_nats)) => EVERY'
- [SELECT_GOAL (unfold_thms_tac ctxt
- (rec_Suc :: map_simp :: set_nats @ @{thms image_Un image_UN UN_simps(10)})),
- rtac @{thm Un_cong}, rtac refl,
- CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_cong}))
- (fn i => EVERY' [rtac @{thm UN_cong[OF refl]},
- REPEAT_DETERM o etac allE, etac (mk_conjunctN n i)]) (n downto 1)])
- (rec_Sucs ~~ (map_simps ~~ set_naturalss))] 1
- end;
-
-fun mk_set_natural_tac hset_def col_natural =
- EVERY' (map rtac [ext, (o_apply RS trans), (hset_def RS trans), sym,
- (o_apply RS trans), (@{thm image_cong} OF [hset_def, refl] RS trans),
- (@{thm image_UN} RS trans), (refl RS @{thm UN_cong}), col_natural]) 1;
-
-fun mk_col_bd_tac m j cts rec_0s rec_Sucs sbd_Card_order sbd_Cinfinite set_sbdss =
- let
- val n = length rec_0s;
- in
- EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
- REPEAT_DETERM o rtac allI,
- CONJ_WRAP' (fn rec_0 => EVERY' (map rtac [ordIso_ordLeq_trans,
- @{thm card_of_ordIso_subst}, rec_0, @{thm Card_order_empty}, sbd_Card_order])) rec_0s,
- REPEAT_DETERM o rtac allI,
- CONJ_WRAP' (fn (rec_Suc, set_sbds) => EVERY'
- [rtac ordIso_ordLeq_trans, rtac @{thm card_of_ordIso_subst}, rtac rec_Suc,
- rtac (sbd_Cinfinite RSN (3, @{thm Un_Cinfinite_bound})), rtac (nth set_sbds (j - 1)),
- REPEAT_DETERM_N (n - 1) o rtac (sbd_Cinfinite RSN (3, @{thm Un_Cinfinite_bound})),
- EVERY' (map2 (fn i => fn set_sbd => EVERY' [rtac @{thm UNION_Cinfinite_bound},
- rtac set_sbd, rtac ballI, REPEAT_DETERM o etac allE,
- etac (mk_conjunctN n i), rtac sbd_Cinfinite]) (1 upto n) (drop m set_sbds))])
- (rec_Sucs ~~ set_sbdss)] 1
- end;
-
-fun mk_set_bd_tac sbd_Cinfinite hset_def col_bd =
- EVERY' (map rtac [ordIso_ordLeq_trans, @{thm card_of_ordIso_subst}, hset_def,
- ctrans, @{thm UNION_Cinfinite_bound}, ordIso_ordLeq_trans, @{thm card_of_nat},
- @{thm natLeq_ordLeq_cinfinite}, sbd_Cinfinite, ballI, col_bd, sbd_Cinfinite,
- ctrans, @{thm infinite_ordLeq_cexp}, sbd_Cinfinite, @{thm cexp_ordLeq_ccexp}]) 1;
-
-fun mk_in_bd_tac isNode_hset isNode_hsets carT_def card_of_carT mor_image Rep_inverse mor_hsets
- sbd_Cnotzero sbd_Card_order mor_Rep coalgT mor_T_final tcoalg =
- let
- val n = length isNode_hsets;
- val in_hin_tac = rtac CollectI THEN'
- CONJ_WRAP' (fn mor_hset => EVERY' (map etac
- [mor_hset OF [coalgT, mor_T_final] RS sym RS ord_eq_le_trans,
- arg_cong RS sym RS ord_eq_le_trans,
- mor_hset OF [tcoalg, mor_Rep, UNIV_I] RS ord_eq_le_trans])) mor_hsets;
- in
- EVERY' [rtac (Thm.permute_prems 0 1 @{thm ordLeq_transitive}), rtac ctrans,
- rtac @{thm card_of_image}, rtac ordIso_ordLeq_trans,
- rtac @{thm card_of_ordIso_subst}, rtac @{thm sym[OF proj_image]}, rtac ctrans,
- rtac @{thm card_of_image}, rtac ctrans, rtac card_of_carT, rtac @{thm cexp_mono2_Cnotzero},
- rtac @{thm cexp_ordLeq_ccexp}, rtac @{thm csum_Cnotzero2}, rtac @{thm ctwo_Cnotzero},
- rtac @{thm Cnotzero_cexp}, rtac sbd_Cnotzero, rtac sbd_Card_order,
- rtac @{thm card_of_mono1}, rtac subsetI, rtac @{thm image_eqI}, rtac sym,
- rtac Rep_inverse, REPEAT_DETERM o eresolve_tac [CollectE, conjE],
- rtac set_mp, rtac equalityD2, rtac @{thm sym[OF proj_image]}, rtac imageE,
- rtac set_rev_mp, rtac mor_image, rtac mor_Rep, rtac UNIV_I, rtac equalityD2,
- rtac @{thm proj_image}, rtac @{thm image_eqI}, atac,
- ftac (carT_def RS equalityD1 RS set_mp),
- REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac,
- rtac (carT_def RS equalityD2 RS set_mp), rtac CollectI, REPEAT_DETERM o rtac exI,
- rtac conjI, rtac refl, rtac conjI, etac conjI, etac conjI, etac conjI, rtac conjI,
- rtac ballI, dtac bspec, atac, REPEAT_DETERM o etac conjE, rtac conjI,
- CONJ_WRAP_GEN' (etac disjE) (fn (i, isNode_hset) =>
- EVERY' [rtac (mk_disjIN n i), rtac isNode_hset, atac, atac, atac, in_hin_tac])
- (1 upto n ~~ isNode_hsets),
- CONJ_WRAP' (fn isNode_hset =>
- EVERY' [rtac ballI, rtac isNode_hset, atac, ftac CollectD, etac @{thm SuccD},
- etac bspec, atac, in_hin_tac])
- isNode_hsets,
- atac, rtac isNode_hset, atac, atac, atac, in_hin_tac] 1
- end;
-
-fun mk_bd_card_order_tac sbd_card_order =
- EVERY' (map rtac [@{thm card_order_ccexp}, sbd_card_order, sbd_card_order]) 1;
-
-fun mk_bd_cinfinite_tac sbd_Cinfinite =
- EVERY' (map rtac [@{thm cinfinite_ccexp}, @{thm ctwo_ordLeq_Cinfinite},
- sbd_Cinfinite, sbd_Cinfinite]) 1;
-
-fun mk_pickWP_assms_tac set_incl_hsets set_incl_hins map_eq =
- let
- val m = length set_incl_hsets;
- in
- EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac CollectI,
- EVERY' (map (fn thm => rtac conjI THEN' etac (thm RS @{thm subset_trans})) set_incl_hsets),
- CONJ_WRAP' (fn thm => rtac thm THEN' REPEAT_DETERM_N m o atac) set_incl_hins,
- REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac CollectI,
- EVERY' (map (fn thm => rtac conjI THEN' etac (thm RS @{thm subset_trans})) set_incl_hsets),
- CONJ_WRAP' (fn thm => rtac thm THEN' REPEAT_DETERM_N m o atac) set_incl_hins,
- REPEAT_DETERM o eresolve_tac [CollectE, conjE], etac map_eq]
- end;
-
-fun mk_coalg_thePull_tac m coalg_def map_wpulls set_naturalss pickWP_assms_tacs
- {context = ctxt, prems = _} =
- unfold_thms_tac ctxt [coalg_def] THEN
- CONJ_WRAP' (fn (map_wpull, (pickWP_assms_tac, set_naturals)) =>
- EVERY' [rtac ballI, dtac @{thm set_mp[OF equalityD1[OF thePull_def]]},
- REPEAT_DETERM o eresolve_tac [CollectE, @{thm prod_caseE}],
- hyp_subst_tac, rtac rev_mp, rtac (map_wpull RS @{thm pickWP(1)}),
- EVERY' (map (etac o mk_conjunctN m) (1 upto m)),
- pickWP_assms_tac,
- SELECT_GOAL (unfold_thms_tac ctxt @{thms o_apply prod.cases}), rtac impI,
- REPEAT_DETERM o eresolve_tac [CollectE, conjE],
- rtac CollectI,
- REPEAT_DETERM_N m o (rtac conjI THEN' rtac subset_UNIV),
- CONJ_WRAP' (fn set_natural =>
- EVERY' [rtac ord_eq_le_trans, rtac trans, rtac set_natural,
- rtac trans_fun_cong_image_id_id_apply, atac])
- (drop m set_naturals)])
- (map_wpulls ~~ (pickWP_assms_tacs ~~ set_naturalss)) 1;
-
-fun mk_mor_thePull_nth_tac conv pick m mor_def map_wpulls map_comps pickWP_assms_tacs
- {context = ctxt, prems = _} =
- let
- val n = length map_comps;
- in
- unfold_thms_tac ctxt [mor_def] THEN
- EVERY' [rtac conjI,
- CONJ_WRAP' (K (rtac ballI THEN' rtac UNIV_I)) (1 upto n),
- CONJ_WRAP' (fn (map_wpull, (pickWP_assms_tac, map_comp)) =>
- EVERY' [rtac ballI, dtac @{thm set_mp[OF equalityD1[OF thePull_def]]},
- REPEAT_DETERM o eresolve_tac [CollectE, @{thm prod_caseE}, conjE],
- hyp_subst_tac,
- SELECT_GOAL (unfold_thms_tac ctxt @{thms o_apply prod.cases}),
- rtac (map_comp RS trans),
- SELECT_GOAL (unfold_thms_tac ctxt (conv :: @{thms o_id id_o})),
- rtac (map_wpull RS pick), REPEAT_DETERM_N m o atac,
- pickWP_assms_tac])
- (map_wpulls ~~ (pickWP_assms_tacs ~~ map_comps))] 1
- end;
-
-val mk_mor_thePull_fst_tac = mk_mor_thePull_nth_tac @{thm fst_conv} @{thm pickWP(2)};
-val mk_mor_thePull_snd_tac = mk_mor_thePull_nth_tac @{thm snd_conv} @{thm pickWP(3)};
-
-fun mk_mor_thePull_pick_tac mor_def unfolds map_comps {context = ctxt, prems = _} =
- unfold_thms_tac ctxt [mor_def, @{thm thePull_def}] THEN rtac conjI 1 THEN
- CONJ_WRAP' (K (rtac ballI THEN' rtac UNIV_I)) unfolds 1 THEN
- CONJ_WRAP' (fn (unfold, map_comp) =>
- EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, @{thm prod_caseE}, conjE],
- hyp_subst_tac,
- SELECT_GOAL (unfold_thms_tac ctxt (unfold :: map_comp :: @{thms comp_def id_def})),
- rtac refl])
- (unfolds ~~ map_comps) 1;
-
-fun mk_pick_col_tac m j cts rec_0s rec_Sucs unfolds set_naturalss map_wpulls pickWP_assms_tacs
- {context = ctxt, prems = _} =
- let
- val n = length rec_0s;
- val ks = n downto 1;
- in
- EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
- REPEAT_DETERM o rtac allI,
- CONJ_WRAP' (fn thm => EVERY'
- [rtac impI, rtac ord_eq_le_trans, rtac thm, rtac @{thm empty_subsetI}]) rec_0s,
- REPEAT_DETERM o rtac allI,
- CONJ_WRAP' (fn (rec_Suc, ((unfold, set_naturals), (map_wpull, pickWP_assms_tac))) =>
- EVERY' [rtac impI, dtac @{thm set_mp[OF equalityD1[OF thePull_def]]},
- REPEAT_DETERM o eresolve_tac [CollectE, @{thm prod_caseE}],
- hyp_subst_tac, rtac rev_mp, rtac (map_wpull RS @{thm pickWP(1)}),
- EVERY' (map (etac o mk_conjunctN m) (1 upto m)),
- pickWP_assms_tac,
- rtac impI, REPEAT_DETERM o eresolve_tac [CollectE, conjE],
- rtac ord_eq_le_trans, rtac rec_Suc,
- rtac @{thm Un_least},
- SELECT_GOAL (unfold_thms_tac ctxt [unfold, nth set_naturals (j - 1),
- @{thm prod.cases}]),
- etac ord_eq_le_trans_trans_fun_cong_image_id_id_apply,
- CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_least})) (fn (i, set_natural) =>
- EVERY' [rtac @{thm UN_least},
- SELECT_GOAL (unfold_thms_tac ctxt [unfold, set_natural, @{thm prod.cases}]),
- etac imageE, hyp_subst_tac, REPEAT_DETERM o etac allE,
- dtac (mk_conjunctN n i), etac mp, etac set_mp, atac])
- (ks ~~ rev (drop m set_naturals))])
- (rec_Sucs ~~ ((unfolds ~~ set_naturalss) ~~ (map_wpulls ~~ pickWP_assms_tacs)))] 1
- end;
-
-fun mk_wpull_tac m coalg_thePull mor_thePull_fst mor_thePull_snd mor_thePull_pick
- mor_unique pick_cols hset_defs =
- EVERY' [rtac (@{thm wpull_def} RS iffD2), REPEAT_DETERM o rtac allI, rtac impI,
- REPEAT_DETERM o etac conjE, rtac bexI, rtac conjI,
- rtac box_equals, rtac mor_unique,
- rtac coalg_thePull, REPEAT_DETERM_N (m - 1) o etac conjI, atac,
- rtac mor_thePull_pick, REPEAT_DETERM_N (m - 1) o etac conjI, atac,
- rtac mor_thePull_fst, REPEAT_DETERM_N (m - 1) o etac conjI, atac,
- rtac @{thm set_mp[OF equalityD2[OF thePull_def]]}, rtac CollectI,
- rtac @{thm prod_caseI}, etac conjI, etac conjI, atac, rtac o_apply, rtac @{thm fst_conv},
- rtac box_equals, rtac mor_unique,
- rtac coalg_thePull, REPEAT_DETERM_N (m - 1) o etac conjI, atac,
- rtac mor_thePull_pick, REPEAT_DETERM_N (m - 1) o etac conjI, atac,
- rtac mor_thePull_snd, REPEAT_DETERM_N (m - 1) o etac conjI, atac,
- rtac @{thm set_mp[OF equalityD2[OF thePull_def]]}, rtac CollectI,
- rtac @{thm prod_caseI}, etac conjI, etac conjI, atac, rtac o_apply, rtac @{thm snd_conv},
- rtac CollectI,
- CONJ_WRAP' (fn (pick, def) =>
- EVERY' [rtac (def RS ord_eq_le_trans), rtac @{thm UN_least},
- rtac pick, REPEAT_DETERM_N (m - 1) o etac conjI, atac,
- rtac @{thm set_mp[OF equalityD2[OF thePull_def]]}, rtac CollectI,
- rtac @{thm prod_caseI}, etac conjI, etac conjI, atac])
- (pick_cols ~~ hset_defs)] 1;
-
-fun mk_wit_tac n dtor_ctors set_simp wit coind_wits {context = ctxt, prems = _} =
- ALLGOALS (TRY o (eresolve_tac coind_wits THEN' rtac refl)) THEN
- REPEAT_DETERM (atac 1 ORELSE
- EVERY' [dtac set_rev_mp, rtac equalityD1, resolve_tac set_simp,
- K (unfold_thms_tac ctxt dtor_ctors),
- REPEAT_DETERM_N n o etac UnE,
- REPEAT_DETERM o
- (TRY o REPEAT_DETERM o etac UnE THEN' TRY o etac @{thm UN_E} THEN'
- (eresolve_tac wit ORELSE'
- (dresolve_tac wit THEN'
- (etac FalseE ORELSE'
- EVERY' [hyp_subst_tac, dtac set_rev_mp, rtac equalityD1, resolve_tac set_simp,
- K (unfold_thms_tac ctxt dtor_ctors), REPEAT_DETERM_N n o etac UnE]))))] 1);
-
-fun mk_coind_wit_tac induct unfolds set_nats wits {context = ctxt, prems = _} =
- rtac induct 1 THEN ALLGOALS (TRY o rtac impI THEN' TRY o hyp_subst_tac) THEN
- unfold_thms_tac ctxt (unfolds @ set_nats @ @{thms image_id id_apply}) THEN
- ALLGOALS (REPEAT_DETERM o etac imageE THEN' TRY o hyp_subst_tac) THEN
- ALLGOALS (TRY o
- FIRST' [rtac TrueI, rtac refl, etac (refl RSN (2, mp)), dresolve_tac wits THEN' etac FalseE])
-
-fun mk_srel_simp_tac in_Jsrels i in_srel map_comp map_cong map_simp set_simps dtor_inject dtor_ctor
- set_naturals set_incls set_set_inclss =
- let
- val m = length set_incls;
- val n = length set_set_inclss;
- val (passive_set_naturals, active_set_naturals) = chop m set_naturals;
- val in_Jsrel = nth in_Jsrels (i - 1);
- val if_tac =
- EVERY' [dtac (in_Jsrel RS iffD1), REPEAT_DETERM o eresolve_tac [exE, conjE, CollectE],
- rtac (in_srel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
- EVERY' (map2 (fn set_natural => fn set_incl =>
- EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac set_natural,
- rtac ord_eq_le_trans, rtac trans_fun_cong_image_id_id_apply,
- etac (set_incl RS @{thm subset_trans})])
- passive_set_naturals set_incls),
- CONJ_WRAP' (fn (in_Jsrel, (set_natural, set_set_incls)) =>
- EVERY' [rtac ord_eq_le_trans, rtac set_natural, rtac @{thm image_subsetI},
- rtac (in_Jsrel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
- CONJ_WRAP' (fn thm => etac (thm RS @{thm subset_trans}) THEN' atac) set_set_incls,
- rtac conjI, rtac refl, rtac refl])
- (in_Jsrels ~~ (active_set_naturals ~~ set_set_inclss)),
- CONJ_WRAP' (fn conv =>
- EVERY' [rtac trans, rtac map_comp, rtac trans, rtac map_cong,
- REPEAT_DETERM_N m o rtac @{thm fun_cong[OF o_id]},
- REPEAT_DETERM_N n o EVERY' (map rtac [trans, o_apply, conv]),
- rtac trans, rtac sym, rtac map_simp, rtac (dtor_inject RS iffD2), atac])
- @{thms fst_conv snd_conv}];
- val only_if_tac =
- EVERY' [dtac (in_srel RS iffD1), REPEAT_DETERM o eresolve_tac [exE, conjE, CollectE],
- rtac (in_Jsrel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
- CONJ_WRAP' (fn (set_simp, passive_set_natural) =>
- EVERY' [rtac ord_eq_le_trans, rtac set_simp, rtac @{thm Un_least},
- rtac ord_eq_le_trans, rtac box_equals, rtac passive_set_natural,
- rtac (dtor_ctor RS sym RS arg_cong), rtac trans_fun_cong_image_id_id_apply, atac,
- CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_least}))
- (fn (active_set_natural, in_Jsrel) => EVERY' [rtac ord_eq_le_trans,
- rtac @{thm UN_cong[OF _ refl]}, rtac @{thm box_equals[OF _ _ refl]},
- rtac active_set_natural, rtac (dtor_ctor RS sym RS arg_cong), rtac @{thm UN_least},
- dtac set_rev_mp, etac @{thm image_mono}, etac imageE,
- dtac @{thm ssubst_mem[OF pair_collapse]}, dtac (in_Jsrel RS iffD1),
- dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE,
- dtac (Thm.permute_prems 0 1 @{thm ssubst_mem}), atac,
- hyp_subst_tac, REPEAT_DETERM o eresolve_tac [CollectE, conjE], atac])
- (rev (active_set_naturals ~~ in_Jsrels))])
- (set_simps ~~ passive_set_naturals),
- rtac conjI,
- REPEAT_DETERM_N 2 o EVERY'[rtac (dtor_inject RS iffD1), rtac trans, rtac map_simp,
- rtac box_equals, rtac map_comp, rtac (dtor_ctor RS sym RS arg_cong), rtac trans,
- rtac map_cong, REPEAT_DETERM_N m o rtac @{thm fun_cong[OF o_id]},
- EVERY' (map (fn in_Jsrel => EVERY' [rtac trans, rtac o_apply, dtac set_rev_mp, atac,
- dtac @{thm ssubst_mem[OF pair_collapse]}, dtac (in_Jsrel RS iffD1),
- dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE, atac]) in_Jsrels),
- atac]]
- in
- EVERY' [rtac iffI, if_tac, only_if_tac] 1
- end;
-
-end;
--- a/src/HOL/Codatatype/Tools/bnf_gfp_util.ML Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,197 +0,0 @@
-(* Title: HOL/BNF/Tools/bnf_gfp_util.ML
- Author: Dmitriy Traytel, TU Muenchen
- Copyright 2012
-
-Library for the codatatype construction.
-*)
-
-signature BNF_GFP_UTIL =
-sig
- val mk_rec_simps: int -> thm -> thm list -> thm list list
-
- val dest_listT: typ -> typ
-
- val mk_Cons: term -> term -> term
- val mk_Shift: term -> term -> term
- val mk_Succ: term -> term -> term
- val mk_Times: term * term -> term
- val mk_append: term * term -> term
- val mk_congruent: term -> term -> term
- val mk_clists: term -> term
- val mk_diag: term -> term
- val mk_equiv: term -> term -> term
- val mk_fromCard: term -> term -> term
- val mk_list_rec: term -> term -> term
- val mk_nat_rec: term -> term -> term
- val mk_pickWP: term -> term -> term -> term -> term -> term
- val mk_prefCl: term -> term
- val mk_proj: term -> term
- val mk_quotient: term -> term -> term
- val mk_shift: term -> term -> term
- val mk_size: term -> term
- val mk_thePull: term -> term -> term -> term -> term
- val mk_toCard: term -> term -> term
- val mk_undefined: typ -> term
- val mk_univ: term -> term
-
- val mk_specN: int -> thm -> thm
-
- val mk_InN_Field: int -> int -> thm
- val mk_InN_inject: int -> int -> thm
- val mk_InN_not_InM: int -> int -> thm
-end;
-
-structure BNF_GFP_Util : BNF_GFP_UTIL =
-struct
-
-open BNF_Util
-
-val mk_append = HOLogic.mk_binop @{const_name append};
-
-fun mk_equiv B R =
- Const (@{const_name equiv}, fastype_of B --> fastype_of R --> HOLogic.boolT) $ B $ R;
-
-fun mk_Sigma (A, B) =
- let
- val AT = fastype_of A;
- val BT = fastype_of B;
- val ABT = mk_relT (HOLogic.dest_setT AT, HOLogic.dest_setT (range_type BT));
- in Const (@{const_name Sigma}, AT --> BT --> ABT) $ A $ B end;
-
-fun mk_diag A =
- let
- val AT = fastype_of A;
- val AAT = mk_relT (HOLogic.dest_setT AT, HOLogic.dest_setT AT);
- in Const (@{const_name diag}, AT --> AAT) $ A end;
-
-fun mk_Times (A, B) =
- let val AT = HOLogic.dest_setT (fastype_of A);
- in mk_Sigma (A, Term.absdummy AT B) end;
-
-fun dest_listT (Type (@{type_name list}, [T])) = T
- | dest_listT T = raise TYPE ("dest_setT: set type expected", [T], []);
-
-fun mk_Succ Kl kl =
- let val T = fastype_of kl;
- in
- Const (@{const_name Succ},
- HOLogic.mk_setT T --> T --> HOLogic.mk_setT (dest_listT T)) $ Kl $ kl
- end;
-
-fun mk_Shift Kl k =
- let val T = fastype_of Kl;
- in
- Const (@{const_name Shift}, T --> dest_listT (HOLogic.dest_setT T) --> T) $ Kl $ k
- end;
-
-fun mk_shift lab k =
- let val T = fastype_of lab;
- in
- Const (@{const_name shift}, T --> dest_listT (Term.domain_type T) --> T) $ lab $ k
- end;
-
-fun mk_prefCl A =
- Const (@{const_name prefCl}, fastype_of A --> HOLogic.boolT) $ A;
-
-fun mk_clists r =
- let val T = fastype_of r;
- in Const (@{const_name clists}, T --> mk_relT (`I (HOLogic.listT (fst (dest_relT T))))) $ r end;
-
-fun mk_toCard A r =
- let
- val AT = fastype_of A;
- val rT = fastype_of r;
- in
- Const (@{const_name toCard},
- AT --> rT --> HOLogic.dest_setT AT --> fst (dest_relT rT)) $ A $ r
- end;
-
-fun mk_fromCard A r =
- let
- val AT = fastype_of A;
- val rT = fastype_of r;
- in
- Const (@{const_name fromCard},
- AT --> rT --> fst (dest_relT rT) --> HOLogic.dest_setT AT) $ A $ r
- end;
-
-fun mk_Cons x xs =
- let val T = fastype_of xs;
- in Const (@{const_name Cons}, dest_listT T --> T --> T) $ x $ xs end;
-
-fun mk_size t = HOLogic.size_const (fastype_of t) $ t;
-
-fun mk_quotient A R =
- let val T = fastype_of A;
- in Const (@{const_name quotient}, T --> fastype_of R --> HOLogic.mk_setT T) $ A $ R end;
-
-fun mk_proj R =
- let val ((AT, BT), T) = `dest_relT (fastype_of R);
- in Const (@{const_name proj}, T --> AT --> HOLogic.mk_setT BT) $ R end;
-
-fun mk_univ f =
- let val ((AT, BT), T) = `dest_funT (fastype_of f);
- in Const (@{const_name univ}, T --> HOLogic.mk_setT AT --> BT) $ f end;
-
-fun mk_congruent R f =
- Const (@{const_name congruent}, fastype_of R --> fastype_of f --> HOLogic.boolT) $ R $ f;
-
-fun mk_undefined T = Const (@{const_name undefined}, T);
-
-fun mk_nat_rec Zero Suc =
- let val T = fastype_of Zero;
- in Const (@{const_name nat_rec}, T --> fastype_of Suc --> HOLogic.natT --> T) $ Zero $ Suc end;
-
-fun mk_list_rec Nil Cons =
- let
- val T = fastype_of Nil;
- val (U, consT) = `(Term.domain_type) (fastype_of Cons);
- in
- Const (@{const_name list_rec}, T --> consT --> HOLogic.listT U --> T) $ Nil $ Cons
- end;
-
-fun mk_thePull B1 B2 f1 f2 =
- let
- val fT1 = fastype_of f1;
- val fT2 = fastype_of f2;
- val BT1 = domain_type fT1;
- val BT2 = domain_type fT2;
- in
- Const (@{const_name thePull}, HOLogic.mk_setT BT1 --> HOLogic.mk_setT BT2 --> fT1 --> fT2 -->
- mk_relT (BT1, BT2)) $ B1 $ B2 $ f1 $ f2
- end;
-
-fun mk_pickWP A f1 f2 b1 b2 =
- let
- val fT1 = fastype_of f1;
- val fT2 = fastype_of f2;
- val AT = domain_type fT1;
- val BT1 = range_type fT1;
- val BT2 = range_type fT2;
- in
- Const (@{const_name pickWP}, HOLogic.mk_setT AT --> fT1 --> fT2 --> BT1 --> BT2 --> AT) $
- A $ f1 $ f2 $ b1 $ b2
- end;
-
-fun mk_InN_not_InM 1 _ = @{thm Inl_not_Inr}
- | mk_InN_not_InM n m =
- if n > m then mk_InN_not_InM m n RS @{thm not_sym}
- else mk_InN_not_InM (n - 1) (m - 1) RS @{thm not_arg_cong_Inr};
-
-fun mk_InN_Field 1 1 = @{thm TrueE[OF TrueI]}
- | mk_InN_Field _ 1 = @{thm Inl_Field_csum}
- | mk_InN_Field 2 2 = @{thm Inr_Field_csum}
- | mk_InN_Field n m = mk_InN_Field (n - 1) (m - 1) RS @{thm Inr_Field_csum};
-
-fun mk_InN_inject 1 _ = @{thm TrueE[OF TrueI]}
- | mk_InN_inject _ 1 = @{thm Inl_inject}
- | mk_InN_inject 2 2 = @{thm Inr_inject}
- | mk_InN_inject n m = @{thm Inr_inject} RS mk_InN_inject (n - 1) (m - 1);
-
-fun mk_specN 0 thm = thm
- | mk_specN n thm = mk_specN (n - 1) (thm RS spec);
-
-fun mk_rec_simps n rec_thm defs = map (fn i =>
- map (fn def => def RS rec_thm RS mk_nthI n i RS fun_cong) defs) (1 upto n);
-
-end;
--- a/src/HOL/Codatatype/Tools/bnf_lfp.ML Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1838 +0,0 @@
-(* Title: HOL/BNF/Tools/bnf_lfp.ML
- Author: Dmitriy Traytel, TU Muenchen
- Author: Andrei Popescu, TU Muenchen
- Copyright 2012
-
-Datatype construction.
-*)
-
-signature BNF_LFP =
-sig
- val bnf_lfp: mixfix list -> (string * sort) list option -> binding list ->
- typ list * typ list list -> BNF_Def.BNF list -> local_theory ->
- (term list * term list * term list * term list * thm * thm list * thm list * thm list *
- thm list * thm list) * local_theory
-end;
-
-structure BNF_LFP : BNF_LFP =
-struct
-
-open BNF_Def
-open BNF_Util
-open BNF_Tactics
-open BNF_FP
-open BNF_FP_Sugar
-open BNF_LFP_Util
-open BNF_LFP_Tactics
-
-(*all BNFs have the same lives*)
-fun bnf_lfp mixfixes resBs bs (resDs, Dss) bnfs lthy =
- let
- 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 b = Binding.name (mk_common_name (map Binding.name_of bs));
-
- (* TODO: check if m, n, etc., are sane *)
-
- val deads = fold (union (op =)) Dss resDs;
- val names_lthy = fold Variable.declare_typ deads lthy;
-
- (* tvars *)
- val (((((((passiveAs, activeAs), allAs)), (passiveBs, activeBs)),
- activeCs), passiveXs), passiveYs) = names_lthy
- |> mk_TFrees live
- |> apfst (`(chop m))
- ||> mk_TFrees live
- ||>> apfst (chop m)
- ||>> mk_TFrees n
- ||>> mk_TFrees m
- ||> fst o mk_TFrees m;
-
- 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';
-
- (* typs *)
- val dead_poss =
- (case resBs of
- NONE => map SOME deads @ replicate m NONE
- | SOME Ts => map (fn T => if member (op =) deads (TFree T) then SOME (TFree T) else NONE) Ts);
- 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 (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';
- val prodBsAs = map2 (curry HOLogic.mk_prodT) activeBs activeAs;
- val prodFTs = mk_FTs (passiveAs @ prodBsAs);
- val prod_sTs = map2 (curry (op -->)) prodFTs activeAs;
-
- (* terms *)
- val mapsAsAs = map4 mk_map_of_bnf Dss Ass Ass bnfs;
- val mapsAsBs = map4 mk_map_of_bnf Dss Ass Bss bnfs;
- val mapsBsAs = map4 mk_map_of_bnf Dss Bss Ass bnfs;
- val mapsBsCs' = map4 mk_map_of_bnf Dss Bss Css' bnfs;
- val mapsAsCs' = map4 mk_map_of_bnf Dss Ass Css' bnfs;
- val map_fsts = map4 mk_map_of_bnf Dss (replicate n (passiveAs @ prodBsAs)) Bss bnfs;
- val map_fsts_rev = map4 mk_map_of_bnf Dss Bss (replicate n (passiveAs @ prodBsAs)) bnfs;
- fun mk_setss Ts = map3 mk_sets_of_bnf (map (replicate live) Dss)
- (map (replicate live) (replicate n Ts)) bnfs;
- val setssAs = mk_setss allAs;
- val bds = map3 mk_bd_of_bnf Dss Ass bnfs;
- val witss = map wits_of_bnf bnfs;
-
- val (((((((((((((((((((zs, zs'), As), Bs), Bs_copy), B's), B''s), ss), prod_ss), s's), s''s),
- fs), fs_copy), inv_fs), self_fs), gs), all_gs), (xFs, xFs')), (yFs, yFs')),
- names_lthy) = lthy
- |> mk_Frees' "z" activeAs
- ||>> mk_Frees "A" ATs
- ||>> mk_Frees "B" BTs
- ||>> mk_Frees "B" BTs
- ||>> mk_Frees "B'" B'Ts
- ||>> mk_Frees "B''" B''Ts
- ||>> mk_Frees "s" sTs
- ||>> mk_Frees "prods" prod_sTs
- ||>> mk_Frees "s'" s'Ts
- ||>> mk_Frees "s''" s''Ts
- ||>> mk_Frees "f" fTs
- ||>> mk_Frees "f" fTs
- ||>> mk_Frees "f" inv_fTs
- ||>> mk_Frees "f" self_fTs
- ||>> mk_Frees "g" gTs
- ||>> mk_Frees "g" all_gTs
- ||>> mk_Frees' "x" FTsAs
- ||>> mk_Frees' "y" FTsBs;
-
- val passive_UNIVs = map HOLogic.mk_UNIV passiveAs;
- val active_UNIVs = map HOLogic.mk_UNIV activeAs;
- val prod_UNIVs = map HOLogic.mk_UNIV prodBsAs;
- val passive_ids = map HOLogic.id_const passiveAs;
- val active_ids = map HOLogic.id_const activeAs;
- val fsts = map fst_const prodBsAs;
-
- (* thms *)
- val bd_card_orders = map bd_card_order_of_bnf bnfs;
- val bd_Card_orders = map bd_Card_order_of_bnf bnfs;
- val bd_Card_order = hd bd_Card_orders;
- val bd_Cinfinite = bd_Cinfinite_of_bnf (hd bnfs);
- val bd_Cnotzeros = map bd_Cnotzero_of_bnf bnfs;
- val bd_Cnotzero = hd bd_Cnotzeros;
- val in_bds = map in_bd_of_bnf bnfs;
- val map_comp's = map map_comp'_of_bnf bnfs;
- val map_congs = map map_cong_of_bnf bnfs;
- val map_ids = map map_id_of_bnf bnfs;
- val map_id's = map map_id'_of_bnf bnfs;
- val map_wpulls = map map_wpull_of_bnf bnfs;
- val set_bdss = map set_bd_of_bnf bnfs;
- val set_natural'ss = map set_natural'_of_bnf bnfs;
-
- val timer = time (timer "Extracted terms & thms");
-
- (* nonemptiness check *)
- fun new_wit X wit = 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 :: _ => error ("Cannot define empty datatype " ^ quote (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_comp =
- 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;
- in
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (x :: fs @ all_gs) (mk_Trueprop_eq (lhs, rhs)))
- (K (mk_map_comp_id_tac map_comp))
- |> Thm.close_derivation
- end;
-
- val map_comp_id_thms = map5 mk_map_comp_id xFs mapsAsBs mapsBsCs' mapsAsCs' map_comp's;
-
- (*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_congL x mapAsAs sets map_cong 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 = map4 mk_prem (drop m sets) self_fs zs zs';
- val goal = mk_Trueprop_eq (Term.list_comb (mapAsAs, passive_ids @ self_fs) $ x, x);
- in
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (x :: self_fs) (Logic.list_implies (prems, goal)))
- (K (mk_map_congL_tac m map_cong map_id'))
- |> Thm.close_derivation
- end;
-
- val map_congL_thms = map5 mk_map_congL xFs mapsAsAs setssAs map_congs map_id's;
- 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 = Binding.suffix_name ("_" ^ algN) b;
- val alg_name = Binding.name_of alg_bind;
- val alg_def_bind = (Thm.def_binding alg_bind, []);
-
- (*forall i = 1 ... n: (\<forall>x \<in> Fi_in A1 .. Am B1 ... Bn. si x \<in> Bi)*)
- val alg_spec =
- let
- val algT = Library.foldr (op -->) (ATs @ BTs @ sTs, HOLogic.boolT);
-
- val ins = map3 mk_in (replicate n (As @ 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 lhs = Term.list_comb (Free (alg_name, algT), As @ Bs @ ss);
- val rhs = Library.foldr1 HOLogic.mk_conj (map5 mk_alg_conjunct Bs ss ins xFs xFs')
- in
- mk_Trueprop_eq (lhs, rhs)
- end;
-
- val ((alg_free, (_, alg_def_free)), (lthy, lthy_old)) =
- lthy
- |> Specification.definition (SOME (alg_bind, NONE, NoSyn), (alg_def_bind, alg_spec))
- ||> `Local_Theory.restore;
-
- val phi = Proof_Context.export_morphism lthy_old lthy;
- val alg = fst (Term.dest_Const (Morphism.term phi alg_free));
- val alg_def = Morphism.thm phi alg_def_free;
-
- fun mk_alg As Bs ss =
- let
- val args = As @ 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 alg_set_thms =
- let
- val alg_prem = HOLogic.mk_Trueprop (mk_alg As Bs ss);
- fun mk_prem x set B = HOLogic.mk_Trueprop (mk_subset (set $ x) B);
- fun mk_concl s x B = HOLogic.mk_Trueprop (HOLogic.mk_mem (s $ x, B));
- val premss = map2 ((fn x => fn sets => map2 (mk_prem x) sets (As @ Bs))) xFs setssAs;
- val concls = map3 mk_concl ss xFs Bs;
- val goals = map3 (fn x => fn prems => fn concl =>
- fold_rev Logic.all (x :: As @ Bs @ ss)
- (Logic.list_implies (alg_prem :: prems, concl))) xFs premss concls;
- in
- map (fn goal =>
- Skip_Proof.prove lthy [] [] goal (K (mk_alg_set_tac alg_def)) |> Thm.close_derivation)
- goals
- end;
-
- fun mk_talg ATs BTs = mk_alg (map HOLogic.mk_UNIV ATs) (map HOLogic.mk_UNIV BTs);
-
- val talg_thm =
- let
- val goal = fold_rev Logic.all ss
- (HOLogic.mk_Trueprop (mk_talg passiveAs activeAs ss))
- in
- Skip_Proof.prove lthy [] [] goal
- (K (stac alg_def 1 THEN CONJ_WRAP (K (EVERY' [rtac ballI, rtac UNIV_I] 1)) ss))
- |> Thm.close_derivation
- end;
-
- val timer = time (timer "Algebra definition & thms");
-
- val alg_not_empty_thms =
- let
- val alg_prem =
- HOLogic.mk_Trueprop (mk_alg passive_UNIVs Bs ss);
- val concls = map (HOLogic.mk_Trueprop o mk_not_empty) Bs;
- val goals =
- map (fn concl =>
- fold_rev Logic.all (Bs @ ss) (Logic.mk_implies (alg_prem, concl))) concls;
- in
- map2 (fn goal => fn alg_set =>
- Skip_Proof.prove lthy [] []
- goal (K (mk_alg_not_empty_tac alg_set alg_set_thms wit_thms))
- |> Thm.close_derivation)
- goals alg_set_thms
- end;
-
- val timer = time (timer "Proved nonemptiness");
-
- (* morphism *)
-
- val mor_bind = Binding.suffix_name ("_" ^ morN) b;
- val mor_name = Binding.name_of mor_bind;
- val mor_def_bind = (Thm.def_binding mor_bind, []);
-
- (*fbetw) forall i = 1 ... n: (\<forall>x \<in> Bi. f x \<in> B'i)*)
- (*mor) forall i = 1 ... n: (\<forall>x \<in> Fi_in UNIV ... UNIV B1 ... Bn.
- f (s1 x) = s1' (Fi_map id ... id f1 ... fn x))*)
- val mor_spec =
- let
- val morT = Library.foldr (op -->) (BTs @ sTs @ B'Ts @ s'Ts @ fTs, HOLogic.boolT);
-
- 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 lhs = Term.list_comb (Free (mor_name, morT), Bs @ ss @ B's @ s's @ fs);
- val rhs = HOLogic.mk_conj
- (Library.foldr1 HOLogic.mk_conj (map5 mk_fbetw fs Bs B's zs zs'),
- Library.foldr1 HOLogic.mk_conj
- (map8 mk_mor setssAs mapsAsBs fs ss s's FTsAs xFs xFs'))
- in
- mk_Trueprop_eq (lhs, rhs)
- end;
-
- val ((mor_free, (_, mor_def_free)), (lthy, lthy_old)) =
- lthy
- |> Specification.definition (SOME (mor_bind, NONE, NoSyn), (mor_def_bind, mor_spec))
- ||> `Local_Theory.restore;
-
- val phi = Proof_Context.export_morphism lthy_old lthy;
- val mor = fst (Term.dest_Const (Morphism.term phi mor_free));
- val mor_def = 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 (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 = fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs)
- (Logic.mk_implies (prem, HOLogic.mk_Trueprop (mk_subset (mk_image f $ B1) B2)));
- val image_goals = map3 mk_image_goal fs Bs B's;
- 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 =
- fold_rev Logic.all (x :: Bs @ ss @ B's @ s's @ fs)
- (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 = map7 mk_elim_goal setssAs mapsAsBs fs ss s's xFs FTsAs;
- fun prove goal =
- Skip_Proof.prove lthy [] [] goal (K (mk_mor_elim_tac mor_def)) |> Thm.close_derivation;
- in
- (map prove image_goals, map prove elim_goals)
- end;
-
- val mor_incl_thm =
- let
- val prems = map2 (HOLogic.mk_Trueprop oo mk_subset) Bs Bs_copy;
- val concl = HOLogic.mk_Trueprop (mk_mor Bs ss Bs_copy ss active_ids);
- in
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (Bs @ ss @ Bs_copy) (Logic.list_implies (prems, concl)))
- (K (mk_mor_incl_tac mor_def map_id's))
- |> Thm.close_derivation
- 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));
- in
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (Bs @ ss @ B's @ s's @ B''s @ s''s @ fs @ gs)
- (Logic.list_implies (prems, concl)))
- (K (mk_mor_comp_tac mor_def set_natural'ss map_comp_id_thms))
- |> Thm.close_derivation
- end;
-
- val mor_inv_thm =
- let
- fun mk_inv_prem f inv_f B B' = HOLogic.mk_conj (mk_subset (mk_image inv_f $ B') B,
- HOLogic.mk_conj (mk_inver inv_f f B, mk_inver f inv_f B'));
- val prems = map HOLogic.mk_Trueprop
- ([mk_mor Bs ss B's s's fs,
- mk_alg passive_UNIVs Bs ss,
- mk_alg passive_UNIVs B's s's] @
- map4 mk_inv_prem fs inv_fs Bs B's);
- val concl = HOLogic.mk_Trueprop (mk_mor B's s's Bs ss inv_fs);
- in
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs @ inv_fs)
- (Logic.list_implies (prems, concl)))
- (K (mk_mor_inv_tac alg_def mor_def
- set_natural'ss morE_thms map_comp_id_thms map_congL_thms))
- |> Thm.close_derivation
- 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);
- in
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs @ fs_copy)
- (Logic.list_implies (prems, concl)))
- (K ((hyp_subst_tac THEN' atac) 1))
- |> Thm.close_derivation
- 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;
- in
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all ss (HOLogic.mk_Trueprop
- (mk_mor (map HOLogic.mk_UNIV FTsAs) maps active_UNIVs ss ss)))
- (K (mk_mor_str_tac ks mor_def))
- |> Thm.close_derivation
- end;
-
- val mor_convol_thm =
- let
- val maps = map3 (fn s => fn prod_s => fn mapx =>
- mk_convol (HOLogic.mk_comp (s, Term.list_comb (mapx, passive_ids @ fsts)), prod_s))
- s's prod_ss map_fsts;
- in
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (s's @ prod_ss) (HOLogic.mk_Trueprop
- (mk_mor prod_UNIVs maps (map HOLogic.mk_UNIV activeBs) s's fsts)))
- (K (mk_mor_convol_tac ks mor_def))
- |> Thm.close_derivation
- 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 (map4 mk_conjunct mapsAsBs fs ss s's);
- in
- Skip_Proof.prove lthy [] [] (fold_rev Logic.all (ss @ s's @ fs) (mk_Trueprop_eq (lhs, rhs)))
- (K (mk_mor_UNIV_tac m morE_thms mor_def))
- |> Thm.close_derivation
- end;
-
- val timer = time (timer "Morphism definition & thms");
-
- (* isomorphism *)
-
- (*mor Bs1 ss1 Bs2 ss2 fs \<and> (\<exists>gs. mor Bs2 ss2 Bs1 ss1 fs \<and>
- forall i = 1 ... n. (inver gs[i] fs[i] Bs1[i] \<and> inver fs[i] gs[i] Bs2[i]))*)
- fun mk_iso Bs1 ss1 Bs2 ss2 fs gs =
- let
- val ex_inv_mor = list_exists_free gs
- (HOLogic.mk_conj (mk_mor Bs2 ss2 Bs1 ss1 gs,
- Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_conj)
- (map3 mk_inver gs fs Bs1) (map3 mk_inver fs gs Bs2))));
- in
- HOLogic.mk_conj (mk_mor Bs1 ss1 Bs2 ss2 fs, ex_inv_mor)
- end;
-
- val iso_alt_thm =
- let
- val prems = map HOLogic.mk_Trueprop
- [mk_alg passive_UNIVs Bs ss,
- mk_alg passive_UNIVs B's s's]
- val concl = mk_Trueprop_eq (mk_iso Bs ss B's s's fs inv_fs,
- HOLogic.mk_conj (mk_mor Bs ss B's s's fs,
- Library.foldr1 HOLogic.mk_conj (map3 mk_bij_betw fs Bs B's)));
- in
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (Bs @ ss @ B's @ s's @ fs) (Logic.list_implies (prems, concl)))
- (K (mk_iso_alt_tac mor_image_thms mor_inv_thm))
- |> Thm.close_derivation
- end;
-
- val timer = time (timer "Isomorphism definition & thms");
-
- (* algebra copies *)
-
- val (copy_alg_thm, ex_copy_alg_thm) =
- let
- val prems = map HOLogic.mk_Trueprop
- (mk_alg passive_UNIVs Bs ss :: map3 mk_bij_betw inv_fs B's Bs);
- val inver_prems = map HOLogic.mk_Trueprop
- (map3 mk_inver inv_fs fs Bs @ map3 mk_inver fs inv_fs B's);
- val all_prems = prems @ inver_prems;
- fun mk_s f s mapT y y' = Term.absfree y' (f $ (s $
- (Term.list_comb (mapT, passive_ids @ inv_fs) $ y)));
-
- val alg = HOLogic.mk_Trueprop
- (mk_alg passive_UNIVs B's (map5 mk_s fs ss mapsBsAs yFs yFs'));
- val copy_str_thm = Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (Bs @ ss @ B's @ inv_fs @ fs)
- (Logic.list_implies (all_prems, alg)))
- (K (mk_copy_str_tac set_natural'ss alg_def alg_set_thms))
- |> Thm.close_derivation;
-
- val iso = HOLogic.mk_Trueprop
- (mk_iso B's (map5 mk_s fs ss mapsBsAs yFs yFs') Bs ss inv_fs fs_copy);
- val copy_alg_thm = Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (Bs @ ss @ B's @ inv_fs @ fs)
- (Logic.list_implies (all_prems, iso)))
- (K (mk_copy_alg_tac set_natural'ss alg_set_thms mor_def iso_alt_thm copy_str_thm))
- |> Thm.close_derivation;
-
- val ex = HOLogic.mk_Trueprop
- (list_exists_free s's
- (HOLogic.mk_conj (mk_alg passive_UNIVs B's s's,
- mk_iso B's s's Bs ss inv_fs fs_copy)));
- val ex_copy_alg_thm = Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (Bs @ ss @ B's @ inv_fs @ fs)
- (Logic.list_implies (prems, ex)))
- (K (mk_ex_copy_alg_tac n copy_str_thm copy_alg_thm))
- |> Thm.close_derivation;
- in
- (copy_alg_thm, ex_copy_alg_thm)
- end;
-
- val timer = time (timer "Copy thms");
-
-
- (* bounds *)
-
- val sum_Card_order = if n = 1 then bd_Card_order else @{thm Card_order_csum};
- val sum_Cnotzero = if n = 1 then bd_Cnotzero else bd_Cnotzero RS @{thm csum_Cnotzero1};
- val sum_Cinfinite = if n = 1 then bd_Cinfinite else bd_Cinfinite RS @{thm Cinfinite_csum1};
- fun mk_set_bd_sums i bd_Card_order bds =
- if n = 1 then bds
- else map (fn thm => bd_Card_order RS mk_ordLeq_csum n i thm) bds;
- val set_bd_sumss = map3 mk_set_bd_sums ks bd_Card_orders set_bdss;
-
- fun mk_in_bd_sum i Co Cnz bd =
- if n = 1 then bd
- else Cnz RS ((Co RS mk_ordLeq_csum n i (Co RS @{thm ordLeq_refl})) RS
- (bd RS @{thm ordLeq_transitive[OF _
- cexp_mono2_Cnotzero[OF _ csum_Cnotzero2[OF ctwo_Cnotzero]]]}));
- val in_bd_sums = map4 mk_in_bd_sum ks bd_Card_orders bd_Cnotzeros in_bds;
-
- val sum_bd = Library.foldr1 (uncurry mk_csum) bds;
- val suc_bd = mk_cardSuc sum_bd;
- 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 = if n = 1 then bd_Card_order RS @{thm cardSuc_Card_order}
- else @{thm cardSuc_Card_order[OF Card_order_csum]};
- val suc_bd_Cinfinite = if n = 1 then bd_Cinfinite RS @{thm Cinfinite_cardSuc}
- else bd_Cinfinite RS @{thm Cinfinite_cardSuc[OF Cinfinite_csum1]};
- 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 Asuc_bd_Cnotzero = Asuc_bd_Cinfinite RS @{thm Cinfinite_Cnotzero};
-
- val suc_bd_Asuc_bd = @{thm ordLess_ordLeq_trans[OF
- ordLess_ctwo_cexp
- cexp_mono1_Cnotzero[OF _ ctwo_Cnotzero]]} OF
- [suc_bd_Card_order, basis_Asuc, suc_bd_Card_order];
-
- val Asuc_bdT = fst (dest_relT (fastype_of (mk_Asuc_bd As)));
- 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 (((((((idxs, Asi_name), (idx, idx')), (jdx, jdx')), II_Bs), II_ss), Asuc_fs),
- names_lthy) = names_lthy
- |> 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
- ||>> mk_Frees "IIB" II_BTs
- ||>> mk_Frees "IIs" II_sTs
- ||>> mk_Frees "f" (map (fn T => Asuc_bdT --> T) activeAs);
-
- 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))));
- in
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all idxs (Logic.list_implies ([prem], concl)))
- (K (mk_bd_limit_tac n suc_bd_Cinfinite))
- |> Thm.close_derivation
- 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 As Asi i sets Ts s k =
- HOLogic.mk_binop @{const_name "sup"}
- (mk_minG Asi i k, mk_image s $ mk_in (As @ map (mk_minG Asi i) ks) sets Ts);
-
- fun mk_min_algs As ss =
- let
- val BTs = map (range_type o fastype_of) ss;
- val Ts = map (HOLogic.dest_setT o fastype_of) As @ 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
- (map4 (mk_minH_component As 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 As 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
- (map4 (mk_minH_component As min_algs idx) setssAs FTsAs ss ks)));
- val goal = fold_rev Logic.all (idx :: As @ ss)
- (Logic.mk_implies (HOLogic.mk_Trueprop i_field, concl));
-
- val min_algs_thm = Skip_Proof.prove lthy [] [] goal
- (K (mk_min_algs_tac suc_bd_worel in_cong'_thms))
- |> Thm.close_derivation;
-
- val min_algs_thms = map (fn k => min_algs_thm RS mk_nthI n k) ks;
-
- fun mk_mono_goal min_alg =
- fold_rev Logic.all (As @ ss) (HOLogic.mk_Trueprop (mk_relChain suc_bd
- (Term.absfree idx' min_alg)));
-
- val monos =
- map2 (fn goal => fn min_algs =>
- Skip_Proof.prove lthy [] [] goal (K (mk_min_algs_mono_tac min_algs))
- |> Thm.close_derivation)
- (map mk_mono_goal min_algss) min_algs_thms;
-
- val Asuc_bd = mk_Asuc_bd As;
-
- 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 = certifyT lthy suc_bdT;
- val card_ct = certify lthy (Term.absfree idx' card_conjunction);
-
- val card_of = singleton (Proof_Context.export names_lthy lthy)
- (Skip_Proof.prove lthy [] []
- (HOLogic.mk_Trueprop (HOLogic.mk_imp (i_field, card_conjunction)))
- (K (mk_min_algs_card_of_tac card_cT card_ct
- m suc_bd_worel min_algs_thms in_bd_sums
- sum_Card_order sum_Cnotzero suc_bd_Card_order suc_bd_Cinfinite suc_bd_Cnotzero
- suc_bd_Asuc_bd Asuc_bd_Cinfinite Asuc_bd_Cnotzero)))
- |> Thm.close_derivation;
-
- val least_prem = HOLogic.mk_Trueprop (mk_alg As Bs ss);
- val least_conjunction = Library.foldr1 HOLogic.mk_conj (map2 mk_subset min_algss Bs);
- val least_cT = certifyT lthy suc_bdT;
- val least_ct = certify lthy (Term.absfree idx' least_conjunction);
-
- val least = singleton (Proof_Context.export names_lthy lthy)
- (Skip_Proof.prove lthy [] []
- (Logic.mk_implies (least_prem,
- HOLogic.mk_Trueprop (HOLogic.mk_imp (i_field, least_conjunction))))
- (K (mk_min_algs_least_tac least_cT least_ct
- suc_bd_worel min_algs_thms alg_set_thms)))
- |> Thm.close_derivation;
- in
- (min_algs_thms, monos, card_of, least)
- end;
-
- val timer = time (timer "min_algs definition & thms");
-
- fun min_alg_bind i = Binding.suffix_name
- ("_" ^ min_algN ^ (if n = 1 then "" else string_of_int i)) b;
- val min_alg_name = Binding.name_of o min_alg_bind;
- val min_alg_def_bind = rpair [] o Thm.def_binding o min_alg_bind;
-
- fun min_alg_spec i =
- let
- val min_algT =
- Library.foldr (op -->) (ATs @ sTs, HOLogic.mk_setT (nth activeAs (i - 1)));
-
- val lhs = Term.list_comb (Free (min_alg_name i, min_algT), As @ ss);
- val rhs = mk_UNION (field_suc_bd)
- (Term.absfree idx' (mk_nthN n (mk_min_algs As ss $ idx) i));
- in
- mk_Trueprop_eq (lhs, rhs)
- end;
-
- val ((min_alg_frees, (_, min_alg_def_frees)), (lthy, lthy_old)) =
- lthy
- |> fold_map (fn i => Specification.definition
- (SOME (min_alg_bind i, NONE, NoSyn), (min_alg_def_bind i, min_alg_spec i))) ks
- |>> apsnd split_list o split_list
- ||> `Local_Theory.restore;
-
- 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 (Morphism.thm phi) min_alg_def_frees;
-
- fun mk_min_alg As ss i =
- let
- val T = HOLogic.mk_setT (range_type (fastype_of (nth ss (i - 1))))
- val args = As @ ss;
- val Ts = map fastype_of args;
- val min_algT = Library.foldr (op -->) (Ts, T);
- in
- Term.list_comb (Const (nth min_algs (i - 1), min_algT), args)
- end;
-
- val (alg_min_alg_thm, card_of_min_alg_thms, least_min_alg_thms, mor_incl_min_alg_thm) =
- let
- val min_algs = map (mk_min_alg As ss) ks;
-
- val goal = fold_rev Logic.all (As @ ss) (HOLogic.mk_Trueprop (mk_alg As min_algs ss));
- val alg_min_alg = Skip_Proof.prove lthy [] [] goal
- (K (mk_alg_min_alg_tac m alg_def min_alg_defs suc_bd_limit_thm sum_Cinfinite
- set_bd_sumss min_algs_thms min_algs_mono_thms))
- |> Thm.close_derivation;
-
- val Asuc_bd = mk_Asuc_bd As;
- fun mk_card_of_thm min_alg def = Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (As @ ss)
- (HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of min_alg) Asuc_bd)))
- (K (mk_card_of_min_alg_tac def card_of_min_algs_thm
- suc_bd_Card_order suc_bd_Asuc_bd Asuc_bd_Cinfinite))
- |> Thm.close_derivation;
-
- val least_prem = HOLogic.mk_Trueprop (mk_alg As Bs ss);
- fun mk_least_thm min_alg B def = Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (As @ Bs @ ss)
- (Logic.mk_implies (least_prem, HOLogic.mk_Trueprop (mk_subset min_alg B))))
- (K (mk_least_min_alg_tac def least_min_algs_thm))
- |> Thm.close_derivation;
-
- val leasts = map3 mk_least_thm min_algs Bs min_alg_defs;
-
- val incl_prem = HOLogic.mk_Trueprop (mk_alg passive_UNIVs Bs ss);
- val incl_min_algs = map (mk_min_alg passive_UNIVs ss) ks;
- val incl = Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (Bs @ ss)
- (Logic.mk_implies (incl_prem,
- HOLogic.mk_Trueprop (mk_mor incl_min_algs ss Bs ss active_ids))))
- (K (EVERY' (rtac mor_incl_thm :: map etac leasts) 1))
- |> Thm.close_derivation;
- 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 = Binding.suffix_name ("_" ^ IITN) b;
-
- val ((IIT_name, (IIT_glob_info, IIT_loc_info)), lthy) =
- typedef false NONE (IIT_bind, params, NoSyn)
- (HOLogic.mk_UNIV II_repT) NONE (EVERY' [rtac exI, rtac 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 (((((((iidx, iidx'), init_xs), (init_xFs, init_xFs')),
- init_fs), init_fs_copy), init_phis), names_lthy) = names_lthy
- |> 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 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 passive_UNIVs 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;
-
- fun str_init_bind i = Binding.suffix_name ("_" ^ str_initN ^ (if n = 1 then "" else
- string_of_int i)) b;
- val str_init_name = Binding.name_of o str_init_bind;
- val str_init_def_bind = rpair [] o Thm.def_binding o str_init_bind;
-
- fun str_init_spec i =
- let
- val T = nth init_FTs (i - 1);
- 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 str_initT = T --> IIT --> Asuc_bdT;
-
- val lhs = Term.list_comb (Free (str_init_name i, str_initT), [init_xF, iidx]);
- val rhs = select_s $ (Term.list_comb (map, map_args) $ init_xF);
- in
- mk_Trueprop_eq (lhs, rhs)
- end;
-
- val ((str_init_frees, (_, str_init_def_frees)), (lthy, lthy_old)) =
- lthy
- |> fold_map (fn i => Specification.definition
- (SOME (str_init_bind i, NONE, NoSyn), (str_init_def_bind i, str_init_spec i))) ks
- |>> apsnd split_list o split_list
- ||> `Local_Theory.restore;
-
- 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 (Morphism.thm phi) str_init_def_frees;
-
- val car_inits = map (mk_min_alg passive_UNIVs str_inits) ks;
-
- (*TODO: replace with instantiate? (problem: figure out right type instantiation)*)
- val alg_init_thm = Skip_Proof.prove lthy [] []
- (HOLogic.mk_Trueprop (mk_alg passive_UNIVs car_inits str_inits))
- (K (rtac alg_min_alg_thm 1))
- |> Thm.close_derivation;
-
- val alg_select_thm = Skip_Proof.prove lthy [] []
- (HOLogic.mk_Trueprop (mk_Ball II
- (Term.absfree iidx' (mk_alg passive_UNIVs select_Bs select_ss))))
- (mk_alg_select_tac Abs_IIT_inverse_thm)
- |> Thm.close_derivation;
-
- val mor_select_thm =
- let
- val alg_prem = HOLogic.mk_Trueprop (mk_alg passive_UNIVs Bs ss);
- val i_prem = HOLogic.mk_Trueprop (HOLogic.mk_mem (iidx, II));
- val mor_prem = HOLogic.mk_Trueprop (mk_mor select_Bs select_ss Bs ss Asuc_fs);
- val prems = [alg_prem, i_prem, mor_prem];
- val concl = HOLogic.mk_Trueprop
- (mk_mor car_inits str_inits Bs ss
- (map (fn f => HOLogic.mk_comp (f, mk_rapp iidx Asuc_bdT)) Asuc_fs));
- in
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (iidx :: Bs @ ss @ Asuc_fs) (Logic.list_implies (prems, concl)))
- (K (mk_mor_select_tac mor_def mor_cong_thm mor_comp_thm mor_incl_min_alg_thm alg_def
- alg_select_thm alg_set_thms set_natural'ss str_init_defs))
- |> Thm.close_derivation
- end;
-
- val (init_ex_mor_thm, init_unique_mor_thms) =
- let
- val prem = HOLogic.mk_Trueprop (mk_alg passive_UNIVs Bs ss);
- val concl = HOLogic.mk_Trueprop
- (list_exists_free init_fs (mk_mor car_inits str_inits Bs ss init_fs));
- val ex_mor = Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (Bs @ ss) (Logic.mk_implies (prem, concl)))
- (mk_init_ex_mor_tac Abs_IIT_inverse_thm ex_copy_alg_thm alg_min_alg_thm
- card_of_min_alg_thms mor_comp_thm mor_select_thm mor_incl_min_alg_thm)
- |> Thm.close_derivation;
-
- 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 (map3 mk_fun_eq init_fs init_fs_copy init_xs));
- val unique_mor = Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (init_xs @ Bs @ ss @ init_fs @ init_fs_copy)
- (Logic.list_implies (prems @ mor_prems, unique)))
- (K (mk_init_unique_mor_tac m alg_def alg_init_thm least_min_alg_thms
- in_mono'_thms alg_set_thms morE_thms map_congs))
- |> Thm.close_derivation;
- in
- (ex_mor, 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
- (map6 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));
- in
- Skip_Proof.prove lthy [] []
- (fold_rev Logic.all init_phis (Logic.mk_implies (prem, concl)))
- (K (mk_init_induct_tac m alg_def alg_init_thm least_min_alg_thms alg_set_thms))
- |> Thm.close_derivation
- end;
-
- val timer = time (timer "Initiality definition & thms");
-
- val ((T_names, (T_glob_infos, T_loc_infos)), lthy) =
- lthy
- |> fold_map3 (fn b => fn mx => fn car_init => typedef false NONE (b, params, mx) car_init NONE
- (EVERY' [rtac ssubst, rtac @{thm ex_in_conv}, resolve_tac alg_not_empty_thms,
- rtac 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_casess = map #Rep_cases T_loc_infos;
- val Rep_injects = map #Rep_inject T_loc_infos;
- val Rep_inverses = map #Rep_inverse T_loc_infos;
- val Abs_inverses = map #Abs_inverse T_loc_infos;
-
- fun mk_inver_thm mk_tac rep abs X thm =
- Skip_Proof.prove lthy [] []
- (HOLogic.mk_Trueprop (mk_inver rep abs X))
- (K (EVERY' [rtac ssubst, rtac @{thm inver_def}, rtac ballI, mk_tac thm] 1))
- |> Thm.close_derivation;
-
- val inver_Reps = map4 (mk_inver_thm rtac) Abs_Ts Rep_Ts (map HOLogic.mk_UNIV Ts) Rep_inverses;
- val inver_Abss = map4 (mk_inver_thm etac) Rep_Ts Abs_Ts car_inits Abs_inverses;
-
- 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 rec_sTs = map (Term.typ_subst_atomic (activeBs ~~ Ts)) prod_sTs;
- val rec_maps = map (Term.subst_atomic_types (activeBs ~~ Ts)) map_fsts;
- val rec_maps_rev = map (Term.subst_atomic_types (activeBs ~~ Ts)) map_fsts_rev;
- val rec_fsts = map (Term.subst_atomic_types (activeBs ~~ Ts)) fsts;
-
- val (((((((((Izs1, Izs1'), (Izs2, Izs2')), (xFs, xFs')), yFs), (AFss, AFss')),
- (fold_f, fold_f')), fs), rec_ss), names_lthy) = names_lthy
- |> mk_Frees' "z1" Ts
- ||>> mk_Frees' "z2" Ts'
- ||>> mk_Frees' "x" FTs
- ||>> mk_Frees "y" FTs'
- ||>> mk_Freess' "z" setFTss
- ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "f") foldT
- ||>> mk_Frees "f" fTs
- ||>> mk_Frees "s" rec_sTs;
-
- val Izs = map2 retype_free Ts zs;
- val phis = map2 retype_free (map mk_pred1T Ts) init_phis;
- val phi2s = map2 retype_free (map2 mk_pred2T Ts Ts') init_phis;
-
- fun ctor_bind i = Binding.suffix_name ("_" ^ ctorN) (nth bs (i - 1));
- val ctor_name = Binding.name_of o ctor_bind;
- val ctor_def_bind = rpair [] o Thm.def_binding o ctor_bind;
-
- fun ctor_spec i abs str map_FT_init x x' =
- let
- val ctorT = nth FTs (i - 1) --> nth Ts (i - 1);
-
- val lhs = Free (ctor_name i, ctorT);
- val rhs = Term.absfree x' (abs $ (str $
- (Term.list_comb (map_FT_init, map HOLogic.id_const passiveAs @ Rep_Ts) $ x)));
- in
- mk_Trueprop_eq (lhs, rhs)
- end;
-
- val ((ctor_frees, (_, ctor_def_frees)), (lthy, lthy_old)) =
- lthy
- |> fold_map6 (fn i => fn abs => fn str => fn mapx => fn x => fn x' =>
- Specification.definition
- (SOME (ctor_bind i, NONE, NoSyn), (ctor_def_bind i, ctor_spec i abs str mapx x x')))
- ks Abs_Ts str_inits map_FT_inits xFs xFs'
- |>> apsnd split_list o split_list
- ||> `Local_Theory.restore;
-
- 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 (Morphism.thm phi) ctor_def_frees;
-
- val (mor_Rep_thm, mor_Abs_thm) =
- let
- val copy = alg_init_thm RS copy_alg_thm;
- fun mk_bij inj Rep cases = @{thm bij_betwI'} OF [inj, Rep, cases];
- val bijs = map3 mk_bij Rep_injects Reps Rep_casess;
- val mor_Rep =
- Skip_Proof.prove lthy [] []
- (HOLogic.mk_Trueprop (mk_mor UNIVs ctors car_inits str_inits Rep_Ts))
- (mk_mor_Rep_tac ctor_defs copy bijs inver_Abss inver_Reps)
- |> Thm.close_derivation;
-
- val inv = mor_inv_thm OF [mor_Rep, talg_thm, alg_init_thm];
- val mor_Abs =
- Skip_Proof.prove lthy [] []
- (HOLogic.mk_Trueprop (mk_mor car_inits str_inits UNIVs ctors Abs_Ts))
- (K (mk_mor_Abs_tac inv inver_Abss inver_Reps))
- |> Thm.close_derivation;
- 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 = Binding.suffix_name ("_" ^ ctor_foldN) (nth bs (i - 1));
- val fold_name = Binding.name_of o fold_bind;
- val fold_def_bind = rpair [] o Thm.def_binding o fold_bind;
-
- fun fold_spec i T AT =
- let
- val foldT = Library.foldr (op -->) (sTs, T --> AT);
-
- val lhs = Term.list_comb (Free (fold_name i, foldT), ss);
- val rhs = mk_nthN n foldx i;
- in
- mk_Trueprop_eq (lhs, rhs)
- end;
-
- val ((fold_frees, (_, fold_def_frees)), (lthy, lthy_old)) =
- lthy
- |> fold_map3 (fn i => fn T => fn AT =>
- Specification.definition
- (SOME (fold_bind i, NONE, NoSyn), (fold_def_bind i, fold_spec i T AT)))
- ks Ts activeAs
- |>> apsnd split_list o split_list
- ||> `Local_Theory.restore;
-
- 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_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 (Morphism.thm phi) fold_def_frees;
-
- val mor_fold_thm =
- let
- val ex_mor = talg_thm RS init_ex_mor_thm;
- val mor_cong = mor_cong_thm OF (map (mk_nth_conv n) ks);
- val mor_comp = mor_Rep_thm RS mor_comp_thm;
- val cT = certifyT lthy foldT;
- val ct = certify lthy fold_fun
- in
- singleton (Proof_Context.export names_lthy lthy)
- (Skip_Proof.prove lthy [] []
- (HOLogic.mk_Trueprop (mk_mor UNIVs ctors active_UNIVs ss (map (mk_fold Ts ss) ks)))
- (K (mk_mor_fold_tac cT ct fold_defs ex_mor (mor_comp RS mor_cong))))
- |> Thm.close_derivation
- end;
-
- val ctor_fold_thms = map (fn morE => rule_by_tactic lthy
- ((rtac CollectI THEN' CONJ_WRAP' (K (rtac @{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 unique_mor = Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (ss @ fs) (Logic.mk_implies (prem, unique)))
- (K (mk_fold_unique_mor_tac type_defs init_unique_mor_thms Reps
- mor_comp_thm mor_Abs_thm mor_fold_thm))
- |> Thm.close_derivation;
- in
- `split_conj_thm unique_mor
- end;
-
- val ctor_fold_unique_thms =
- split_conj_thm (mk_conjIN n RS
- (mor_UNIV_thm RS @{thm ssubst[of _ _ "%x. x"]} 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 = Binding.suffix_name ("_" ^ dtorN) (nth bs (i - 1));
- val dtor_name = Binding.name_of o dtor_bind;
- val dtor_def_bind = rpair [] o Thm.def_binding o dtor_bind;
-
- fun dtor_spec i FT T =
- let
- val dtorT = T --> FT;
-
- val lhs = Free (dtor_name i, dtorT);
- val rhs = mk_fold Ts map_ctors i;
- in
- mk_Trueprop_eq (lhs, rhs)
- end;
-
- val ((dtor_frees, (_, dtor_def_frees)), (lthy, lthy_old)) =
- lthy
- |> fold_map3 (fn i => fn FT => fn T =>
- Specification.definition
- (SOME (dtor_bind i, NONE, NoSyn), (dtor_def_bind i, dtor_spec i FT T))) ks FTs Ts
- |>> apsnd split_list o split_list
- ||> `Local_Theory.restore;
-
- 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 (Morphism.thm phi) dtor_def_frees;
-
- val ctor_o_dtor_thms = map2 (fold_thms 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 = map3 mk_goal dtors ctors FTs;
- in
- map5 (fn goal => fn dtor_def => fn foldx => fn map_comp_id => fn map_congL =>
- Skip_Proof.prove lthy [] [] goal
- (K (mk_dtor_o_ctor_tac dtor_def foldx map_comp_id map_congL ctor_o_fold_thms))
- |> Thm.close_derivation)
- goals dtor_defs ctor_fold_thms map_comp_id_thms map_congL_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 fst_rec_pair_thms =
- let
- val mor = mor_comp_thm OF [mor_fold_thm, mor_convol_thm];
- in
- map2 (fn unique => fn fold_ctor =>
- trans OF [mor RS unique, fold_ctor]) fold_unique_mor_thms fold_ctor_thms
- end;
-
- fun rec_bind i = Binding.suffix_name ("_" ^ ctor_recN) (nth bs (i - 1));
- val rec_name = Binding.name_of o rec_bind;
- val rec_def_bind = rpair [] o Thm.def_binding o rec_bind;
-
- fun rec_spec i T AT =
- let
- val recT = Library.foldr (op -->) (rec_sTs, T --> AT);
- val maps = map3 (fn ctor => fn prod_s => fn mapx =>
- mk_convol (HOLogic.mk_comp (ctor, Term.list_comb (mapx, passive_ids @ rec_fsts)), prod_s))
- ctors rec_ss rec_maps;
-
- val lhs = Term.list_comb (Free (rec_name i, recT), rec_ss);
- val rhs = HOLogic.mk_comp (snd_const (HOLogic.mk_prodT (T, AT)), mk_fold Ts maps i);
- in
- mk_Trueprop_eq (lhs, rhs)
- end;
-
- val ((rec_frees, (_, rec_def_frees)), (lthy, lthy_old)) =
- lthy
- |> fold_map3 (fn i => fn T => fn AT =>
- Specification.definition
- (SOME (rec_bind i, NONE, NoSyn), (rec_def_bind i, rec_spec i T AT)))
- ks Ts activeAs
- |>> apsnd split_list o split_list
- ||> `Local_Theory.restore;
-
- val phi = Proof_Context.export_morphism lthy_old lthy;
- val recs = map (Morphism.term phi) rec_frees;
- val rec_names = map (fst o dest_Const) recs;
- fun mk_rec ss i = Term.list_comb (Const (nth rec_names (i - 1), Library.foldr (op -->)
- (map fastype_of ss, nth Ts (i - 1) --> range_type (fastype_of (nth ss (i - 1))))), ss);
- val rec_defs = map (Morphism.thm phi) rec_def_frees;
-
- val convols = map2 (fn T => fn i => mk_convol (HOLogic.id_const T, mk_rec rec_ss i)) Ts ks;
- val ctor_rec_thms =
- let
- fun mk_goal i rec_s rec_map ctor x =
- let
- val lhs = mk_rec rec_ss i $ (ctor $ x);
- val rhs = rec_s $ (Term.list_comb (rec_map, passive_ids @ convols) $ x);
- in
- fold_rev Logic.all (x :: rec_ss) (mk_Trueprop_eq (lhs, rhs))
- end;
- val goals = map5 mk_goal ks rec_ss rec_maps_rev ctors xFs;
- in
- map2 (fn goal => fn foldx =>
- Skip_Proof.prove lthy [] [] goal (mk_rec_tac rec_defs foldx fst_rec_pair_thms)
- |> Thm.close_derivation)
- goals ctor_fold_thms
- end;
-
- val timer = time (timer "rec definitions & thms");
-
- val (ctor_induct_thm, induct_params) =
- let
- fun mk_prem phi ctor sets x =
- let
- fun mk_IH phi set z =
- let
- val prem = HOLogic.mk_Trueprop (HOLogic.mk_mem (z, set $ x));
- val concl = HOLogic.mk_Trueprop (phi $ z);
- in
- Logic.all z (Logic.mk_implies (prem, concl))
- end;
-
- val IHs = map3 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 = map4 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);
- in
- (Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (phis @ Izs) goal)
- (K (mk_ctor_induct_tac m set_natural'ss init_induct_thm morE_thms mor_Abs_thm
- Rep_inverses Abs_inverses Reps))
- |> Thm.close_derivation,
- rev (Term.add_tfrees goal []))
- end;
-
- val cTs = map (SOME o certifyT lthy o TFree) induct_params;
-
- val weak_ctor_induct_thms =
- let fun insts i = (replicate (i - 1) TrueI) @ (@{thm 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 = HOLogic.mk_Trueprop (HOLogic.mk_mem (z1, (set $ x)));
- val prem2 = HOLogic.mk_Trueprop (HOLogic.mk_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 = map5 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 = map7 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
- (map3 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 = map3 (SOME o certify lthy ooo mk_t) phi2s (Izs1 ~~ Izs1') (Izs2 ~~ Izs2');
- val goal = Logic.list_implies (prems, concl);
- in
- (singleton (Proof_Context.export names_lthy lthy)
- (Skip_Proof.prove lthy [] [] goal
- (mk_ctor_induct2_tac cTs cts ctor_induct_thm weak_ctor_induct_thms))
- |> Thm.close_derivation,
- rev (Term.add_tfrees goal []))
- end;
-
- val timer = time (timer "induction");
-
- (*register new datatypes as BNFs*)
- val lthy = if m = 0 then lthy else
- let
- val fTs = map2 (curry op -->) passiveAs passiveBs;
- val f1Ts = map2 (curry op -->) passiveAs passiveYs;
- val f2Ts = map2 (curry op -->) passiveBs passiveYs;
- val p1Ts = map2 (curry op -->) passiveXs passiveAs;
- val p2Ts = map2 (curry op -->) passiveXs passiveBs;
- val uTs = map2 (curry op -->) Ts Ts';
- val B1Ts = map HOLogic.mk_setT passiveAs;
- val B2Ts = map HOLogic.mk_setT passiveBs;
- val AXTs = map HOLogic.mk_setT passiveXs;
- val XTs = mk_Ts passiveXs;
- val YTs = mk_Ts passiveYs;
- val IRTs = map2 (curry mk_relT) passiveAs passiveBs;
- val IphiTs = map2 mk_pred2T passiveAs passiveBs;
-
- val (((((((((((((((fs, fs'), fs_copy), us),
- B1s), B2s), AXs), (xs, xs')), f1s), f2s), p1s), p2s), (ys, ys')), IRs), Iphis),
- names_lthy) = names_lthy
- |> mk_Frees' "f" fTs
- ||>> mk_Frees "f" fTs
- ||>> mk_Frees "u" uTs
- ||>> mk_Frees "B1" B1Ts
- ||>> mk_Frees "B2" B2Ts
- ||>> mk_Frees "A" AXTs
- ||>> mk_Frees' "x" XTs
- ||>> mk_Frees "f1" f1Ts
- ||>> mk_Frees "f2" f2Ts
- ||>> mk_Frees "p1" p1Ts
- ||>> mk_Frees "p2" p2Ts
- ||>> mk_Frees' "y" passiveAs
- ||>> mk_Frees "R" IRTs
- ||>> mk_Frees "P" IphiTs;
-
- 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 fs_copy_maps = map (mk_map Ts fs_copy Ts' ctor's pmapsABT') ks;
- val Yctors = mk_ctors passiveYs;
- val f1s_maps = map (mk_map Ts f1s YTs Yctors (mk_passive_maps passiveAs passiveYs)) ks;
- val f2s_maps = map (mk_map Ts' f2s YTs Yctors (mk_passive_maps passiveBs passiveYs)) ks;
- val p1s_maps = map (mk_map XTs p1s Ts ctors (mk_passive_maps passiveXs passiveAs)) ks;
- val p2s_maps = map (mk_map XTs p2s Ts' ctor's (mk_passive_maps passiveXs passiveBs)) ks;
-
- val map_simp_thms =
- let
- fun mk_goal fs_map map ctor ctor' = fold_rev Logic.all fs
- (mk_Trueprop_eq (HOLogic.mk_comp (fs_map, ctor),
- HOLogic.mk_comp (ctor', Term.list_comb (map, fs @ fs_maps))));
- val goals = map4 mk_goal fs_maps map_FTFT's ctors ctor's;
- val maps =
- map4 (fn goal => fn foldx => fn map_comp_id => fn map_cong =>
- Skip_Proof.prove lthy [] [] goal (K (mk_map_tac m n foldx map_comp_id map_cong))
- |> Thm.close_derivation)
- goals ctor_fold_thms map_comp_id_thms map_congs;
- in
- map (fn thm => thm RS @{thm pointfreeE}) maps
- end;
-
- val (map_unique_thms, map_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 = map4 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_maps));
- val unique = Skip_Proof.prove lthy [] []
- (fold_rev Logic.all (us @ fs) (Logic.list_implies (prems, goal)))
- (K (mk_map_unique_tac m mor_def fold_unique_mor_thm map_comp_id_thms map_congs))
- |> Thm.close_derivation;
- in
- `split_conj_thm unique
- end;
-
- val timer = time (timer "map functions for the new datatypes");
-
- val bd = mk_cpow sum_bd;
- val bd_Cinfinite = sum_Cinfinite RS @{thm Cinfinite_cpow};
- fun mk_cpow_bd thm = @{thm ordLeq_transitive} OF
- [thm, sum_Card_order RS @{thm cpow_greater_eq}];
- val set_bd_cpowss = map (map mk_cpow_bd) set_bd_sumss;
-
- val timer = time (timer "bounds for the new datatypes");
-
- val ls = 1 upto m;
- val setsss = map (mk_setss o mk_set_Ts) passiveAs;
- val map_setss = map (fn T => map2 (fn Ds =>
- mk_map_of_bnf Ds (passiveAs @ Ts) (mk_set_Ts T)) Dss bnfs) 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 = map5 (fn l => fn T => map3 (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_simp_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 =
- map3 (fn sets => map4 (mk_goal sets) ctors sets) setss_by_range colss map_setss;
- val setss = map (map2 (fn foldx => fn goal =>
- Skip_Proof.prove lthy [] [] goal (K (mk_set_tac foldx)) |> Thm.close_derivation)
- ctor_fold_thms) goalss;
-
- fun mk_simp_goal pas_set act_sets sets ctor z set =
- Logic.all z (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 =>
- map4 (fn Fsets => mk_simp_goal (nth Fsets (i - 1)) (drop m Fsets) sets)
- FTs_setss ctors xFs sets)
- ls setss_by_range;
-
- val set_simpss = map3 (fn i => map3 (fn set_nats => fn goal => fn set =>
- Skip_Proof.prove lthy [] [] goal
- (K (mk_set_simp_tac set (nth set_nats (i - 1)) (drop m set_nats)))
- |> Thm.close_derivation)
- set_natural'ss) ls simp_goalss setss;
- in
- set_simpss
- end;
-
- fun mk_set_thms set_simp = (@{thm xt1(3)} OF [set_simp, @{thm Un_upper1}]) ::
- map (fn i => (@{thm xt1(3)} OF [set_simp, @{thm Un_upper2}]) RS
- (mk_Un_upper n i RS subset_trans) RSN
- (2, @{thm UN_upper} RS subset_trans))
- (1 upto n);
- val Fset_set_thmsss = transpose (map (map mk_set_thms) set_simp_thmss);
-
- val timer = time (timer "set functions for the new datatypes");
-
- val cxs = map (SOME o certify lthy) Izs;
- val setss_by_bnf' =
- map (map (Term.subst_atomic_types (passiveAs ~~ passiveBs))) setss_by_bnf;
- val setss_by_range' = transpose setss_by_bnf';
-
- val set_natural_thmss =
- let
- fun mk_set_natural f map z set set' =
- HOLogic.mk_eq (mk_image f $ (set $ z), set' $ (map $ z));
-
- fun mk_cphi f map z set set' = certify lthy
- (Term.absfree (dest_Free z) (mk_set_natural f map z set set'));
-
- val csetss = map (map (certify lthy)) setss_by_range';
-
- val cphiss = map3 (fn f => fn sets => fn sets' =>
- (map4 (mk_cphi f) fs_maps Izs sets sets')) fs setss_by_range setss_by_range';
-
- val inducts = map (fn cphis =>
- Drule.instantiate' cTs (map SOME cphis @ cxs) ctor_induct_thm) cphiss;
-
- val goals =
- map3 (fn f => fn sets => fn sets' =>
- HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
- (map4 (mk_set_natural f) fs_maps Izs sets sets')))
- fs setss_by_range setss_by_range';
-
- fun mk_tac induct = mk_set_nat_tac m (rtac induct) set_natural'ss map_simp_thms;
- val thms =
- map5 (fn goal => fn csets => fn set_simps => fn induct => fn i =>
- singleton (Proof_Context.export names_lthy lthy)
- (Skip_Proof.prove lthy [] [] goal (mk_tac induct csets set_simps i))
- |> Thm.close_derivation)
- goals csetss set_simp_thmss inducts ls;
- in
- map split_conj_thm thms
- end;
-
- val set_bd_thmss =
- let
- fun mk_set_bd z set = mk_ordLeq (mk_card_of (set $ z)) bd;
-
- fun mk_cphi z set = certify lthy (Term.absfree (dest_Free z) (mk_set_bd z set));
-
- val cphiss = map (map2 mk_cphi Izs) setss_by_range;
-
- val inducts = map (fn cphis =>
- Drule.instantiate' cTs (map SOME cphis @ cxs) ctor_induct_thm) cphiss;
-
- val goals =
- map (fn sets =>
- HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
- (map2 mk_set_bd Izs sets))) setss_by_range;
-
- fun mk_tac induct = mk_set_bd_tac m (rtac induct) bd_Cinfinite set_bd_cpowss;
- val thms =
- map4 (fn goal => fn set_simps => fn induct => fn i =>
- singleton (Proof_Context.export names_lthy lthy)
- (Skip_Proof.prove lthy [] [] goal (mk_tac induct set_simps i))
- |> Thm.close_derivation)
- goals set_simp_thmss inducts ls;
- in
- map split_conj_thm thms
- end;
-
- val map_cong_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_cong sets z fmap gmap =
- HOLogic.mk_imp
- (Library.foldr1 HOLogic.mk_conj (map5 (mk_prem z) sets fs fs_copy ys ys'),
- HOLogic.mk_eq (fmap $ z, gmap $ z));
-
- fun mk_cphi sets z fmap gmap =
- certify lthy (Term.absfree (dest_Free z) (mk_map_cong sets z fmap gmap));
-
- val cphis = map4 mk_cphi setss_by_bnf Izs fs_maps fs_copy_maps;
-
- val induct = Drule.instantiate' cTs (map SOME cphis @ cxs) ctor_induct_thm;
-
- val goal =
- HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
- (map4 mk_map_cong setss_by_bnf Izs fs_maps fs_copy_maps));
-
- val thm = singleton (Proof_Context.export names_lthy lthy)
- (Skip_Proof.prove lthy [] [] goal
- (mk_mcong_tac (rtac induct) Fset_set_thmsss map_congs map_simp_thms))
- |> Thm.close_derivation;
- in
- split_conj_thm thm
- end;
-
- val in_incl_min_alg_thms =
- let
- fun mk_prem z sets =
- HOLogic.mk_mem (z, mk_in As sets (fastype_of z));
-
- fun mk_incl z sets i =
- HOLogic.mk_imp (mk_prem z sets, HOLogic.mk_mem (z, mk_min_alg As ctors i));
-
- fun mk_cphi z sets i =
- certify lthy (Term.absfree (dest_Free z) (mk_incl z sets i));
-
- val cphis = map3 mk_cphi Izs setss_by_bnf ks;
-
- val induct = Drule.instantiate' cTs (map SOME cphis @ cxs) ctor_induct_thm;
-
- val goal =
- HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
- (map3 mk_incl Izs setss_by_bnf ks));
-
- val thm = singleton (Proof_Context.export names_lthy lthy)
- (Skip_Proof.prove lthy [] [] goal
- (mk_incl_min_alg_tac (rtac induct) Fset_set_thmsss alg_set_thms alg_min_alg_thm))
- |> Thm.close_derivation;
- in
- split_conj_thm thm
- end;
-
- val Xsetss = map (map (Term.subst_atomic_types (passiveAs ~~ passiveXs))) setss_by_bnf;
-
- val map_wpull_thms =
- let
- val cTs = map (SOME o certifyT lthy o TFree) induct2_params;
- val cxs = map (SOME o certify lthy) (interleave Izs1 Izs2);
-
- fun mk_prem z1 z2 sets1 sets2 map1 map2 =
- HOLogic.mk_conj
- (HOLogic.mk_mem (z1, mk_in B1s sets1 (fastype_of z1)),
- HOLogic.mk_conj
- (HOLogic.mk_mem (z2, mk_in B2s sets2 (fastype_of z2)),
- HOLogic.mk_eq (map1 $ z1, map2 $ z2)));
-
- val prems = map6 mk_prem Izs1 Izs2 setss_by_bnf setss_by_bnf' f1s_maps f2s_maps;
-
- fun mk_concl z1 z2 sets map1 map2 T x x' =
- mk_Bex (mk_in AXs sets T) (Term.absfree x'
- (HOLogic.mk_conj (HOLogic.mk_eq (map1 $ x, z1), HOLogic.mk_eq (map2 $ x, z2))));
-
- val concls = map8 mk_concl Izs1 Izs2 Xsetss p1s_maps p2s_maps XTs xs xs';
-
- val goals = map2 (curry HOLogic.mk_imp) prems concls;
-
- fun mk_cphi z1 z2 goal = certify lthy (Term.absfree z1 (Term.absfree z2 goal));
-
- val cphis = map3 mk_cphi Izs1' Izs2' goals;
-
- val induct = Drule.instantiate' cTs (map SOME cphis @ cxs) ctor_induct2_thm;
-
- val goal = Logic.list_implies (map HOLogic.mk_Trueprop
- (map8 mk_wpull AXs B1s B2s f1s f2s (replicate m NONE) p1s p2s),
- HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj goals));
-
- val thm = singleton (Proof_Context.export names_lthy lthy)
- (Skip_Proof.prove lthy [] [] goal
- (K (mk_lfp_map_wpull_tac m (rtac induct) map_wpulls map_simp_thms
- (transpose set_simp_thmss) Fset_set_thmsss ctor_inject_thms)))
- |> Thm.close_derivation;
- in
- split_conj_thm thm
- end;
-
- val timer = time (timer "helpers for BNF properties");
-
- val map_id_tacs = map (K o mk_map_id_tac map_ids) map_unique_thms;
- val map_comp_tacs =
- map2 (K oo mk_map_comp_tac map_comp's map_simp_thms) map_unique_thms ks;
- val map_cong_tacs = map (mk_map_cong_tac m) map_cong_thms;
- val set_nat_tacss = map (map (K o mk_set_natural_tac)) (transpose set_natural_thmss);
- val bd_co_tacs = replicate n (K (mk_bd_card_order_tac bd_card_orders));
- val bd_cinf_tacs = replicate n (K (rtac (bd_Cinfinite RS conjunct1) 1));
- val set_bd_tacss = map (map (fn thm => K (rtac thm 1))) (transpose set_bd_thmss);
- val in_bd_tacs = map2 (K oo mk_in_bd_tac sum_Card_order suc_bd_Cnotzero)
- in_incl_min_alg_thms card_of_min_alg_thms;
- val map_wpull_tacs = map (K o mk_wpull_tac) map_wpull_thms;
-
- val srel_O_Gr_tacs = replicate n (simple_srel_O_Gr_tac o #context);
-
- val tacss = map10 zip_axioms map_id_tacs map_comp_tacs map_cong_tacs set_nat_tacss
- bd_co_tacs bd_cinf_tacs set_bd_tacss in_bd_tacs map_wpull_tacs srel_O_Gr_tacs;
-
- 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
- map3 (fn ctor => fn i => map close_wit o minimize_wits o maps (mk_wit witss ctor i))
- ctors (0 upto n - 1) witss
- end;
-
- fun wit_tac _ = mk_wit_tac n (flat set_simp_thmss) (maps wit_thms_of_bnf bnfs);
-
- val policy = user_policy Derive_All_Facts_Note_Most;
-
- val (Ibnfs, lthy) =
- fold_map6 (fn tacs => fn b => fn mapx => fn sets => fn T => fn wits => fn lthy =>
- bnf_def Dont_Inline policy I tacs wit_tac (SOME deads)
- (((((b, fold_rev Term.absfree fs' mapx), sets), absdummy T bd), wits), NONE) lthy
- |> register_bnf (Local_Theory.full_name lthy b))
- tacss bs fs_maps setss_by_bnf Ts ctor_witss lthy;
-
- val fold_maps = fold_thms lthy (map (fn bnf =>
- mk_unabs_def m (map_def_of_bnf bnf RS @{thm meta_eq_to_obj_eq})) Ibnfs);
-
- val fold_sets = fold_thms lthy (maps (fn bnf =>
- map (fn thm => thm RS @{thm meta_eq_to_obj_eq}) (set_defs_of_bnf bnf)) Ibnfs);
-
- val timer = time (timer "registered new datatypes as BNFs");
-
- val srels = map2 (fn Ds => mk_srel_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs;
- val Isrels = map (mk_srel_of_bnf deads passiveAs passiveBs) Ibnfs;
- val rels = map2 (fn Ds => mk_rel_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs;
- val Irels = map (mk_rel_of_bnf deads passiveAs passiveBs) Ibnfs;
-
- val IrelRs = map (fn Isrel => Term.list_comb (Isrel, IRs)) Isrels;
- val relRs = map (fn srel => Term.list_comb (srel, IRs @ IrelRs)) srels;
- val Ipredphis = map (fn Isrel => Term.list_comb (Isrel, Iphis)) Irels;
- val predphis = map (fn srel => Term.list_comb (srel, Iphis @ Ipredphis)) rels;
-
- val in_srels = map in_srel_of_bnf bnfs;
- val in_Isrels = map in_srel_of_bnf Ibnfs;
- val srel_defs = map srel_def_of_bnf bnfs;
- val Isrel_defs = map srel_def_of_bnf Ibnfs;
- val Irel_defs = map rel_def_of_bnf Ibnfs;
-
- val set_incl_thmss = map (map (fold_sets o hd)) Fset_set_thmsss;
- val set_set_incl_thmsss = map (transpose o map (map fold_sets o tl)) Fset_set_thmsss;
- val folded_map_simp_thms = map fold_maps map_simp_thms;
- val folded_set_simp_thmss = map (map fold_sets) set_simp_thmss;
- val folded_set_simp_thmss' = transpose folded_set_simp_thmss;
-
- val Isrel_simp_thms =
- let
- fun mk_goal xF yF ctor ctor' IrelR relR = fold_rev Logic.all (xF :: yF :: IRs)
- (mk_Trueprop_eq (HOLogic.mk_mem (HOLogic.mk_prod (ctor $ xF, ctor' $ yF), IrelR),
- HOLogic.mk_mem (HOLogic.mk_prod (xF, yF), relR)));
- val goals = map6 mk_goal xFs yFs ctors ctor's IrelRs relRs;
- in
- map12 (fn i => fn goal => fn in_srel => fn map_comp => fn map_cong =>
- fn map_simp => fn set_simps => fn ctor_inject => fn ctor_dtor =>
- fn set_naturals => fn set_incls => fn set_set_inclss =>
- Skip_Proof.prove lthy [] [] goal
- (K (mk_srel_simp_tac in_Isrels i in_srel map_comp map_cong map_simp set_simps
- ctor_inject ctor_dtor set_naturals set_incls set_set_inclss))
- |> Thm.close_derivation)
- ks goals in_srels map_comp's map_congs folded_map_simp_thms folded_set_simp_thmss'
- ctor_inject_thms ctor_dtor_thms set_natural'ss set_incl_thmss set_set_incl_thmsss
- end;
-
- val Irel_simp_thms =
- let
- fun mk_goal xF yF ctor ctor' Ipredphi predphi = fold_rev Logic.all (xF :: yF :: Iphis)
- (mk_Trueprop_eq (Ipredphi $ (ctor $ xF) $ (ctor' $ yF), predphi $ xF $ yF));
- val goals = map6 mk_goal xFs yFs ctors ctor's Ipredphis predphis;
- in
- map3 (fn goal => fn srel_def => fn Isrel_simp =>
- Skip_Proof.prove lthy [] [] goal
- (mk_rel_simp_tac srel_def Irel_defs Isrel_defs Isrel_simp)
- |> Thm.close_derivation)
- goals srel_defs Isrel_simp_thms
- end;
-
- val timer = time (timer "additional properties");
-
- val ls' = if m = 1 then [0] else ls
-
- val Ibnf_common_notes =
- [(map_uniqueN, [fold_maps map_unique_thm])]
- |> map (fn (thmN, thms) =>
- ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), [(thms, [])]));
-
- val Ibnf_notes =
- [(map_simpsN, map single folded_map_simp_thms),
- (set_inclN, set_incl_thmss),
- (set_set_inclN, map flat set_set_incl_thmsss),
- (srel_simpN, map single Isrel_simp_thms),
- (rel_simpN, map single Irel_simp_thms)] @
- map2 (fn i => fn thms => (mk_set_simpsN i, map single thms)) ls' folded_set_simp_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; lthy |> Local_Theory.notes (Ibnf_common_notes @ Ibnf_notes) |> snd
- end;
-
- val common_notes =
- [(ctor_inductN, [ctor_induct_thm]),
- (ctor_induct2N, [ctor_induct2_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_fold_uniqueN, ctor_fold_unique_thms),
- (ctor_foldsN, ctor_fold_thms),
- (ctor_injectN, ctor_inject_thms),
- (ctor_recsN, ctor_rec_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)
- in
- ((dtors, ctors, folds, recs, ctor_induct_thm, dtor_ctor_thms, ctor_dtor_thms, ctor_inject_thms,
- ctor_fold_thms, ctor_rec_thms),
- lthy |> Local_Theory.notes (common_notes @ notes) |> snd)
- end;
-
-val _ =
- Outer_Syntax.local_theory @{command_spec "data_raw"}
- "define BNF-based inductive datatypes (low-level)"
- (Parse.and_list1
- ((Parse.binding --| @{keyword ":"}) -- (Parse.typ --| @{keyword "="} -- Parse.typ)) >>
- (snd oo fp_bnf_cmd bnf_lfp o apsnd split_list o split_list));
-
-val _ =
- Outer_Syntax.local_theory @{command_spec "data"} "define BNF-based inductive datatypes"
- (parse_datatype_cmd true bnf_lfp);
-
-end;
--- a/src/HOL/Codatatype/Tools/bnf_lfp_tactics.ML Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,835 +0,0 @@
-(* Title: HOL/BNF/Tools/bnf_lfp_tactics.ML
- Author: Dmitriy Traytel, TU Muenchen
- Author: Andrei Popescu, TU Muenchen
- Copyright 2012
-
-Tactics for the datatype construction.
-*)
-
-signature BNF_LFP_TACTICS =
-sig
- val mk_alg_min_alg_tac: int -> thm -> thm list -> thm -> thm -> thm list list -> thm list ->
- thm list -> tactic
- val mk_alg_not_empty_tac: thm -> thm list -> thm list -> tactic
- val mk_alg_select_tac: thm -> {prems: 'a, context: Proof.context} -> tactic
- val mk_alg_set_tac: thm -> tactic
- val mk_bd_card_order_tac: thm list -> tactic
- val mk_bd_limit_tac: int -> thm -> tactic
- val mk_card_of_min_alg_tac: thm -> thm -> thm -> thm -> thm -> tactic
- val mk_copy_alg_tac: thm list list -> thm list -> thm -> thm -> thm -> tactic
- val mk_copy_str_tac: thm list list -> thm -> thm list -> tactic
- val mk_ctor_induct_tac: int -> thm list list -> thm -> thm list -> thm -> thm list -> thm list ->
- thm list -> tactic
- val mk_ctor_induct2_tac: ctyp option list -> cterm option list -> thm -> thm list ->
- {prems: 'a, context: Proof.context} -> tactic
- val mk_dtor_o_ctor_tac: thm -> thm -> thm -> thm -> thm list -> tactic
- val mk_ex_copy_alg_tac: int -> thm -> thm -> tactic
- val mk_in_bd_tac: thm -> thm -> thm -> thm -> tactic
- val mk_incl_min_alg_tac: (int -> tactic) -> thm list list list -> thm list -> thm ->
- {prems: 'a, context: Proof.context} -> tactic
- val mk_init_ex_mor_tac: thm -> thm -> thm -> thm list -> thm -> thm -> thm ->
- {prems: 'a, context: Proof.context} -> tactic
- val mk_init_induct_tac: int -> thm -> thm -> thm list -> thm list -> tactic
- val mk_init_unique_mor_tac: int -> thm -> thm -> thm list -> thm list -> thm list -> thm list ->
- thm list -> tactic
- val mk_iso_alt_tac: thm list -> thm -> tactic
- val mk_fold_unique_mor_tac: thm list -> thm list -> thm list -> thm -> thm -> thm -> tactic
- val mk_least_min_alg_tac: thm -> thm -> tactic
- val mk_lfp_map_wpull_tac: int -> (int -> tactic) -> thm list -> thm list -> thm list list ->
- thm list list list -> thm list -> tactic
- val mk_map_comp_tac: thm list -> thm list -> thm -> int -> tactic
- val mk_map_id_tac: thm list -> thm -> tactic
- val mk_map_tac: int -> int -> thm -> thm -> thm -> tactic
- val mk_map_unique_tac: int -> thm -> thm -> thm list -> thm list -> tactic
- val mk_mcong_tac: (int -> tactic) -> thm list list list -> thm list -> thm list ->
- {prems: 'a, context: Proof.context} -> tactic
- val mk_min_algs_card_of_tac: ctyp -> cterm -> int -> thm -> thm list -> thm list -> thm -> thm ->
- thm -> thm -> thm -> thm -> thm -> thm -> tactic
- val mk_min_algs_least_tac: ctyp -> cterm -> thm -> thm list -> thm list -> tactic
- val mk_min_algs_mono_tac: thm -> tactic
- val mk_min_algs_tac: thm -> thm list -> tactic
- val mk_mor_Abs_tac: thm -> thm list -> thm list -> tactic
- val mk_mor_Rep_tac: thm list -> thm -> thm list -> thm list -> thm list ->
- {prems: 'a, context: Proof.context} -> tactic
- val mk_mor_UNIV_tac: int -> thm list -> thm -> tactic
- val mk_mor_comp_tac: thm -> thm list list -> thm list -> tactic
- val mk_mor_convol_tac: 'a list -> thm -> tactic
- val mk_mor_elim_tac: thm -> tactic
- val mk_mor_incl_tac: thm -> thm list -> tactic
- val mk_mor_inv_tac: thm -> thm -> thm list list -> thm list -> thm list -> thm list -> tactic
- val mk_mor_fold_tac: ctyp -> cterm -> thm list -> thm -> thm -> tactic
- val mk_mor_select_tac: thm -> thm -> thm -> thm -> thm -> thm -> thm list -> thm list list ->
- thm list -> tactic
- val mk_mor_str_tac: 'a list -> thm -> tactic
- val mk_rec_tac: thm list -> thm -> thm list -> {prems: 'a, context: Proof.context} -> tactic
- val mk_set_bd_tac: int -> (int -> tactic) -> thm -> thm list list -> thm list -> int ->
- {prems: 'a, context: Proof.context} -> tactic
- val mk_set_nat_tac: int -> (int -> tactic) -> thm list list -> thm list -> cterm list ->
- thm list -> int -> {prems: 'a, context: Proof.context} -> tactic
- val mk_set_natural_tac: thm -> tactic
- val mk_set_simp_tac: thm -> thm -> thm list -> tactic
- val mk_set_tac: thm -> tactic
- val mk_srel_simp_tac: thm list -> int -> thm -> thm -> thm -> thm -> thm list -> thm ->
- thm -> thm list -> thm list -> thm list list -> tactic
- val mk_wit_tac: int -> thm list -> thm list -> tactic
- val mk_wpull_tac: thm -> tactic
-end;
-
-structure BNF_LFP_Tactics : BNF_LFP_TACTICS =
-struct
-
-open BNF_Tactics
-open BNF_LFP_Util
-open BNF_Util
-
-val fst_snd_convs = @{thms fst_conv snd_conv};
-val id_apply = @{thm id_apply};
-val meta_mp = @{thm meta_mp};
-val ord_eq_le_trans = @{thm ord_eq_le_trans};
-val subset_trans = @{thm subset_trans};
-val trans_fun_cong_image_id_id_apply = @{thm trans[OF fun_cong[OF image_id] id_apply]};
-
-fun mk_alg_set_tac alg_def =
- dtac (alg_def RS iffD1) 1 THEN
- REPEAT_DETERM (etac conjE 1) THEN
- EVERY' [etac bspec, rtac CollectI] 1 THEN
- REPEAT_DETERM (etac conjI 1) THEN atac 1;
-
-fun mk_alg_not_empty_tac alg_set alg_sets wits =
- (EVERY' [rtac notI, hyp_subst_tac, ftac alg_set] THEN'
- REPEAT_DETERM o FIRST'
- [rtac subset_UNIV,
- EVERY' [rtac @{thm subset_emptyI}, eresolve_tac wits],
- EVERY' [rtac subsetI, rtac FalseE, eresolve_tac wits],
- EVERY' [rtac subsetI, dresolve_tac wits, hyp_subst_tac,
- FIRST' (map (fn thm => rtac thm THEN' atac) alg_sets)]] THEN'
- etac @{thm emptyE}) 1;
-
-fun mk_mor_elim_tac mor_def =
- (dtac (subst OF [mor_def]) THEN'
- REPEAT o etac conjE THEN'
- TRY o rtac @{thm image_subsetI} THEN'
- etac bspec THEN'
- atac) 1;
-
-fun mk_mor_incl_tac mor_def map_id's =
- (stac mor_def THEN'
- rtac conjI THEN'
- CONJ_WRAP' (K (EVERY' [rtac ballI, etac set_mp, stac id_apply, atac])) map_id's THEN'
- CONJ_WRAP' (fn thm =>
- (EVERY' [rtac ballI, rtac trans, rtac id_apply, stac thm, rtac refl])) map_id's) 1;
-
-fun mk_mor_comp_tac mor_def set_natural's map_comp_ids =
- let
- val fbetw_tac = EVERY' [rtac ballI, stac o_apply, etac bspec, etac bspec, atac];
- fun mor_tac (set_natural', map_comp_id) =
- EVERY' [rtac ballI, stac o_apply, rtac trans,
- rtac trans, dtac @{thm rev_bspec}, atac, etac arg_cong,
- REPEAT o eresolve_tac [CollectE, conjE], etac bspec, rtac CollectI] THEN'
- CONJ_WRAP' (fn thm =>
- FIRST' [rtac subset_UNIV,
- (EVERY' [rtac ord_eq_le_trans, rtac thm, rtac @{thm image_subsetI},
- etac bspec, etac set_mp, atac])]) set_natural' THEN'
- rtac (map_comp_id RS arg_cong);
- in
- (dtac (mor_def RS subst) THEN' dtac (mor_def RS subst) THEN' stac mor_def THEN'
- REPEAT o etac conjE THEN'
- rtac conjI THEN'
- CONJ_WRAP' (K fbetw_tac) set_natural's THEN'
- CONJ_WRAP' mor_tac (set_natural's ~~ map_comp_ids)) 1
- end;
-
-fun mk_mor_inv_tac alg_def mor_def set_natural's morEs map_comp_ids map_congLs =
- let
- val fbetw_tac = EVERY' [rtac ballI, etac set_mp, etac imageI];
- fun Collect_tac set_natural' =
- CONJ_WRAP' (fn thm =>
- FIRST' [rtac subset_UNIV,
- (EVERY' [rtac ord_eq_le_trans, rtac thm, rtac subset_trans,
- etac @{thm image_mono}, atac])]) set_natural';
- fun mor_tac (set_natural', ((morE, map_comp_id), map_congL)) =
- EVERY' [rtac ballI, ftac @{thm rev_bspec}, atac,
- REPEAT o eresolve_tac [CollectE, conjE], rtac sym, rtac trans, rtac sym,
- etac @{thm inverE}, etac bspec, rtac CollectI, Collect_tac set_natural',
- rtac trans, etac (morE RS arg_cong), rtac CollectI, Collect_tac set_natural',
- rtac trans, rtac (map_comp_id RS arg_cong), rtac (map_congL RS arg_cong),
- REPEAT_DETERM_N (length morEs) o
- (EVERY' [rtac subst, rtac @{thm inver_pointfree}, etac @{thm inver_mono}, atac])];
- in
- (stac mor_def THEN'
- dtac (alg_def RS iffD1) THEN'
- dtac (alg_def RS iffD1) THEN'
- REPEAT o etac conjE THEN'
- rtac conjI THEN'
- CONJ_WRAP' (K fbetw_tac) set_natural's THEN'
- CONJ_WRAP' mor_tac (set_natural's ~~ (morEs ~~ map_comp_ids ~~ map_congLs))) 1
- end;
-
-fun mk_mor_str_tac ks mor_def =
- (stac mor_def THEN' rtac conjI THEN'
- CONJ_WRAP' (K (EVERY' [rtac ballI, rtac UNIV_I])) ks THEN'
- CONJ_WRAP' (K (EVERY' [rtac ballI, rtac refl])) ks) 1;
-
-fun mk_mor_convol_tac ks mor_def =
- (stac mor_def THEN' rtac conjI THEN'
- CONJ_WRAP' (K (EVERY' [rtac ballI, rtac UNIV_I])) ks THEN'
- CONJ_WRAP' (K (EVERY' [rtac ballI, rtac trans, rtac @{thm fst_convol'}, rtac o_apply])) ks) 1;
-
-fun mk_mor_UNIV_tac m morEs mor_def =
- let
- val n = length morEs;
- fun mor_tac morE = EVERY' [rtac ext, rtac trans, rtac o_apply, rtac trans, etac morE,
- rtac CollectI, CONJ_WRAP' (K (rtac subset_UNIV)) (1 upto m + n),
- rtac sym, rtac o_apply];
- in
- EVERY' [rtac iffI, CONJ_WRAP' mor_tac morEs,
- stac mor_def, rtac conjI, CONJ_WRAP' (K (rtac ballI THEN' rtac UNIV_I)) morEs,
- REPEAT_DETERM o etac conjE, REPEAT_DETERM_N n o dtac (@{thm fun_eq_iff} RS subst),
- CONJ_WRAP' (K (EVERY' [rtac ballI, REPEAT_DETERM o etac allE, rtac trans,
- etac (o_apply RS subst), rtac o_apply])) morEs] 1
- end;
-
-fun mk_iso_alt_tac mor_images mor_inv =
- let
- val n = length mor_images;
- fun if_wrap_tac thm =
- EVERY' [rtac ssubst, rtac @{thm bij_betw_iff_ex}, rtac exI, rtac conjI,
- rtac @{thm inver_surj}, etac thm, etac thm, atac, etac conjI, atac]
- val if_tac =
- EVERY' [etac thin_rl, etac thin_rl, REPEAT o eresolve_tac [conjE, exE],
- rtac conjI, atac, CONJ_WRAP' if_wrap_tac mor_images];
- val only_if_tac =
- EVERY' [rtac conjI, etac conjunct1, EVERY' (map (fn thm =>
- EVERY' [rtac exE, rtac @{thm bij_betw_ex_weakE}, etac (conjunct2 RS thm)])
- (map (mk_conjunctN n) (1 upto n))), REPEAT o rtac exI, rtac conjI, rtac mor_inv,
- etac conjunct1, atac, atac, REPEAT_DETERM_N n o atac,
- CONJ_WRAP' (K (etac conjunct2)) mor_images];
- in
- (rtac iffI THEN' if_tac THEN' only_if_tac) 1
- end;
-
-fun mk_copy_str_tac set_natural's alg_def alg_sets =
- let
- val n = length alg_sets;
- val bij_betw_inv_tac =
- EVERY' [etac thin_rl, REPEAT_DETERM_N n o EVERY' [dtac @{thm bij_betwI}, atac, atac],
- REPEAT_DETERM_N (2 * n) o etac thin_rl, REPEAT_DETERM_N (n - 1) o etac conjI, atac];
- fun set_tac thms =
- EVERY' [rtac ord_eq_le_trans, resolve_tac thms, rtac subset_trans,
- etac @{thm image_mono}, rtac equalityD1, etac @{thm bij_betw_imageE}];
- val copy_str_tac =
- CONJ_WRAP' (fn (thms, thm) =>
- EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac set_mp,
- rtac equalityD1, etac @{thm bij_betw_imageE}, rtac imageI, etac thm,
- REPEAT_DETERM o rtac subset_UNIV, REPEAT_DETERM_N n o (set_tac thms)])
- (set_natural's ~~ alg_sets);
- in
- (rtac rev_mp THEN' DETERM o bij_betw_inv_tac THEN' rtac impI THEN'
- stac alg_def THEN' copy_str_tac) 1
- end;
-
-fun mk_copy_alg_tac set_natural's alg_sets mor_def iso_alt copy_str =
- let
- val n = length alg_sets;
- val fbetw_tac = CONJ_WRAP' (K (etac @{thm bij_betwE})) alg_sets;
- fun set_tac thms =
- EVERY' [rtac ord_eq_le_trans, resolve_tac thms, rtac subset_trans,
- REPEAT_DETERM o etac conjE, etac @{thm image_mono},
- rtac equalityD1, etac @{thm bij_betw_imageE}];
- val mor_tac =
- CONJ_WRAP' (fn (thms, thm) =>
- EVERY' [rtac ballI, etac CollectE, etac @{thm inverE}, etac thm,
- REPEAT_DETERM o rtac subset_UNIV, REPEAT_DETERM_N n o (set_tac thms)])
- (set_natural's ~~ alg_sets);
- in
- (rtac (iso_alt RS @{thm ssubst[of _ _ "%x. x"]}) THEN'
- etac copy_str THEN' REPEAT_DETERM o atac THEN'
- rtac conjI THEN' stac mor_def THEN' rtac conjI THEN' fbetw_tac THEN' mor_tac THEN'
- CONJ_WRAP' (K atac) alg_sets) 1
- end;
-
-fun mk_ex_copy_alg_tac n copy_str copy_alg =
- EVERY' [REPEAT_DETERM_N n o rtac exI, rtac conjI, etac copy_str,
- REPEAT_DETERM_N n o atac,
- REPEAT_DETERM_N n o etac @{thm bij_betw_inver2},
- REPEAT_DETERM_N n o etac @{thm bij_betw_inver1}, etac copy_alg,
- REPEAT_DETERM_N n o atac,
- REPEAT_DETERM_N n o etac @{thm bij_betw_inver2},
- REPEAT_DETERM_N n o etac @{thm bij_betw_inver1}] 1;
-
-fun mk_bd_limit_tac n bd_Cinfinite =
- EVERY' [REPEAT_DETERM o etac conjE, rtac rev_mp, rtac @{thm Cinfinite_limit_finite},
- REPEAT_DETERM_N n o rtac @{thm finite.insertI}, rtac @{thm finite.emptyI},
- REPEAT_DETERM_N n o etac @{thm insert_subsetI}, rtac @{thm empty_subsetI},
- rtac bd_Cinfinite, rtac impI, etac bexE, rtac bexI,
- CONJ_WRAP' (fn i =>
- EVERY' [etac bspec, REPEAT_DETERM_N i o rtac @{thm insertI2}, rtac @{thm insertI1}])
- (0 upto n - 1),
- atac] 1;
-
-fun mk_min_algs_tac worel in_congs =
- let
- val minG_tac = EVERY' [rtac @{thm UN_cong}, rtac refl, dtac bspec, atac, etac arg_cong];
- fun minH_tac thm =
- EVERY' [rtac @{thm Un_cong}, minG_tac, rtac @{thm image_cong}, rtac thm,
- REPEAT_DETERM_N (length in_congs) o minG_tac, rtac refl];
- in
- (rtac (worel RS (@{thm wo_rel.worec_fixpoint} RS fun_cong)) THEN' rtac ssubst THEN'
- rtac meta_eq_to_obj_eq THEN' rtac (worel RS @{thm wo_rel.adm_wo_def}) THEN'
- REPEAT_DETERM_N 3 o rtac allI THEN' rtac impI THEN'
- CONJ_WRAP_GEN' (EVERY' [rtac Pair_eqI, rtac conjI]) minH_tac in_congs) 1
- end;
-
-fun mk_min_algs_mono_tac min_algs = EVERY' [stac @{thm relChain_def}, rtac allI, rtac allI,
- rtac impI, rtac @{thm case_split}, rtac @{thm xt1(3)}, rtac min_algs, etac @{thm FieldI2},
- rtac subsetI, rtac UnI1, rtac @{thm UN_I}, etac @{thm underS_I}, atac, atac,
- rtac equalityD1, dtac @{thm notnotD}, hyp_subst_tac, rtac refl] 1;
-
-fun mk_min_algs_card_of_tac cT ct m worel min_algs in_bds bd_Card_order bd_Cnotzero
- suc_Card_order suc_Cinfinite suc_Cnotzero suc_Asuc Asuc_Cinfinite Asuc_Cnotzero =
- let
- val induct = worel RS
- Drule.instantiate' [SOME cT] [NONE, SOME ct] @{thm well_order_induct_imp};
- val src = 1 upto m + 1;
- val dest = (m + 1) :: (1 upto m);
- val absorbAs_tac = if m = 0 then K (all_tac)
- else EVERY' [rtac @{thm ordIso_transitive}, rtac @{thm csum_cong1},
- rtac @{thm ordIso_transitive},
- BNF_Tactics.mk_rotate_eq_tac (rtac @{thm ordIso_refl} THEN'
- FIRST' [rtac @{thm card_of_Card_order}, rtac @{thm Card_order_csum},
- rtac @{thm Card_order_cexp}])
- @{thm ordIso_transitive} @{thm csum_assoc} @{thm csum_com} @{thm csum_cong}
- src dest,
- rtac @{thm csum_absorb1}, rtac Asuc_Cinfinite, rtac ctrans, rtac @{thm ordLeq_csum1},
- FIRST' [rtac @{thm Card_order_csum}, rtac @{thm card_of_Card_order}],
- rtac @{thm ordLeq_cexp1}, rtac suc_Cnotzero, rtac @{thm Card_order_csum}];
-
- val minG_tac = EVERY' [rtac @{thm UNION_Cinfinite_bound}, rtac @{thm ordLess_imp_ordLeq},
- rtac @{thm ordLess_transitive}, rtac @{thm card_of_underS}, rtac suc_Card_order,
- atac, rtac suc_Asuc, rtac ballI, etac allE, dtac mp, etac @{thm underS_E},
- dtac mp, etac @{thm underS_Field}, REPEAT o etac conjE, atac, rtac Asuc_Cinfinite]
-
- fun mk_minH_tac (min_alg, in_bd) = EVERY' [rtac @{thm ordIso_ordLeq_trans},
- rtac @{thm card_of_ordIso_subst}, etac min_alg, rtac @{thm Un_Cinfinite_bound},
- minG_tac, rtac ctrans, rtac @{thm card_of_image}, rtac ctrans, rtac in_bd, rtac ctrans,
- rtac @{thm cexp_mono1_Cnotzero}, rtac @{thm csum_mono1},
- REPEAT_DETERM_N m o rtac @{thm csum_mono2},
- CONJ_WRAP_GEN' (rtac @{thm csum_cinfinite_bound}) (K minG_tac) min_algs,
- REPEAT_DETERM o FIRST'
- [rtac @{thm card_of_Card_order}, rtac @{thm Card_order_csum}, rtac Asuc_Cinfinite],
- rtac @{thm csum_Cnotzero2}, rtac @{thm ctwo_Cnotzero}, rtac bd_Card_order,
- rtac @{thm ordIso_ordLeq_trans}, rtac @{thm cexp_cong1_Cnotzero}, absorbAs_tac,
- rtac @{thm csum_absorb1}, rtac Asuc_Cinfinite, rtac @{thm ctwo_ordLeq_Cinfinite},
- rtac Asuc_Cinfinite, rtac bd_Card_order,
- rtac @{thm csum_Cnotzero2}, rtac @{thm ctwo_Cnotzero}, rtac Asuc_Cnotzero,
- rtac @{thm ordIso_imp_ordLeq}, rtac @{thm cexp_cprod_ordLeq},
- TRY o rtac @{thm csum_Cnotzero2}, rtac @{thm ctwo_Cnotzero}, rtac suc_Cinfinite,
- rtac bd_Cnotzero, rtac @{thm cardSuc_ordLeq}, rtac bd_Card_order, rtac Asuc_Cinfinite];
- in
- (rtac induct THEN'
- rtac impI THEN'
- CONJ_WRAP' mk_minH_tac (min_algs ~~ in_bds)) 1
- end;
-
-fun mk_min_algs_least_tac cT ct worel min_algs alg_sets =
- let
- val induct = worel RS
- Drule.instantiate' [SOME cT] [NONE, SOME ct] @{thm well_order_induct_imp};
-
- val minG_tac = EVERY' [rtac @{thm UN_least}, etac allE, dtac mp, etac @{thm underS_E},
- dtac mp, etac @{thm underS_Field}, REPEAT_DETERM o etac conjE, atac];
-
- fun mk_minH_tac (min_alg, alg_set) = EVERY' [rtac ord_eq_le_trans, etac min_alg,
- rtac @{thm Un_least}, minG_tac, rtac @{thm image_subsetI},
- REPEAT_DETERM o eresolve_tac [CollectE, conjE], etac alg_set,
- REPEAT_DETERM o FIRST' [atac, etac subset_trans THEN' minG_tac]];
- in
- (rtac induct THEN'
- rtac impI THEN'
- CONJ_WRAP' mk_minH_tac (min_algs ~~ alg_sets)) 1
- end;
-
-fun mk_alg_min_alg_tac m alg_def min_alg_defs bd_limit bd_Cinfinite
- set_bdss min_algs min_alg_monos =
- let
- val n = length min_algs;
- fun mk_cardSuc_UNION_tac set_bds (mono, def) = EVERY'
- [rtac bexE, rtac @{thm cardSuc_UNION_Cinfinite}, rtac bd_Cinfinite, rtac mono,
- etac (def RSN (2, @{thm subset_trans[OF _ equalityD1]})), resolve_tac set_bds];
- fun mk_conjunct_tac (set_bds, (min_alg, min_alg_def)) =
- EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, conjE],
- EVERY' (map (mk_cardSuc_UNION_tac set_bds) (min_alg_monos ~~ min_alg_defs)), rtac bexE,
- rtac bd_limit, REPEAT_DETERM_N (n - 1) o etac conjI, atac,
- rtac (min_alg_def RS @{thm set_mp[OF equalityD2]}),
- rtac @{thm UN_I}, REPEAT_DETERM_N (m + 3 * n) o etac thin_rl, atac, rtac set_mp,
- rtac equalityD2, rtac min_alg, atac, rtac UnI2, rtac @{thm image_eqI}, rtac refl,
- rtac CollectI, REPEAT_DETERM_N m o dtac asm_rl, REPEAT_DETERM_N n o etac thin_rl,
- REPEAT_DETERM o etac conjE,
- CONJ_WRAP' (K (FIRST' [atac,
- EVERY' [etac subset_trans, rtac subsetI, rtac @{thm UN_I},
- etac @{thm underS_I}, atac, atac]]))
- set_bds];
- in
- (rtac (alg_def RS @{thm ssubst[of _ _ "%x. x"]}) THEN'
- CONJ_WRAP' mk_conjunct_tac (set_bdss ~~ (min_algs ~~ min_alg_defs))) 1
- end;
-
-fun mk_card_of_min_alg_tac min_alg_def card_of suc_Card_order suc_Asuc Asuc_Cinfinite =
- EVERY' [stac min_alg_def, rtac @{thm UNION_Cinfinite_bound},
- rtac @{thm ordIso_ordLeq_trans}, rtac @{thm card_of_Field_ordIso}, rtac suc_Card_order,
- rtac @{thm ordLess_imp_ordLeq}, rtac suc_Asuc, rtac ballI, dtac rev_mp, rtac card_of,
- REPEAT_DETERM o etac conjE, atac, rtac Asuc_Cinfinite] 1;
-
-fun mk_least_min_alg_tac min_alg_def least =
- EVERY' [stac min_alg_def, rtac @{thm UN_least}, dtac least, dtac mp, atac,
- REPEAT_DETERM o etac conjE, atac] 1;
-
-fun mk_alg_select_tac Abs_inverse {context = ctxt, prems = _} =
- EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac] 1 THEN
- unfold_thms_tac ctxt (Abs_inverse :: fst_snd_convs) THEN atac 1;
-
-fun mk_mor_select_tac mor_def mor_cong mor_comp mor_incl_min_alg alg_def alg_select
- alg_sets set_natural's str_init_defs =
- let
- val n = length alg_sets;
- val fbetw_tac =
- CONJ_WRAP' (K (EVERY' [rtac ballI, etac @{thm rev_bspec}, etac CollectE, atac])) alg_sets;
- val mor_tac =
- CONJ_WRAP' (fn thm => EVERY' [rtac ballI, rtac thm]) str_init_defs;
- fun alg_epi_tac ((alg_set, str_init_def), set_natural') =
- EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac CollectI,
- rtac ballI, ftac (alg_select RS bspec), stac str_init_def, etac alg_set,
- REPEAT_DETERM o FIRST' [rtac subset_UNIV,
- EVERY' [rtac ord_eq_le_trans, resolve_tac set_natural', rtac subset_trans,
- etac @{thm image_mono}, rtac @{thm image_Collect_subsetI}, etac bspec, atac]]];
- in
- (rtac mor_cong THEN' REPEAT_DETERM_N n o (rtac sym THEN' rtac @{thm o_id}) THEN'
- rtac (Thm.permute_prems 0 1 mor_comp) THEN' etac (Thm.permute_prems 0 1 mor_comp) THEN'
- stac mor_def THEN' rtac conjI THEN' fbetw_tac THEN' mor_tac THEN' rtac mor_incl_min_alg THEN'
- stac alg_def THEN' CONJ_WRAP' alg_epi_tac ((alg_sets ~~ str_init_defs) ~~ set_natural's)) 1
- end;
-
-fun mk_init_ex_mor_tac Abs_inverse copy_alg_ex alg_min_alg card_of_min_algs
- mor_comp mor_select mor_incl_min_alg {context = ctxt, prems = _} =
- let
- val n = length card_of_min_algs;
- val card_of_ordIso_tac = EVERY' [rtac ssubst, rtac @{thm card_of_ordIso},
- rtac @{thm ordIso_symmetric}, rtac conjunct1, rtac conjunct2, atac];
- fun internalize_tac card_of = EVERY' [rtac subst, rtac @{thm internalize_card_of_ordLeq2},
- rtac @{thm ordLeq_ordIso_trans}, rtac card_of, rtac subst,
- rtac @{thm Card_order_iff_ordIso_card_of}, rtac @{thm Card_order_cexp}];
- in
- (rtac rev_mp THEN'
- REPEAT_DETERM_N (2 * n) o (rtac mp THEN' rtac @{thm ex_mono} THEN' rtac impI) THEN'
- REPEAT_DETERM_N (n + 1) o etac thin_rl THEN' rtac (alg_min_alg RS copy_alg_ex) THEN'
- REPEAT_DETERM_N n o atac THEN'
- REPEAT_DETERM_N n o card_of_ordIso_tac THEN'
- EVERY' (map internalize_tac card_of_min_algs) THEN'
- rtac impI THEN'
- REPEAT_DETERM o eresolve_tac [exE, conjE] THEN'
- REPEAT_DETERM o rtac exI THEN'
- rtac mor_select THEN' atac THEN' rtac CollectI THEN'
- REPEAT_DETERM o rtac exI THEN'
- rtac conjI THEN' rtac refl THEN' atac THEN'
- K (unfold_thms_tac ctxt (Abs_inverse :: fst_snd_convs)) THEN'
- etac mor_comp THEN' etac mor_incl_min_alg) 1
- end;
-
-fun mk_init_unique_mor_tac m
- alg_def alg_min_alg least_min_algs in_monos alg_sets morEs map_congs =
- let
- val n = length least_min_algs;
- val ks = (1 upto n);
-
- fun mor_tac morE in_mono = EVERY' [etac morE, rtac set_mp, rtac in_mono,
- REPEAT_DETERM_N n o rtac @{thm Collect_restrict}, rtac CollectI,
- REPEAT_DETERM_N (m + n) o (TRY o rtac conjI THEN' atac)];
- fun cong_tac map_cong = EVERY' [rtac (map_cong RS arg_cong),
- REPEAT_DETERM_N m o rtac refl,
- REPEAT_DETERM_N n o (etac @{thm prop_restrict} THEN' atac)];
-
- fun mk_alg_tac (alg_set, (in_mono, (morE, map_cong))) = EVERY' [rtac ballI, rtac CollectI,
- REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac conjI, rtac (alg_min_alg RS alg_set),
- REPEAT_DETERM_N m o rtac subset_UNIV,
- REPEAT_DETERM_N n o (etac subset_trans THEN' rtac @{thm Collect_restrict}),
- rtac trans, mor_tac morE in_mono,
- rtac trans, cong_tac map_cong,
- rtac sym, mor_tac morE in_mono];
-
- fun mk_unique_tac (k, least_min_alg) =
- select_prem_tac n (etac @{thm prop_restrict}) k THEN' rtac least_min_alg THEN'
- stac alg_def THEN'
- CONJ_WRAP' mk_alg_tac (alg_sets ~~ (in_monos ~~ (morEs ~~ map_congs)));
- in
- CONJ_WRAP' mk_unique_tac (ks ~~ least_min_algs) 1
- end;
-
-fun mk_init_induct_tac m alg_def alg_min_alg least_min_algs alg_sets =
- let
- val n = length least_min_algs;
-
- fun mk_alg_tac alg_set = EVERY' [rtac ballI, rtac CollectI,
- REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac conjI, rtac (alg_min_alg RS alg_set),
- REPEAT_DETERM_N m o rtac subset_UNIV,
- REPEAT_DETERM_N n o (etac subset_trans THEN' rtac @{thm Collect_restrict}),
- rtac mp, etac bspec, rtac CollectI,
- REPEAT_DETERM_N m o (rtac conjI THEN' atac),
- CONJ_WRAP' (K (etac subset_trans THEN' rtac @{thm Collect_restrict})) alg_sets,
- CONJ_WRAP' (K (rtac ballI THEN' etac @{thm prop_restrict} THEN' atac)) alg_sets];
-
- fun mk_induct_tac least_min_alg =
- rtac ballI THEN' etac @{thm prop_restrict} THEN' rtac least_min_alg THEN'
- stac alg_def THEN'
- CONJ_WRAP' mk_alg_tac alg_sets;
- in
- CONJ_WRAP' mk_induct_tac least_min_algs 1
- end;
-
-fun mk_mor_Rep_tac ctor_defs copy bijs inver_Abss inver_Reps {context = ctxt, prems = _} =
- (K (unfold_thms_tac ctxt ctor_defs) THEN' rtac conjunct1 THEN' rtac copy THEN'
- EVERY' (map (fn bij => EVERY' [rtac bij, atac, etac bexI, rtac UNIV_I]) bijs) THEN'
- EVERY' (map rtac inver_Abss) THEN'
- EVERY' (map rtac inver_Reps)) 1;
-
-fun mk_mor_Abs_tac inv inver_Abss inver_Reps =
- (rtac inv THEN'
- EVERY' (map2 (fn inver_Abs => fn inver_Rep =>
- EVERY' [rtac conjI, rtac subset_UNIV, rtac conjI, rtac inver_Rep, rtac inver_Abs])
- inver_Abss inver_Reps)) 1;
-
-fun mk_mor_fold_tac cT ct fold_defs ex_mor mor =
- (EVERY' (map stac fold_defs) THEN' EVERY' [rtac rev_mp, rtac ex_mor, rtac impI] THEN'
- REPEAT_DETERM_N (length fold_defs) o etac exE THEN'
- rtac (Drule.instantiate' [SOME cT] [SOME ct] @{thm someI}) THEN' etac mor) 1;
-
-fun mk_fold_unique_mor_tac type_defs init_unique_mors Reps mor_comp mor_Abs mor_fold =
- let
- fun mk_unique type_def =
- EVERY' [rtac @{thm surj_fun_eq}, rtac (type_def RS @{thm type_definition.Abs_image}),
- rtac ballI, resolve_tac init_unique_mors,
- EVERY' (map (fn thm => atac ORELSE' rtac thm) Reps),
- rtac mor_comp, rtac mor_Abs, atac,
- rtac mor_comp, rtac mor_Abs, rtac mor_fold];
- in
- CONJ_WRAP' mk_unique type_defs 1
- end;
-
-fun mk_dtor_o_ctor_tac dtor_def foldx map_comp_id map_congL ctor_o_folds =
- EVERY' [stac dtor_def, rtac ext, rtac trans, rtac o_apply, rtac trans, rtac foldx,
- rtac trans, rtac map_comp_id, rtac trans, rtac map_congL,
- EVERY' (map (fn thm => rtac ballI THEN' rtac (trans OF [thm RS fun_cong, id_apply]))
- ctor_o_folds),
- rtac sym, rtac id_apply] 1;
-
-fun mk_rec_tac rec_defs foldx fst_recs {context = ctxt, prems = _}=
- unfold_thms_tac ctxt
- (rec_defs @ map (fn thm => thm RS @{thm convol_expand_snd}) fst_recs) THEN
- EVERY' [rtac trans, rtac o_apply, rtac trans, rtac (foldx RS @{thm arg_cong[of _ _ snd]}),
- rtac @{thm snd_convol'}] 1;
-
-fun mk_ctor_induct_tac m set_natural'ss init_induct morEs mor_Abs Rep_invs Abs_invs Reps =
- let
- val n = length set_natural'ss;
- val ks = 1 upto n;
-
- fun mk_IH_tac Rep_inv Abs_inv set_natural' =
- DETERM o EVERY' [dtac meta_mp, rtac (Rep_inv RS arg_cong RS subst), etac bspec,
- dtac set_rev_mp, rtac equalityD1, rtac set_natural', etac imageE,
- hyp_subst_tac, rtac (Abs_inv RS ssubst), etac set_mp, atac, atac];
-
- fun mk_closed_tac (k, (morE, set_natural's)) =
- EVERY' [select_prem_tac n (dtac asm_rl) k, rtac ballI, rtac impI,
- rtac (mor_Abs RS morE RS arg_cong RS ssubst), atac,
- REPEAT_DETERM o eresolve_tac [CollectE, conjE], dtac @{thm meta_spec},
- EVERY' (map3 mk_IH_tac Rep_invs Abs_invs (drop m set_natural's)), atac];
-
- fun mk_induct_tac (Rep, Rep_inv) =
- EVERY' [rtac (Rep_inv RS arg_cong RS subst), etac (Rep RSN (2, bspec))];
- in
- (rtac mp THEN' rtac impI THEN'
- DETERM o CONJ_WRAP_GEN' (etac conjE THEN' rtac conjI) mk_induct_tac (Reps ~~ Rep_invs) THEN'
- rtac init_induct THEN'
- DETERM o CONJ_WRAP' mk_closed_tac (ks ~~ (morEs ~~ set_natural'ss))) 1
- end;
-
-fun mk_ctor_induct2_tac cTs cts ctor_induct weak_ctor_inducts {context = ctxt, prems = _} =
- let
- val n = length weak_ctor_inducts;
- val ks = 1 upto n;
- fun mk_inner_induct_tac induct i =
- EVERY' [rtac allI, fo_rtac induct ctxt,
- select_prem_tac n (dtac @{thm meta_spec2}) i,
- REPEAT_DETERM_N n o
- EVERY' [dtac meta_mp THEN_ALL_NEW Goal.norm_hhf_tac,
- REPEAT_DETERM o dtac @{thm meta_spec}, etac (spec RS meta_mp), atac],
- atac];
- in
- EVERY' [rtac rev_mp, rtac (Drule.instantiate' cTs cts ctor_induct),
- EVERY' (map2 mk_inner_induct_tac weak_ctor_inducts ks), rtac impI,
- REPEAT_DETERM o eresolve_tac [conjE, allE],
- CONJ_WRAP' (K atac) ks] 1
- end;
-
-fun mk_map_tac m n foldx map_comp_id map_cong =
- EVERY' [rtac ext, rtac trans, rtac o_apply, rtac trans, rtac foldx, rtac trans, rtac o_apply,
- rtac trans, rtac (map_comp_id RS arg_cong), rtac trans, rtac (map_cong RS arg_cong),
- REPEAT_DETERM_N m o rtac refl,
- REPEAT_DETERM_N n o (EVERY' (map rtac [trans, o_apply, id_apply])),
- rtac sym, rtac o_apply] 1;
-
-fun mk_map_unique_tac m mor_def fold_unique_mor map_comp_ids map_congs =
- let
- val n = length map_congs;
- fun mk_mor (comp_id, cong) = EVERY' [rtac ballI, rtac trans, etac @{thm pointfreeE},
- rtac sym, rtac trans, rtac o_apply, rtac trans, rtac (comp_id RS arg_cong),
- rtac (cong RS arg_cong),
- REPEAT_DETERM_N m o rtac refl,
- REPEAT_DETERM_N n o (EVERY' (map rtac [trans, o_apply, id_apply]))];
- in
- EVERY' [rtac fold_unique_mor, rtac ssubst, rtac mor_def, rtac conjI,
- CONJ_WRAP' (K (EVERY' [rtac ballI, rtac UNIV_I])) map_congs,
- CONJ_WRAP' mk_mor (map_comp_ids ~~ map_congs)] 1
- end;
-
-fun mk_set_tac foldx = EVERY' [rtac ext, rtac trans, rtac o_apply,
- rtac trans, rtac foldx, rtac sym, rtac o_apply] 1;
-
-fun mk_set_simp_tac set set_natural' set_natural's =
- let
- val n = length set_natural's;
- fun mk_UN thm = rtac (thm RS @{thm arg_cong[of _ _ Union]} RS trans) THEN'
- rtac @{thm Union_image_eq};
- in
- EVERY' [rtac (set RS @{thm pointfreeE} RS trans), rtac @{thm Un_cong},
- rtac (trans OF [set_natural', trans_fun_cong_image_id_id_apply]),
- REPEAT_DETERM_N (n - 1) o rtac @{thm Un_cong},
- EVERY' (map mk_UN set_natural's)] 1
- end;
-
-fun mk_set_nat_tac m induct_tac set_natural'ss
- map_simps csets set_simps i {context = ctxt, prems = _} =
- let
- val n = length map_simps;
-
- fun useIH set_nat = EVERY' [rtac trans, rtac @{thm image_UN}, rtac trans, rtac @{thm UN_cong},
- rtac refl, Goal.assume_rule_tac ctxt, rtac sym, rtac trans, rtac @{thm UN_cong},
- rtac set_nat, rtac refl, rtac @{thm UN_simps(10)}];
-
- fun mk_set_nat cset map_simp set_simp set_nats =
- EVERY' [rtac trans, rtac @{thm image_cong}, rtac set_simp, rtac refl,
- rtac sym, rtac (trans OF [map_simp RS HOL_arg_cong cset, set_simp RS trans]),
- rtac sym, EVERY' (map rtac (trans :: @{thms image_Un Un_cong})),
- rtac sym, rtac (nth set_nats (i - 1)),
- REPEAT_DETERM_N (n - 1) o EVERY' (map rtac (trans :: @{thms image_Un Un_cong})),
- EVERY' (map useIH (drop m set_nats))];
- in
- (induct_tac THEN' EVERY' (map4 mk_set_nat csets map_simps set_simps set_natural'ss)) 1
- end;
-
-fun mk_set_bd_tac m induct_tac bd_Cinfinite set_bdss set_simps i {context = ctxt, prems = _} =
- let
- val n = length set_simps;
-
- fun useIH set_bd = EVERY' [rtac @{thm UNION_Cinfinite_bound}, rtac set_bd, rtac ballI,
- Goal.assume_rule_tac ctxt, rtac bd_Cinfinite];
-
- fun mk_set_nat set_simp set_bds =
- EVERY' [rtac @{thm ordIso_ordLeq_trans}, rtac @{thm card_of_ordIso_subst}, rtac set_simp,
- rtac (bd_Cinfinite RSN (3, @{thm Un_Cinfinite_bound})), rtac (nth set_bds (i - 1)),
- REPEAT_DETERM_N (n - 1) o rtac (bd_Cinfinite RSN (3, @{thm Un_Cinfinite_bound})),
- EVERY' (map useIH (drop m set_bds))];
- in
- (induct_tac THEN' EVERY' (map2 mk_set_nat set_simps set_bdss)) 1
- end;
-
-fun mk_mcong_tac induct_tac set_setsss map_congs map_simps {context = ctxt, prems = _} =
- let
- fun use_asm thm = EVERY' [etac bspec, etac set_rev_mp, rtac thm];
-
- fun useIH set_sets = EVERY' [rtac mp, Goal.assume_rule_tac ctxt,
- CONJ_WRAP' (fn thm =>
- EVERY' [rtac ballI, etac bspec, etac set_rev_mp, etac thm]) set_sets];
-
- fun mk_map_cong map_simp map_cong set_setss =
- EVERY' [rtac impI, REPEAT_DETERM o etac conjE,
- rtac trans, rtac map_simp, rtac trans, rtac (map_cong RS arg_cong),
- EVERY' (map use_asm (map hd set_setss)),
- EVERY' (map useIH (transpose (map tl set_setss))),
- rtac sym, rtac map_simp];
- in
- (induct_tac THEN' EVERY' (map3 mk_map_cong map_simps map_congs set_setsss)) 1
- end;
-
-fun mk_incl_min_alg_tac induct_tac set_setsss alg_sets alg_min_alg {context = ctxt, prems = _} =
- let
- fun use_asm thm = etac (thm RS subset_trans);
-
- fun useIH set_sets = EVERY' [rtac subsetI, rtac mp, Goal.assume_rule_tac ctxt,
- rtac CollectI, CONJ_WRAP' (fn thm => EVERY' [etac (thm RS subset_trans), atac]) set_sets];
-
- fun mk_incl alg_set set_setss =
- EVERY' [rtac impI, REPEAT_DETERM o eresolve_tac [CollectE, conjE],
- rtac (alg_min_alg RS alg_set),
- EVERY' (map use_asm (map hd set_setss)),
- EVERY' (map useIH (transpose (map tl set_setss)))];
- in
- (induct_tac THEN' EVERY' (map2 mk_incl alg_sets set_setsss)) 1
- end;
-
-fun mk_lfp_map_wpull_tac m induct_tac wpulls map_simps set_simpss set_setsss ctor_injects =
- let
- val n = length wpulls;
- val ks = 1 upto n;
- val ls = 1 upto m;
-
- fun use_pass_asm thm = rtac conjI THEN' etac (thm RS subset_trans);
- fun use_act_asm thm = etac (thm RS subset_trans) THEN' atac;
-
- fun useIH set_sets i = EVERY' [rtac ssubst, rtac @{thm wpull_def},
- REPEAT_DETERM_N m o etac thin_rl, select_prem_tac n (dtac asm_rl) i,
- rtac allI, rtac allI, rtac impI, REPEAT_DETERM o etac conjE,
- REPEAT_DETERM o dtac @{thm meta_spec},
- dtac meta_mp, atac,
- dtac meta_mp, atac, etac mp,
- rtac conjI, rtac CollectI, CONJ_WRAP' use_act_asm set_sets,
- rtac conjI, rtac CollectI, CONJ_WRAP' use_act_asm set_sets,
- atac];
-
- fun mk_subset thm = EVERY' [rtac ord_eq_le_trans, rtac thm, rtac @{thm Un_least}, atac,
- REPEAT_DETERM_N (n - 1) o rtac @{thm Un_least},
- REPEAT_DETERM_N n o
- EVERY' [rtac @{thm UN_least}, rtac CollectE, etac set_rev_mp, atac,
- REPEAT_DETERM o etac conjE, atac]];
-
- fun mk_wpull wpull map_simp set_simps set_setss ctor_inject =
- EVERY' [rtac impI, REPEAT_DETERM o eresolve_tac [CollectE, conjE],
- rtac rev_mp, rtac wpull,
- EVERY' (map (fn i => REPEAT_DETERM_N (i - 1) o etac thin_rl THEN' atac) ls),
- EVERY' (map2 useIH (transpose (map tl set_setss)) ks),
- rtac impI, REPEAT_DETERM_N (m + n) o etac thin_rl,
- dtac @{thm subst[OF wpull_def, of "%x. x"]}, etac allE, etac allE, etac impE,
- rtac conjI, rtac CollectI, EVERY' (map (use_pass_asm o hd) set_setss),
- CONJ_WRAP' (K (rtac subset_refl)) ks,
- rtac conjI, rtac CollectI, EVERY' (map (use_pass_asm o hd) set_setss),
- CONJ_WRAP' (K (rtac subset_refl)) ks,
- rtac subst, rtac ctor_inject, rtac trans, rtac sym, rtac map_simp,
- rtac trans, atac, rtac map_simp, REPEAT_DETERM o eresolve_tac [CollectE, conjE, bexE],
- hyp_subst_tac, rtac bexI, rtac conjI, rtac map_simp, rtac map_simp, rtac CollectI,
- CONJ_WRAP' mk_subset set_simps];
- in
- (induct_tac THEN' EVERY' (map5 mk_wpull wpulls map_simps set_simpss set_setsss ctor_injects)) 1
- end;
-
-(* BNF tactics *)
-
-fun mk_map_id_tac map_ids unique =
- (rtac sym THEN' rtac unique THEN'
- EVERY' (map (fn thm =>
- EVERY' [rtac trans, rtac @{thm id_o}, rtac trans, rtac sym, rtac @{thm o_id},
- rtac (thm RS sym RS arg_cong)]) map_ids)) 1;
-
-fun mk_map_comp_tac map_comps map_simps unique iplus1 =
- let
- val i = iplus1 - 1;
- val unique' = Thm.permute_prems 0 i unique;
- val map_comps' = drop i map_comps @ take i map_comps;
- val map_simps' = drop i map_simps @ take i map_simps;
- fun mk_comp comp simp =
- EVERY' [rtac ext, rtac trans, rtac o_apply, rtac trans, rtac o_apply,
- rtac trans, rtac (simp RS arg_cong), rtac trans, rtac simp,
- rtac trans, rtac (comp RS arg_cong), rtac sym, rtac o_apply];
- in
- (rtac sym THEN' rtac unique' THEN' EVERY' (map2 mk_comp map_comps' map_simps')) 1
- end;
-
-fun mk_set_natural_tac set_nat =
- EVERY' (map rtac [ext, trans, o_apply, sym, trans, o_apply, set_nat]) 1;
-
-fun mk_in_bd_tac sum_Card_order sucbd_Cnotzero incl card_of_min_alg =
- EVERY' [rtac ctrans, rtac @{thm card_of_mono1}, rtac subsetI, etac rev_mp,
- rtac incl, rtac ctrans, rtac card_of_min_alg, rtac @{thm cexp_mono2_Cnotzero},
- rtac @{thm cardSuc_ordLeq_cpow}, rtac sum_Card_order, rtac @{thm csum_Cnotzero2},
- rtac @{thm ctwo_Cnotzero}, rtac sucbd_Cnotzero] 1;
-
-fun mk_bd_card_order_tac bd_card_orders =
- (rtac @{thm card_order_cpow} THEN'
- CONJ_WRAP_GEN' (rtac @{thm card_order_csum}) rtac bd_card_orders) 1;
-
-fun mk_wpull_tac wpull =
- EVERY' [rtac ssubst, rtac @{thm wpull_def}, rtac allI, rtac allI,
- rtac wpull, REPEAT_DETERM o atac] 1;
-
-fun mk_wit_tac n set_simp wit =
- REPEAT_DETERM (atac 1 ORELSE
- EVERY' [dtac set_rev_mp, rtac equalityD1, resolve_tac set_simp,
- REPEAT_DETERM o
- (TRY o REPEAT_DETERM o etac UnE THEN' TRY o etac @{thm UN_E} THEN'
- (eresolve_tac wit ORELSE'
- (dresolve_tac wit THEN'
- (etac FalseE ORELSE'
- EVERY' [hyp_subst_tac, dtac set_rev_mp, rtac equalityD1, resolve_tac set_simp,
- REPEAT_DETERM_N n o etac UnE]))))] 1);
-
-fun mk_srel_simp_tac in_Isrels i in_srel map_comp map_cong map_simp set_simps ctor_inject
- ctor_dtor set_naturals set_incls set_set_inclss =
- let
- val m = length set_incls;
- val n = length set_set_inclss;
-
- val (passive_set_naturals, active_set_naturals) = chop m set_naturals;
- val in_Isrel = nth in_Isrels (i - 1);
- val le_arg_cong_ctor_dtor = ctor_dtor RS arg_cong RS ord_eq_le_trans;
- val eq_arg_cong_ctor_dtor = ctor_dtor RS arg_cong RS trans;
- val if_tac =
- EVERY' [dtac (in_Isrel RS iffD1), REPEAT_DETERM o eresolve_tac [exE, conjE, CollectE],
- rtac (in_srel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
- EVERY' (map2 (fn set_natural => fn set_incl =>
- EVERY' [rtac conjI, rtac ord_eq_le_trans, rtac set_natural,
- rtac ord_eq_le_trans, rtac trans_fun_cong_image_id_id_apply,
- rtac (set_incl RS subset_trans), etac le_arg_cong_ctor_dtor])
- passive_set_naturals set_incls),
- CONJ_WRAP' (fn (in_Isrel, (set_natural, set_set_incls)) =>
- EVERY' [rtac ord_eq_le_trans, rtac set_natural, rtac @{thm image_subsetI},
- rtac (in_Isrel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
- CONJ_WRAP' (fn thm =>
- EVERY' (map etac [thm RS subset_trans, le_arg_cong_ctor_dtor]))
- set_set_incls,
- rtac conjI, rtac refl, rtac refl])
- (in_Isrels ~~ (active_set_naturals ~~ set_set_inclss)),
- CONJ_WRAP' (fn conv =>
- EVERY' [rtac trans, rtac map_comp, rtac trans, rtac map_cong,
- REPEAT_DETERM_N m o rtac @{thm fun_cong[OF o_id]},
- REPEAT_DETERM_N n o EVERY' (map rtac [trans, o_apply, conv]),
- rtac (ctor_inject RS iffD1), rtac trans, rtac sym, rtac map_simp,
- etac eq_arg_cong_ctor_dtor])
- fst_snd_convs];
- val only_if_tac =
- EVERY' [dtac (in_srel RS iffD1), REPEAT_DETERM o eresolve_tac [exE, conjE, CollectE],
- rtac (in_Isrel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
- CONJ_WRAP' (fn (set_simp, passive_set_natural) =>
- EVERY' [rtac ord_eq_le_trans, rtac set_simp, rtac @{thm Un_least},
- rtac ord_eq_le_trans, rtac @{thm box_equals[OF _ refl]},
- rtac passive_set_natural, rtac trans_fun_cong_image_id_id_apply, atac,
- CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_least}))
- (fn (active_set_natural, in_Isrel) => EVERY' [rtac ord_eq_le_trans,
- rtac @{thm UN_cong[OF _ refl]}, rtac active_set_natural, rtac @{thm UN_least},
- dtac set_rev_mp, etac @{thm image_mono}, etac imageE,
- dtac @{thm ssubst_mem[OF pair_collapse]}, dtac (in_Isrel RS iffD1),
- dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE,
- dtac (Thm.permute_prems 0 1 @{thm ssubst_mem}), atac,
- hyp_subst_tac, REPEAT_DETERM o eresolve_tac [CollectE, conjE], atac])
- (rev (active_set_naturals ~~ in_Isrels))])
- (set_simps ~~ passive_set_naturals),
- rtac conjI,
- REPEAT_DETERM_N 2 o EVERY' [rtac trans, rtac map_simp, rtac (ctor_inject RS iffD2),
- rtac trans, rtac map_comp, rtac trans, rtac map_cong,
- REPEAT_DETERM_N m o rtac @{thm fun_cong[OF o_id]},
- EVERY' (map (fn in_Isrel => EVERY' [rtac trans, rtac o_apply, dtac set_rev_mp, atac,
- dtac @{thm ssubst_mem[OF pair_collapse]}, dtac (in_Isrel RS iffD1),
- dtac @{thm someI_ex}, REPEAT_DETERM o etac conjE, atac]) in_Isrels),
- atac]]
- in
- EVERY' [rtac iffI, if_tac, only_if_tac] 1
- end;
-
-end;
--- a/src/HOL/Codatatype/Tools/bnf_lfp_util.ML Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,76 +0,0 @@
-(* Title: HOL/BNF/Tools/bnf_lfp_util.ML
- Author: Dmitriy Traytel, TU Muenchen
- Author: Jasmin Blanchette, TU Muenchen
- Copyright 2012
-
-Library for the datatype construction.
-*)
-
-signature BNF_LFP_UTIL =
-sig
- val HOL_arg_cong: cterm -> thm
-
- val mk_bij_betw: term -> term -> term -> term
- val mk_cardSuc: term -> term
- val mk_convol: term * term -> term
- val mk_cpow: term -> term
- val mk_inver: term -> term -> term -> term
- val mk_not_empty: term -> term
- val mk_not_eq: term -> term -> term
- val mk_rapp: term -> typ -> term
- val mk_relChain: term -> term -> term
- val mk_underS: term -> term
- val mk_worec: term -> term -> term
-end;
-
-structure BNF_LFP_Util : BNF_LFP_UTIL =
-struct
-
-open BNF_Util
-
-fun HOL_arg_cong ct = Drule.instantiate'
- (map SOME (Thm.dest_ctyp (Thm.ctyp_of_term ct))) [NONE, NONE, SOME ct] arg_cong;
-
-(*reverse application*)
-fun mk_rapp arg T = Term.absdummy (fastype_of arg --> T) (Bound 0 $ arg);
-
-fun mk_underS r =
- let val T = fst (dest_relT (fastype_of r));
- in Const (@{const_name rel.underS}, mk_relT (T, T) --> T --> HOLogic.mk_setT T) $ r end;
-
-fun mk_worec r f =
- let val (A, AB) = apfst domain_type (dest_funT (fastype_of f));
- in Const (@{const_name wo_rel.worec}, mk_relT (A, A) --> (AB --> AB) --> AB) $ r $ f end;
-
-fun mk_relChain r f =
- let val (A, AB) = `domain_type (fastype_of f);
- in Const (@{const_name relChain}, mk_relT (A, A) --> AB --> HOLogic.boolT) $ r $ f end;
-
-fun mk_cardSuc r =
- let val T = fst (dest_relT (fastype_of r));
- in Const (@{const_name cardSuc}, mk_relT (T, T) --> mk_relT (`I (HOLogic.mk_setT T))) $ r end;
-
-fun mk_cpow r =
- let val T = fst (dest_relT (fastype_of r));
- in Const (@{const_name cpow}, mk_relT (T, T) --> mk_relT (`I (HOLogic.mk_setT T))) $ r end;
-
-fun mk_bij_betw f A B =
- Const (@{const_name bij_betw},
- fastype_of f --> fastype_of A --> fastype_of B --> HOLogic.boolT) $ f $ A $ B;
-
-fun mk_inver f g A =
- Const (@{const_name inver}, fastype_of f --> fastype_of g --> fastype_of A --> HOLogic.boolT) $
- f $ g $ A;
-
-fun mk_not_eq x y = HOLogic.mk_not (HOLogic.mk_eq (x, y));
-
-fun mk_not_empty B = mk_not_eq B (HOLogic.mk_set (HOLogic.dest_setT (fastype_of B)) []);
-
-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;
-
-end;
--- a/src/HOL/Codatatype/Tools/bnf_tactics.ML Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,125 +0,0 @@
-(* Title: HOL/BNF/Tools/bnf_tactics.ML
- Author: Dmitriy Traytel, TU Muenchen
- Author: Jasmin Blanchette, TU Muenchen
- Copyright 2012
-
-General tactics for bounded natural functors.
-*)
-
-signature BNF_TACTICS =
-sig
- val ss_only: thm list -> simpset
-
- val prefer_tac: int -> tactic
- val select_prem_tac: int -> (int -> tactic) -> int -> int -> tactic
- val fo_rtac: thm -> Proof.context -> int -> tactic
- val subst_asm_tac: Proof.context -> thm list -> int -> tactic
- val subst_tac: Proof.context -> thm list -> int -> tactic
- val substs_tac: Proof.context -> thm list -> int -> tactic
- val unfold_thms_tac: Proof.context -> thm list -> tactic
- val mk_unfold_thms_then_tac: Proof.context -> thm list -> ('a -> tactic) -> 'a -> tactic
-
- val mk_flatten_assoc_tac: (int -> tactic) -> thm -> thm -> thm -> tactic
- val mk_rotate_eq_tac: (int -> tactic) -> thm -> thm -> thm -> thm -> ''a list -> ''a list ->
- int -> tactic
-
- val mk_Abs_bij_thm: Proof.context -> thm -> thm -> thm
- val mk_Abs_inj_thm: thm -> thm
-
- val simple_srel_O_Gr_tac: Proof.context -> tactic
- val mk_rel_simp_tac: thm -> thm list -> thm list -> thm -> {prems: 'a, context: Proof.context} ->
- tactic
-
- val mk_map_comp_id_tac: thm -> tactic
- val mk_map_cong_tac: int -> thm -> {prems: 'a, context: Proof.context} -> tactic
- val mk_map_congL_tac: int -> thm -> thm -> tactic
-end;
-
-structure BNF_Tactics : BNF_TACTICS =
-struct
-
-open BNF_Util
-
-fun ss_only thms = Simplifier.clear_ss HOL_basic_ss addsimps thms;
-
-(* FIXME: why not in "Pure"? *)
-fun prefer_tac i = defer_tac i THEN PRIMITIVE (Thm.permute_prems 0 ~1);
-
-fun select_prem_tac n tac k = DETERM o (EVERY' [REPEAT_DETERM_N (k - 1) o etac thin_rl,
- tac, REPEAT_DETERM_N (n - k) o etac thin_rl]);
-
-(*stolen from Christian Urban's Cookbook*)
-fun fo_rtac thm = Subgoal.FOCUS (fn {concl, ...} =>
- let
- val concl_pat = Drule.strip_imp_concl (cprop_of thm)
- val insts = Thm.first_order_match (concl_pat, concl)
- in
- rtac (Drule.instantiate_normalize insts thm) 1
- end);
-
-fun unfold_thms_tac ctxt thms = Local_Defs.unfold_tac ctxt (distinct Thm.eq_thm_prop thms);
-
-fun mk_unfold_thms_then_tac lthy defs tac x = unfold_thms_tac lthy defs THEN tac x;
-
-(*unlike "unfold_thms_tac", succeeds when the RHS contains schematic variables not in the LHS*)
-fun subst_asm_tac ctxt = EqSubst.eqsubst_asm_tac ctxt [0];
-fun subst_tac ctxt = EqSubst.eqsubst_tac ctxt [0];
-fun substs_tac ctxt = REPEAT_DETERM oo subst_tac ctxt;
-
-
-(* Theorems for open typedefs with UNIV as representing set *)
-
-fun mk_Abs_inj_thm inj = inj OF (replicate 2 UNIV_I);
-fun mk_Abs_bij_thm ctxt Abs_inj_thm surj = rule_by_tactic ctxt ((rtac surj THEN' etac exI) 1)
- (Abs_inj_thm RS @{thm bijI});
-
-
-
-(* General tactic generators *)
-
-(*applies assoc rule to the lhs of an equation as long as possible*)
-fun mk_flatten_assoc_tac refl_tac trans assoc cong = rtac trans 1 THEN
- REPEAT_DETERM (CHANGED ((FIRST' [rtac trans THEN' rtac assoc, rtac cong THEN' refl_tac]) 1)) THEN
- refl_tac 1;
-
-(*proves two sides of an equation to be equal assuming both are flattened and rhs can be obtained
-from lhs by the given permutation of monoms*)
-fun mk_rotate_eq_tac refl_tac trans assoc com cong =
- let
- fun gen_tac [] [] = K all_tac
- | gen_tac [x] [y] = if x = y then refl_tac else error "mk_rotate_eq_tac: different lists"
- | gen_tac (x :: xs) (y :: ys) = if x = y
- then rtac cong THEN' refl_tac THEN' gen_tac xs ys
- else rtac trans THEN' rtac com THEN'
- K (mk_flatten_assoc_tac refl_tac trans assoc cong) THEN'
- gen_tac (xs @ [x]) (y :: ys)
- | gen_tac _ _ = error "mk_rotate_eq_tac: different lists";
- in
- gen_tac
- end;
-
-fun simple_srel_O_Gr_tac ctxt =
- unfold_thms_tac ctxt @{thms Collect_fst_snd_mem_eq Collect_pair_mem_eq} THEN rtac refl 1;
-
-fun mk_rel_simp_tac srel_def IJrel_defs IJsrel_defs srel_simp {context = ctxt, prems = _} =
- unfold_thms_tac ctxt IJrel_defs THEN
- subst_tac ctxt [unfold_thms ctxt (IJrel_defs @ IJsrel_defs @
- @{thms Collect_pair_mem_eq mem_Collect_eq fst_conv snd_conv}) srel_simp] 1 THEN
- unfold_thms_tac ctxt (srel_def ::
- @{thms Collect_fst_snd_mem_eq mem_Collect_eq pair_mem_Collect_split fst_conv snd_conv
- split_conv}) THEN
- rtac refl 1;
-
-fun mk_map_comp_id_tac map_comp =
- (rtac trans THEN' rtac map_comp THEN' REPEAT_DETERM o stac @{thm o_id} THEN' rtac refl) 1;
-
-fun mk_map_cong_tac m map_cong {context = ctxt, prems = _} =
- EVERY' [rtac mp, rtac map_cong,
- CONJ_WRAP' (K (rtac ballI THEN' Goal.assume_rule_tac ctxt)) (1 upto m)] 1;
-
-fun mk_map_congL_tac passive map_cong map_id' =
- (rtac trans THEN' rtac map_cong THEN' EVERY' (replicate passive (rtac refl))) 1 THEN
- REPEAT_DETERM (EVERY' [rtac trans, etac bspec, atac, rtac sym, rtac @{thm id_apply}] 1) THEN
- rtac map_id' 1;
-
-end;
--- a/src/HOL/Codatatype/Tools/bnf_util.ML Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,619 +0,0 @@
-(* Title: HOL/BNF/Tools/bnf_util.ML
- Author: Dmitriy Traytel, TU Muenchen
- Copyright 2012
-
-Library for bounded natural functors.
-*)
-
-signature BNF_UTIL =
-sig
- val map3: ('a -> 'b -> 'c -> 'd) -> 'a list -> 'b list -> 'c list -> 'd list
- val map4: ('a -> 'b -> 'c -> 'd -> 'e) -> 'a list -> 'b list -> 'c list -> 'd list -> 'e list
- val map5: ('a -> 'b -> 'c -> 'd -> 'e -> 'f) ->
- 'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list
- val map6: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g) ->
- 'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list
- val map7: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h) ->
- 'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list
- val map8: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i) ->
- 'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list -> 'i list
- val map9: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i -> 'j) ->
- 'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list ->
- 'i list -> 'j list
- val map10: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i -> 'j -> 'k) ->
- 'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list ->
- 'i list -> 'j list -> 'k list
- val map11: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i -> 'j -> 'k -> 'l) ->
- 'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list ->
- 'i list -> 'j list -> 'k list -> 'l list
- val map12: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i -> 'j -> 'k -> 'l -> 'm) ->
- 'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h list ->
- 'i list -> 'j list -> 'k list -> 'l list -> 'm list
- val fold_map2: ('a -> 'b -> 'c -> 'd * 'c) -> 'a list -> 'b list -> 'c -> 'd list * 'c
- val fold_map3: ('a -> 'b -> 'c -> 'd -> 'e * 'd) ->
- 'a list -> 'b list -> 'c list -> 'd -> 'e list * 'd
- val fold_map4: ('a -> 'b -> 'c -> 'd -> 'e -> 'f * 'e) ->
- 'a list -> 'b list -> 'c list -> 'd list -> 'e -> 'f list * 'e
- val fold_map5: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g * 'f) ->
- 'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f -> 'g list * 'f
- val fold_map6: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h * 'g) ->
- 'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g -> 'h list * 'g
- val fold_map7: ('a -> 'b -> 'c -> 'd -> 'e -> 'f -> 'g -> 'h -> 'i * 'h) ->
- 'a list -> 'b list -> 'c list -> 'd list -> 'e list -> 'f list -> 'g list -> 'h -> 'i list * 'h
- val interleave: 'a list -> 'a list -> 'a list
- val transpose: 'a list list -> 'a list list
- val seq_conds: (bool -> 'a -> 'b) -> int -> int -> 'a list -> 'b list
-
- val mk_fresh_names: Proof.context -> int -> string -> string list * Proof.context
- val mk_TFrees: int -> Proof.context -> typ list * Proof.context
- val mk_TFreess: int list -> Proof.context -> typ list list * Proof.context
- val mk_TFrees': sort list -> Proof.context -> typ list * Proof.context
- val mk_Frees: string -> typ list -> Proof.context -> term list * Proof.context
- val mk_Freess: string -> typ list list -> Proof.context -> term list list * Proof.context
- val mk_Freesss: string -> typ list list list -> Proof.context ->
- term list list list * Proof.context
- val mk_Freessss: string -> typ list list list list -> Proof.context ->
- term list list list list * Proof.context
- val mk_Frees': string -> typ list -> Proof.context ->
- (term list * (string * typ) list) * Proof.context
- val mk_Freess': string -> typ list list -> Proof.context ->
- (term list list * (string * typ) list list) * Proof.context
- val nonzero_string_of_int: int -> string
-
- val strip_typeN: int -> typ -> typ list * typ
-
- val mk_predT: typ list -> typ
- val mk_pred1T: typ -> typ
- val mk_pred2T: typ -> typ -> typ
- val mk_optionT: typ -> typ
- val mk_relT: typ * typ -> typ
- val dest_relT: typ -> typ * typ
- val mk_sumT: typ * typ -> typ
-
- val ctwo: term
- val fst_const: typ -> term
- val snd_const: typ -> term
- val Id_const: typ -> term
-
- val mk_Ball: term -> term -> term
- val mk_Bex: term -> term -> term
- val mk_Card_order: term -> term
- val mk_Field: term -> term
- val mk_Gr: term -> term -> term
- val mk_IfN: typ -> term list -> term list -> term
- val mk_Trueprop_eq: term * term -> term
- val mk_UNION: term -> term -> term
- val mk_Union: typ -> term
- val mk_card_binop: string -> (typ * typ -> typ) -> term -> term -> term
- val mk_card_of: term -> term
- val mk_card_order: term -> term
- val mk_ccexp: term -> term -> term
- val mk_cexp: term -> term -> term
- val mk_cinfinite: term -> term
- val mk_collect: term list -> typ -> term
- val mk_converse: term -> term
- val mk_cprod: term -> term -> term
- val mk_csum: term -> term -> term
- val mk_dir_image: term -> term -> term
- val mk_image: term -> term
- val mk_in: term list -> term list -> typ -> term
- val mk_ordLeq: term -> term -> term
- val mk_rel_comp: term * term -> term
- val mk_subset: term -> term -> term
- val mk_wpull: term -> term -> term -> term -> term -> (term * term) option -> term -> term -> term
-
- val list_all_free: term list -> term -> term
- val list_exists_free: term list -> term -> term
-
- (*parameterized terms*)
- val mk_nthN: int -> term -> int -> term
-
- (*parameterized thms*)
- val mk_Un_upper: int -> int -> thm
- val mk_conjIN: int -> thm
- val mk_conjunctN: int -> int -> thm
- val conj_dests: int -> thm -> thm list
- val mk_disjIN: int -> int -> thm
- val mk_nthI: int -> int -> thm
- val mk_nth_conv: int -> int -> thm
- val mk_ordLeq_csum: int -> int -> thm -> thm
- val mk_UnIN: int -> int -> thm
-
- val ctrans: thm
- val o_apply: thm
- val set_mp: thm
- val set_rev_mp: thm
- val subset_UNIV: thm
- val Pair_eqD: thm
- val Pair_eqI: thm
- val mk_sym: thm -> thm
- val mk_trans: thm -> thm -> thm
- val mk_unabs_def: int -> thm -> thm
-
- val is_refl: thm -> bool
- val no_refl: thm list -> thm list
- val no_reflexive: thm list -> thm list
-
- val fold_thms: Proof.context -> thm list -> thm -> thm
- val unfold_thms: Proof.context -> thm list -> thm -> thm
-
- val mk_permute: ''a list -> ''a list -> 'b list -> 'b list
- val find_indices: ''a list -> ''a list -> int list
-
- val certifyT: Proof.context -> typ -> ctyp
- val certify: Proof.context -> term -> cterm
-
- val parse_binding_colon: Token.T list -> binding * Token.T list
- val parse_opt_binding_colon: Token.T list -> binding * Token.T list
-
- val typedef: bool -> binding option -> binding * (string * sort) list * mixfix -> term ->
- (binding * binding) option -> tactic -> local_theory -> (string * Typedef.info) * local_theory
-
- val WRAP: ('a -> tactic) -> ('a -> tactic) -> 'a list -> tactic -> tactic
- val WRAP': ('a -> int -> tactic) -> ('a -> int -> tactic) -> 'a list -> (int -> tactic) -> int ->
- tactic
- val CONJ_WRAP_GEN: tactic -> ('a -> tactic) -> 'a list -> tactic
- val CONJ_WRAP_GEN': (int -> tactic) -> ('a -> int -> tactic) -> 'a list -> int -> tactic
- val CONJ_WRAP: ('a -> tactic) -> 'a list -> tactic
- val CONJ_WRAP': ('a -> int -> tactic) -> 'a list -> int -> tactic
-end;
-
-structure BNF_Util : BNF_UTIL =
-struct
-
-(* Library proper *)
-
-fun map3 _ [] [] [] = []
- | map3 f (x1::x1s) (x2::x2s) (x3::x3s) = f x1 x2 x3 :: map3 f x1s x2s x3s
- | map3 _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun map4 _ [] [] [] [] = []
- | map4 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) = f x1 x2 x3 x4 :: map4 f x1s x2s x3s x4s
- | map4 _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun map5 _ [] [] [] [] [] = []
- | map5 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) =
- f x1 x2 x3 x4 x5 :: map5 f x1s x2s x3s x4s x5s
- | map5 _ _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun map6 _ [] [] [] [] [] [] = []
- | map6 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) =
- f x1 x2 x3 x4 x5 x6 :: map6 f x1s x2s x3s x4s x5s x6s
- | map6 _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun map7 _ [] [] [] [] [] [] [] = []
- | map7 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) =
- f x1 x2 x3 x4 x5 x6 x7 :: map7 f x1s x2s x3s x4s x5s x6s x7s
- | map7 _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun map8 _ [] [] [] [] [] [] [] [] = []
- | map8 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) (x8::x8s) =
- f x1 x2 x3 x4 x5 x6 x7 x8 :: map8 f x1s x2s x3s x4s x5s x6s x7s x8s
- | map8 _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun map9 _ [] [] [] [] [] [] [] [] [] = []
- | map9 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s)
- (x6::x6s) (x7::x7s) (x8::x8s) (x9::x9s) =
- f x1 x2 x3 x4 x5 x6 x7 x8 x9 :: map9 f x1s x2s x3s x4s x5s x6s x7s x8s x9s
- | map9 _ _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun map10 _ [] [] [] [] [] [] [] [] [] [] = []
- | map10 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s)
- (x6::x6s) (x7::x7s) (x8::x8s) (x9::x9s) (x10::x10s) =
- f x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 :: map10 f x1s x2s x3s x4s x5s x6s x7s x8s x9s x10s
- | map10 _ _ _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun map11 _ [] [] [] [] [] [] [] [] [] [] [] = []
- | map11 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s)
- (x6::x6s) (x7::x7s) (x8::x8s) (x9::x9s) (x10::x10s) (x11::x11s) =
- f x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 :: map11 f x1s x2s x3s x4s x5s x6s x7s x8s x9s x10s x11s
- | map11 _ _ _ _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun map12 _ [] [] [] [] [] [] [] [] [] [] [] [] = []
- | map12 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s)
- (x6::x6s) (x7::x7s) (x8::x8s) (x9::x9s) (x10::x10s) (x11::x11s) (x12::x12s) =
- f x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 ::
- map12 f x1s x2s x3s x4s x5s x6s x7s x8s x9s x10s x11s x12s
- | map12 _ _ _ _ _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun fold_map2 _ [] [] acc = ([], acc)
- | fold_map2 f (x1::x1s) (x2::x2s) acc =
- let
- val (x, acc') = f x1 x2 acc;
- val (xs, acc'') = fold_map2 f x1s x2s acc';
- in (x :: xs, acc'') end
- | fold_map2 _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun fold_map3 _ [] [] [] acc = ([], acc)
- | fold_map3 f (x1::x1s) (x2::x2s) (x3::x3s) acc =
- let
- val (x, acc') = f x1 x2 x3 acc;
- val (xs, acc'') = fold_map3 f x1s x2s x3s acc';
- in (x :: xs, acc'') end
- | fold_map3 _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun fold_map4 _ [] [] [] [] acc = ([], acc)
- | fold_map4 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) acc =
- let
- val (x, acc') = f x1 x2 x3 x4 acc;
- val (xs, acc'') = fold_map4 f x1s x2s x3s x4s acc';
- in (x :: xs, acc'') end
- | fold_map4 _ _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun fold_map5 _ [] [] [] [] [] acc = ([], acc)
- | fold_map5 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) acc =
- let
- val (x, acc') = f x1 x2 x3 x4 x5 acc;
- val (xs, acc'') = fold_map5 f x1s x2s x3s x4s x5s acc';
- in (x :: xs, acc'') end
- | fold_map5 _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun fold_map6 _ [] [] [] [] [] [] acc = ([], acc)
- | fold_map6 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) acc =
- let
- val (x, acc') = f x1 x2 x3 x4 x5 x6 acc;
- val (xs, acc'') = fold_map6 f x1s x2s x3s x4s x5s x6s acc';
- in (x :: xs, acc'') end
- | fold_map6 _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-fun fold_map7 _ [] [] [] [] [] [] [] acc = ([], acc)
- | fold_map7 f (x1::x1s) (x2::x2s) (x3::x3s) (x4::x4s) (x5::x5s) (x6::x6s) (x7::x7s) acc =
- let
- val (x, acc') = f x1 x2 x3 x4 x5 x6 x7 acc;
- val (xs, acc'') = fold_map7 f x1s x2s x3s x4s x5s x6s x7s acc';
- in (x :: xs, acc'') end
- | fold_map7 _ _ _ _ _ _ _ _ _ = raise ListPair.UnequalLengths;
-
-(*stolen from ~~/src/HOL/Tools/SMT/smt_utils.ML*)
-fun certify ctxt = Thm.cterm_of (Proof_Context.theory_of ctxt);
-fun certifyT ctxt = Thm.ctyp_of (Proof_Context.theory_of ctxt);
-
-val parse_binding_colon = Parse.binding --| @{keyword ":"};
-val parse_opt_binding_colon = Scan.optional parse_binding_colon Binding.empty;
-
-(*TODO: is this really different from Typedef.add_typedef_global?*)
-fun typedef def opt_name typ set opt_morphs tac lthy =
- let
- val ((name, info), (lthy, lthy_old)) =
- lthy
- |> Typedef.add_typedef def opt_name typ set opt_morphs tac
- ||> `Local_Theory.restore;
- val phi = Proof_Context.export_morphism lthy_old lthy;
- in
- ((name, Typedef.transform_info phi info), lthy)
- end;
-
-(*Tactical WRAP surrounds a static given tactic (core) with two deterministic chains of tactics*)
-fun WRAP gen_before gen_after xs core_tac =
- fold_rev (fn x => fn tac => gen_before x THEN tac THEN gen_after x) xs core_tac;
-
-fun WRAP' gen_before gen_after xs core_tac =
- fold_rev (fn x => fn tac => gen_before x THEN' tac THEN' gen_after x) xs core_tac;
-
-fun CONJ_WRAP_GEN conj_tac gen_tac xs =
- let val (butlast, last) = split_last xs;
- in WRAP (fn thm => conj_tac THEN gen_tac thm) (K all_tac) butlast (gen_tac last) end;
-
-fun CONJ_WRAP_GEN' conj_tac gen_tac xs =
- let val (butlast, last) = split_last xs;
- in WRAP' (fn thm => conj_tac THEN' gen_tac thm) (K (K all_tac)) butlast (gen_tac last) end;
-
-(*not eta-converted because of monotype restriction*)
-fun CONJ_WRAP gen_tac = CONJ_WRAP_GEN (rtac conjI 1) gen_tac;
-fun CONJ_WRAP' gen_tac = CONJ_WRAP_GEN' (rtac conjI) gen_tac;
-
-
-
-(* Term construction *)
-
-(** Fresh variables **)
-
-fun nonzero_string_of_int 0 = ""
- | nonzero_string_of_int n = string_of_int n;
-
-val mk_TFrees' = apfst (map TFree) oo Variable.invent_types;
-
-fun mk_TFrees n = mk_TFrees' (replicate n HOLogic.typeS);
-val mk_TFreess = fold_map mk_TFrees;
-
-fun mk_names n x = if n = 1 then [x] else map (fn i => x ^ string_of_int i) (1 upto n);
-
-fun mk_fresh_names ctxt = (fn xs => Variable.variant_fixes xs ctxt) oo mk_names;
-fun mk_Frees x Ts ctxt = mk_fresh_names ctxt (length Ts) x |>> (fn xs => map2 (curry Free) xs Ts);
-fun mk_Freess x Tss = fold_map2 mk_Frees (mk_names (length Tss) x) Tss;
-fun mk_Freesss x Tsss = fold_map2 mk_Freess (mk_names (length Tsss) x) Tsss;
-fun mk_Freessss x Tssss = fold_map2 mk_Freesss (mk_names (length Tssss) x) Tssss;
-fun mk_Frees' x Ts ctxt = mk_fresh_names ctxt (length Ts) x |>> (fn xs => `(map Free) (xs ~~ Ts));
-fun mk_Freess' x Tss = fold_map2 mk_Frees' (mk_names (length Tss) x) Tss #>> split_list;
-
-
-(** Types **)
-
-fun strip_typeN 0 T = ([], T)
- | strip_typeN n (Type (@{type_name fun}, [T, T'])) = strip_typeN (n - 1) T' |>> cons T
- | strip_typeN _ T = raise TYPE ("strip_typeN", [T], []);
-
-fun mk_predT Ts = Ts ---> HOLogic.boolT;
-fun mk_pred1T T = mk_predT [T];
-fun mk_pred2T T U = mk_predT [T, U];
-fun mk_optionT T = Type (@{type_name option}, [T]);
-val mk_relT = HOLogic.mk_setT o HOLogic.mk_prodT;
-val dest_relT = HOLogic.dest_prodT o HOLogic.dest_setT;
-fun mk_sumT (LT, RT) = Type (@{type_name Sum_Type.sum}, [LT, RT]);
-fun mk_partial_funT (ranT, domT) = domT --> mk_optionT ranT;
-
-
-(** Constants **)
-
-fun fst_const T = Const (@{const_name fst}, T --> fst (HOLogic.dest_prodT T));
-fun snd_const T = Const (@{const_name snd}, T --> snd (HOLogic.dest_prodT T));
-fun Id_const T = Const (@{const_name Id}, mk_relT (T, T));
-
-
-(** Operators **)
-
-val mk_Trueprop_eq = HOLogic.mk_Trueprop o HOLogic.mk_eq;
-
-fun mk_IfN _ _ [t] = t
- | mk_IfN T (c :: cs) (t :: ts) =
- Const (@{const_name If}, HOLogic.boolT --> T --> T --> T) $ c $ t $ mk_IfN T cs ts;
-
-fun mk_converse R =
- let
- val RT = dest_relT (fastype_of R);
- val RST = mk_relT (snd RT, fst RT);
- in Const (@{const_name converse}, fastype_of R --> RST) $ R end;
-
-fun mk_rel_comp (R, S) =
- let
- val RT = fastype_of R;
- val ST = fastype_of S;
- val RST = mk_relT (fst (dest_relT RT), snd (dest_relT ST));
- in Const (@{const_name relcomp}, RT --> ST --> RST) $ R $ S end;
-
-fun mk_Gr A f =
- let val ((AT, BT), FT) = `dest_funT (fastype_of f);
- in Const (@{const_name Gr}, HOLogic.mk_setT AT --> FT --> mk_relT (AT, BT)) $ A $ f end;
-
-fun mk_image f =
- let val (T, U) = dest_funT (fastype_of f);
- in Const (@{const_name image},
- (T --> U) --> (HOLogic.mk_setT T) --> (HOLogic.mk_setT U)) $ f end;
-
-fun mk_Ball X f =
- Const (@{const_name Ball}, fastype_of X --> fastype_of f --> HOLogic.boolT) $ X $ f;
-
-fun mk_Bex X f =
- Const (@{const_name Bex}, fastype_of X --> fastype_of f --> HOLogic.boolT) $ X $ f;
-
-fun mk_UNION X f =
- let val (T, U) = dest_funT (fastype_of f);
- in Const (@{const_name SUPR}, fastype_of X --> (T --> U) --> U) $ X $ f end;
-
-fun mk_Union T =
- Const (@{const_name Sup}, HOLogic.mk_setT (HOLogic.mk_setT T) --> HOLogic.mk_setT T);
-
-fun mk_Field r =
- let val T = fst (dest_relT (fastype_of r));
- in Const (@{const_name Field}, mk_relT (T, T) --> HOLogic.mk_setT T) $ r end;
-
-fun mk_card_order bd =
- let
- val T = fastype_of bd;
- val AT = fst (dest_relT T);
- in
- Const (@{const_name card_order_on}, HOLogic.mk_setT AT --> T --> HOLogic.boolT) $
- (HOLogic.mk_UNIV AT) $ bd
- end;
-
-fun mk_Card_order bd =
- let
- val T = fastype_of bd;
- val AT = fst (dest_relT T);
- in
- Const (@{const_name card_order_on}, HOLogic.mk_setT AT --> T --> HOLogic.boolT) $
- mk_Field bd $ bd
- end;
-
-fun mk_cinfinite bd =
- Const (@{const_name cinfinite}, fastype_of bd --> HOLogic.boolT) $ bd;
-
-fun mk_ordLeq t1 t2 =
- HOLogic.mk_mem (HOLogic.mk_prod (t1, t2),
- Const (@{const_name ordLeq}, mk_relT (fastype_of t1, fastype_of t2)));
-
-fun mk_card_of A =
- let
- val AT = fastype_of A;
- val T = HOLogic.dest_setT AT;
- in
- Const (@{const_name card_of}, AT --> mk_relT (T, T)) $ A
- end;
-
-fun mk_dir_image r f =
- let val (T, U) = dest_funT (fastype_of f);
- in Const (@{const_name dir_image}, mk_relT (T, T) --> (T --> U) --> mk_relT (U, U)) $ r $ f end;
-
-(*FIXME: "x"?*)
-(*(nth sets i) must be of type "T --> 'ai set"*)
-fun mk_in As sets T =
- let
- fun in_single set A =
- let val AT = fastype_of A;
- in Const (@{const_name less_eq},
- AT --> AT --> HOLogic.boolT) $ (set $ Free ("x", T)) $ A end;
- in
- if length sets > 0
- then HOLogic.mk_Collect ("x", T, foldr1 (HOLogic.mk_conj) (map2 in_single sets As))
- else HOLogic.mk_UNIV T
- end;
-
-fun mk_wpull A B1 B2 f1 f2 pseudo p1 p2 =
- let
- val AT = fastype_of A;
- val BT1 = fastype_of B1;
- val BT2 = fastype_of B2;
- val FT1 = fastype_of f1;
- val FT2 = fastype_of f2;
- val PT1 = fastype_of p1;
- val PT2 = fastype_of p2;
- val T1 = HOLogic.dest_setT BT1;
- val T2 = HOLogic.dest_setT BT2;
- val domP = domain_type PT1;
- val ranF = range_type FT1;
- val _ = if is_some pseudo orelse
- (HOLogic.dest_setT AT = domP andalso
- domain_type FT1 = T1 andalso
- domain_type FT2 = T2 andalso
- domain_type PT2 = domP andalso
- range_type PT1 = T1 andalso
- range_type PT2 = T2 andalso
- range_type FT2 = ranF)
- then () else raise TYPE ("mk_wpull", [BT1, BT2, FT1, FT2, PT1, PT2], []);
- in
- (case pseudo of
- NONE => Const (@{const_name wpull},
- AT --> BT1 --> BT2 --> FT1 --> FT2 --> PT1 --> PT2 --> HOLogic.boolT) $
- A $ B1 $ B2 $ f1 $ f2 $ p1 $ p2
- | SOME (e1, e2) => Const (@{const_name wppull},
- AT --> BT1 --> BT2 --> FT1 --> FT2 --> fastype_of e1 --> fastype_of e2 -->
- PT1 --> PT2 --> HOLogic.boolT) $
- A $ B1 $ B2 $ f1 $ f2 $ e1 $ e2 $ p1 $ p2)
- end;
-
-fun mk_subset t1 t2 =
- Const (@{const_name less_eq}, (fastype_of t1) --> (fastype_of t2) --> HOLogic.boolT) $ t1 $ t2;
-
-fun mk_card_binop binop typop t1 t2 =
- let
- val (T1, relT1) = `(fst o dest_relT) (fastype_of t1);
- val (T2, relT2) = `(fst o dest_relT) (fastype_of t2);
- in
- Const (binop, relT1 --> relT2 --> mk_relT (typop (T1, T2), typop (T1, T2))) $ t1 $ t2
- end;
-
-val mk_csum = mk_card_binop @{const_name csum} mk_sumT;
-val mk_cprod = mk_card_binop @{const_name cprod} HOLogic.mk_prodT;
-val mk_cexp = mk_card_binop @{const_name cexp} mk_partial_funT;
-val mk_ccexp = mk_card_binop @{const_name ccexp} mk_partial_funT;
-val ctwo = @{term ctwo};
-
-fun mk_collect xs defT =
- let val T = (case xs of [] => defT | (x::_) => fastype_of x);
- in Const (@{const_name collect}, HOLogic.mk_setT T --> T) $ (HOLogic.mk_set T xs) end;
-
-fun mk_permute src dest xs = map (nth xs o (fn x => find_index ((curry op =) x) src)) dest;
-
-val list_all_free =
- fold_rev (fn free => fn P =>
- let val (x, T) = Term.dest_Free free;
- in HOLogic.all_const T $ Term.absfree (x, T) P end);
-
-val list_exists_free =
- fold_rev (fn free => fn P =>
- let val (x, T) = Term.dest_Free free;
- in HOLogic.exists_const T $ Term.absfree (x, T) P end);
-
-fun find_indices xs ys = map_filter I
- (map_index (fn (i, y) => if member (op =) xs y then SOME i else NONE) ys);
-
-fun mk_trans thm1 thm2 = trans OF [thm1, thm2];
-fun mk_sym thm = sym OF [thm];
-
-(*TODO: antiquote heavily used theorems once*)
-val ctrans = @{thm ordLeq_transitive};
-val o_apply = @{thm o_apply};
-val set_mp = @{thm set_mp};
-val set_rev_mp = @{thm set_rev_mp};
-val subset_UNIV = @{thm subset_UNIV};
-val Pair_eqD = @{thm iffD1[OF Pair_eq]};
-val Pair_eqI = @{thm iffD2[OF Pair_eq]};
-
-fun mk_nthN 1 t 1 = t
- | mk_nthN _ t 1 = HOLogic.mk_fst t
- | mk_nthN 2 t 2 = HOLogic.mk_snd t
- | mk_nthN n t m = mk_nthN (n - 1) (HOLogic.mk_snd t) (m - 1);
-
-fun mk_nth_conv n m =
- let
- fun thm b = if b then @{thm fst_snd} else @{thm snd_snd}
- fun mk_nth_conv _ 1 1 = refl
- | mk_nth_conv _ _ 1 = @{thm fst_conv}
- | mk_nth_conv _ 2 2 = @{thm snd_conv}
- | mk_nth_conv b _ 2 = @{thm snd_conv} RS thm b
- | mk_nth_conv b n m = mk_nth_conv false (n - 1) (m - 1) RS thm b;
- in mk_nth_conv (not (m = n)) n m end;
-
-fun mk_nthI 1 1 = @{thm TrueE[OF TrueI]}
- | mk_nthI n m = fold (curry op RS) (replicate (m - 1) @{thm sndI})
- (if m = n then @{thm TrueE[OF TrueI]} else @{thm fstI});
-
-fun mk_conjunctN 1 1 = @{thm TrueE[OF TrueI]}
- | mk_conjunctN _ 1 = conjunct1
- | mk_conjunctN 2 2 = conjunct2
- | mk_conjunctN n m = conjunct2 RS (mk_conjunctN (n - 1) (m - 1));
-
-fun conj_dests n thm = map (fn k => thm RS mk_conjunctN n k) (1 upto n);
-
-fun mk_conjIN 1 = @{thm TrueE[OF TrueI]}
- | mk_conjIN n = mk_conjIN (n - 1) RSN (2, conjI);
-
-fun mk_disjIN 1 1 = @{thm TrueE[OF TrueI]}
- | mk_disjIN _ 1 = disjI1
- | mk_disjIN 2 2 = disjI2
- | mk_disjIN n m = (mk_disjIN (n - 1) (m - 1)) RS disjI2;
-
-fun mk_ordLeq_csum 1 1 thm = thm
- | mk_ordLeq_csum _ 1 thm = @{thm ordLeq_transitive} OF [thm, @{thm ordLeq_csum1}]
- | mk_ordLeq_csum 2 2 thm = @{thm ordLeq_transitive} OF [thm, @{thm ordLeq_csum2}]
- | mk_ordLeq_csum n m thm = @{thm ordLeq_transitive} OF
- [mk_ordLeq_csum (n - 1) (m - 1) thm, @{thm ordLeq_csum2[OF Card_order_csum]}];
-
-local
- fun mk_Un_upper' 0 = subset_refl
- | mk_Un_upper' 1 = @{thm Un_upper1}
- | mk_Un_upper' k = Library.foldr (op RS o swap)
- (replicate (k - 1) @{thm subset_trans[OF Un_upper1]}, @{thm Un_upper1});
-in
- fun mk_Un_upper 1 1 = subset_refl
- | mk_Un_upper n 1 = mk_Un_upper' (n - 2) RS @{thm subset_trans[OF Un_upper1]}
- | mk_Un_upper n m = mk_Un_upper' (n - m) RS @{thm subset_trans[OF Un_upper2]};
-end;
-
-local
- fun mk_UnIN' 0 = @{thm UnI2}
- | mk_UnIN' m = mk_UnIN' (m - 1) RS @{thm UnI1};
-in
- fun mk_UnIN 1 1 = @{thm TrueE[OF TrueI]}
- | mk_UnIN n 1 = Library.foldr1 (op RS o swap) (replicate (n - 1) @{thm UnI1})
- | mk_UnIN n m = mk_UnIN' (n - m)
-end;
-
-fun interleave xs ys = flat (map2 (fn x => fn y => [x, y]) xs ys);
-
-fun transpose [] = []
- | transpose ([] :: xss) = transpose xss
- | transpose xss = map hd xss :: transpose (map tl xss);
-
-fun seq_conds f n k xs =
- if k = n then
- map (f false) (take (k - 1) xs)
- else
- let val (negs, pos) = split_last (take k xs) in
- map (f false) negs @ [f true pos]
- end;
-
-fun mk_unabs_def 0 thm = thm
- | mk_unabs_def n thm = mk_unabs_def (n - 1) thm RS @{thm spec[OF iffD1[OF fun_eq_iff]]};
-
-fun is_refl thm =
- op aconv (HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of thm)))
- handle TERM _ => false;
-
-val no_refl = filter_out is_refl;
-val no_reflexive = filter_out Thm.is_reflexive;
-
-fun fold_thms ctxt thms = Local_Defs.fold ctxt (distinct Thm.eq_thm_prop thms);
-fun unfold_thms ctxt thms = Local_Defs.unfold ctxt (distinct Thm.eq_thm_prop thms);
-
-end;
--- a/src/HOL/Codatatype/Tools/bnf_wrap.ML Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,665 +0,0 @@
-(* Title: HOL/BNF/Tools/bnf_wrap.ML
- Author: Jasmin Blanchette, TU Muenchen
- Copyright 2012
-
-Wrapping existing datatypes.
-*)
-
-signature BNF_WRAP =
-sig
- val mk_half_pairss: 'a list -> ('a * 'a) list list
- val mk_ctr: typ list -> term -> term
- val mk_disc_or_sel: typ list -> term -> term
- val base_name_of_ctr: term -> string
- val wrap_datatype: ({prems: thm list, context: Proof.context} -> tactic) list list ->
- ((bool * term list) * term) *
- (binding list * (binding list list * (binding * term) list list)) -> local_theory ->
- (term list * term list list * thm list * thm list * thm list * thm list list * thm list *
- thm list list) * local_theory
- val parse_wrap_options: bool parser
- val parse_bound_term: (binding * string) parser
-end;
-
-structure BNF_Wrap : BNF_WRAP =
-struct
-
-open BNF_Util
-open BNF_Wrap_Tactics
-
-val isN = "is_";
-val unN = "un_";
-fun mk_unN 1 1 suf = unN ^ suf
- | mk_unN _ l suf = unN ^ suf ^ string_of_int l;
-
-val case_congN = "case_cong";
-val case_eqN = "case_eq";
-val casesN = "cases";
-val collapseN = "collapse";
-val disc_excludeN = "disc_exclude";
-val disc_exhaustN = "disc_exhaust";
-val discsN = "discs";
-val distinctN = "distinct";
-val exhaustN = "exhaust";
-val expandN = "expand";
-val injectN = "inject";
-val nchotomyN = "nchotomy";
-val selsN = "sels";
-val splitN = "split";
-val split_asmN = "split_asm";
-val weak_case_cong_thmsN = "weak_case_cong";
-
-val std_binding = @{binding _};
-
-val induct_simp_attrs = @{attributes [induct_simp]};
-val cong_attrs = @{attributes [cong]};
-val iff_attrs = @{attributes [iff]};
-val safe_elim_attrs = @{attributes [elim!]};
-val simp_attrs = @{attributes [simp]};
-
-fun pad_list x n xs = xs @ replicate (n - length xs) x;
-
-fun unflat_lookup eq ys zs = map (map (fn x => nth zs (find_index (curry eq x) ys)));
-
-fun mk_half_pairss' _ [] = []
- | mk_half_pairss' indent (x :: xs) =
- indent @ fold_rev (cons o single o pair x) xs (mk_half_pairss' ([] :: indent) xs);
-
-fun mk_half_pairss xs = mk_half_pairss' [[]] xs;
-
-fun join_halves n half_xss other_half_xss =
- let
- val xsss =
- map2 (map2 append) (Library.chop_groups n half_xss)
- (transpose (Library.chop_groups n other_half_xss))
- val xs = interleave (flat half_xss) (flat other_half_xss);
- in (xs, xsss |> `transpose) end;
-
-fun mk_undefined T = Const (@{const_name undefined}, T);
-
-fun mk_ctr Ts t =
- let val Type (_, Ts0) = body_type (fastype_of t) in
- Term.subst_atomic_types (Ts0 ~~ Ts) t
- end;
-
-fun mk_disc_or_sel Ts t =
- Term.subst_atomic_types (snd (Term.dest_Type (domain_type (fastype_of t))) ~~ Ts) t;
-
-fun base_name_of_ctr c =
- Long_Name.base_name (case head_of c of
- Const (s, _) => s
- | Free (s, _) => s
- | _ => error "Cannot extract name of constructor");
-
-fun rapp u t = betapply (t, u);
-
-fun eta_expand_arg xs f_xs = fold_rev Term.lambda xs f_xs;
-
-fun prepare_wrap_datatype prep_term (((no_dests, raw_ctrs), raw_case),
- (raw_disc_bindings, (raw_sel_bindingss, raw_sel_defaultss))) no_defs_lthy =
- let
- (* TODO: sanity checks on arguments *)
- (* TODO: case syntax *)
-
- val n = length raw_ctrs;
- val ks = 1 upto n;
-
- val _ = if n > 0 then () else error "No constructors specified";
-
- val ctrs0 = map (prep_term no_defs_lthy) raw_ctrs;
- val case0 = prep_term no_defs_lthy raw_case;
- val sel_defaultss =
- pad_list [] n (map (map (apsnd (prep_term no_defs_lthy))) raw_sel_defaultss);
-
- val Type (dataT_name, As0) = body_type (fastype_of (hd ctrs0));
- val data_b = Binding.qualified_name dataT_name;
- val data_b_name = Binding.name_of data_b;
-
- val (As, B) =
- no_defs_lthy
- |> mk_TFrees' (map Type.sort_of_atyp As0)
- ||> the_single o fst o mk_TFrees 1;
-
- val dataT = Type (dataT_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;
-
- val raw_disc_bindings' = pad_list Binding.empty n raw_disc_bindings;
-
- fun can_really_rely_on_disc k =
- not (Binding.eq_name (nth raw_disc_bindings' (k - 1), Binding.empty)) orelse
- nth ms (k - 1) = 0;
- fun can_rely_on_disc k =
- can_really_rely_on_disc k orelse (k = 1 andalso not (can_really_rely_on_disc 2));
- fun can_omit_disc_binding k m =
- n = 1 orelse m = 0 orelse (n = 2 andalso can_rely_on_disc (3 - k));
-
- val std_disc_binding =
- Binding.qualify false data_b_name o Binding.name o prefix isN o base_name_of_ctr;
-
- val disc_bindings =
- raw_disc_bindings'
- |> map4 (fn k => fn m => fn ctr => fn disc =>
- Option.map (Binding.qualify false data_b_name)
- (if Binding.eq_name (disc, Binding.empty) then
- if can_omit_disc_binding k m then NONE else SOME (std_disc_binding ctr)
- else if Binding.eq_name (disc, std_binding) then
- SOME (std_disc_binding ctr)
- else
- SOME disc)) ks ms ctrs0;
-
- val no_discs = map is_none disc_bindings;
- val no_discs_at_all = forall I no_discs;
-
- fun std_sel_binding m l = Binding.name o mk_unN m l o base_name_of_ctr;
-
- val sel_bindingss =
- pad_list [] n raw_sel_bindingss
- |> map3 (fn ctr => fn m => map2 (fn l => fn sel =>
- Binding.qualify false data_b_name
- (if Binding.eq_name (sel, Binding.empty) orelse Binding.eq_name (sel, std_binding) then
- std_sel_binding m l ctr
- else
- sel)) (1 upto m) o pad_list Binding.empty m) ctrs0 ms;
-
- fun mk_case Ts T =
- let
- val (bindings, body) = strip_type (fastype_of case0)
- val Type (_, Ts0) = List.last bindings
- in Term.subst_atomic_types ((body, T) :: (Ts0 ~~ Ts)) case0 end;
-
- val casex = mk_case As B;
- val case_Ts = map (fn Ts => Ts ---> B) ctr_Tss;
-
- val (((((((xss, xss'), yss), fs), gs), [u', v']), (p, p')), names_lthy) = no_defs_lthy |>
- mk_Freess' "x" ctr_Tss
- ||>> mk_Freess "y" ctr_Tss
- ||>> mk_Frees "f" case_Ts
- ||>> mk_Frees "g" case_Ts
- ||>> (apfst (map (rpair dataT)) oo Variable.variant_fixes) [data_b_name, data_b_name ^ "'"]
- ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "P") HOLogic.boolT;
-
- val u = Free u';
- val v = Free v';
- 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;
-
- val eta_fs = map2 eta_expand_arg xss xfs;
- val eta_gs = map2 eta_expand_arg xss xgs;
-
- val fcase = Term.list_comb (casex, eta_fs);
- val gcase = Term.list_comb (casex, eta_gs);
-
- val ufcase = fcase $ u;
- val vfcase = fcase $ v;
- val vgcase = gcase $ v;
-
- fun mk_u_eq_u () = HOLogic.mk_eq (u, u);
-
- val u_eq_v = mk_Trueprop_eq (u, v);
-
- 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
- (case nth disc_bindings (k - 1) of
- NONE => nth exist_xs_u_eq_ctrs (k - 1)
- | SOME b => get_udisc b (k - 1));
-
- val (all_sels_distinct, discs, selss, udiscs, uselss, vdiscs, vselss, disc_defs, sel_defs,
- sel_defss, lthy') =
- if no_dests then
- (true, [], [], [], [], [], [], [], [], [], no_defs_lthy)
- else
- let
- fun disc_free b = Free (Binding.name_of b, mk_pred1T dataT);
-
- 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_default T t =
- let
- val Ts0 = map TFree (Term.add_tfreesT (fastype_of t) []);
- val Ts = map TFree (Term.add_tfreesT T []);
- in Term.subst_atomic_types (Ts0 ~~ Ts) t end;
-
- fun mk_sel_case_args b proto_sels T =
- map2 (fn Ts => fn k =>
- (case AList.lookup (op =) proto_sels k of
- NONE =>
- (case AList.lookup Binding.eq_name (rev (nth sel_defaultss (k - 1))) b of
- NONE => fold_rev (Term.lambda o curry Free Name.uu) Ts (mk_undefined T)
- | SOME t => mk_default (Ts ---> T) t)
- | SOME (xs, x) => fold_rev Term.lambda xs x)) 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 no_defs_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 no_defs_lthy T) ^
- " vs. " ^ quote (Syntax.string_of_typ no_defs_lthy T')));
- in
- mk_Trueprop_eq (Free (Binding.name_of b, dataT --> T) $ u,
- Term.list_comb (mk_case As T, mk_sel_case_args b proto_sels T) $ u)
- end;
-
- val sel_bindings = flat sel_bindingss;
- val uniq_sel_bindings = distinct Binding.eq_name sel_bindings;
- val all_sels_distinct = (length uniq_sel_bindings = length sel_bindings);
-
- val sel_binding_index =
- if all_sels_distinct then 1 upto length sel_bindings
- else map (fn b => find_index (curry Binding.eq_name b) uniq_sel_bindings) sel_bindings;
-
- val proto_sels = flat (map3 (fn k => fn xs => map (fn x => (k, (xs, x)))) ks xss xss);
- val sel_infos =
- AList.group (op =) (sel_binding_index ~~ proto_sels)
- |> sort (int_ord o pairself fst)
- |> map snd |> curry (op ~~) uniq_sel_bindings;
- val sel_bindings = map fst sel_infos;
-
- 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)) =
- no_defs_lthy
- |> apfst split_list o fold_map4 (fn k => fn m => fn exist_xs_u_eq_ctr =>
- fn NONE =>
- if n = 1 then pair (Term.lambda u (mk_u_eq_u ()), unique_disc_no_def)
- else if m = 0 then pair (Term.lambda u exist_xs_u_eq_ctr, refl)
- else pair (alternate_disc k, alternate_disc_no_def)
- | SOME b => Specification.definition (SOME (b, NONE, NoSyn),
- ((Thm.def_binding b, []), disc_spec b exist_xs_u_eq_ctr)) #>> apsnd snd)
- ks ms 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.restore;
-
- 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;
-
- val udiscs = map (rapp u) discs;
- val uselss = map (map (rapp u)) selss;
-
- val vdiscs = map (rapp v) discs;
- val vselss = map (map (rapp v)) selss;
- in
- (all_sels_distinct, discs, selss, udiscs, uselss, vdiscs, vselss, 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 (u, xctr)]) in
- fold_rev Logic.all [p, u] (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
- map4 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 (xss ~~ xctrs))
- end;
-
- val cases_goal =
- map3 (fn xs => fn xctr => fn xf =>
- fold_rev Logic.all (fs @ xs) (mk_Trueprop_eq (fcase $ xctr, xf))) xss xctrs xfs;
-
- val goalss = [exhaust_goal] :: inject_goalss @ half_distinct_goalss @ [cases_goal];
-
- fun after_qed thmss lthy =
- let
- val ([exhaust_thm], (inject_thmss, (half_distinct_thmss, [case_thms]))) =
- (hd thmss, apsnd (chop (n * n)) (chop n (tl thmss)));
-
- val inject_thms = flat inject_thmss;
-
- val Tinst = map (pairself (certifyT lthy)) (map Logic.varifyT_global As ~~ As);
-
- fun inst_thm t thm =
- Drule.instantiate' [] [SOME (certify lthy t)]
- (Thm.instantiate (Tinst, []) (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;
-
- 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
- Skip_Proof.prove lthy [] [] goal (fn _ => mk_nchotomy_tac n exhaust_thm)
- end;
-
- val (all_sel_thms, sel_thmss, disc_thmss, disc_thms, discI_thms, disc_exclude_thms,
- disc_exhaust_thms, collapse_thms, expand_thms, case_eq_thms) =
- if no_dests then
- ([], [], [], [], [], [], [], [], [], [])
- else
- let
- fun make_sel_thm xs' case_thm sel_def =
- zero_var_indexes (Drule.gen_all (Drule.rename_bvars' (map (SOME o fst) xs')
- (Drule.forall_intr_vars (case_thm RS (sel_def RS trans)))));
-
- fun has_undefined_rhs thm =
- (case snd (HOLogic.dest_eq (HOLogic.dest_Trueprop (prop_of thm))) of
- Const (@{const_name undefined}, _) => true
- | _ => false);
-
- val sel_thmss = map3 (map oo make_sel_thm) xss' case_thms sel_defss;
-
- val all_sel_thms =
- (if all_sels_distinct andalso forall null sel_defaultss 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_u_eq_u (), the_single exist_xs_u_eq_ctrs);
- in
- Skip_Proof.prove lthy [] [] goal (fn _ => mk_unique_disc_def_tac m uexhaust_thm)
- |> singleton (Proof_Context.export names_lthy lthy)
- |> Thm.close_derivation
- 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));
- in
- Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
- mk_alternate_disc_def_tac ctxt k (nth disc_defs (2 - k))
- (nth distinct_thms (2 - k)) uexhaust_thm)
- |> singleton (Proof_Context.export names_lthy lthy)
- |> Thm.close_derivation
- end;
-
- val has_alternate_disc_def =
- exists (fn def => Thm.eq_thm_prop (def, alternate_disc_no_def)) disc_defs;
-
- 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
- map3 mk_thms discI_thms not_discI_thms distinct_thmsss' |> `transpose
- end;
-
- val disc_thms = flat (map2 (fn true => K [] | false => I) no_discs disc_thmss);
-
- val (disc_exclude_thms, (disc_exclude_thmsss', disc_exclude_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 = Skip_Proof.prove lthy [] [] goal (K tac);
-
- val infos = ms ~~ discD_thms ~~ udiscs;
- val half_pairss = mk_half_pairss infos;
-
- val half_goalss = map mk_goal half_pairss;
- val half_thmss =
- map3 (fn [] => K (K []) | [goal] => fn [(((m, discD), _), _)] =>
- fn disc_thm => [prove (mk_half_disc_exclude_tac 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 (prove o mk_other_half_disc_exclude_tac)) half_thmss
- other_half_goalss;
- in
- join_halves n half_thmss other_half_thmss
- |>> has_alternate_disc_def ? K []
- end;
-
- val disc_exhaust_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
- Skip_Proof.prove lthy [] [] goal (fn _ =>
- mk_disc_exhaust_tac n exhaust_thm discI_thms)
- end;
-
- val disc_exhaust_thms =
- if has_alternate_disc_def orelse no_discs_at_all then [] else [disc_exhaust_thm];
-
- val (collapse_thms, collapse_thm_opts) =
- let
- fun mk_goal ctr udisc usels =
- let
- val prem = HOLogic.mk_Trueprop udisc;
- val concl =
- mk_Trueprop_eq ((null usels ? swap) (Term.list_comb (ctr, usels), u));
- in
- if prem aconv concl then NONE
- else SOME (Logic.all u (Logic.mk_implies (prem, concl)))
- end;
- val goals = map3 mk_goal ctrs udiscs uselss;
- in
- map4 (fn m => fn discD => fn sel_thms => Option.map (fn goal =>
- Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
- mk_collapse_tac ctxt m discD sel_thms)
- |> perhaps (try (fn thm => refl RS thm)))) ms discD_thms sel_thmss goals
- |> `(map_filter I)
- end;
-
- val expand_thms =
- 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 uncollapse_thms =
- map (fn NONE => Drule.dummy_thm | SOME thm => thm RS sym) collapse_thm_opts;
-
- val goal =
- Library.foldr Logic.list_implies
- (map5 mk_prems ks udiscs uselss vdiscs vselss, u_eq_v);
- in
- [Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
- mk_expand_tac ctxt n ms (inst_thm u disc_exhaust_thm)
- (inst_thm v disc_exhaust_thm) uncollapse_thms disc_exclude_thmsss
- disc_exclude_thmsss')]
- |> Proof_Context.export names_lthy lthy
- end;
-
- val case_eq_thms =
- let
- fun mk_body f usels = Term.list_comb (f, usels);
- val goal = mk_Trueprop_eq (ufcase, mk_IfN B udiscs (map2 mk_body fs uselss));
- in
- [Skip_Proof.prove lthy [] [] goal (fn {context = ctxt, ...} =>
- mk_case_eq_tac ctxt n uexhaust_thm case_thms disc_thmss' sel_thmss)]
- |> Proof_Context.export names_lthy lthy
- end;
- in
- (all_sel_thms, sel_thmss, disc_thmss, disc_thms, discI_thms, disc_exclude_thms,
- disc_exhaust_thms, collapse_thms, expand_thms, case_eq_thms)
- end;
-
- val (case_cong_thm, weak_case_cong_thm) =
- let
- fun mk_prem xctr xs f g =
- fold_rev Logic.all xs (Logic.mk_implies (mk_Trueprop_eq (v, xctr),
- mk_Trueprop_eq (f, g)));
-
- val goal =
- Logic.list_implies (u_eq_v :: map4 mk_prem xctrs xss fs gs,
- mk_Trueprop_eq (ufcase, vgcase));
- val weak_goal = Logic.mk_implies (u_eq_v, mk_Trueprop_eq (ufcase, vfcase));
- in
- (Skip_Proof.prove lthy [] [] goal (fn _ => mk_case_cong_tac uexhaust_thm case_thms),
- Skip_Proof.prove lthy [] [] weak_goal (K (etac arg_cong 1)))
- |> pairself (singleton (Proof_Context.export names_lthy lthy))
- end;
-
- val (split_thm, split_asm_thm) =
- let
- fun mk_conjunct xctr xs f_xs =
- list_all_free xs (HOLogic.mk_imp (HOLogic.mk_eq (u, xctr), q $ f_xs));
- fun mk_disjunct xctr xs f_xs =
- list_exists_free xs (HOLogic.mk_conj (HOLogic.mk_eq (u, xctr),
- HOLogic.mk_not (q $ f_xs)));
-
- val lhs = q $ ufcase;
-
- val goal =
- mk_Trueprop_eq (lhs, Library.foldr1 HOLogic.mk_conj (map3 mk_conjunct xctrs xss xfs));
- val asm_goal =
- mk_Trueprop_eq (lhs, HOLogic.mk_not (Library.foldr1 HOLogic.mk_disj
- (map3 mk_disjunct xctrs xss xfs)));
-
- val split_thm =
- Skip_Proof.prove lthy [] [] goal
- (fn _ => mk_split_tac uexhaust_thm case_thms inject_thmss distinct_thmsss)
- |> singleton (Proof_Context.export names_lthy lthy)
- val split_asm_thm =
- Skip_Proof.prove lthy [] [] asm_goal (fn {context = ctxt, ...} =>
- mk_split_asm_tac ctxt split_thm)
- |> singleton (Proof_Context.export names_lthy lthy)
- in
- (split_thm, split_asm_thm)
- 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 dataT_name));
-
- val notes =
- [(case_congN, [case_cong_thm], []),
- (case_eqN, case_eq_thms, []),
- (casesN, case_thms, simp_attrs),
- (collapseN, collapse_thms, simp_attrs),
- (discsN, disc_thms, simp_attrs),
- (disc_excludeN, disc_exclude_thms, []),
- (disc_exhaustN, disc_exhaust_thms, [exhaust_case_names_attr]),
- (distinctN, distinct_thms, simp_attrs @ induct_simp_attrs),
- (exhaustN, [exhaust_thm], [exhaust_case_names_attr, cases_type_attr]),
- (expandN, expand_thms, []),
- (injectN, inject_thms, iff_attrs @ induct_simp_attrs),
- (nchotomyN, [nchotomy_thm], []),
- (selsN, all_sel_thms, simp_attrs),
- (splitN, [split_thm], []),
- (split_asmN, [split_asm_thm], []),
- (weak_case_cong_thmsN, [weak_case_cong_thm], cong_attrs)]
- |> filter_out (null o #2)
- |> map (fn (thmN, thms, attrs) =>
- ((Binding.qualify true data_b_name (Binding.name thmN), attrs), [(thms, [])]));
-
- val notes' =
- [(map (fn th => th RS notE) distinct_thms, safe_elim_attrs)]
- |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])]));
- in
- ((discs, selss, inject_thms, distinct_thms, case_thms, disc_thmss, discI_thms, sel_thmss),
- lthy |> Local_Theory.notes (notes' @ notes) |> snd)
- end;
- in
- (goalss, after_qed, lthy')
- end;
-
-fun wrap_datatype tacss = (fn (goalss, after_qed, lthy) =>
- map2 (map2 (Skip_Proof.prove lthy [] [])) goalss tacss
- |> (fn thms => after_qed thms lthy)) oo prepare_wrap_datatype (K I);
-
-val wrap_datatype_cmd = (fn (goalss, after_qed, lthy) =>
- Proof.theorem NONE (snd oo after_qed) (map (map (rpair [])) goalss) lthy) oo
- prepare_wrap_datatype Syntax.read_term;
-
-fun parse_bracket_list parser = @{keyword "["} |-- Parse.list parser --| @{keyword "]"};
-
-val parse_bindings = parse_bracket_list Parse.binding;
-val parse_bindingss = parse_bracket_list parse_bindings;
-
-val parse_bound_term = (Parse.binding --| @{keyword ":"}) -- Parse.term;
-val parse_bound_terms = parse_bracket_list parse_bound_term;
-val parse_bound_termss = parse_bracket_list parse_bound_terms;
-
-val parse_wrap_options =
- Scan.optional (@{keyword "("} |-- (@{keyword "no_dests"} >> K true) --| @{keyword ")"}) false;
-
-val _ =
- Outer_Syntax.local_theory_to_proof @{command_spec "wrap_data"} "wraps an existing datatype"
- ((parse_wrap_options -- (@{keyword "["} |-- Parse.list Parse.term --| @{keyword "]"}) --
- Parse.term -- Scan.optional (parse_bindings -- Scan.optional (parse_bindingss --
- Scan.optional parse_bound_termss []) ([], [])) ([], ([], [])))
- >> wrap_datatype_cmd);
-
-end;
--- a/src/HOL/Codatatype/Tools/bnf_wrap_tactics.ML Fri Sep 21 16:34:40 2012 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,122 +0,0 @@
-(* Title: HOL/BNF/Tools/bnf_wrap_tactics.ML
- Author: Jasmin Blanchette, TU Muenchen
- Copyright 2012
-
-Tactics for wrapping datatypes.
-*)
-
-signature BNF_WRAP_TACTICS =
-sig
- val mk_alternate_disc_def_tac: Proof.context -> int -> thm -> thm -> thm -> tactic
- val mk_case_cong_tac: thm -> thm list -> tactic
- val mk_case_eq_tac: Proof.context -> int -> thm -> thm list -> thm list list -> thm list list ->
- tactic
- val mk_collapse_tac: Proof.context -> int -> thm -> thm list -> tactic
- val mk_disc_exhaust_tac: int -> thm -> thm list -> tactic
- val mk_expand_tac: Proof.context -> int -> int list -> thm -> thm -> thm list ->
- thm list list list -> thm list list list -> tactic
- val mk_half_disc_exclude_tac: int -> thm -> thm -> tactic
- val mk_nchotomy_tac: int -> thm -> tactic
- val mk_other_half_disc_exclude_tac: thm -> tactic
- val mk_split_tac: thm -> thm list -> thm list list -> thm list list list -> tactic
- val mk_split_asm_tac: Proof.context -> thm -> tactic
- val mk_unique_disc_def_tac: int -> thm -> tactic
-end;
-
-structure BNF_Wrap_Tactics : BNF_WRAP_TACTICS =
-struct
-
-open BNF_Util
-open BNF_Tactics
-
-val meta_mp = @{thm meta_mp};
-
-fun if_P_or_not_P_OF pos thm = thm RS (if pos then @{thm if_P} else @{thm if_not_P});
-
-fun mk_nchotomy_tac n exhaust =
- (rtac allI THEN' rtac exhaust THEN'
- EVERY' (maps (fn k => [rtac (mk_disjIN n k), REPEAT_DETERM o rtac exI, atac]) (1 upto n))) 1;
-
-fun mk_unique_disc_def_tac m uexhaust =
- EVERY' [rtac iffI, rtac uexhaust, REPEAT_DETERM_N m o rtac exI, atac, rtac refl] 1;
-
-fun mk_alternate_disc_def_tac ctxt k other_disc_def distinct uexhaust =
- EVERY' ([subst_tac ctxt [other_disc_def], rtac @{thm iffI_np}, REPEAT_DETERM o etac exE,
- hyp_subst_tac, SELECT_GOAL (unfold_thms_tac ctxt [not_ex]), REPEAT_DETERM o rtac allI,
- rtac distinct, rtac uexhaust] @
- (([etac notE, REPEAT_DETERM o rtac exI, atac], [REPEAT_DETERM o rtac exI, atac])
- |> k = 1 ? swap |> op @)) 1;
-
-fun mk_half_disc_exclude_tac m discD disc' =
- (dtac discD THEN' REPEAT_DETERM_N m o etac exE THEN' hyp_subst_tac THEN' rtac disc') 1;
-
-fun mk_other_half_disc_exclude_tac half = (etac @{thm contrapos_pn} THEN' etac half) 1;
-
-fun mk_disc_exhaust_tac n exhaust discIs =
- (rtac exhaust THEN'
- EVERY' (map2 (fn k => fn discI =>
- dtac discI THEN' select_prem_tac n (etac meta_mp) k THEN' atac) (1 upto n) discIs)) 1;
-
-fun mk_collapse_tac ctxt m discD sels =
- (dtac discD THEN'
- (if m = 0 then
- atac
- else
- REPEAT_DETERM_N m o etac exE THEN' hyp_subst_tac THEN'
- SELECT_GOAL (unfold_thms_tac ctxt sels) THEN' rtac refl)) 1;
-
-fun mk_expand_tac ctxt n ms udisc_exhaust vdisc_exhaust uncollapses disc_excludesss
- disc_excludesss' =
- if ms = [0] then
- rtac (@{thm trans_sym} OF (replicate 2 (the_single uncollapses RS sym))) 1
- else
- let
- val ks = 1 upto n;
- val maybe_atac = if n = 1 then K all_tac else atac;
- in
- (rtac udisc_exhaust THEN'
- EVERY' (map5 (fn k => fn m => fn disc_excludess => fn disc_excludess' => fn uuncollapse =>
- EVERY' [if m = 0 then K all_tac else subst_tac ctxt [uuncollapse] THEN' maybe_atac,
- rtac sym, rtac vdisc_exhaust,
- EVERY' (map4 (fn k' => fn disc_excludes => fn disc_excludes' => fn vuncollapse =>
- EVERY'
- (if k' = k then
- if m = 0 then
- [hyp_subst_tac, rtac refl]
- else
- [subst_tac ctxt [vuncollapse], maybe_atac,
- if n = 1 then K all_tac else EVERY' [dtac meta_mp, atac, dtac meta_mp, atac],
- REPEAT_DETERM_N (Int.max (0, m - 1)) o etac conjE, asm_simp_tac (ss_only [])]
- else
- [dtac (the_single (if k = n then disc_excludes else disc_excludes')),
- etac (if k = n then @{thm iff_contradict(1)} else @{thm iff_contradict(2)}),
- atac, atac]))
- ks disc_excludess disc_excludess' uncollapses)])
- ks ms disc_excludesss disc_excludesss' uncollapses)) 1
- end;
-
-fun mk_case_eq_tac ctxt n uexhaust cases discss' selss =
- (rtac uexhaust THEN'
- EVERY' (map3 (fn casex => fn if_discs => fn sels =>
- EVERY' [hyp_subst_tac, SELECT_GOAL (unfold_thms_tac ctxt (if_discs @ sels)), rtac casex])
- cases (map2 (seq_conds if_P_or_not_P_OF n) (1 upto n) discss') selss)) 1;
-
-fun mk_case_cong_tac uexhaust cases =
- (rtac uexhaust THEN'
- EVERY' (maps (fn casex => [dtac sym, asm_simp_tac (ss_only [casex])]) cases)) 1;
-
-val naked_ctxt = Proof_Context.init_global @{theory HOL};
-
-fun mk_split_tac uexhaust cases injectss distinctsss =
- rtac uexhaust 1 THEN
- ALLGOALS (fn k => (hyp_subst_tac THEN'
- simp_tac (ss_only (@{thms simp_thms} @ cases @ nth injectss (k - 1) @
- flat (nth distinctsss (k - 1))))) k) THEN
- ALLGOALS (blast_tac naked_ctxt);
-
-val split_asm_thms = @{thms imp_conv_disj de_Morgan_conj de_Morgan_disj not_not not_ex};
-
-fun mk_split_asm_tac ctxt split =
- rtac (split RS trans) 1 THEN unfold_thms_tac ctxt split_asm_thms THEN rtac refl 1;
-
-end;
--- a/src/HOL/ROOT Fri Sep 21 16:34:40 2012 +0200
+++ b/src/HOL/ROOT Fri Sep 21 16:45:06 2012 +0200
@@ -610,13 +610,13 @@
"document/root.tex"
"document/root.bib"
-session "HOL-Codatatype" in Codatatype = "HOL-Cardinals" +
- description {* New (Co)datatype Package *}
+session "HOL-BNF" in BNF = "HOL-Cardinals" +
+ description {* Bounded Natural Functors for (Co)datatypes *}
options [document = false]
- theories Codatatype
+ theories BNF
-session "HOL-Codatatype-Examples" in "Codatatype/Examples" = "HOL-Codatatype" +
- description {* Examples for the New (Co)datatype Package *}
+session "HOL-BNF-Examples" in "BNF/Examples" = "HOL-BNF" +
+ description {* Examples for Bounded Natural Functors *}
options [document = false]
theories
HFset