--- a/CONTRIBUTORS Tue Aug 28 17:51:43 2012 +0200
+++ b/CONTRIBUTORS Tue Aug 28 17:53:08 2012 +0200
@@ -6,6 +6,12 @@
Contributions to this Isabelle version
--------------------------------------
+* August 2012: Dmitriy Traytel, Andrei Popescu, Jasmin Blanchette, TUM
+ New (co)datatype package.
+
+* August 2012: Andrei Popescu and Dmitriy Traytel, TUM
+ Theories of ordinals and cardinals.
+
* July 2012: Makarius Wenzel, Université Paris-Sud / LRI
Advanced support for Isabelle sessions and build management, notably
"isabelle build".
@@ -161,7 +167,7 @@
* November 2009: Lukas Bulwahn, TUM
Predicate Compiler: a compiler for inductive predicates to
equational specifications.
-
+
* November 2009: Sascha Boehme, TUM and Burkhart Wolff, LRI Paris
HOL-Boogie: an interactive prover back-end for Boogie and VCC.
--- a/NEWS Tue Aug 28 17:51:43 2012 +0200
+++ b/NEWS Tue Aug 28 17:53:08 2012 +0200
@@ -41,6 +41,12 @@
*** HOL ***
+* HOL/Codatatype: New (co)datatype package with support for mixed,
+nested recursion and interesting non-free datatypes.
+
+* HOL/Ordinals_and_Cardinals: Theories of ordinals and cardinals
+(supersedes the AFP entry of the same name).
+
* Library/Debug.thy and Library/Parallel.thy: debugging and parallel
execution for code generated towards Isabelle/ML.
@@ -51,7 +57,7 @@
- added an optimisation for equality premises.
It is switched on by default, and can be switched off by setting
- the configuration quickcheck_optimise_equality to false.
+ the configuration quickcheck_optimise_equality to false.
* The SMT solver Z3 has now by default a restricted set of directly
supported features. For the full set of features (div/mod, nonlinear
@@ -4557,7 +4563,7 @@
type and generates all solutions by prolog-style backwards proof using
the declared rules.
-This setup also deals with rules like
+This setup also deals with rules like
"is_measure f ==> is_measure (list_size f)"
@@ -5329,7 +5335,7 @@
rewrites now. That is, trace_simp_depth_limit is set to 1 by
default. Thus there is less danger of being flooded by the trace. The
trace indicates where parts have been suppressed.
-
+
* Provers/classical: removed obsolete classical version of elim_format
attribute; classical elim/dest rules are now treated uniformly when
manipulating the claset.
@@ -5368,7 +5374,7 @@
Examples are in the directory MetisExamples. WARNING: the
Isabelle/HOL-Metis integration does not yet work properly with
multi-threading.
-
+
* Command 'sledgehammer' invokes external automatic theorem provers as
background processes. It generates calls to the "metis" method if
successful. These can be pasted into the proof. Users do not have to
@@ -5734,7 +5740,7 @@
users are encouraged to use the new package.
* Method "lexicographic_order" automatically synthesizes termination
-relations as lexicographic combinations of size measures.
+relations as lexicographic combinations of size measures.
* Case-expressions allow arbitrary constructor-patterns (including
"_") and take their order into account, like in functional
@@ -5818,7 +5824,7 @@
(K x) r The K-combinator that is internally used is called K_record.
INCOMPATIBILITY: Usage of the plain update functions has to be
adapted.
-
+
* Class "semiring_0" now contains annihilation axioms x * 0 = 0 and 0
* x = 0, which are required for a semiring. Richer structures do not
inherit from semiring_0 anymore, because this property is a theorem
@@ -6714,19 +6720,19 @@
mult_neg_le now named mult_nonpos_nonpos
* The following lemmas in Ring_and_Field have been added to the simplifier:
-
+
zero_le_square
- not_square_less_zero
+ not_square_less_zero
The following lemmas have been deleted from Real/RealPow:
-
+
realpow_zero_zero
realpow_two
realpow_less
zero_le_power
realpow_two_le
abs_realpow_two
- realpow_two_abs
+ realpow_two_abs
* Theory Parity: added rules for simplifying exponents.
--- a/doc-src/IsarImplementation/Integration.thy Tue Aug 28 17:51:43 2012 +0200
+++ b/doc-src/IsarImplementation/Integration.thy Tue Aug 28 17:53:08 2012 +0200
@@ -6,7 +6,7 @@
section {* Isar toplevel \label{sec:isar-toplevel} *}
-text {* The Isar toplevel may be considered the centeral hub of the
+text {* The Isar toplevel may be considered the central hub of the
Isabelle/Isar system, where all key components and sub-systems are
integrated into a single read-eval-print loop of Isar commands,
which also incorporates the underlying ML compiler.
--- a/etc/isar-keywords.el Tue Aug 28 17:51:43 2012 +0200
+++ b/etc/isar-keywords.el Tue Aug 28 17:53:08 2012 +0200
@@ -32,6 +32,10 @@
"axiomatization"
"axioms"
"back"
+ "bnf_codata"
+ "bnf_data"
+ "bnf_def"
+ "bnf_of_typ"
"boogie_end"
"boogie_open"
"boogie_status"
@@ -170,6 +174,7 @@
"print_ast_translation"
"print_attributes"
"print_binds"
+ "print_bnfs"
"print_bundles"
"print_cases"
"print_claset"
@@ -387,6 +392,7 @@
"print_antiquotations"
"print_attributes"
"print_binds"
+ "print_bnfs"
"print_bundles"
"print_cases"
"print_claset"
@@ -463,6 +469,9 @@
"attribute_setup"
"axiomatization"
"axioms"
+ "bnf_codata"
+ "bnf_data"
+ "bnf_of_typ"
"boogie_end"
"boogie_open"
"bundle"
@@ -568,6 +577,7 @@
(defconst isar-keywords-theory-goal
'("ax_specification"
+ "bnf_def"
"boogie_vc"
"code_pred"
"corollary"
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Codatatype/BNF_Comp.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,20 @@
+(* Title: HOL/Codatatype/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
+keywords
+ "bnf_of_typ" :: thy_decl
+uses
+ "Tools/bnf_comp_tactics.ML"
+ "Tools/bnf_comp.ML"
+ "Tools/bnf_fp_util.ML"
+begin
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Codatatype/BNF_Def.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,22 @@
+(* Title: HOL/Codatatype/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_Library
+keywords
+ "print_bnfs" :: diag
+and
+ "bnf_def" :: thy_goal
+uses
+ "Tools/bnf_util.ML"
+ "Tools/bnf_tactics.ML"
+ "Tools/bnf_def.ML"
+begin
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Codatatype/BNF_GFP.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,20 @@
+(* Title: HOL/Codatatype/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_Comp
+keywords
+ "bnf_codata" :: thy_decl
+uses
+ "Tools/bnf_gfp_util.ML"
+ "Tools/bnf_gfp_tactics.ML"
+ "Tools/bnf_gfp.ML"
+begin
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Codatatype/BNF_LFP.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,20 @@
+(* Title: HOL/Codatatype/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_Comp
+keywords
+ "bnf_data" :: thy_decl
+uses
+ "Tools/bnf_lfp_util.ML"
+ "Tools/bnf_lfp_tactics.ML"
+ "Tools/bnf_lfp.ML"
+begin
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Codatatype/BNF_Library.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,824 @@
+(* Title: HOL/Codatatype/BNF_Library.thy
+ Author: Dmitriy Traytel, TU Muenchen
+ Copyright 2012
+
+Library for bounded natural functors.
+*)
+
+header {* Library for Bounded Natural Functors *}
+
+theory BNF_Library
+imports "../Ordinals_and_Cardinals/Cardinal_Arithmetic" "~~/src/HOL/Library/List_Prefix"
+ Equiv_Relations_More
+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
+
+lemma mem_Collect_eq_split: "{(x, y). (x, y) \<in> X} = X"
+by simp
+
+lemma image_comp: "image (f o g) = image f o image g"
+by (rule ext) (auto simp only: o_apply image_def)
+
+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
+
+definition collect where
+ "collect F x = (\<Union>f \<in> F. f x)"
+
+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 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 auto
+
+lemma subset_emptyI: "(\<And>x. x \<in> A \<Longrightarrow> False) \<Longrightarrow> A \<subseteq> {}"
+by auto
+
+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 auto
+
+lemma UN_image_subset: "\<Union>f ` g x \<subseteq> X = (g x \<subseteq> {x. f x \<subseteq> X})"
+by auto
+
+lemma image_Collect_subsetI:
+ "(\<And>x. P x \<Longrightarrow> f x \<in> B) \<Longrightarrow> f ` {x. P x} \<subseteq> B"
+by auto
+
+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 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_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
+
+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
+
+
+section{* Weak pullbacks: *}
+
+definition csquare where
+"csquare A f1 f2 p1 p2 \<longleftrightarrow> (\<forall> a \<in> A. f1 (p1 a) = f2 (p2 a))"
+
+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))"
+
+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 wpull_id: "wpull UNIV B1 B2 id id id id"
+unfolding wpull_def by simp
+
+
+(* 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))"
+
+
+(* 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 auto
+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
+
+
+(* Operators: *)
+definition diag where "diag A \<equiv> {(a,a) | a. a \<in> A}"
+definition "Gr A f = {(a,f 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_def': "Id = {(a,b). a = b}"
+by auto
+
+lemma Id_alt: "Id = Gr UNIV id"
+unfolding Gr_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_UNIV_id: "f = id \<Longrightarrow> (Gr UNIV f)^-1 O Gr UNIV f = Gr UNIV f"
+unfolding Gr_def by auto
+
+lemma Gr_fst_snd: "(Gr R fst)^-1 O Gr R snd = R"
+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 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 wpull_Gr:
+"wpull (Gr A f) A (f ` A) f id fst snd"
+unfolding wpull_def Gr_def by auto
+
+lemma Gr_incl: "Gr A f \<subseteq> A <*> B \<longleftrightarrow> f ` A \<subseteq> B"
+unfolding Gr_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)
+
+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 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
+
+
+(* Relation composition as a weak pseudo-pullback *)
+
+(* pick middle *)
+definition "pickM P Q a c = (SOME b. (a,b) \<in> P \<and> (b,c) \<in> Q)"
+
+lemma pickM:
+assumes "(a,c) \<in> P O Q"
+shows "(a, pickM P Q a c) \<in> P \<and> (pickM P Q a c, c) \<in> Q"
+unfolding pickM_def apply(rule someI_ex)
+using assms unfolding relcomp_def by auto
+
+definition fstO where "fstO P Q ac = (fst ac, pickM P Q (fst ac) (snd ac))"
+definition sndO where "sndO P Q ac = (pickM P Q (fst ac) (snd ac), snd ac)"
+
+lemma fstO_in: "ac \<in> P O Q \<Longrightarrow> fstO P Q ac \<in> P"
+by (metis assms fstO_def pickM surjective_pairing)
+
+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"
+by (metis assms sndO_def pickM surjective_pairing)
+
+lemma csquare_fstO_sndO:
+"csquare (P O Q) snd fst (fstO P Q) (sndO P Q)"
+unfolding csquare_def fstO_def sndO_def using pickM by auto
+
+lemma wppull_fstO_sndO:
+shows "wppull (P O Q) P Q snd fst fst snd (fstO P Q) (sndO P Q)"
+using pickM unfolding wppull_def fstO_def sndO_def relcomp_def by auto
+
+lemma subst_Pair: "P x y \<Longrightarrow> a = (x, y) \<Longrightarrow> P (fst a) (snd a)"
+by simp
+
+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 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 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 Collect_restrict: "{x. x \<in> X \<and> P x} \<subseteq> X"
+by auto
+
+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 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
+
+
+subsection {* Convolution product *}
+
+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 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 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
+
+lemma convol_expand_snd: "fst o f = g \<Longrightarrow> <g, snd o f> = f"
+unfolding convol_def by auto
+
+subsection{* Facts about functions *}
+
+lemma pointfreeE: "f o g = f' o g' \<Longrightarrow> f (g x) = f' (g' x)"
+unfolding o_def fun_eq_iff by simp
+
+lemma pointfree_idE: "f o g = id \<Longrightarrow> f (g x) = x"
+unfolding o_def fun_eq_iff by simp
+
+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 gg imageI inj_f the_inv_into_f_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 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 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 (metis bij_betw_iff_ex bij_betw_imageE)
+
+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 by auto (metis rev_image_eqI)
+
+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 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 by (metis Card_order_trans insert(5) insertE y(2) z)
+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)`)
+
+(*Extended List_Prefix*)
+
+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))"
+
+lemmas sh_def = Shift_def shift_def
+
+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 ssubst_mem: "\<lbrakk>t = s; s \<in> X\<rbrakk> \<Longrightarrow> t \<in> X" by simp
+
+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 sum_case_cong: "p = q \<Longrightarrow> sum_case f g p = sum_case f g q"
+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 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_base: "\<lbrakk>\<And>x. s = x \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
+by 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> \<forall>x. s = f (Inr x) \<longrightarrow> P"
+by (metis obj_sumE)
+
+lemma not_arg_cong_Inr: "x \<noteq> y \<Longrightarrow> Inr x \<noteq> Inr y"
+by auto
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Codatatype/Basic_BNFs.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,1528 @@
+(* Title: HOL/Codatatype/Basic_BNFs.thy
+ Author: Dmitriy Traytel, TU Muenchen
+ Author: Andrei Popescu, TU Muenchen
+ 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 Basic_BNFs
+imports BNF_Def "~~/src/HOL/Quotient_Examples/FSet" "~~/src/HOL/Library/Multiset" Countable_Set
+begin
+
+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"]
+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)
+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)
+done
+
+lemma ID_pred[simp]: "ID_pred \<phi> = \<phi>"
+unfolding ID_pred_def ID_rel_def Gr_def fun_eq_iff by auto
+
+bnf_def DEADID = "id :: 'a \<Rightarrow> 'a" [] "\<lambda>_:: 'a. natLeq +c |UNIV :: 'a set|" ["SOME x :: 'a. True"]
+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)
+by (rule card_of_Card_order)
+
+lemma DEADID_pred[simp]: "DEADID_pred = (op =)"
+unfolding DEADID_pred_def DEADID.rel_Id by simp
+
+ML {*
+
+signature BASIC_BNFS =
+sig
+ val ID_bnf: BNF_Def.BNF
+ val ID_rel_def: thm
+ val ID_pred_def: thm
+
+ val DEADID_bnf: BNF_Def.BNF
+end;
+
+structure Basic_BNFs : BASIC_BNFS =
+struct
+
+ val ID_bnf = the (BNF_Def.bnf_of @{context} "ID");
+ val DEADID_bnf = the (BNF_Def.bnf_of @{context} "DEADID");
+
+ val rel_def = BNF_Def.rel_def_of_bnf ID_bnf;
+ val ID_rel_def = rel_def RS sym;
+ val ID_pred_def =
+ Local_Defs.unfold @{context} [rel_def] (BNF_Def.pred_def_of_bnf ID_bnf) RS sym;
+
+end;
+*}
+
+definition sum_setl :: "'a + 'b \<Rightarrow> 'a set" where
+"sum_setl x = (case x of Inl z => {z} | _ => {})"
+
+definition sum_setr :: "'a + 'b \<Rightarrow> 'b set" where
+"sum_setr x = (case x of Inr z => {z} | _ => {})"
+
+lemmas sum_set_defs = sum_setl_def[abs_def] sum_setr_def[abs_def]
+
+bnf_def sum = sum_map [sum_setl, sum_setr] "\<lambda>_::'a + 'b. natLeq" [Inl, Inr]
+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> sum_setl x \<Longrightarrow> f1 z = g1 z" and
+ a2: "\<And>z. z \<in> sum_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: sum_setl_def)
+ next
+ case Inr thus ?thesis using a2 by (clarsimp simp: sum_setr_def)
+ qed
+next
+ fix f1 f2
+ show "sum_setl o sum_map f1 f2 = image f1 o sum_setl"
+ by (rule ext, unfold o_apply) (simp add: sum_setl_def split: sum.split)
+next
+ fix f1 f2
+ show "sum_setr o sum_map f1 f2 = image f2 o sum_setr"
+ by (rule ext, unfold o_apply) (simp add: sum_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 "|sum_setl x| \<le>o natLeq"
+ apply (rule ordLess_imp_ordLeq)
+ apply (rule finite_iff_ordLess_natLeq[THEN iffD1])
+ by (simp add: sum_setl_def split: sum.split)
+next
+ fix x
+ show "|sum_setr x| \<le>o natLeq"
+ apply (rule ordLess_imp_ordLeq)
+ apply (rule finite_iff_ordLess_natLeq[THEN iffD1])
+ by (simp add: sum_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. sum_setl x \<subseteq> A1 \<and> sum_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. sum_setl x \<subseteq> A1 \<and> sum_setr x \<subseteq> A2}
+ {x. sum_setl x \<subseteq> B11 \<and> sum_setr x \<subseteq> B12} {x. sum_setl x \<subseteq> B21 \<and> sum_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: sum_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
+qed (auto simp: sum_set_defs)
+
+lemma sum_pred[simp]:
+ "sum_pred \<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))"
+unfolding sum_setl_def sum_setr_def sum_pred_def sum_rel_def Gr_def relcomp_unfold converse_unfold
+by (fastforce split: sum.splits)+
+
+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]
+
+bnf_def prod = map_pair [fsts, snds] "\<lambda>_::'a \<times> 'b. ctwo *c natLeq" [Pair]
+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
+qed simp+
+
+lemma prod_pred[simp]:
+"prod_pred \<phi> \<psi> p1 p2 = (case p1 of (a1, b1) \<Rightarrow> case p2 of (a2, b2) \<Rightarrow> (\<phi> a1 a2 \<and> \<psi> b1 b2))"
+unfolding prod_set_defs prod_pred_def prod_rel_def Gr_def relcomp_unfold converse_unfold by auto
+(* TODO: pred characterization for each basic BNF *)
+
+(* 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
+
+bnf_def "fun" = "op \<circ>" [range] "\<lambda>_:: 'a \<Rightarrow> 'b. natLeq +c |UNIV :: 'a set|"
+ ["%c x::'b::type. c::'a::type"]
+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
+qed auto
+
+lemma fun_pred[simp]: "fun_pred \<phi> f g = (\<forall>x. \<phi> (f x) (g x))"
+unfolding fun_rel_def fun_pred_def Gr_def relcomp_unfold converse_unfold
+by (auto intro!: exI dest!: in_mono)
+
+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 list = 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 auto
+
+bnf_def deadlist = "map id" [] "\<lambda>_::'a list. |lists (UNIV :: 'a set)|" ["[]"]
+by (auto simp add: cinfinite_def wpull_def infinite_UNIV_listI map.id
+ ordLeq_transitive ctwo_def card_of_card_order_on Field_card_of card_of_mono1 ordLeq_cexp2)
+
+(* 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)
+
+bnf_def fset = map_fset [fset] "\<lambda>_::'a fset. natLeq" ["{||}"]
+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
+qed auto
+
+lemma fset_pred[simp]: "fset_pred 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)))" (is "?L = ?R")
+proof
+ assume ?L thus ?R unfolding fset_rel_def fset_pred_def
+ 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)
+next
+ assume ?R
+ def R' \<equiv> "the_inv fset (Collect (split R) \<inter> (fset a \<times> fset b))" (is "the_inv fset ?R'")
+ have "finite ?R'" by (intro finite_Int[OF disjI2] finite_cartesian_product) auto
+ hence *: "fset R' = ?R'" unfolding R'_def by (intro fset_to_fset)
+ show ?L unfolding fset_rel_def fset_pred_def 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 `?R`]
+ 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 `?R`]
+ by (auto simp add: fset_cong[symmetric] image_def Int_def split: prod.splits)
+ qed (auto simp add: *)
+qed
+
+(* 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
+
+bnf_def cset = cIm [rcset] "\<lambda>_::'a cset. natLeq" ["cEmp"]
+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
+qed (unfold cEmp_def, auto)
+
+
+(* Multisets *)
+
+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_count_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_count_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 mset_map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a multiset \<Rightarrow> 'b multiset" where
+"mset_map h = Abs_multiset \<circ> mmap h \<circ> count"
+
+bnf_def mset = mset_map [set_of] "\<lambda>_::'a multiset. natLeq" ["{#}"]
+unfolding mset_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)
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Codatatype/Codatatype.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,14 @@
+(* Title: HOL/Codatatype/Codatatype.thy
+ Author: Dmitriy Traytel, TU Muenchen
+ Copyright 2012
+
+The (co)datatype package.
+*)
+
+header {* The (Co)datatype Package *}
+
+theory Codatatype
+imports BNF_LFP BNF_GFP
+begin
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Codatatype/Countable_Set.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,327 @@
+(* Title: HOL/Codatatype/Countable_Set.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+(At most) countable sets.
+*)
+
+header {* (At Most) Countable Sets *}
+
+theory Countable_Set
+imports "../Ordinals_and_Cardinals/Cardinal_Arithmetic" "~~/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)
+
+
+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/Codatatype/Equiv_Relations_More.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,161 @@
+(* Title: HOL/Codatatype/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/Codatatype/Examples/HFset.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,60 @@
+(* Title: Codatatype_Examples/HFset.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Hereditary sets.
+*)
+
+header {* Hereditary Sets *}
+
+theory HFset
+imports "../Codatatype"
+begin
+
+
+section {* Datatype definition *}
+
+bnf_data hfset: 'hfset = "'hfset fset"
+
+
+section {* Customization of terms *}
+
+subsection{* Constructors *}
+
+definition "Fold hs \<equiv> hfset_fld hs"
+
+lemma hfset_simps[simp]:
+"\<And>hs1 hs2. Fold hs1 = Fold hs2 \<longrightarrow> hs1 = hs2"
+unfolding Fold_def hfset.fld_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.fld_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.fld_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 *}
+
+lemma hfset_rec:
+"hfset_rec R (Fold hs) = R (map_fset <id, hfset_rec R> hs)"
+using hfset.rec unfolding Fold_def .
+
+(* The iterator has a simpler form: *)
+lemma hfset_iter:
+"hfset_iter R (Fold hs) = R (map_fset (hfset_iter R) hs)"
+using hfset.iter unfolding Fold_def .
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Codatatype/Examples/Infinite_Derivation_Trees/Gram_Lang.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,1366 @@
+(* Title: 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> coit id S"
+
+lemma deftr_simps[simp]:
+"root (deftr n) = n"
+"cont (deftr n) = image (id \<oplus> deftr) (S n)"
+using coit(1)[of id S n] coit(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> coit 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 coit(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 coit(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 coit(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> coit 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 coit(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 coit(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 coit(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/Codatatype/Examples/Infinite_Derivation_Trees/Parallel.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,143 @@
+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> coit 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 coit(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 coit(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/Codatatype/Examples/Infinite_Derivation_Trees/Prelim.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,66 @@
+(* Title: Gram_Tree.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Preliminaries
+*)
+
+
+theory Prelim
+imports "../../Codatatype/Codatatype"
+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/Codatatype/Examples/Infinite_Derivation_Trees/Tree.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,326 @@
+(* Title: Gram_Tree.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Trees with nonterminal internal nodes and terminal leafs.
+*)
+
+
+header {* Trees with nonterminal internal nodes and terminal leafs *}
+
+
+theory Tree
+imports Prelim
+begin
+
+typedecl N typedecl T
+
+bnf_codata Tree: 'Tree = "N \<times> (T + 'Tree) fset"
+
+
+section {* Sugar notations for Tree *}
+
+subsection{* Setup for map, set, pred *}
+
+(* These should be eventually inferred from compositionality *)
+
+lemma TreeBNF_map:
+"TreeBNF_map f (n,as) = (n, map_fset (id \<oplus> f) as)"
+unfolding TreeBNF_map_def id_apply
+sum_map_def by simp
+
+lemma TreeBNF_map':
+"TreeBNF_map f n_as = (fst n_as, map_fset (id \<oplus> f) (snd n_as))"
+using TreeBNF_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 TreeBNF_pred: "TreeBNF_pred \<phi> (n1,as1) (n2,as2) \<longleftrightarrow> n1 = n2 \<and> llift2 \<phi> as1 as2"
+unfolding llift2_def TreeBNF.pred_unfold
+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_fld (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.fld_unf 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.fld_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.fld_exhaust[of tr])
+ fix x assume "tr = Tree_fld 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.pred_coinduct[of \<phi> tr1 tr2] phi]) proof clarify
+ fix tr1 tr2 assume \<phi>: "\<phi> tr1 tr2"
+ show "TreeBNF_pred \<phi> (Tree_unf tr1) (Tree_unf tr2)"
+ apply(cases rule: Tree.fld_exhaust[of tr1], cases rule: Tree.fld_exhaust[of tr2])
+ apply (simp add: Tree.unf_fld)
+ apply(case_tac x, case_tac xa, simp)
+ unfolding TreeBNF_pred 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.unf_fld[simp]
+declare Tree.fld_unf[simp]
+
+lemma Tree_unf_NNode[simp]:
+"Tree_unf (NNode n as) = (n,as)"
+unfolding NNode_def Tree.unf_fld ..
+
+lemma Tree_unf_root_ccont:
+"Tree_unf tr = (root tr, ccont tr)"
+unfolding root_def ccont_def
+by (metis (lifting) NNode_asNNode Tree_unf_NNode)
+
+(* Coiteration *)
+definition TTree_coit ::
+"('b \<Rightarrow> N) \<Rightarrow> ('b \<Rightarrow> (T + 'b) fset) \<Rightarrow> 'b \<Rightarrow> Tree"
+where "TTree_coit rt ct \<equiv> Tree_coiter <rt,ct>"
+
+lemma Tree_coit_coit:
+"Tree_coiter s = TTree_coit (fst o s) (snd o s)"
+apply(rule ext)
+unfolding TTree_coit_def by simp
+
+theorem TTree_coit:
+"root (TTree_coit rt ct b) = rt b"
+"ccont (TTree_coit rt ct b) = map_fset (id \<oplus> TTree_coit rt ct) (ct b)"
+using Tree.coiter[of "<rt,ct>" b] unfolding Tree_coit_coit fst_convol snd_convol
+unfolding TreeBNF_map' fst_convol' snd_convol'
+unfolding Tree_unf_root_ccont by simp_all
+
+(* Corecursion, stronger than coitation *)
+definition TTree_corec ::
+"('b \<Rightarrow> N) \<Rightarrow> ('b \<Rightarrow> (T + (Tree + 'b)) fset) \<Rightarrow> 'b \<Rightarrow> Tree"
+where "TTree_corec rt ct \<equiv> Tree_corec <rt,ct>"
+
+lemma Tree_corec_corec:
+"Tree_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.corec[of "<rt,ct>" b] unfolding Tree_corec_corec fst_convol snd_convol
+unfolding TreeBNF_map' fst_convol' snd_convol'
+unfolding Tree_unf_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 "coit rt ct \<equiv> TTree_coit 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 coit:
+"root (coit rt ct b) = rt b"
+"finite (ct b) \<Longrightarrow> cont (coit rt ct b) = image (id \<oplus> coit rt ct) (ct b)"
+using TTree_coit[of rt "the_inv fset \<circ> ct" b] unfolding coit_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/Codatatype/Examples/Lambda_Term.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,259 @@
+(* Title: Codatatype_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 "../Codatatype"
+begin
+
+
+section {* Datatype definition *}
+
+bnf_data trm: 'trm = "'a + 'trm \<times> 'trm + 'a \<times> 'trm + ('a \<times> 'trm) fset \<times> 'trm"
+
+
+section {* Customization of terms *}
+
+subsection{* Set and map *}
+
+lemma trmBNF_set2_Lt: "trmBNF_set2 (Inr (Inr (Inr (xts, t)))) = snd ` (fset xts) \<union> {t}"
+unfolding trmBNF_set2_def sum_set_defs prod_set_defs collect_def[abs_def]
+by auto
+
+lemma trmBNF_set2_Var: "\<And>x. trmBNF_set2 (Inl x) = {}"
+and trmBNF_set2_App:
+"\<And>t1 t2. trmBNF_set2 (Inr (Inl t1t2)) = {fst t1t2, snd t1t2}"
+and trmBNF_set2_Lam:
+"\<And>x t. trmBNF_set2 (Inr (Inr (Inl (x, t)))) = {t}"
+unfolding trmBNF_set2_def sum_set_defs prod_set_defs collect_def[abs_def]
+by auto
+
+lemma trmBNF_map:
+"\<And> a1. trmBNF_map f1 f2 (Inl a1) = Inl (f1 a1)"
+"\<And> a2 b2. trmBNF_map f1 f2 (Inr (Inl (a2,b2))) = Inr (Inl (f2 a2, f2 b2))"
+"\<And> a1 a2. trmBNF_map f1 f2 (Inr (Inr (Inl (a1,a2)))) = Inr (Inr (Inl (f1 a1, f2 a2)))"
+"\<And> a1a2s a2.
+ trmBNF_map f1 f2 (Inr (Inr (Inr (a1a2s, a2)))) =
+ Inr (Inr (Inr (map_fset (\<lambda> (a1', a2'). (f1 a1', f2 a2')) a1a2s, f2 a2)))"
+unfolding trmBNF_map_def collect_def[abs_def] map_pair_def by auto
+
+
+subsection{* Constructors *}
+
+definition "Var x \<equiv> trm_fld (Inl x)"
+definition "App t1 t2 \<equiv> trm_fld (Inr (Inl (t1,t2)))"
+definition "Lam x t \<equiv> trm_fld (Inr (Inr (Inl (x,t))))"
+definition "Lt xts t \<equiv> trm_fld (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.fld_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.fld_exhaust[of t])
+ fix x assume "t = trm_fld 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.fld_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> trmBNF_set2 u \<Longrightarrow> phi t"
+ show "phi (trm_fld 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 trmBNF_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 trmBNF_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 trmBNF_set2_Lt xts_t fset_fset_member image_def by auto
+ qed
+ qed
+ qed
+qed
+
+
+subsection{* Recursion and iteration *}
+
+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_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.rec trmBNF_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.rec trmBNF_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.rec trmBNF_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.rec trmBNF_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 "trmiter var app lam lt \<equiv> trm_iter (sumJoinI4 var app lam lt)"
+
+lemma trmiter_Var[simp]:
+"trmiter var app lam lt (Var x) = var x"
+unfolding trmiter_def Var_def trm.iter trmBNF_map(1) by simp
+
+lemma trmiter_App[simp]:
+"trmiter var app lam lt (App t1 t2) =
+ app (trmiter var app lam lt t1) (trmiter var app lam lt t2)"
+unfolding trmiter_def App_def trm.iter trmBNF_map(2) by simp
+
+lemma trmiter_Lam[simp]:
+"trmiter var app lam lt (Lam x t) = lam x (trmiter var app lam lt t)"
+unfolding trmiter_def Lam_def trm.iter trmBNF_map(3) by simp
+
+lemma trmiter_Lt[simp]:
+"trmiter var app lam lt (Lt xts t) =
+ lt (map_fset (\<lambda> (x,t). (x,trmiter var app lam lt t)) xts) (trmiter var app lam lt t)"
+unfolding trmiter_def Lt_def trm.iter trmBNF_map(4) by simp
+
+
+subsection{* Example: The set of all variables varsOf and free variables fvarsOf of a term: *}
+
+definition "varsOf = trmiter
+(\<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 = trmiter
+(\<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/Codatatype/Examples/ListF.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,171 @@
+(* Title: Codatatype_Examples/ListF.thy
+ Author: Dmitriy Traytel, TU Muenchen
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Finite lists.
+*)
+
+header {* Finite Lists *}
+
+theory ListF
+imports "../Codatatype"
+begin
+
+bnf_data listF: 'list = "unit + 'a \<times> 'list"
+
+definition "NilF = listF_fld (Inl ())"
+definition "Conss a as \<equiv> listF_fld (Inr (a, as))"
+
+lemma listF_map_NilF[simp]: "listF_map f NilF = NilF"
+unfolding listF_map_def listFBNF_map_def NilF_def listF.iter by simp
+
+lemma listF_map_Conss[simp]:
+ "listF_map f (Conss x xs) = Conss (f x) (listF_map f xs)"
+unfolding listF_map_def listFBNF_map_def Conss_def listF.iter by simp
+
+lemma listF_set_NilF[simp]: "listF_set NilF = {}"
+unfolding listF_set_def NilF_def listF.iter listFBNF_set1_def listFBNF_set2_def
+ sum_set_defs listFBNF_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.iter listFBNF_set1_def listFBNF_set2_def
+ sum_set_defs prod_set_defs listFBNF_map_def collect_def[abs_def] by simp
+
+lemma iter_sum_case_NilF: "listF_iter (sum_case f g) NilF = f ()"
+unfolding NilF_def listF.iter listFBNF_map_def by simp
+
+
+lemma iter_sum_case_Conss:
+ "listF_iter (sum_case f g) (Conss y ys) = g (y, listF_iter (sum_case f g) ys)"
+unfolding Conss_def listF.iter listFBNF_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.fld_induct)
+ fix xs :: "unit + 'a \<times> 'a listF"
+ assume raw_IH: "\<And>a. a \<in> listFBNF_set2 xs \<Longrightarrow> P a"
+ show "P (listF_fld 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_fld xs = Conss y ys"
+ unfolding Conss_def listF.fld_inject by (blast intro: prod.exhaust)
+ hence "ys \<in> listFBNF_set2 xs"
+ unfolding listFBNF_set2_def Conss_def listF.fld_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.fld_inject)
+
+definition Singll ("[[_]]") where
+ [simp]: "Singll a \<equiv> Conss a NilF"
+
+definition appendd (infixr "@@" 65) where
+ "appendd \<equiv> listF_iter (sum_case (\<lambda> _. id) (\<lambda> (a,f) bs. Conss a (f bs)))"
+
+definition "lrev \<equiv> listF_iter (sum_case (\<lambda> _. NilF) (\<lambda> (b,bs). bs @@ [[b]]))"
+
+lemma lrev_NilF[simp]: "lrev NilF = NilF"
+unfolding lrev_def by (simp add: iter_sum_case_NilF)
+
+lemma lrev_Conss[simp]: "lrev (Conss y ys) = lrev ys @@ [[y]]"
+unfolding lrev_def by (simp add: iter_sum_case_Conss)
+
+lemma NilF_appendd[simp]: "NilF @@ ys = ys"
+unfolding appendd_def by (simp add: iter_sum_case_NilF)
+
+lemma Conss_append[simp]: "Conss x xs @@ ys = Conss x (xs @@ ys)"
+unfolding appendd_def by (simp add: iter_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/Codatatype/Examples/Misc_Codata.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,88 @@
+(* Title: Codatatype_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 "../Codatatype"
+begin
+
+ML {* quick_and_dirty := false *}
+
+ML {* PolyML.fullGC (); *}
+
+bnf_codata simple: 'a = "unit + unit + unit + unit"
+
+bnf_codata stream: 's = "'a \<times> 's"
+
+bnf_codata llist: 'llist = "unit + 'a \<times> 'llist"
+
+bnf_codata some_passive: 'a = "'a + 'b + 'c + 'd + 'e"
+
+(*
+ ('a, 'b1, 'b2) F1 = 'a * 'b1 + 'a * 'b2
+ ('a, 'b1, 'b2) F2 = unit + 'b1 * 'b2
+*)
+
+bnf_codata F1: 'b1 = "'a \<times> 'b1 + 'a \<times> 'b2"
+and F2: 'b2 = "unit + 'b1 * 'b2"
+
+bnf_codata EXPR: 'E = "'T + 'T \<times> 'E"
+and TERM: 'T = "'F + 'F \<times> 'T"
+and FACTOR: 'F = "'a + 'b + 'E"
+
+bnf_codata llambda:
+ 'trm = "string +
+ 'trm \<times> 'trm +
+ string \<times> 'trm +
+ (string \<times> 'trm) fset \<times> 'trm"
+
+bnf_codata par_llambda:
+ 'trm = "'a +
+ 'trm \<times> 'trm +
+ 'a \<times> 'trm +
+ ('a \<times> 'trm) fset \<times> 'trm"
+
+(*
+ 'a tree = Empty | Node of 'a * 'a forest ('b = unit + 'a * 'c)
+ 'a forest = Nil | Cons of 'a tree * 'a forest ('c = unit + 'b * 'c)
+*)
+
+bnf_codata tree: 'tree = "unit + 'a \<times> 'forest"
+and forest: 'forest = "unit + 'tree \<times> 'forest"
+
+bnf_codata CPS: 'a = "'b + 'b \<Rightarrow> 'a"
+
+bnf_codata fun_rhs: 'a = "'b1 \<Rightarrow> 'b2 \<Rightarrow> 'b3 \<Rightarrow> 'b4 \<Rightarrow> 'b5 \<Rightarrow> 'b6 \<Rightarrow> 'b7 \<Rightarrow> 'b8 \<Rightarrow> 'b9 \<Rightarrow> 'a"
+
+bnf_codata fun_rhs': 'a = "'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> 'a"
+
+bnf_codata some_killing: 'a = "'b \<Rightarrow> 'd \<Rightarrow> ('a + 'c)"
+and in_here: 'c = "'d \<times> 'b + 'e"
+
+bnf_codata some_killing': 'a = "'b \<Rightarrow> 'd \<Rightarrow> ('a + 'c)"
+and in_here': 'c = "'d + 'e"
+
+bnf_codata some_killing'': 'a = "'b \<Rightarrow> 'c"
+and in_here'': 'c = "'d \<times> 'b + 'e"
+
+bnf_codata less_killing: 'a = "'b \<Rightarrow> 'c"
+
+(* SLOW, MEMORY-HUNGRY
+bnf_codata K1': 'K1 = "'K2 + 'a list"
+and K2': 'K2 = "'K3 + 'c fset"
+and K3': 'K3 = "'K3 + 'K4 + 'K4 \<times> 'K5"
+and K4': 'K4 = "'K5 + 'a list list list"
+and K5': 'K5 = "'K6"
+and K6': 'K6 = "'K7"
+and K7': 'K7 = "'K8"
+and K8': 'K8 = "'K1 list"
+*)
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Codatatype/Examples/Misc_Data.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,161 @@
+(* Title: Codatatype_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 "../Codatatype"
+begin
+
+ML {* quick_and_dirty := false *}
+
+ML {* PolyML.fullGC (); *}
+
+bnf_data simple: 'a = "unit + unit + unit + unit"
+
+bnf_data mylist: 'list = "unit + 'a \<times> 'list"
+
+bnf_data some_passive: 'a = "'a + 'b + 'c + 'd + 'e"
+
+bnf_data lambda:
+ 'trm = "string +
+ 'trm \<times> 'trm +
+ string \<times> 'trm +
+ (string \<times> 'trm) fset \<times> 'trm"
+
+bnf_data par_lambda:
+ 'trm = "'a +
+ 'trm \<times> 'trm +
+ 'a \<times> 'trm +
+ ('a \<times> 'trm) fset \<times> 'trm"
+
+(*
+ ('a, 'b1, 'b2) F1 = 'a * 'b1 + 'a * 'b2
+ ('a, 'b1, 'b2) F2 = unit + 'b1 * 'b2
+*)
+
+bnf_data F1: 'b1 = "'a \<times> 'b1 + 'a \<times> 'b2"
+and F2: 'b2 = "unit + 'b1 * 'b2"
+
+(*
+ 'a tree = Empty | Node of 'a * 'a forest ('b = unit + 'a * 'c)
+ 'a forest = Nil | Cons of 'a tree * 'a forest ('c = unit + 'b * 'c)
+*)
+
+bnf_data tree: 'tree = "unit + 'a \<times> 'forest"
+and forest: 'forest = "unit + 'tree \<times> 'forest"
+
+(*
+ 'a tree = Empty | Node of 'a branch * 'a branch ('b = unit + 'c * 'c)
+' a branch = Branch of 'a * 'a tree ('c = 'a * 'b)
+*)
+
+bnf_data tree': 'tree = "unit + 'branch \<times> 'branch"
+and branch: 'branch = "'a \<times> 'tree"
+
+(*
+ exp = Sum of term * exp ('c = 'd + 'd * 'c)
+ term = Prod of factor * term ('d = 'e + 'e * 'd)
+ factor = C 'a | V 'b | Paren exp ('e = 'a + 'b + 'c)
+*)
+
+bnf_data EXPR: 'E = "'T + 'T \<times> 'E"
+and TERM: 'T = "'F + 'F \<times> 'T"
+and FACTOR: 'F = "'a + 'b + 'E"
+
+bnf_data some_killing: 'a = "'b \<Rightarrow> 'd \<Rightarrow> ('a + 'c)"
+and in_here: 'c = "'d \<times> 'b + 'e"
+
+bnf_data nofail1: 'a = "'a \<times> 'b + 'b"
+bnf_data nofail2: 'a = "('a \<times> 'b \<times> 'a \<times> 'b) list"
+bnf_data nofail3: 'a = "'b \<times> ('a \<times> 'b \<times> 'a \<times> 'b) fset"
+bnf_data nofail4: 'a = "('a \<times> ('a \<times> 'b \<times> 'a \<times> 'b) fset) list"
+
+(*
+bnf_data fail: 'a = "'a \<times> 'b \<times> 'a \<times> 'b list"
+bnf_data fail: 'a = "'a \<times> 'b \<times> 'a \<times> 'b"
+bnf_data fail: 'a = "'a \<times> 'b + 'a"
+bnf_data fail: 'a = "'a \<times> 'b"
+*)
+
+bnf_data L1: 'L1 = "'L2 list"
+and L2: 'L2 = "'L1 fset + 'L2"
+
+bnf_data K1: 'K1 = "'K2"
+and K2: 'K2 = "'K3"
+and K3: 'K3 = "'K1 list"
+
+bnf_data t1: 't1 = "'t3 + 't2"
+and t2: 't2 = "'t1"
+and t3: 't3 = "unit"
+
+bnf_data t1': 't1 = "'t2 + 't3"
+and t2': 't2 = "'t1"
+and t3': 't3 = "unit"
+
+(*
+bnf_data fail1: 'L1 = "'L2"
+and fail2: 'L2 = "'L3"
+and fail2: 'L3 = "'L1"
+
+bnf_data fail1: 'L1 = "'L2 list \<times> 'L2"
+and fail2: 'L2 = "'L2 fset \<times> 'L3"
+and fail2: 'L3 = "'L1"
+
+bnf_data fail1: 'L1 = "'L2 list \<times> 'L2"
+and fail2: 'L2 = "'L1 fset \<times> 'L1"
+*)
+(* SLOW
+bnf_data K1': 'K1 = "'K2 + 'a list"
+and K2': 'K2 = "'K3 + 'c fset"
+and K3': 'K3 = "'K3 + 'K4 + 'K4 \<times> 'K5"
+and K4': 'K4 = "'K5 + 'a list list list"
+and K5': 'K5 = "'K6"
+and K6': 'K6 = "'K7"
+and K7': 'K7 = "'K8"
+and K8': 'K8 = "'K1 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:
+bnf_data t1: 't1 = "'t2 * 't3 + 't2 * 't4"
+and t2: 't2 = "unit"
+and t3: 't3 = 't4
+and t4: 't4 = 't1
+*)
+
+bnf_data k1: 'k1 = "'k2 * 'k3 + 'k2 * 'k4"
+and k2: 'k2 = unit
+and k3: 'k3 = 'k4
+and k4: 'k4 = unit
+
+bnf_data tt1: 'tt1 = "'tt3 * 'tt2 + 'tt2 * 'tt4"
+and tt2: 'tt2 = unit
+and tt3: 'tt3 = 'tt1
+and tt4: 'tt4 = unit
+(* SLOW
+bnf_data s1: 's1 = "'s2 * 's3 * 's4 + 's3 + 's2 * 's6 + 's4 * 's2 + 's2 * 's2"
+and s2: 's2 = "'s7 * 's5 + 's5 * 's4 * 's6"
+and s3: 's3 = "'s1 * 's7 * 's2 + 's3 * 's3 + 's4 * 's5"
+and s4: 's4 = 's5
+and s5: 's5 = unit
+and s6: 's6 = "'s6 + 's1 * 's2 + 's6"
+and s7: 's7 = "'s8 + 's5"
+and s8: 's8 = nat
+*)
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Codatatype/Examples/Process.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,742 @@
+(* Title: Codatatype_Examples/Process.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Processes.
+*)
+
+header {* Processes *}
+
+theory Process
+imports "../Codatatype"
+begin
+
+bnf_codata process: 'p = "'a * 'p + 'p * 'p"
+(* codatatype
+ 'a process = Action (prefOf :: 'a) (contOf :: 'a process) |
+ Choice (ch1Of :: 'a process) (ch2Of :: 'a process)
+*)
+
+(* Read: prefix of, continuation of, choice 1 of, choice 2 of *)
+
+section {* Customization *}
+
+subsection{* Setup for map, set, pred *}
+
+(* These should be eventually inferred from compositionality *)
+
+lemma processBNF_map[simp]:
+"processBNF_map fa fp (Inl (a ,p)) = Inl (fa a, fp p)"
+"processBNF_map fa fp (Inr (p1, p2)) = Inr (fp p1, fp p2)"
+unfolding processBNF_map_def by auto
+
+lemma processBNF_pred[simp]:
+"processBNF_pred (op =) \<phi> (Inl (a,p)) (Inl (a',p')) \<longleftrightarrow> a = a' \<and> \<phi> p p'"
+"processBNF_pred (op =) \<phi> (Inr (p,q)) (Inr (p',q')) \<longleftrightarrow> \<phi> p p' \<and> \<phi> q q'"
+"\<not> processBNF_pred (op =) \<phi> (Inl ap) (Inr q1q2)"
+"\<not> processBNF_pred (op =) \<phi> (Inr q1q2) (Inl ap)"
+by (auto simp: diag_def processBNF.pred_unfold)
+
+
+subsection{* Constructors *}
+
+definition Action :: "'a \<Rightarrow> 'a process \<Rightarrow> 'a process"
+where "Action a p \<equiv> process_fld (Inl (a,p))"
+
+definition Choice :: "'a process \<Rightarrow> 'a process \<Rightarrow> 'a process"
+where "Choice p1 p2 \<equiv> process_fld (Inr (p1,p2))"
+
+lemmas ctor_defs = Action_def Choice_def
+
+
+subsection {* Discriminators *}
+
+(* One discriminator for each constructor. By the constructor exhaustiveness,
+one of them is of course redundant, so for n constructors we only need n-1
+discriminators. However, keeping n discriminators seems more uniform. *)
+
+definition isAction :: "'a process \<Rightarrow> bool"
+where "isAction q \<equiv> \<exists> a p. q = Action a p"
+
+definition isChoice :: "'a process \<Rightarrow> bool"
+where "isChoice q \<equiv> \<exists> p1 p2. q = Choice p1 p2"
+
+lemmas discr_defs = isAction_def isChoice_def
+
+
+subsection {* Pre-selectors *}
+
+(* These are mere auxiliaries *)
+
+definition "asAction q \<equiv> SOME ap. q = Action (fst ap) (snd ap)"
+definition "asChoice q \<equiv> SOME p1p2. q = Choice (fst p1p2) (snd p1p2)"
+lemmas pre_sel_defs = asAction_def asChoice_def
+
+
+subsection {* Selectors *}
+
+(* One for each pair (constructor, constructor argument) *)
+
+(* For Action: *)
+definition prefOf :: "'a process \<Rightarrow> 'a" where "prefOf q = fst (asAction q)"
+definition contOf :: "'a process \<Rightarrow> 'a process" where "contOf q = snd (asAction q)"
+
+(* For Choice: *)
+definition ch1Of :: "'a process \<Rightarrow> 'a process" where "ch1Of q = fst (asChoice q)"
+definition ch2Of :: "'a process \<Rightarrow> 'a process" where "ch2Of q = snd (asChoice q)"
+
+lemmas sel_defs = prefOf_def contOf_def ch1Of_def ch2Of_def
+
+
+subsection {* Basic properties *}
+
+(* Selectors versus discriminators *)
+lemma isAction_asAction:
+"isAction q \<longleftrightarrow> q = Action (fst (asAction q)) (snd (asAction q))"
+(is "?L \<longleftrightarrow> ?R")
+proof
+ assume ?L
+ then obtain a p where q: "q = Action a p" unfolding isAction_def by auto
+ show ?R unfolding asAction_def q by (rule someI[of _ "(a,p)"]) simp
+qed(unfold isAction_def, auto)
+
+theorem isAction_prefOf_contOf:
+"isAction q \<longleftrightarrow> q = Action (prefOf q) (contOf q)"
+using isAction_asAction unfolding prefOf_def contOf_def .
+
+lemma isChoice_asChoice:
+"isChoice q \<longleftrightarrow> q = Choice (fst (asChoice q)) (snd (asChoice q))"
+(is "?L \<longleftrightarrow> ?R")
+proof
+ assume ?L
+ then obtain p1 p2 where q: "q = Choice p1 p2" unfolding isChoice_def by auto
+ show ?R unfolding asChoice_def q by (rule someI[of _ "(p1,p2)"]) simp
+qed(unfold isChoice_def, auto)
+
+theorem isChoice_ch1Of_ch2Of:
+"isChoice q \<longleftrightarrow> q = Choice (ch1Of q) (ch2Of q)"
+using isChoice_asChoice unfolding ch1Of_def ch2Of_def .
+
+(* Constructors *)
+theorem process_simps[simp]:
+"Action a p = Action a' p' \<longleftrightarrow> a = a' \<and> p = p'"
+"Choice p1 p2 = Choice p1' p2' \<longleftrightarrow> p1 = p1' \<and> p2 = p2'"
+(* *)
+"Action a p \<noteq> Choice p1 p2"
+"Choice p1 p2 \<noteq> Action a p"
+unfolding ctor_defs process.fld_inject by auto
+
+theorem process_cases[elim, case_names Action Choice]:
+assumes Action: "\<And> a p. q = Action a p \<Longrightarrow> phi"
+and Choice: "\<And> p1 p2. q = Choice p1 p2 \<Longrightarrow> phi"
+shows phi
+proof(cases rule: process.fld_exhaust[of q])
+ fix x assume "q = process_fld x"
+ thus ?thesis
+ apply(cases x)
+ apply(case_tac a) using Action unfolding ctor_defs apply blast
+ apply(case_tac b) using Choice unfolding ctor_defs apply blast
+ done
+qed
+
+(* Constructors versus discriminators *)
+theorem isAction_isChoice:
+"isAction p \<or> isChoice p"
+unfolding isAction_def isChoice_def by (cases rule: process_cases) auto
+
+theorem isAction_Action[simp]: "isAction (Action a p)"
+unfolding isAction_def by auto
+
+theorem isAction_Choice[simp]: "\<not> isAction (Choice p1 p2)"
+unfolding isAction_def by auto
+
+theorem isChoice_Choice[simp]: "isChoice (Choice p1 p2)"
+unfolding isChoice_def by auto
+
+theorem isChoice_Action[simp]: "\<not> isChoice (Action a p)"
+unfolding isChoice_def by auto
+
+theorem not_isAction_isChoice: "\<not> (isAction p \<and> isChoice p)"
+by (cases rule: process_cases[of p]) auto
+
+(* Constructors versus selectors *)
+theorem dest_ctor[simp]:
+"prefOf (Action a p) = a"
+"contOf (Action a p) = p"
+"ch1Of (Choice p1 p2) = p1"
+"ch2Of (Choice p1 p2) = p2"
+using isAction_Action[of a p]
+ isChoice_Choice[of p1 p2]
+unfolding isAction_prefOf_contOf
+ isChoice_ch1Of_ch2Of by auto
+
+theorem ctor_dtor[simp]:
+"\<And> p. isAction p \<Longrightarrow> Action (prefOf p) (contOf p) = p"
+"\<And> p. isChoice p \<Longrightarrow> Choice (ch1Of p) (ch2Of p) = p"
+unfolding isAction_def isChoice_def by 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.pred_coinduct, of \<phi>, OF _ phi], clarify)
+ fix p p' assume \<phi>: "\<phi> p p'"
+ show "processBNF_pred (op =) \<phi> (process_unf p) (process_unf p')"
+ proof(cases rule: process_cases[of p])
+ case (Action a q) note p = Action
+ hence "isAction p'" using iss[OF \<phi>] by (cases rule: process_cases[of p'], auto)
+ then obtain a' q' where p': "p' = Action a' q'" by (cases rule: process_cases[of p'], auto)
+ have 0: "a = a' \<and> \<phi> q q'" using Act[OF \<phi>[unfolded p p']] .
+ have unf: "process_unf p = Inl (a,q)" "process_unf p' = Inl (a',q')"
+ unfolding p p' Action_def process.unf_fld by simp_all
+ show ?thesis using 0 unfolding unf by simp
+ next
+ case (Choice p1 p2) note p = Choice
+ hence "isChoice p'" using iss[OF \<phi>] by (cases rule: process_cases[of p'], auto)
+ then obtain p1' p2' where p': "p' = Choice p1' p2'"
+ by (cases rule: process_cases[of p'], auto)
+ have 0: "\<phi> p1 p1' \<and> \<phi> p2 p2'" using Ch[OF \<phi>[unfolded p p']] .
+ have unf: "process_unf p = Inr (p1,p2)" "process_unf p' = Inr (p1',p2')"
+ unfolding p p' Choice_def process.unf_fld by simp_all
+ show ?thesis using 0 unfolding unf by simp
+ qed
+qed
+
+(* Stronger coinduction, up to equality: *)
+theorem process_coind_upto[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.pred_coinduct_upto, of \<phi>, OF _ phi], clarify)
+ fix p p' assume \<phi>: "\<phi> p p'"
+ show "processBNF_pred (op =) (\<lambda>a b. \<phi> a b \<or> a = b) (process_unf p) (process_unf p')"
+ proof(cases rule: process_cases[of p])
+ case (Action a q) note p = Action
+ hence "isAction p'" using iss[OF \<phi>] by (cases rule: process_cases[of p'], auto)
+ then obtain a' q' where p': "p' = Action a' q'" by (cases rule: process_cases[of p'], auto)
+ have 0: "a = a' \<and> (\<phi> q q' \<or> q = q')" using Act[OF \<phi>[unfolded p p']] .
+ have unf: "process_unf p = Inl (a,q)" "process_unf p' = Inl (a',q')"
+ unfolding p p' Action_def process.unf_fld by simp_all
+ show ?thesis using 0 unfolding unf by simp
+ next
+ case (Choice p1 p2) note p = Choice
+ hence "isChoice p'" using iss[OF \<phi>] by (cases rule: process_cases[of p'], auto)
+ then obtain p1' p2' where p': "p' = Choice p1' p2'"
+ by (cases rule: process_cases[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 unf: "process_unf p = Inr (p1,p2)" "process_unf p' = Inr (p1',p2')"
+ unfolding p p' Choice_def process.unf_fld by simp_all
+ show ?thesis using 0 unfolding unf by simp
+ qed
+qed
+
+
+subsection {* Coiteration and corecursion *}
+
+(* Preliminaries: *)
+definition
+"join22 isA pr co isC c1 c2 \<equiv>
+ \<lambda> P. if isA P then Inl (pr P, co P)
+ else if isC P then Inr (c1 P, c2 P)
+ else undefined"
+
+declare process.unf_fld[simp]
+declare process.fld_unf[simp]
+
+lemma unf_Action[simp]:
+"process_unf (Action a p) = Inl (a,p)"
+unfolding Action_def process.unf_fld ..
+
+lemma unf_Choice[simp]:
+"process_unf (Choice p1 p2) = Inr (p1,p2)"
+unfolding Choice_def process.unf_fld ..
+
+lemma isAction_unf:
+assumes "isAction p"
+shows "process_unf p = Inl (prefOf p, contOf p)"
+using assms unfolding isAction_def by auto
+
+lemma isChoice_unf:
+assumes "isChoice p"
+shows "process_unf p = Inr (ch1Of p, ch2Of p)"
+using assms unfolding isChoice_def by auto
+
+lemma unf_join22:
+"process_unf p = join22 isAction prefOf contOf isChoice ch1Of ch2Of p"
+unfolding join22_def
+using isAction_unf isChoice_unf not_isAction_isChoice isAction_isChoice by auto
+
+lemma isA_join22:
+assumes "isA P"
+shows "join22 isA pr co isC c1 c2 P = Inl (pr P, co P)"
+using assms unfolding join22_def by auto
+
+lemma isC_join22:
+assumes "\<not> isA P" and "isC P"
+shows "join22 isA pr co isC c1 c2 P = Inr (c1 P, c2 P)"
+using assms unfolding join22_def by auto
+
+(* Coiteration *)
+definition pcoiter ::
+"('b \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> 'b)
+ \<Rightarrow>
+ ('b \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'b)
+ \<Rightarrow>
+ 'b \<Rightarrow> 'a process"
+where "pcoiter isA pr co isC c1 c2 \<equiv> process_coiter (join22 isA pr co isC c1 c2)"
+
+lemma unf_prefOf:
+assumes "process_unf q = Inl (a,p)"
+shows "prefOf q = a"
+using assms by (cases rule: process_cases[of q]) auto
+
+lemma unf_contOf:
+assumes "process_unf q = Inl (a,p)"
+shows "contOf q = p"
+using assms by (cases rule: process_cases[of q]) auto
+
+lemma unf_ch1Of:
+assumes "process_unf q = Inr (p1,p2)"
+shows "ch1Of q = p1"
+using assms by (cases rule: process_cases[of q]) auto
+
+lemma unf_ch2Of:
+assumes "process_unf q = Inr (p1,p2)"
+shows "ch2Of q = p2"
+using assms by (cases rule: process_cases[of q]) auto
+
+theorem pcoiter:
+"\<And>P. isA P \<Longrightarrow>
+ pcoiter isA pr co isC c1 c2 P =
+ Action (pr P)
+ (pcoiter isA pr co isC c1 c2 (co P))"
+"\<And>P. \<lbrakk>\<not> isA P; isC P\<rbrakk> \<Longrightarrow>
+ pcoiter isA pr co isC c1 c2 P =
+ Choice (pcoiter isA pr co isC c1 c2 (c1 P))
+ (pcoiter isA pr co isC c1 c2 (c2 P))"
+proof-
+ fix P
+ let ?f = "pcoiter isA pr co isC c1 c2" let ?s = "join22 isA pr co isC c1 c2"
+ assume isA: "isA P"
+ have unf: "process_unf (process_coiter ?s P) = Inl (pr P, ?f (co P))"
+ using process.coiter[of ?s P]
+ unfolding isA_join22[of isA P "pr" co isC c1 c2, OF isA]
+ processBNF_map id_apply pcoiter_def .
+ thus "?f P = Action (pr P) (?f (co P))"
+ unfolding pcoiter_def Action_def using process.fld_unf by metis
+next
+ fix P
+ let ?f = "pcoiter isA pr co isC c1 c2" let ?s = "join22 isA pr co isC c1 c2"
+ assume isA: "\<not> isA P" and isC: "isC P"
+ have unf: "process_unf (process_coiter ?s P) = Inr (?f (c1 P), ?f (c2 P))"
+ using process.coiter[of ?s P]
+ unfolding isC_join22[of isA P isC "pr" co c1 c2, OF isA isC]
+ processBNF_map id_apply pcoiter_def .
+ thus "?f P = Choice (?f (c1 P)) (?f (c2 P))"
+ unfolding pcoiter_def Choice_def using process.fld_unf by metis
+qed
+
+(* Corecursion, more general than coiteration (often unnecessarily more general): *)
+definition pcorec ::
+"('b \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> 'a process + 'b)
+ \<Rightarrow>
+ ('b \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'a process + 'b) \<Rightarrow> ('b \<Rightarrow> 'a process + 'b)
+ \<Rightarrow>
+ 'b \<Rightarrow> 'a process"
+where
+"pcorec isA pr co isC c1 c2 \<equiv> process_corec (join22 isA pr co isC c1 c2)"
+
+theorem pcorec_Action:
+assumes isA: "isA P"
+shows
+"case co P of
+ Inl p \<Rightarrow> pcorec isA pr co isC c1 c2 P = Action (pr P) p
+ |Inr Q \<Rightarrow> pcorec isA pr co isC c1 c2 P =
+ Action (pr P)
+ (pcorec isA pr co isC c1 c2 Q)"
+proof-
+ let ?f = "pcorec isA pr co isC c1 c2" let ?s = "join22 isA pr co isC c1 c2"
+ show ?thesis
+ proof(cases "co P")
+ case (Inl p)
+ have "process_unf (process_corec ?s P) = Inl (pr P, p)"
+ using process.corec[of ?s P]
+ unfolding isA_join22[of isA P "pr" co isC c1 c2, OF isA]
+ processBNF_map id_apply pcorec_def Inl by simp
+ thus ?thesis unfolding Inl pcorec_def Action_def using process.fld_unf by (simp, metis)
+ next
+ case (Inr Q)
+ have "process_unf (process_corec ?s P) = Inl (pr P, ?f Q)"
+ using process.corec[of ?s P]
+ unfolding isA_join22[of isA P "pr" co isC c1 c2, OF isA]
+ processBNF_map id_apply pcorec_def Inr by simp
+ thus ?thesis unfolding Inr pcorec_def Action_def using process.fld_unf by (simp, metis)
+ qed
+qed
+
+theorem pcorec_Choice:
+assumes isA: "\<not> isA P" and isC: "isC P"
+shows
+"case (c1 P,c2 P) of
+ (Inl p1, Inl p2) \<Rightarrow> pcorec isA pr co isC c1 c2 P =
+ Choice p1 p2
+ |(Inl p1, Inr Q2) \<Rightarrow> pcorec isA pr co isC c1 c2 P =
+ Choice p1
+ (pcorec isA pr co isC c1 c2 Q2)
+ |(Inr Q1, Inl p2) \<Rightarrow> pcorec isA pr co isC c1 c2 P =
+ Choice (pcorec isA pr co isC c1 c2 Q1)
+ p2
+ |(Inr Q1, Inr Q2) \<Rightarrow> pcorec isA pr co isC c1 c2 P =
+ Choice (pcorec isA pr co isC c1 c2 Q1)
+ (pcorec isA pr co isC c1 c2 Q2)"
+proof-
+ let ?f = "pcorec isA pr co isC c1 c2" let ?s = "join22 isA pr co isC c1 c2"
+ show ?thesis
+ proof(cases "c1 P")
+ case (Inl p1) note c1 = Inl
+ show ?thesis
+ proof(cases "c2 P")
+ case (Inl p2) note c2 = Inl
+ have "process_unf (process_corec ?s P) = Inr (p1, p2)"
+ using process.corec[of ?s P]
+ unfolding isC_join22[of isA P isC "pr" co c1 c2, OF isA isC]
+ processBNF_map id_apply pcorec_def c1 c2 by simp
+ thus ?thesis unfolding c1 c2 pcorec_def Choice_def using process.fld_unf by (simp, metis)
+ next
+ case (Inr Q2) note c2 = Inr
+ have "process_unf (process_corec ?s P) = Inr (p1, ?f Q2)"
+ using process.corec[of ?s P]
+ unfolding isC_join22[of isA P isC "pr" co c1 c2, OF isA isC]
+ processBNF_map id_apply pcorec_def c1 c2 by simp
+ thus ?thesis unfolding c1 c2 pcorec_def Choice_def using process.fld_unf by (simp, metis)
+ qed
+ next
+ case (Inr Q1) note c1 = Inr
+ show ?thesis
+ proof(cases "c2 P")
+ case (Inl p2) note c2 = Inl
+ have "process_unf (process_corec ?s P) = Inr (?f Q1, p2)"
+ using process.corec[of ?s P]
+ unfolding isC_join22[of isA P isC "pr" co c1 c2, OF isA isC]
+ processBNF_map id_apply pcorec_def c1 c2 by simp
+ thus ?thesis unfolding c1 c2 pcorec_def Choice_def using process.fld_unf by (simp, metis)
+ next
+ case (Inr Q2) note c2 = Inr
+ have "process_unf (process_corec ?s P) = Inr (?f Q1, ?f Q2)"
+ using process.corec[of ?s P]
+ unfolding isC_join22[of isA P isC "pr" co c1 c2, OF isA isC]
+ processBNF_map id_apply pcorec_def c1 c2 by simp
+ thus ?thesis unfolding c1 c2 pcorec_def Choice_def using process.fld_unf by (simp, metis)
+ qed
+ qed
+qed
+
+theorems pcorec = pcorec_Action pcorec_Choice
+
+
+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>
+ pcoiter
+ (\<lambda> P. True)
+ (\<lambda> P. ''a'')
+ (\<lambda> P. P)
+ undefined
+ undefined
+ undefined
+ ()"
+
+lemma BX: "BX = Action ''a'' BX"
+unfolding BX_def
+using pcoiter(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 "isC \<equiv> \<lambda> K. case K of x \<Rightarrow> True |y \<Rightarrow> False |ax \<Rightarrow> False"
+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 = isC_def c1_def c2_def
+
+definition "F \<equiv> pcoiter isA pr co isC 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 pcoiter(2)[of isA x isC "pr" co c1 c2]
+ pcoiter(1)[of isA y "pr" co isC c1 c2]
+ pcoiter(1)[of isA ax "pr" co isC 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 isCH where
+"isCH sys (VAR X) =
+ (case sys X of CH T1 T2 \<Rightarrow> True |PROC p \<Rightarrow> isChoice p |_ \<Rightarrow> False)"
+|
+"isCH sys (PROC p) = isChoice p"
+|
+"isCH sys (ACT a T) = False"
+|
+"isCH sys (CH T1 T2) = True"
+
+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"
+
+lemma guarded_isACT_isCH:
+assumes g: "guarded sys"
+shows "isACT sys T \<or> isCH sys T"
+proof(induct T)
+ case (VAR X)
+ thus ?case
+ using g isAction_isChoice unfolding guarded_def by (cases "sys X", auto)
+qed(insert isAction_isChoice assms, unfold guarded_def, auto)
+
+definition
+"solution sys \<equiv>
+ pcoiter
+ (isACT sys)
+ (PREF sys)
+ (CONT sys)
+ (isCH 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 pcoiter(1)[of "isACT sys" T "PREF sys" "CONT sys"
+ "isCH sys" "CH1 sys" "CH2 sys"] assms by simp
+
+lemma solution_Choice:
+assumes "\<not> isACT sys T" "isCH sys T"
+shows "solution sys T = Choice (solution sys (CH1 sys T)) (solution sys (CH2 sys T))"
+unfolding solution_def
+using pcoiter(2)[of "isACT sys" T "isCH sys" "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 isCH_VAR:
+assumes g: "guarded sys"
+shows "isCH sys (VAR X) \<longleftrightarrow> isCH 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
+ proof(cases "isCH sys (VAR X)")
+ case True
+ hence T: "isCH sys (sys X)" unfolding isCH_VAR[OF g] .
+ show ?thesis
+ unfolding solution_Choice[OF TT T] using solution_Choice[of sys "VAR X"] FFalse True g
+ unfolding guarded_def by (cases "sys X", auto)
+ next
+ case False
+ hence False using FFalse guarded_isACT_isCH[OF g, of "VAR X"] by simp
+ thus ?thesis by simp
+ qed
+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: "isCH sys (PROC p)" and p: "\<not> isAction p"
+ using not_isAction_isChoice by auto
+ hence 1: "\<not> isACT sys (PROC p)" by simp
+ show ?thesis using 0 iss not_isAction_isChoice
+ unfolding solution_Choice[OF 1 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: "isCH sys (PROC (Choice p q))" by simp
+ hence 1: "\<not> isACT sys (PROC (Choice p q))" using not_isAction_isChoice by auto
+ show ?case using Choice unfolding solution_Choice[OF 1 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) isCH.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/Codatatype/Examples/Stream.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,159 @@
+(* Title: Codatatype_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
+
+bnf_codata stream: 's = "'a \<times> 's"
+
+(* selectors for streams *)
+definition "hdd as \<equiv> fst (stream_unf as)"
+definition "tll as \<equiv> snd (stream_unf as)"
+
+lemma coiter_pair_fun_hdd[simp]: "hdd (stream_coiter (f \<odot> g) t) = f t"
+unfolding hdd_def pair_fun_def stream.coiter by simp
+
+lemma coiter_pair_fun_tll[simp]: "tll (stream_coiter (f \<odot> g) t) =
+ stream_coiter (f \<odot> g) (g t)"
+unfolding tll_def pair_fun_def stream.coiter by simp
+
+(* infinite trees: *)
+coinductive infiniteTr where
+"\<lbrakk>tr' \<in> listF_set (sub tr); infiniteTr tr'\<rbrakk> \<Longrightarrow> infiniteTr tr"
+
+lemma infiniteTr_coind_upto[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_coiter
+ (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_coind_upto[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_coind_upto[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_coiter (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 streamBNF_pred[simp]: "streamBNF_pred \<phi>1 \<phi>2 a b = (\<phi>1 (fst a) (fst b) \<and> \<phi>2 (snd a) (snd b))"
+by (auto simp: streamBNF.pred_unfold)
+
+lemmas stream_coind = mp[OF stream.pred_coinduct, unfolded streamBNF_pred[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_coiter ((%(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_coiter ((%x. 1) \<odot> id) ()"
+definition twos :: "nat stream" where [simp]: "twos = stream_coiter ((%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 phi="%x1 x2. \<exists>x. x1 = ones \<oplus> ones \<and> x2 = twos"])
+ auto
+
+lemma "n \<cdot> twos = ns (2 * n)"
+by (intro stream_coind[where phi="%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 phi="%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 phi="%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 phi="%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 phi="%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/Codatatype/Examples/TreeFI.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,81 @@
+(* Title: Codatatype_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
+
+bnf_codata treeFI: 'tree = "'a \<times> 'tree listF"
+
+lemma treeFIBNF_listF_set[simp]: "treeFIBNF_set2 (i, xs) = listF_set xs"
+unfolding treeFIBNF_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_unf tr)"
+definition "sub tr \<equiv> snd (treeFI_unf tr)"
+
+lemma unf[simp]: "treeFI_unf 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 coiter_pair_fun_lab: "lab (treeFI_coiter (f \<odot> g) t) = f t"
+unfolding lab_def pair_fun_def treeFI.coiter treeFIBNF_map_def by simp
+
+lemma coiter_pair_fun_sub: "sub (treeFI_coiter (f \<odot> g) t) = listF_map (treeFI_coiter (f \<odot> g)) (g t)"
+unfolding sub_def pair_fun_def treeFI.coiter treeFIBNF_map_def by simp
+
+(* Tree reverse:*)
+definition "trev \<equiv> treeFI_coiter (lab \<odot> lrev o sub)"
+
+lemma trev_simps1[simp]: "lab (trev t) = lab t"
+unfolding trev_def by (simp add: coiter_pair_fun_lab)
+
+lemma trev_simps2[simp]: "sub (trev t) = listF_map trev (lrev (sub t))"
+unfolding trev_def by (simp add: coiter_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.unf_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 "treeFIBNF_map id fst ?z = treeFI_unf a" "treeFIBNF_map id snd ?z = treeFI_unf b"
+ unfolding treeFIBNF_map_def by auto
+ moreover have "\<forall>(x, y) \<in> treeFIBNF_set2 ?z. phi x y"
+ proof safe
+ fix z1 z2
+ assume "(z1, z2) \<in> treeFIBNF_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.
+ (treeFIBNF_map id fst z = treeFI_unf a \<and>
+ treeFIBNF_map id snd z = treeFI_unf b) \<and>
+ (\<forall>x y. (x, y) \<in> treeFIBNF_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/Codatatype/Examples/TreeFsetI.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,56 @@
+(* Title: Codatatype_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 "../Codatatype"
+begin
+
+definition pair_fun (infixr "\<odot>" 50) where
+ "f \<odot> g \<equiv> \<lambda>x. (f x, g x)"
+
+bnf_codata treeFsetI: 't = "'a \<times> 't fset"
+
+(* selectors for trees *)
+definition "lab t \<equiv> fst (treeFsetI_unf t)"
+definition "sub t \<equiv> snd (treeFsetI_unf t)"
+
+lemma unf[simp]: "treeFsetI_unf t = (lab t, sub t)"
+unfolding lab_def sub_def by simp
+
+lemma coiter_pair_fun_lab: "lab (treeFsetI_coiter (f \<odot> g) t) = f t"
+unfolding lab_def pair_fun_def treeFsetI.coiter treeFsetIBNF_map_def by simp
+
+lemma coiter_pair_fun_sub: "sub (treeFsetI_coiter (f \<odot> g) t) = map_fset (treeFsetI_coiter (f \<odot> g)) (g t)"
+unfolding sub_def pair_fun_def treeFsetI.coiter treeFsetIBNF_map_def by simp
+
+(* tree map (contrived example): *)
+definition "tmap f \<equiv> treeFsetI_coiter (f o lab \<odot> sub)"
+
+lemma tmap_simps1[simp]: "lab (tmap f t) = f (lab t)"
+unfolding tmap_def by (simp add: coiter_pair_fun_lab)
+
+lemma trev_simps2[simp]: "sub (tmap f t) = map_fset (tmap f) (sub t)"
+unfolding tmap_def by (simp add: coiter_pair_fun_sub)
+
+lemma treeFsetIBNF_pred[simp]: "treeFsetIBNF_pred 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: treeFsetIBNF.pred_unfold)
+done
+
+lemmas treeFsetI_coind = mp[OF treeFsetI.pred_coinduct]
+
+lemma "tmap (f o g) x = tmap f (tmap g x)"
+by (intro treeFsetI_coind[where phi="%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/Codatatype/README.html Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,58 @@
+<!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>Codatatype Package</title>
+</head>
+
+<body>
+
+<h3><i>Codatatype</i>: A (co)datatype package based on bounded natural functors
+(BNFs)</h3>
+
+<p>
+The <i>Codatatype</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>Codatatype.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> file registers
+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. Future versions are expected
+to support multiple constructors and selectors per (co)datatype (instead of a
+single <i>fld</i> or <i>unf</i> constant) and provide a nicer syntax for
+(co)datatype and (co)recursive function definitions. Please contact
+any 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/Codatatype/Tools/bnf_comp.ML Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,834 @@
+(* Title: HOL/Codatatype/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_thms
+ val empty_unfold: unfold_thms
+ val map_unfolds_of: unfold_thms -> thm list
+ val set_unfoldss_of: unfold_thms -> thm list list
+ val rel_unfolds_of: unfold_thms -> thm list
+ val pred_unfolds_of: unfold_thms -> thm list
+
+ val default_comp_sort: (string * sort) list list -> (string * sort) list
+ val bnf_of_typ: BNF_Def.const_policy -> binding -> (binding -> binding) ->
+ ((string * sort) list list -> (string * sort) list) -> typ -> unfold_thms * Proof.context ->
+ (BNF_Def.BNF * (typ list * typ list)) * (unfold_thms * Proof.context)
+ val bnf_of_typ_cmd: binding * string -> Proof.context -> Proof.context
+ val seal_bnf: unfold_thms -> binding -> typ list -> BNF_Def.BNF -> Proof.context ->
+ BNF_Def.BNF * local_theory
+ val normalize_bnfs: (int -> binding -> binding) -> ''a list list -> ''a list ->
+ (''a list list -> ''a list) -> BNF_Def.BNF list -> unfold_thms -> Proof.context ->
+ (int list list * ''a list) * (BNF_Def.BNF list * (unfold_thms * Proof.context))
+end;
+
+structure BNF_Comp : BNF_COMP =
+struct
+
+open BNF_Def
+open BNF_Util
+open BNF_Tactics
+open BNF_Comp_Tactics
+
+type unfold_thms = {
+ map_unfolds: thm list,
+ set_unfoldss: thm list list,
+ rel_unfolds: thm list,
+ pred_unfolds: thm list
+};
+
+fun add_to_thms thms NONE = thms
+ | add_to_thms thms (SOME new) = if Thm.is_reflexive new then thms else insert Thm.eq_thm new thms;
+fun adds_to_thms thms NONE = thms
+ | adds_to_thms thms (SOME news) = insert (eq_set Thm.eq_thm) (filter_refl news) thms;
+
+fun mk_unfold_thms maps setss rels preds =
+ {map_unfolds = maps, set_unfoldss = setss, rel_unfolds = rels, pred_unfolds = preds};
+
+val empty_unfold = mk_unfold_thms [] [] [] [];
+
+fun add_to_unfold_opt map_opt sets_opt rel_opt pred_opt
+ {map_unfolds = maps, set_unfoldss = setss, rel_unfolds = rels, pred_unfolds = preds} = {
+ map_unfolds = add_to_thms maps map_opt,
+ set_unfoldss = adds_to_thms setss sets_opt,
+ rel_unfolds = add_to_thms rels rel_opt,
+ pred_unfolds = add_to_thms preds pred_opt};
+
+fun add_to_unfold map sets rel pred =
+ add_to_unfold_opt (SOME map) (SOME sets) (SOME rel) (SOME pred);
+
+val map_unfolds_of = #map_unfolds;
+val set_unfoldss_of = #set_unfoldss;
+val rel_unfolds_of = #rel_unfolds;
+val pred_unfolds_of = #pred_unfolds;
+
+val bdTN = "bdT";
+
+val compN = "comp_"
+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) ^ "_";
+
+val no_thm = refl;
+val Collect_split_box_equals = box_equals RS @{thm Collect_split_cong};
+val abs_pred_sym = sym RS @{thm abs_pred_def};
+val abs_pred_sym_pred_abs = abs_pred_sym RS @{thm pred_def_abs};
+
+(*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, 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', Asets), xs), _(*names_lthy*)) = lthy
+ |> apfst snd o mk_Frees' "f" (map2 (curry (op -->)) 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 comp_map = 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));
+
+ (*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_comp_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 (comp_sets, comp_sets_alt) = map_split mk_comp_set (0 upto ilive - 1);
+
+ (*(inner_1.bd +c ... +c inner_m.bd) *c outer.bd*)
+ val comp_bd = Term.absdummy CCA (mk_cprod
+ (Library.foldr1 (uncurry mk_csum) inner_bds) outer_bd);
+
+ fun comp_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 comp_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_comp_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 comp_set_natural_tacs = map mk_single_comp_set_natural_tac (0 upto ilive - 1);
+
+ fun comp_bd_card_order_tac _ =
+ mk_comp_bd_card_order_tac (map bd_card_order_of_bnf inners) (bd_card_order_of_bnf outer);
+
+ fun comp_bd_cinfinite_tac _ =
+ mk_comp_bd_cinfinite_tac (bd_cinfinite_of_bnf inner) (bd_cinfinite_of_bnf outer);
+
+ val comp_set_alt_thms =
+ if ! quick_and_dirty then
+ replicate ilive no_thm
+ else
+ map (fn goal => Skip_Proof.prove lthy [] [] goal
+ (fn {context, ...} => (mk_comp_set_alt_tac context (collect_set_natural_of_bnf outer))))
+ (map2 (curry (HOLogic.mk_Trueprop o HOLogic.mk_eq)) comp_sets comp_sets_alt);
+
+ fun comp_map_cong_tac _ =
+ mk_comp_map_cong_tac comp_set_alt_thms (map_cong_of_bnf outer) (map map_cong_of_bnf inners);
+
+ val comp_set_bd_tacs =
+ if ! quick_and_dirty then
+ replicate (length comp_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 comp_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 comp_single_set_bd_thm) (0 upto ilive - 1);
+ in
+ map2 (fn comp_set_alt => fn single_set_bds => fn {context, ...} =>
+ mk_comp_set_bd_tac context comp_set_alt single_set_bds)
+ comp_set_alt_thms single_set_bd_thmss
+ end;
+
+ val comp_in_alt_thm =
+ if ! quick_and_dirty then
+ no_thm
+ else
+ let
+ val comp_in = mk_in Asets comp_sets CCA;
+ val comp_in_alt = mk_in (map2 (mk_in Asets) inner_setss CAs) outer_sets CCA;
+ val goal =
+ fold_rev Logic.all Asets
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq (comp_in, comp_in_alt)));
+ in
+ Skip_Proof.prove lthy [] [] goal
+ (fn {context, ...} => mk_comp_in_alt_tac context comp_set_alt_thms)
+ end;
+
+ fun comp_in_bd_tac _ =
+ mk_comp_in_bd_tac comp_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 comp_map_wpull_tac _ =
+ mk_map_wpull_tac comp_in_alt_thm (map map_wpull_of_bnf inners) (map_wpull_of_bnf outer);
+
+ val tacs = [comp_map_id_tac, comp_map_comp_tac, comp_map_cong_tac] @ comp_set_natural_tacs @
+ [comp_bd_card_order_tac, comp_bd_cinfinite_tac] @ comp_set_bd_tacs @
+ [comp_in_bd_tac, comp_map_wpull_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 comp_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') =
+ add_bnf const_policy (K Derive_Some_Facts) qualify tacs wit_tac (SOME (oDs @ flat Dss))
+ ((((b, comp_map), comp_sets), comp_bd), comp_wits) lthy;
+
+ val outer_rel_Gr = rel_Gr_of_bnf outer RS sym;
+ val outer_rel_cong = rel_cong_of_bnf outer;
+
+ val comp_rel_unfold_thm =
+ trans OF [rel_def_of_bnf bnf',
+ trans OF [comp_in_alt_thm RS @{thm subst_rel_def},
+ trans OF [@{thm arg_cong2[of _ _ _ _ relcomp]} OF
+ [trans OF [outer_rel_Gr RS @{thm arg_cong[of _ _ converse]},
+ rel_converse_of_bnf outer RS sym], outer_rel_Gr],
+ trans OF [rel_O_of_bnf outer RS sym, outer_rel_cong OF
+ (map (fn bnf => rel_def_of_bnf bnf RS sym) inners)]]]];
+
+ val comp_pred_unfold_thm = Collect_split_box_equals OF [comp_rel_unfold_thm,
+ pred_def_of_bnf bnf' RS abs_pred_sym,
+ trans OF [outer_rel_cong OF (map (fn bnf => pred_def_of_bnf bnf RS abs_pred_sym) inners),
+ pred_def_of_bnf outer RS abs_pred_sym]];
+
+ val unfold' = add_to_unfold (map_def_of_bnf bnf') (set_defs_of_bnf bnf')
+ comp_rel_unfold_thm comp_pred_unfold_thm unfold;
+ in
+ (bnf', (unfold', lthy'))
+ end;
+
+fun clean_compose_bnf_cmd (outer, inners) lthy =
+ let
+ val outer = the (bnf_of lthy outer)
+ val inners = map (the o bnf_of lthy) inners
+ val b = name_of_bnf outer
+ |> Binding.prefix_name compN
+ |> Binding.suffix_name ("_" ^ implode (map (Binding.name_of o name_of_bnf) inners));
+ in
+ (snd o snd) (clean_compose_bnf Dont_Inline I b outer inners
+ (empty_unfold, lthy))
+ end;
+
+(* Killing live variables *)
+
+fun killN_bnf qualify n bnf (unfold, lthy) = if n = 0 then (bnf, (unfold, lthy)) else
+ let
+ val b = Binding.prefix_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 killN_map = Term.list_comb (mk_map_of_bnf Ds As Bs bnf, map HOLogic.id_const killedAs);
+
+ val bnf_sets = mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf;
+ val killN_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 killN_bd = mk_cprod
+ (Library.foldr1 (uncurry mk_csum) (map (mk_card_of o HOLogic.mk_UNIV) killedAs)) bnf_bd;
+
+ fun killN_map_id_tac _ = rtac (map_id_of_bnf bnf) 1;
+ fun killN_map_comp_tac {context, ...} =
+ Local_Defs.unfold_tac context ((map_comp_of_bnf bnf RS sym) :: @{thms o_assoc id_o o_id}) THEN
+ rtac refl 1;
+ fun killN_map_cong_tac {context, ...} =
+ mk_killN_map_cong_tac context n (live - n) (map_cong_of_bnf bnf);
+ val killN_set_natural_tacs =
+ map (fn thm => fn _ => rtac thm 1) (drop n (set_natural_of_bnf bnf));
+ fun killN_bd_card_order_tac _ = mk_killN_bd_card_order_tac n (bd_card_order_of_bnf bnf);
+ fun killN_bd_cinfinite_tac _ = mk_killN_bd_cinfinite_tac (bd_Cinfinite_of_bnf bnf);
+ val killN_set_bd_tacs =
+ map (fn thm => fn _ => mk_killN_set_bd_tac (bd_Card_order_of_bnf bnf) thm)
+ (drop n (set_bd_of_bnf bnf));
+
+ val killN_in_alt_thm =
+ if ! quick_and_dirty then
+ no_thm
+ else
+ let
+ val killN_in = mk_in Asets killN_sets T;
+ val killN_in_alt = mk_in (map HOLogic.mk_UNIV killedAs @ Asets) bnf_sets T;
+ val goal =
+ fold_rev Logic.all Asets (HOLogic.mk_Trueprop (HOLogic.mk_eq (killN_in, killN_in_alt)));
+ in
+ Skip_Proof.prove lthy [] [] goal (K killN_in_alt_tac)
+ end;
+
+ fun killN_in_bd_tac _ =
+ mk_killN_in_bd_tac n (live > n) killN_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 killN_map_wpull_tac _ =
+ mk_map_wpull_tac killN_in_alt_thm [] (map_wpull_of_bnf bnf);
+
+ val tacs = [killN_map_id_tac, killN_map_comp_tac, killN_map_cong_tac] @ killN_set_natural_tacs @
+ [killN_bd_card_order_tac, killN_bd_cinfinite_tac] @ killN_set_bd_tacs @
+ [killN_in_bd_tac, killN_map_wpull_tac];
+
+ val wits = mk_wits_of_bnf (replicate nwits Ds) (replicate nwits As) bnf;
+
+ val killN_wits = map (fn t => fold absfree (Term.add_frees t []) t)
+ (map (fn (I, wit) => Term.list_comb (wit, map (nth xs) I)) wits);
+
+ fun wit_tac _ = mk_simple_wit_tac (wit_thms_of_bnf bnf);
+
+ val (bnf', lthy') =
+ add_bnf Smart_Inline (K Derive_Some_Facts) qualify tacs wit_tac (SOME (killedAs @ Ds))
+ ((((b, killN_map), killN_sets), Term.absdummy T killN_bd), killN_wits) lthy;
+
+ val rel_Gr = rel_Gr_of_bnf bnf RS sym;
+
+ val killN_rel_unfold_thm =
+ trans OF [rel_def_of_bnf bnf',
+ trans OF [killN_in_alt_thm RS @{thm subst_rel_def},
+ trans OF [@{thm arg_cong2[of _ _ _ _ relcomp]} OF
+ [trans OF [rel_Gr RS @{thm arg_cong[of _ _ converse]}, rel_converse_of_bnf bnf RS sym],
+ rel_Gr],
+ trans OF [rel_O_of_bnf bnf RS sym, rel_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})]]]];
+
+ val killN_pred_unfold_thm = Collect_split_box_equals OF
+ [Local_Defs.unfold lthy' @{thms Id_def'} killN_rel_unfold_thm,
+ pred_def_of_bnf bnf' RS abs_pred_sym, pred_def_of_bnf bnf RS abs_pred_sym];
+
+ val unfold' = add_to_unfold (map_def_of_bnf bnf') (set_defs_of_bnf bnf')
+ killN_rel_unfold_thm killN_pred_unfold_thm unfold;
+ in
+ (bnf', (unfold', lthy'))
+ end;
+
+fun killN_bnf_cmd (n, raw_bnf) lthy =
+ (snd o snd) (killN_bnf I n (the (bnf_of lthy raw_bnf)) (empty_unfold, lthy));
+
+(* Adding dummy live variables *)
+
+fun liftN_bnf qualify n bnf (unfold, lthy) = if n = 0 then (bnf, (unfold, lthy)) else
+ let
+ val b = Binding.prefix_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 liftN_map =
+ fold_rev Term.absdummy (map2 (curry (op -->)) newAs newBs) (mk_map_of_bnf Ds As Bs bnf);
+
+ val bnf_sets = mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf;
+ val liftN_sets = map (fn A => absdummy T (HOLogic.mk_set A [])) newAs @ bnf_sets;
+
+ val liftN_bd = mk_bd_of_bnf Ds As bnf;
+
+ fun liftN_map_id_tac _ = rtac (map_id_of_bnf bnf) 1;
+ fun liftN_map_comp_tac {context, ...} =
+ Local_Defs.unfold_tac context ((map_comp_of_bnf bnf RS sym) :: @{thms o_assoc id_o o_id}) THEN
+ rtac refl 1;
+ fun liftN_map_cong_tac {context, ...} =
+ rtac (map_cong_of_bnf bnf) 1 THEN REPEAT_DETERM_N live (Goal.assume_rule_tac context 1);
+ val liftN_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 liftN_bd_card_order_tac _ = rtac (bd_card_order_of_bnf bnf) 1;
+ fun liftN_bd_cinfinite_tac _ = rtac (bd_cinfinite_of_bnf bnf) 1;
+ val liftN_set_bd_tacs =
+ if ! quick_and_dirty then
+ replicate (n + live) (K all_tac)
+ else
+ replicate n (K (mk_liftN_set_bd_tac (bd_Card_order_of_bnf bnf))) @
+ (map (fn thm => fn _ => rtac thm 1) (set_bd_of_bnf bnf));
+
+ val liftN_in_alt_thm =
+ if ! quick_and_dirty then
+ no_thm
+ else
+ let
+ val liftN_in = mk_in Asets liftN_sets T;
+ val liftN_in_alt = mk_in (drop n Asets) bnf_sets T;
+ val goal =
+ fold_rev Logic.all Asets (HOLogic.mk_Trueprop (HOLogic.mk_eq (liftN_in, liftN_in_alt)))
+ in
+ Skip_Proof.prove lthy [] [] goal (K liftN_in_alt_tac)
+ end;
+
+ fun liftN_in_bd_tac _ =
+ mk_liftN_in_bd_tac n liftN_in_alt_thm (in_bd_of_bnf bnf) (bd_Card_order_of_bnf bnf);
+ fun liftN_map_wpull_tac _ =
+ mk_map_wpull_tac liftN_in_alt_thm [] (map_wpull_of_bnf bnf);
+
+ val tacs = [liftN_map_id_tac, liftN_map_comp_tac, liftN_map_cong_tac] @ liftN_set_natural_tacs @
+ [liftN_bd_card_order_tac, liftN_bd_cinfinite_tac] @ liftN_set_bd_tacs @
+ [liftN_in_bd_tac, liftN_map_wpull_tac];
+
+ val liftN_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') =
+ add_bnf Smart_Inline (K Derive_Some_Facts) qualify tacs wit_tac (SOME Ds)
+ ((((b, liftN_map), liftN_sets), Term.absdummy T liftN_bd), liftN_wits) lthy;
+
+ val liftN_rel_unfold_thm =
+ trans OF [rel_def_of_bnf bnf',
+ trans OF [liftN_in_alt_thm RS @{thm subst_rel_def}, rel_def_of_bnf bnf RS sym]];
+
+ val liftN_pred_unfold_thm = Collect_split_box_equals OF [liftN_rel_unfold_thm,
+ pred_def_of_bnf bnf' RS abs_pred_sym, pred_def_of_bnf bnf RS abs_pred_sym];
+
+ val unfold' = add_to_unfold (map_def_of_bnf bnf') (set_defs_of_bnf bnf')
+ liftN_rel_unfold_thm liftN_pred_unfold_thm unfold;
+ in
+ (bnf', (unfold', lthy'))
+ end;
+
+fun liftN_bnf_cmd (n, raw_bnf) lthy =
+ (snd o snd) (liftN_bnf I n (the (bnf_of lthy raw_bnf)) (empty_unfold, lthy));
+
+(* Changing the order of live variables *)
+
+fun permute_bnf qualify src dest bnf (unfold, lthy) = if src = dest then (bnf, (unfold, lthy)) else
+ let
+ val b = Binding.prefix_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 permute_map = 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))));
+
+ val bnf_sets = mk_sets_of_bnf (replicate live Ds) (replicate live As) bnf;
+ val permute_sets = permute bnf_sets;
+
+ val permute_bd = mk_bd_of_bnf Ds As bnf;
+
+ fun permute_map_id_tac _ = rtac (map_id_of_bnf bnf) 1;
+ fun permute_map_comp_tac _ = rtac (map_comp_of_bnf bnf) 1;
+ fun permute_map_cong_tac {context, ...} =
+ rtac (map_cong_of_bnf bnf) 1 THEN REPEAT_DETERM_N live (Goal.assume_rule_tac context 1);
+ val permute_set_natural_tacs =
+ permute (map (fn thm => fn _ => rtac thm 1) (set_natural_of_bnf bnf));
+ fun permute_bd_card_order_tac _ = rtac (bd_card_order_of_bnf bnf) 1;
+ fun permute_bd_cinfinite_tac _ = rtac (bd_cinfinite_of_bnf bnf) 1;
+ val permute_set_bd_tacs = permute (map (fn thm => fn _ => rtac thm 1) (set_bd_of_bnf bnf));
+
+ val permute_in_alt_thm =
+ if ! quick_and_dirty then
+ no_thm
+ else
+ let
+ val permute_in = mk_in Asets permute_sets T;
+ val permute_in_alt = mk_in (permute_rev Asets) bnf_sets T;
+ val goal =
+ fold_rev Logic.all Asets
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq (permute_in, permute_in_alt)));
+ in
+ Skip_Proof.prove lthy [] [] goal (K (mk_permute_in_alt_tac src dest))
+ end;
+
+ fun permute_in_bd_tac _ =
+ mk_permute_in_bd_tac src dest permute_in_alt_thm (in_bd_of_bnf bnf)
+ (bd_Card_order_of_bnf bnf);
+ fun permute_map_wpull_tac _ =
+ mk_map_wpull_tac permute_in_alt_thm [] (map_wpull_of_bnf bnf);
+
+ val tacs = [permute_map_id_tac, permute_map_comp_tac, permute_map_cong_tac] @
+ permute_set_natural_tacs @ [permute_bd_card_order_tac, permute_bd_cinfinite_tac] @
+ permute_set_bd_tacs @ [permute_in_bd_tac, permute_map_wpull_tac];
+
+ val permute_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') =
+ add_bnf Smart_Inline (K Derive_Some_Facts) qualify tacs wit_tac (SOME Ds)
+ ((((b, permute_map), permute_sets), Term.absdummy T permute_bd), permute_wits) lthy;
+
+ val permute_rel_unfold_thm =
+ trans OF [rel_def_of_bnf bnf',
+ trans OF [permute_in_alt_thm RS @{thm subst_rel_def}, rel_def_of_bnf bnf RS sym]];
+
+ val permute_pred_unfold_thm = Collect_split_box_equals OF [permute_rel_unfold_thm,
+ pred_def_of_bnf bnf' RS abs_pred_sym, pred_def_of_bnf bnf RS abs_pred_sym];
+
+ val unfold' = add_to_unfold (map_def_of_bnf bnf') (set_defs_of_bnf bnf')
+ permute_rel_unfold_thm permute_pred_unfold_thm unfold;
+ in
+ (bnf', (unfold', lthy'))
+ end;
+
+fun permute_bnf_cmd ((src, dest), raw_bnf) lthy =
+ (snd o snd) (permute_bnf I src dest (the (bnf_of lthy raw_bnf))
+ (empty_unfold, lthy));
+
+(* Hide the type of the bound (optimization) and unfold the definitions (nicer to the user) *)
+
+fun seal_bnf unfold 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 = filter_refl (map_unfolds_of unfold);
+ val set_unfoldss = map filter_refl (set_unfoldss_of unfold);
+
+ 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 unfold_maps = fold (Local_Defs.unfold lthy o single) map_unfolds;
+ val unfold_sets = fold (Local_Defs.unfold lthy) set_unfoldss;
+ val unfold_defs = unfold_sets o unfold_maps;
+ 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 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 ((bdT_name, (bdT_glob_info, bdT_loc_info)), (lthy', lthy)) =
+ lthy
+ |> Typedef.add_typedef true NONE (bdT_bind, params, NoSyn)
+ (HOLogic.mk_UNIV bd_repT) NONE (EVERY' [rtac exI, rtac UNIV_I] 1)
+ ||> `Local_Theory.restore;
+
+ val phi = Proof_Context.export_morphism lthy lthy';
+
+ val bnf_bd' = mk_dir_image bnf_bd
+ (Const (#Abs_name bdT_glob_info, bd_repT --> Type (bdT_name, map TFree params)))
+
+ val set_def = Morphism.thm phi (the (#set_def bdT_loc_info));
+ val Abs_inject = Morphism.thm phi (#Abs_inject bdT_loc_info);
+ val Abs_cases = Morphism.thm phi (#Abs_cases bdT_loc_info);
+
+ val Abs_bdT_inj = mk_Abs_inj_thm set_def Abs_inject;
+ val Abs_bdT_bij = mk_Abs_bij_thm lthy' set_def Abs_inject Abs_cases;
+
+ 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_defs thm) THEN'
+ SOLVE o REPEAT_DETERM o (atac ORELSE' Goal.assume_rule_tac ctxt)) 1;
+ val tacs =
+ map mk_tac ([map_id_of_bnf bnf, map_comp_of_bnf bnf, map_cong_of_bnf bnf] @
+ set_natural_of_bnf bnf) @
+ map K [rtac bd_card_order 1, rtac bd_cinfinite 1] @
+ map mk_tac (set_bds @ [in_bd, map_wpull_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_defs (wit_thms_of_bnf bnf));
+
+ val (bnf', lthy') = add_bnf Hardly_Inline (K Derive_All_Facts) I tacs wit_tac NONE
+ ((((b, bnf_map), bnf_sets), Term.absdummy T bnf_bd'), bnf_wits) lthy;
+
+ val defs' = filter_refl (map_def_of_bnf bnf' :: set_defs_of_bnf bnf');
+ val unfold_defs' = unfold_defs o Local_Defs.unfold lthy' defs';
+
+ val rel_def = unfold_defs' (rel_def_of_bnf bnf');
+ val rel_unfold = Local_Defs.unfold lthy'
+ (map unfold_defs (filter_refl (rel_unfolds_of unfold))) rel_def;
+
+ val unfold_defs'' =
+ unfold_defs' o (Local_Defs.unfold lthy' (filter_refl [rel_def_of_bnf bnf']));
+
+ val pred_def = unfold_defs'' (pred_def_of_bnf bnf' RS abs_pred_sym_pred_abs);
+ val pred_unfold = Local_Defs.unfold lthy'
+ (map unfold_defs (filter_refl (pred_unfolds_of unfold))) pred_def;
+
+ fun note thmN thms =
+ snd o Local_Theory.note
+ ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), thms);
+ in
+ (bnf', lthy'
+ |> note rel_unfoldN [rel_unfold]
+ |> note pred_unfoldN [pred_unfold])
+ end;
+
+(* Composition pipeline *)
+
+fun permute_and_kill qualify n src dest bnf =
+ bnf
+ |> permute_bnf qualify src dest
+ #> uncurry (killN_bnf qualify n);
+
+fun lift_and_permute qualify n src dest bnf =
+ bnf
+ |> liftN_bnf qualify n
+ #> uncurry (permute_bnf qualify src dest);
+
+fun normalize_bnfs qualify Ass Ds sort bnfs unfold 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', 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, 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', 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' b sort outer inners oDs Dss tfreess (unfold, lthy) =
+ let
+ val name = Binding.name_of b;
+ fun qualify i bind =
+ let val namei = if i > 0 then name ^ string_of_int i else name;
+ in
+ if member (op =) (#2 (Binding.dest bind)) (namei, true) then qualify' bind
+ else qualify' (Binding.prefix_name namei bind)
+ end;
+
+ val Ass = map (map dest_TFree) tfreess;
+ val Ds = fold (fold Term.add_tfreesT) (oDs :: Dss) [];
+
+ val ((kill_poss, As), (inners', (unfold', lthy'))) =
+ normalize_bnfs qualify Ass Ds sort inners unfold 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 I b outer inners' (unfold', lthy'))
+ end;
+
+fun compose_bnf_cmd (((((b, outer), inners), oDs), Dss), Ass) lthy = (snd o snd)
+ (compose_bnf Dont_Inline I b default_comp_sort (the (bnf_of lthy outer))
+ (map (the o bnf_of lthy) inners)
+ (map (Syntax.read_typ lthy) oDs) (map (map (Syntax.read_typ lthy)) Dss)
+ (map (map (Syntax.read_typ lthy)) Ass) (empty_unfold, lthy));
+
+fun bnf_of_typ _ _ _ _ (T as TFree _) (unfold, lthy) =
+ ((Basic_BNFs.ID_bnf, ([], [T])), (add_to_unfold_opt NONE NONE
+ (SOME Basic_BNFs.ID_rel_def) (SOME Basic_BNFs.ID_pred_def) unfold, lthy))
+ | bnf_of_typ _ _ _ _ (TVar _) _ = error "Unexpected schematic variable"
+ | bnf_of_typ const_policy b qualify' sort (T as Type (C, Ts)) (unfold, lthy) =
+ let val tfrees = Term.add_tfreesT T [];
+ in
+ if null tfrees then
+ ((Basic_BNFs.DEADID_bnf, ([T], [])), (unfold, lthy))
+ else if forall (can Term.dest_TFree) Ts andalso length Ts = length tfrees then let
+ val bnf = the (bnf_of lthy (Long_Name.base_name C));
+ 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);
+ val rel_def = rel_def_of_bnf bnf;
+ val unfold' = add_to_unfold_opt NONE NONE (SOME (rel_def RS sym))
+ (SOME (Local_Defs.unfold lthy [rel_def] (pred_def_of_bnf bnf) RS sym)) unfold;
+ in ((bnf, deads_lives), (unfold', lthy)) end
+ else
+ let
+ (* FIXME: we should allow several BNFs with the same base name *)
+ val Tname = Long_Name.base_name C;
+ val name = Binding.name_of b ^ "_" ^ Tname;
+ fun qualify i bind =
+ let val namei = if i > 0 then name ^ string_of_int i else name;
+ in
+ if member (op =) (#2 (Binding.dest bind)) (namei, true) then qualify' bind
+ else qualify' (Binding.prefix_name namei bind)
+ end;
+ val outer = the (bnf_of lthy Tname);
+ val odead = dead_of_bnf outer;
+ val olive = live_of_bnf outer;
+ val oDs_pos = find_indices [TFree ("dead", [])]
+ (snd (dest_Type
+ (mk_T_of_bnf (replicate odead (TFree ("dead", []))) (replicate olive dummyT) outer)));
+ 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', lthy')) =
+ apfst (apsnd split_list o split_list)
+ (fold_map2 (fn i =>
+ bnf_of_typ Smart_Inline (Binding.name (name ^ string_of_int i)) (qualify i) sort)
+ (if length Ts' = 1 then [0] else (1 upto length Ts'))
+ Ts' (unfold, lthy));
+ in
+ compose_bnf const_policy (qualify 0) b sort outer inners oDs Dss Ass (unfold', lthy')
+ end
+ end;
+
+fun bnf_of_typ_cmd (b, rawT) lthy = (snd o snd)
+ (bnf_of_typ Dont_Inline b I default_comp_sort (Syntax.read_typ lthy rawT)
+ (empty_unfold, lthy));
+
+val _ =
+ Outer_Syntax.local_theory @{command_spec "bnf_of_typ"} "parse a type as composition of BNFs"
+ (((Parse.binding --| Parse.$$$ "=") -- Parse.typ) >> bnf_of_typ_cmd);
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Codatatype/Tools/bnf_comp_tactics.ML Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,394 @@
+(* Title: HOL/Codatatype/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_killN_bd_card_order_tac: int -> thm -> tactic
+ val mk_killN_bd_cinfinite_tac: thm -> tactic
+ val killN_in_alt_tac: tactic
+ val mk_killN_in_bd_tac: int -> bool -> thm -> thm -> thm -> thm -> thm -> tactic
+ val mk_killN_map_cong_tac: Proof.context -> int -> int -> thm -> tactic
+ val mk_killN_set_bd_tac: thm -> thm -> tactic
+
+ val empty_natural_tac: tactic
+ val liftN_in_alt_tac: tactic
+ val mk_liftN_in_bd_tac: int -> thm -> thm -> thm -> tactic
+ val mk_liftN_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
+end;
+
+structure BNF_Comp_Tactics : BNF_COMP_TACTICS =
+struct
+
+open BNF_Util
+open BNF_Tactics
+
+val arg_cong_Union = @{thm arg_cong[of _ _ Union]};
+val o_eq_dest_lhs = @{thm o_eq_dest_lhs};
+val set_mp = @{thm set_mp};
+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 =
+ Local_Defs.unfold_tac ctxt @{thms sym[OF o_assoc]} THEN
+ Local_Defs.unfold_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 @{thm trans[OF image_cong[OF o_apply refl]]}, 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
+ Local_Defs.unfold_tac ctxt [comp_set_alt] THEN
+ rtac @{thm comp_set_bd_Union_o_collect} 1 THEN
+ Local_Defs.unfold_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 =
+ Local_Defs.unfold_tac ctxt (comp_set_alts @ comp_in_alt_thms) THEN
+ Local_Defs.unfold_tac ctxt @{thms set_eq_subset} THEN
+ rtac conjI 1 THEN
+ REPEAT_DETERM (
+ rtac @{thm subsetI} 1 THEN
+ Local_Defs.unfold_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 @{thm subset_UNIV})) ORELSE'
+ atac ORELSE'
+ (rtac @{thm subset_UNIV})) 1)) ORELSE rtac @{thm 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 @{thm Card_order_csum} THEN' rtac @{thm ordLeq_csum2}) ORELSE'
+ (rtac @{thm Card_order_ctwo} THEN' rtac @{thm ordLeq_refl})) THEN'
+ rtac @{thm 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 @{thm ordLeq_csum2} ORELSE' rtac @{thm ordLeq_refl}) THEN'
+ rtac @{thm 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 @{thm ordLeq_csum2} ORELSE' rtac @{thm ordLeq_refl}) THEN'
+ rtac @{thm 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 @{thm csum_Cnotzero2} THEN'
+ rtac @{thm ctwo_Cnotzero} THEN'
+ rtac Gbd_Card_order THEN'
+ rtac @{thm cexp_cprod} THEN'
+ TRY o rtac @{thm csum_Cnotzero2} THEN'
+ rtac @{thm 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
+ Local_Defs.unfold_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_killN_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_killN_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_killN_bd_cinfinite_tac bd_Cinfinite =
+ (rtac @{thm cinfinite_cprod2} THEN'
+ TRY o rtac @{thm csum_Cnotzero1} THEN'
+ rtac @{thm Cnotzero_UNIV} THEN'
+ rtac bd_Cinfinite) 1;
+
+fun mk_killN_set_bd_tac bd_Card_order set_bd =
+ (rtac ctrans THEN'
+ rtac set_bd THEN'
+ rtac @{thm ordLeq_cprod2} THEN'
+ TRY o rtac @{thm csum_Cnotzero1} THEN'
+ rtac @{thm Cnotzero_UNIV} THEN'
+ rtac bd_Card_order) 1
+
+val killN_in_alt_tac =
+ ((rtac @{thm Collect_cong} THEN' rtac @{thm iffI}) 1 THEN
+ REPEAT_DETERM (CHANGED (etac conjE 1)) THEN
+ REPEAT_DETERM (CHANGED ((etac conjI ORELSE'
+ rtac conjI THEN' rtac @{thm subset_UNIV}) 1)) THEN
+ (rtac @{thm 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 @{thm subset_UNIV} 1));
+
+fun mk_killN_in_bd_tac n nontrivial_killN_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_killN_in then
+ rtac @{thm 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 @{thm ordIso_transitive}) 1) THEN
+ (rtac @{thm ordIso_refl} THEN'
+ rtac @{thm Card_order_csum} THEN'
+ rtac @{thm ordIso_transitive} THEN'
+ rtac @{thm csum_assoc} THEN'
+ rtac @{thm ordIso_transitive} THEN'
+ rtac @{thm csum_cong1} THEN'
+ K (mk_flatten_assoc_tac
+ (rtac @{thm ordIso_refl} THEN'
+ FIRST' [rtac @{thm card_of_Card_order}, rtac @{thm Card_order_csum}])
+ @{thm ordIso_transitive} @{thm csum_assoc} @{thm csum_cong}) THEN'
+ rtac @{thm ordIso_refl} THEN'
+ (rtac @{thm card_of_Card_order} ORELSE' rtac @{thm Card_order_csum})) 1
+ else all_tac) THEN
+ (rtac @{thm csum_com} THEN'
+ rtac bd_Card_order THEN'
+ rtac disjI1 THEN'
+ rtac @{thm csum_Cnotzero2} THEN'
+ rtac @{thm ctwo_Cnotzero} THEN'
+ rtac disjI1 THEN'
+ rtac @{thm csum_Cnotzero2} THEN'
+ TRY o rtac @{thm csum_Cnotzero1} THEN'
+ rtac @{thm 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 @{thm card_of_Card_order} ORELSE' rtac @{thm Card_order_csum}) THEN'
+ rtac bd_Cnotzero THEN'
+ rtac @{thm csum_cexp'} THEN'
+ rtac @{thm Cinfinite_cprod2} THEN'
+ TRY o rtac @{thm csum_Cnotzero1} THEN'
+ rtac @{thm Cnotzero_UNIV} THEN'
+ rtac bd_Cinfinite THEN'
+ ((rtac @{thm Card_order_ctwo} THEN' rtac @{thm ordLeq_refl}) ORELSE'
+ (rtac @{thm Card_order_csum} THEN' rtac @{thm ordLeq_csum2})) THEN'
+ rtac @{thm Card_order_ctwo} THEN'
+ rtac disjI1 THEN'
+ rtac @{thm csum_Cnotzero2} THEN'
+ TRY o rtac @{thm csum_Cnotzero1} THEN'
+ rtac @{thm Cnotzero_UNIV} THEN'
+ rtac bd_Card_order THEN'
+ rtac @{thm cexp_cprod_ordLeq} THEN'
+ TRY o rtac @{thm csum_Cnotzero2} THEN'
+ rtac @{thm ctwo_Cnotzero} THEN'
+ rtac @{thm Cinfinite_cprod2} THEN'
+ TRY o rtac @{thm csum_Cnotzero1} THEN'
+ rtac @{thm Cnotzero_UNIV} THEN'
+ rtac bd_Cinfinite THEN'
+ rtac bd_Cnotzero THEN'
+ rtac @{thm ordLeq_cprod2} THEN'
+ TRY o rtac @{thm csum_Cnotzero1} THEN'
+ rtac @{thm Cnotzero_UNIV} THEN'
+ rtac bd_Card_order) 1;
+
+
+
+(* Lift operation *)
+
+val empty_natural_tac = rtac @{thm empty_natural} 1;
+
+fun mk_liftN_set_bd_tac bd_Card_order = (rtac @{thm Card_order_empty} THEN' rtac bd_Card_order) 1;
+
+val liftN_in_alt_tac =
+ ((rtac @{thm Collect_cong} THEN' rtac @{thm 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_liftN_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 @{thm ordLeq_csum2} THEN'
+ (rtac @{thm Card_order_csum} ORELSE' rtac @{thm card_of_Card_order})) 1) THEN
+ (rtac @{thm ordLeq_csum2} THEN'
+ (rtac @{thm Card_order_csum} ORELSE' rtac @{thm card_of_Card_order})) 1) ORELSE
+ (rtac @{thm ordLeq_csum2} THEN' rtac @{thm Card_order_ctwo}) 1) THEN
+ (rtac disjI1 THEN' TRY o rtac @{thm csum_Cnotzero2} THEN' rtac @{thm 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 @{thm card_of_Card_order}, rtac @{thm Card_order_csum}])
+ @{thm 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 @{thm csum_Cnotzero2} THEN'
+ rtac @{thm ctwo_Cnotzero} THEN'
+ rtac disjI1 THEN'
+ TRY o rtac @{thm csum_Cnotzero2} THEN'
+ rtac @{thm ctwo_Cnotzero}) 1;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Codatatype/Tools/bnf_def.ML Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,1217 @@
+(* Title: HOL/Codatatype/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 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 setN: string
+ val relN: string
+ val predN: string
+ val mk_setN: int -> string
+ val rel_unfoldN: string
+ val pred_unfoldN: string
+
+ val mk_T_of_bnf: typ list -> typ list -> BNF -> typ
+ val mk_bd_of_bnf: typ list -> typ list -> BNF -> term
+ val mk_map_of_bnf: typ list -> typ list -> typ list -> BNF -> term
+ val mk_pred_of_bnf: typ list -> typ list -> typ list -> BNF -> term
+ val mk_rel_of_bnf: typ list -> typ list -> typ list -> BNF -> term
+ val mk_sets_of_bnf: typ list list -> typ list list -> BNF -> term list
+ val mk_wits_of_bnf: typ list list -> typ list list -> BNF -> (int list * term) list
+
+ val bd_Card_order_of_bnf: BNF -> thm
+ val bd_Cinfinite_of_bnf: BNF -> thm
+ val bd_Cnotzero_of_bnf: BNF -> thm
+ val bd_card_order_of_bnf: BNF -> thm
+ val bd_cinfinite_of_bnf: BNF -> thm
+ val collect_set_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_rel_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 pred_def_of_bnf: BNF -> thm
+ val rel_Gr_of_bnf: BNF -> thm
+ val rel_Id_of_bnf: BNF -> thm
+ val rel_O_of_bnf: BNF -> thm
+ val rel_cong_of_bnf: BNF -> thm
+ val rel_converse_of_bnf: BNF -> thm
+ val rel_def_of_bnf: BNF -> thm
+ val rel_mono_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 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 * term) list -> (''a list * term) list
+ val wits_of_bnf: BNF -> nonemptiness_witness list
+
+ datatype const_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline
+ datatype fact_policy =
+ Derive_Some_Facts | Derive_All_Facts | Derive_All_Facts_Note_Most | Note_All_Facts_and_Axioms
+ val bnf_note_all: bool Config.T
+ val user_policy: Proof.context -> fact_policy
+
+ val print_bnfs: Proof.context -> unit
+ val add_bnf: 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 -> local_theory ->
+ BNF * local_theory
+
+ val filter_refl: thm list -> thm list
+ val add_bnf_cmd: (((binding * string) * string list) * string) * string list -> local_theory ->
+ Proof.state
+end;
+
+structure BNF_Def : BNF_DEF =
+struct
+
+open BNF_Util
+open BNF_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
+};
+
+fun mk_axioms' ((((((((id, comp), cong), nat), c_o), cinf), set_bd), in_bd), wpull) =
+ {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};
+
+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
+ ||> the_single
+ |> mk_axioms';
+
+fun dest_axioms {map_id, map_comp, map_cong, set_natural,
+ bd_card_order, bd_cinfinite, set_bd, in_bd, map_wpull} =
+ [map_id, map_comp, map_cong] @ set_natural @ [bd_card_order, bd_cinfinite] @
+ set_bd @ [in_bd, map_wpull];
+
+fun map_axioms f
+ {map_id = map_id, map_comp = map_comp, map_cong = map_cong, set_natural = set_natural,
+ bd_card_order = bd_card_order, bd_cinfinite = bd_cinfinite,
+ set_bd = set_bd, in_bd = in_bd, map_wpull = map_wpull} =
+ {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};
+
+val morph_axioms = map_axioms o Morphism.thm;
+
+type defs = {
+ map_def: thm,
+ set_defs: thm list,
+ rel_def: thm,
+ pred_def: thm
+}
+
+fun mk_defs map sets rel pred = {map_def = map, set_defs = sets, rel_def = rel, pred_def = pred};
+
+fun map_defs f {map_def = map, set_defs = sets, rel_def = rel, pred_def = pred} =
+ {map_def = f map, set_defs = List.map f sets, rel_def = f rel, pred_def = f pred};
+
+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_rel: thm lazy,
+ map_comp': thm lazy,
+ map_id': thm lazy,
+ map_wppull: thm lazy,
+ rel_cong: thm lazy,
+ rel_mono: thm lazy,
+ rel_Id: thm lazy,
+ rel_Gr: thm lazy,
+ rel_converse: thm lazy,
+ rel_O: thm lazy,
+ set_natural': thm lazy list
+};
+
+fun mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero
+ collect_set_natural in_cong in_mono in_rel map_comp' map_id' map_wppull
+ rel_cong rel_mono rel_Id rel_Gr rel_converse rel_O set_natural' = {
+ 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_rel = in_rel,
+ map_comp' = map_comp',
+ map_id' = map_id',
+ map_wppull = map_wppull,
+ rel_cong = rel_cong,
+ rel_mono = rel_mono,
+ rel_Id = rel_Id,
+ rel_Gr = rel_Gr,
+ rel_converse = rel_converse,
+ rel_O = rel_O,
+ set_natural' = set_natural'};
+
+fun map_facts f {
+ bd_Card_order,
+ bd_Cinfinite,
+ bd_Cnotzero,
+ collect_set_natural,
+ in_cong,
+ in_mono,
+ in_rel,
+ map_comp',
+ map_id',
+ map_wppull,
+ rel_cong,
+ rel_mono,
+ rel_Id,
+ rel_Gr,
+ rel_converse,
+ rel_O,
+ set_natural'} =
+ {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_rel = Lazy.map f in_rel,
+ map_comp' = Lazy.map f map_comp',
+ map_id' = Lazy.map f map_id',
+ map_wppull = Lazy.map f map_wppull,
+ rel_cong = Lazy.map f rel_cong,
+ rel_mono = Lazy.map f rel_mono,
+ rel_Id = Lazy.map f rel_Id,
+ rel_Gr = Lazy.map f rel_Gr,
+ rel_converse = Lazy.map f rel_converse,
+ rel_O = Lazy.map f rel_O,
+ set_natural' = map (Lazy.map f) set_natural'};
+
+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,
+ pred: term
+};
+
+(* getters *)
+
+fun rep_bnf (BNF bnf) = bnf;
+val name_of_bnf = #name o rep_bnf;
+val T_of_bnf = #T o rep_bnf;
+fun mk_T_of_bnf Ds Ts bnf =
+ let val bnf_rep = rep_bnf bnf
+ in Term.typ_subst_atomic ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#T bnf_rep) end;
+val live_of_bnf = #live o rep_bnf;
+val lives_of_bnf = #lives o rep_bnf;
+val dead_of_bnf = #dead o rep_bnf;
+val deads_of_bnf = #deads o rep_bnf;
+val axioms_of_bnf = #axioms o rep_bnf;
+val facts_of_bnf = #facts o rep_bnf;
+val nwits_of_bnf = #nwits o rep_bnf;
+val wits_of_bnf = #wits o rep_bnf;
+
+(*terms*)
+val map_of_bnf = #map o rep_bnf;
+val sets_of_bnf = #sets o rep_bnf;
+fun mk_map_of_bnf Ds Ts Us bnf =
+ let val bnf_rep = rep_bnf bnf;
+ in
+ Term.subst_atomic_types
+ ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#map bnf_rep)
+ end;
+fun mk_sets_of_bnf Dss Tss bnf =
+ let val bnf_rep = rep_bnf bnf;
+ in
+ map2 (fn (Ds, Ts) => Term.subst_atomic_types
+ ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts))) (Dss ~~ Tss) (#sets bnf_rep)
+ end;
+val bd_of_bnf = #bd o rep_bnf;
+fun mk_bd_of_bnf Ds Ts bnf =
+ let val bnf_rep = rep_bnf bnf;
+ in Term.subst_atomic_types ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#bd bnf_rep) end;
+fun mk_wits_of_bnf Dss Tss bnf =
+ let
+ val bnf_rep = rep_bnf bnf;
+ val wits = map (fn x => (#I x, #wit x)) (#wits bnf_rep);
+ in
+ map2 (fn (Ds, Ts) => apsnd (Term.subst_atomic_types
+ ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)))) (Dss ~~ Tss) wits
+ end;
+val rel_of_bnf = #rel o rep_bnf;
+fun mk_rel_of_bnf Ds Ts Us bnf =
+ let val bnf_rep = rep_bnf bnf;
+ in
+ Term.subst_atomic_types
+ ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#rel bnf_rep)
+ end;
+val pred_of_bnf = #pred o rep_bnf;
+fun mk_pred_of_bnf Ds Ts 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)) (#pred 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_rel_of_bnf = Lazy.force o #in_rel 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 pred_def_of_bnf = #pred_def o #defs o rep_bnf;
+val rel_cong_of_bnf = Lazy.force o #rel_cong o #facts o rep_bnf;
+val rel_mono_of_bnf = Lazy.force o #rel_mono o #facts o rep_bnf;
+val rel_def_of_bnf = #rel_def o #defs o rep_bnf;
+val rel_Id_of_bnf = Lazy.force o #rel_Id o #facts o rep_bnf;
+val rel_Gr_of_bnf = Lazy.force o #rel_Gr o #facts o rep_bnf;
+val rel_converse_of_bnf = Lazy.force o #rel_converse o #facts o rep_bnf;
+val rel_O_of_bnf = Lazy.force o #rel_O o #facts o rep_bnf;
+val set_bd_of_bnf = #set_bd o #axioms o rep_bnf;
+val set_defs_of_bnf = #set_defs o #defs o rep_bnf;
+val set_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 wit_thms_of_bnf = maps #prop o wits_of_bnf;
+val wit_thmss_of_bnf = map #prop o wits_of_bnf;
+
+fun mk_bnf name T live lives lives' dead deads map sets bd axioms defs facts wits rel pred =
+ BNF {name = name, T = T,
+ live = live, lives = lives, lives' = lives', dead = dead, deads = deads,
+ map = map, sets = sets, bd = bd,
+ axioms = axioms, defs = defs, facts = facts,
+ nwits = length wits, wits = wits, rel = rel, pred = pred};
+
+fun 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, pred = pred}) =
+ 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, pred = Morphism.term phi pred};
+
+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_relT (instA, instB)))
+ Vartab.empty;
+ in Envir.subst_term (tyenv, Vartab.empty) rel end;
+
+fun normalize_pred ctxt instTs instA instB pred =
+ let
+ val thy = Proof_Context.theory_of ctxt;
+ val tyenv =
+ Sign.typ_match thy (fastype_of pred,
+ Library.foldr (op -->) (instTs, instA --> instB --> HOLogic.boolT)) Vartab.empty;
+ in Envir.subst_term (tyenv, Vartab.empty) pred end;
+
+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 : term) :: todo) =
+ if exists (fn (J, _) => subset (op =) (J, I)) (done @ todo)
+ then minimize done todo
+ else minimize ((I, wit) :: done) todo;
+ in minimize [] wits end;
+
+fun unfold_defs_tac lthy defs mk_tac context = Local_Defs.unfold_tac lthy defs THEN mk_tac context;
+
+
+
+(* Names *)
+
+fun nonzero_string_of_int 0 = ""
+ | nonzero_string_of_int n = string_of_int n;
+
+val mapN = "map";
+val setN = "set";
+fun mk_setN i = setN ^ nonzero_string_of_int i;
+val bdN = "bd";
+val witN = "wit";
+fun mk_witN i = witN ^ nonzero_string_of_int i;
+val relN = "rel";
+val predN = "pred";
+val rel_unfoldN = relN ^ "_unfold";
+val pred_unfoldN = predN ^ "_unfold";
+
+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_congN = "in_cong";
+val in_monoN = "in_mono";
+val in_relN = "in_rel";
+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_wppullN = "map_wppull";
+val map_wpullN = "map_wpull";
+val rel_congN = "rel_cong";
+val rel_IdN = "rel_Id";
+val rel_GrN = "rel_Gr";
+val rel_converseN = "rel_converse";
+val rel_ON = "rel_comp";
+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_Some_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 ctxt =
+ if Config.get ctxt bnf_note_all then Note_All_Facts_and_Axioms
+ else Derive_All_Facts_Note_Most;
+
+val smart_max_inline_size = 25; (*FUDGE*)
+
+val no_def = Drule.reflexive_thm;
+val no_fact = refl;
+
+fun is_reflexive th =
+ let val t = Thm.prop_of th;
+ in
+ op aconv (Logic.dest_equals t)
+ handle TERM _ => op aconv (HOLogic.dest_eq (HOLogic.dest_Trueprop t))
+ handle TERM _ => false
+ end;
+
+val filter_refl = filter_out is_reflexive;
+
+
+
+(* Define new BNFs *)
+
+fun prepare_bnf const_policy mk_fact_policy qualify prep_term Ds_opt
+ ((((raw_b, raw_map), raw_sets), raw_bd_Abs), raw_wits) 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;
+
+ 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;
+
+ fun maybe_define needed_for_fp (b, rhs) lthy =
+ let
+ val inline =
+ (not needed_for_fp orelse fact_policy = Derive_Some_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, no_def), 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 lthy0 lthy = lthy |> not (pointer_eq (lthy0, lthy)) ? Local_Theory.restore;
+
+ 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)) =
+ no_defs_lthy
+ |> maybe_define false map_bind_def
+ ||>> apfst split_list o fold_map (maybe_define false) set_binds_defs
+ ||>> maybe_define false bd_bind_def
+ ||>> apfst split_list o fold_map (maybe_define false) wit_binds_defs
+ ||> `(maybe_restore no_defs_lthy);
+
+ (*transforms defined frees into consts (and more)*)
+ val phi = Proof_Context.export_morphism lthy 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 one_step_defs = filter_refl (bnf_map_def :: bnf_bd_def :: bnf_set_defs @ bnf_wit_defs);
+
+ val _ = case map_filter (try dest_Free)
+ (bnf_map_term :: bnf_set_terms @ [bnf_bd_term] @ bnf_wit_terms) of
+ [] => ()
+ | frees => Proof_Display.print_consts true lthy (K false) frees;
+
+ val bnf_map = Morphism.term phi bnf_map_term;
+
+ fun iter_split ((Ts, T1), T2) = if length Ts < live then error "Bad map function"
+ else if length Ts = live then ((Ts, T1), T2)
+ else iter_split (split_last Ts, T1 --> T2);
+
+ (*TODO: handle errors*)
+ (*simple shape analysis of a map function*)
+ val (((alphas, betas), CA), _) =
+ apfst (apfst (map_split dest_funT))
+ (iter_split (apfst split_last (strip_type (fastype_of bnf_map))));
+
+ 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;
+
+ (*FIXME: check DUP here, not in after_qed*)
+ val key = Name_Space.full_name Name_Space.default_naming b;
+
+ (*TODO: further checks of type of bnf_map*)
+ (*TODO: check types of bnf_sets*)
+ (*TODO: check type of bnf_bd*)
+
+ 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' t =
+ Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As')) t;
+ fun mk_bnf_T As' T =
+ Term.typ_subst_atomic ((deads ~~ Ds) @ (alphas ~~ As')) T;
+
+ 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 (fn T => fn U => T --> U --> HOLogic.boolT) As' Bs';
+
+ 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 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 ((((((((((((((((((((((((fs, fs_copy), gs), hs), (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 (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;
+
+ val goal_map_id =
+ 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 goal_map_comp =
+ 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)
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq (bnf_map_app_comp, comp_bnf_map_app)))
+ end;
+
+ val goal_map_cong =
+ let
+ fun mk_prem z set f f_copy =
+ Logic.all z (Logic.mk_implies
+ (HOLogic.mk_Trueprop (HOLogic.mk_mem (z, set $ x)),
+ HOLogic.mk_Trueprop (HOLogic.mk_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 goal_set_naturals =
+ 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
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq (set_comp_map, image_comp_set)))
+ end;
+ in
+ map3 mk_goal bnf_sets_As bnf_sets_Bs fs
+ end;
+
+ val goal_card_order_bd = HOLogic.mk_Trueprop (mk_card_order bnf_bd_As);
+
+ val goal_cinfinite_bd = HOLogic.mk_Trueprop (mk_cinfinite bnf_bd_As);
+
+ val goal_set_bds =
+ 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 goal_in_bd =
+ 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 goal_map_wpull =
+ 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 goals =
+ [goal_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];
+
+ 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 =
+ Skip_Proof.prove lthy [] []
+ (HOLogic.mk_Trueprop (mk_Card_order (bnf_bd_As)))
+ (fn _ => mk_Card_order_tac (#bd_card_order axioms));
+
+ 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_Some_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
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq (collect_map, image_collect)));
+ in
+ Skip_Proof.prove lthy [] [] goal
+ (fn {context = ctxt, ...} => mk_collect_set_natural_tac ctxt (#set_natural axioms))
+ 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 goal_in_mono =
+ 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 [] [] goal_in_mono (K (mk_in_mono_tac live))
+ 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 goal_in_cong =
+ 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 [] [] goal_in_cong (K ((TRY o hyp_subst_tac THEN' rtac refl) 1))
+ 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);
+
+ (* relator *)
+
+ (*%R1 .. Rn. Gr (in R1 .. Rn) (map fst .. fst)^-1 O Gr (in R1 .. Rn) (map snd .. snd)*)
+ val rel_rhs =
+ 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 Rs (map (mk_bnf_t RTs) bnf_sets) CRs';
+ in
+ fold_rev Term.absfree Rs'
+ (mk_rel_comp (mk_converse (mk_Gr bnf_in map1), mk_Gr bnf_in map2))
+ end;
+ 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 true rel_bind_def
+ ||> `(maybe_restore lthy);
+
+ (*transforms defined frees into consts*)
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+ val bnf_rel = Morphism.term phi bnf_rel_term;
+
+ fun mk_bnf_rel setRTs CA' CB' = normalize_rel lthy setRTs CA' CB' bnf_rel;
+
+ val relAsBs = mk_bnf_rel setRTs CA' CB';
+ val bnf_rel_def = Morphism.thm phi raw_rel_def;
+ val rel_def_unabs =
+ if fact_policy <> Derive_Some_Facts then
+ mk_unabs_def live (bnf_rel_def RS meta_eq_to_obj_eq)
+ else
+ no_fact;
+
+ val pred_rhs = fold absfree (y' :: x' :: rev Qs') (HOLogic.mk_mem (HOLogic.mk_prod (x, y),
+ Term.list_comb (relAsBs, map3 (fn Q => fn T => fn U =>
+ HOLogic.Collect_const (HOLogic.mk_prodT (T, U)) $ HOLogic.mk_split Q)
+ Qs As' Bs')));
+ val pred_bind_def = (fn () => Binding.suffix_name ("_" ^ predN) b, pred_rhs);
+
+ val ((bnf_pred_term, raw_pred_def), (lthy, lthy_old)) =
+ lthy
+ |> maybe_define true pred_bind_def
+ ||> `(maybe_restore lthy);
+
+ (*transforms defined frees into consts*)
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+ val bnf_pred = Morphism.term phi bnf_pred_term;
+
+ fun mk_bnf_pred QTs CA' CB' = normalize_pred lthy QTs CA' CB' bnf_pred;
+
+ val pred = mk_bnf_pred QTs CA' CB';
+ val bnf_pred_def = Morphism.thm phi raw_pred_def;
+ val pred_def_unabs =
+ if fact_policy <> Derive_Some_Facts then
+ mk_unabs_def (live + 2) (bnf_pred_def RS meta_eq_to_obj_eq)
+ else
+ no_fact;
+
+ 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'))
+ end;
+
+ val map_wppull = mk_lazy mk_map_wppull;
+
+ fun mk_rel_Gr () =
+ let
+ val lhs = Term.list_comb (relAsBs, 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)
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)));
+ in
+ Skip_Proof.prove lthy [] [] goal
+ (mk_rel_Gr_tac bnf_rel_def (#map_id axioms) (#map_cong axioms)
+ (#map_wpull axioms) (Lazy.force in_cong) (Lazy.force map_id')
+ (Lazy.force map_comp') (map Lazy.force set_natural'))
+ end;
+
+ val rel_Gr = mk_lazy mk_rel_Gr;
+
+ fun mk_rel_prems f = map2 (HOLogic.mk_Trueprop oo f) Rs Rs_copy
+ fun mk_rel_concl f = HOLogic.mk_Trueprop
+ (f (Term.list_comb (relAsBs, Rs), Term.list_comb (relAsBs, Rs_copy)));
+
+ fun mk_rel_mono () =
+ let
+ val mono_prems = mk_rel_prems mk_subset;
+ val mono_concl = mk_rel_concl (uncurry mk_subset);
+ in
+ Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (mono_prems, mono_concl)))
+ (mk_rel_mono_tac bnf_rel_def (Lazy.force in_mono))
+ end;
+
+ fun mk_rel_cong () =
+ let
+ val cong_prems = mk_rel_prems (curry HOLogic.mk_eq);
+ val cong_concl = mk_rel_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)
+ end;
+
+ val rel_mono = mk_lazy mk_rel_mono;
+ val rel_cong = mk_lazy mk_rel_cong;
+
+ fun mk_rel_Id () =
+ let val relAsAs = mk_bnf_rel 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_rel_Id_tac live (Lazy.force rel_Gr) (#map_id axioms))
+ end;
+
+ val rel_Id = mk_lazy mk_rel_Id;
+
+ fun mk_rel_converse () =
+ let
+ val relBsAs = mk_bnf_rel setRT's CB' CA';
+ val lhs = Term.list_comb (relBsAs, map mk_converse Rs);
+ val rhs = mk_converse (Term.list_comb (relAsBs, 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_rel_converse_le_tac bnf_rel_def (Lazy.force rel_Id) (#map_cong axioms)
+ (Lazy.force map_comp') (map Lazy.force set_natural'))
+ val goal = fold_rev Logic.all Rs (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)));
+ in
+ Skip_Proof.prove lthy [] [] goal (fn _ => mk_rel_converse_tac le_thm)
+ end;
+
+ val rel_converse = mk_lazy mk_rel_converse;
+
+ fun mk_rel_O () =
+ let
+ val relAsCs = mk_bnf_rel setRTsAsCs CA' CC';
+ val relBsCs = mk_bnf_rel setRTsBsCs CB' CC';
+ val lhs = Term.list_comb (relAsCs, map2 (curry mk_rel_comp) Rs Ss);
+ val rhs = mk_rel_comp (Term.list_comb (relAsBs, Rs), Term.list_comb (relBsCs, Ss));
+ val goal =
+ fold_rev Logic.all (Rs @ Ss) (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)));
+ in
+ Skip_Proof.prove lthy [] [] goal
+ (mk_rel_O_tac bnf_rel_def (Lazy.force rel_Id) (#map_cong axioms)
+ (Lazy.force map_wppull) (Lazy.force map_comp') (map Lazy.force set_natural'))
+ end;
+
+ val rel_O = mk_lazy mk_rel_O;
+
+ fun mk_in_rel () =
+ 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 (relAsBs, 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) (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)));
+ in
+ Skip_Proof.prove lthy [] [] goal (mk_in_rel_tac bnf_rel_def (length bnf_sets))
+ end;
+
+ val in_rel = mk_lazy mk_in_rel;
+
+ val defs = mk_defs bnf_map_def bnf_set_defs rel_def_unabs pred_def_unabs;
+
+ val facts = mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_natural
+ in_cong in_mono in_rel map_comp' map_id' map_wppull
+ rel_cong rel_mono rel_Id rel_Gr rel_converse rel_O set_natural';
+
+ val wits = map2 mk_witness bnf_wits wit_thms;
+
+ val bnf_rel = Term.subst_atomic_types
+ ((Ds ~~ deads) @ (As' ~~ alphas) @ (Bs' ~~ betas)) relAsBs;
+ val bnf_pred = Term.subst_atomic_types
+ ((Ds ~~ deads) @ (As' ~~ alphas) @ (Bs' ~~ betas)) pred;
+
+ val bnf = mk_bnf b CA live alphas betas dead deads bnf_map bnf_sets bnf_bd axioms defs facts
+ wits bnf_rel bnf_pred;
+
+ fun note thmN thms =
+ snd o Local_Theory.note
+ ((qualify (Binding.qualify true (Binding.name_of b) (Binding.name thmN)), []), thms);
+
+ fun lazy_note thmN thms = note thmN (map Lazy.force thms);
+ 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);
+ in
+ note bd_card_orderN [#bd_card_order axioms]
+ #> note bd_cinfiniteN [#bd_cinfinite axioms]
+ #> note bd_Card_orderN [#bd_Card_order facts]
+ #> note bd_CinfiniteN [#bd_Cinfinite facts]
+ #> note bd_CnotzeroN [#bd_Cnotzero facts]
+ #> lazy_note collect_set_naturalN [#collect_set_natural facts]
+ #> note in_bdN [#in_bd axioms]
+ #> lazy_note in_monoN [#in_mono facts]
+ #> lazy_note in_relN [#in_rel facts]
+ #> note map_compN [#map_comp axioms]
+ #> note map_idN [#map_id axioms]
+ #> note map_wpullN [#map_wpull axioms]
+ #> note set_naturalN (#set_natural axioms)
+ #> note set_bdN (#set_bd axioms)
+ #> fold2 note witNs wit_thms
+ end
+ else
+ I)
+ |> (if fact_policy = Note_All_Facts_and_Axioms orelse
+ fact_policy = Derive_All_Facts_Note_Most then
+ lazy_note rel_IdN [#rel_Id facts]
+ #> lazy_note rel_GrN [#rel_Gr facts]
+ #> lazy_note rel_converseN [#rel_converse facts]
+ #> lazy_note rel_ON [#rel_O facts]
+ #> lazy_note map_id'N [#map_id' facts]
+ #> lazy_note map_comp'N [#map_comp' facts]
+ #> note map_congN [#map_cong axioms]
+ #> lazy_note set_natural'N (#set_natural' facts)
+ #> Local_Theory.declaration {syntax = false, pervasive = true}
+ (fn phi => Data.map (Symtab.update_new (key, morph_bnf phi bnf)))
+ else
+ I))
+ end;
+ in
+ (goals, wit_goalss, after_qed, lthy', one_step_defs)
+ end;
+
+fun add_bnf const_policy fact_policy qualify tacs wit_tac Ds =
+ (fn (goals, wit_goalss, after_qed, lthy, defs) =>
+ let
+ val wits_tac = K (TRYALL Goal.conjunction_tac) THEN' unfold_defs_tac lthy defs wit_tac;
+ val wit_goals = wit_goalss |> map Logic.mk_conjunction_balanced;
+ val wit_goal = Logic.mk_conjunction_balanced wit_goals;
+ val wit_thms =
+ Skip_Proof.prove lthy [] [] wit_goal wits_tac
+ |> Conjunction.elim_balanced (length wit_goals)
+ |> map2 (Conjunction.elim_balanced o length) wit_goalss
+ |> map (map (Thm.forall_elim_vars 0));
+ in
+ (map2 (Skip_Proof.prove lthy [] []) goals (map (unfold_defs_tac lthy defs) tacs))
+ |> (fn thms => after_qed (map single thms @ wit_thms) lthy)
+ end) oo prepare_bnf const_policy fact_policy qualify
+ (fn ctxt => singleton (Type_Infer_Context.infer_types ctxt)) Ds;
+
+val add_bnf_cmd = (fn (goals, wit_goals, after_qed, lthy, defs) =>
+ Proof.unfolding ([[(defs, [])]])
+ (Proof.theorem NONE (snd oo after_qed)
+ (map (single o rpair []) goals @ map (map (rpair [])) wit_goals) lthy)) oo
+ prepare_bnf Do_Inline user_policy 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.binding --| Parse.$$$ "=") -- Parse.term --
+ (Parse.$$$ "[" |-- Parse.list Parse.term --| Parse.$$$ "]") -- Parse.term --
+ (Parse.$$$ "[" |-- Parse.list Parse.term --| Parse.$$$ "]")) >> add_bnf_cmd);
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Codatatype/Tools/bnf_fp_util.ML Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,279 @@
+(* Title: HOL/Codatatype/Tools/bnf_fp_util.ML
+ Author: Dmitriy Traytel, TU Muenchen
+ Copyright 2012
+
+Shared library for the datatype and the codatatype construction.
+*)
+
+signature BNF_FP_UTIL =
+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 coN: string
+ val coinductN: string
+ val coiterN: string
+ val coiter_uniqueN: string
+ val corecN: string
+ val fldN: string
+ val fld_coiterN: string
+ val fld_exhaustN: string
+ val fld_induct2N: string
+ val fld_inductN: string
+ val fld_injectN: string
+ val fld_unfN: string
+ val hsetN: string
+ val hset_recN: string
+ val inductN: string
+ val isNodeN: string
+ val iterN: string
+ val iter_uniqueN: string
+ val lsbisN: string
+ val map_simpsN: string
+ val map_uniqueN: string
+ val min_algN: string
+ val morN: string
+ val pred_coinductN: string
+ val pred_coinduct_uptoN: string
+ val recN: string
+ val rel_coinductN: string
+ val rel_coinduct_uptoN: string
+ val rvN: string
+ val set_inclN: string
+ val set_set_inclN: string
+ val strTN: string
+ val str_initN: string
+ val sum_bdN: string
+ val sum_bdTN: string
+ val unfN: string
+ val unf_coinductN: string
+ val unf_coinduct_uptoN: string
+ val unf_exhaustN: string
+ val unf_fldN: string
+ val unf_injectN: string
+ val uniqueN: string
+ val uptoN: 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 typedef: bool -> binding option -> binding * (string * sort) list * mixfix -> term ->
+ (binding * binding) option -> tactic -> local_theory -> (string * Typedef.info) * local_theory
+
+ val split_conj_thm: thm -> thm list
+ val split_conj_prems: int -> thm -> thm
+
+ val list_all_free: term list -> term -> term
+ val list_exists_free: term list -> term -> term
+
+ val mk_Field: term -> term
+ val mk_union: term * term -> term
+
+ val mk_tactics: 'a -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a list
+
+ val fixpoint: ('a * 'a -> bool) -> ('a list -> 'a list) -> 'a list -> 'a list
+ val interleave: 'a list -> 'a list -> 'a list
+
+ val certifyT: Proof.context -> typ -> ctyp
+ val certify: Proof.context -> term -> cterm
+
+ val fp_bnf: (binding list -> typ list list -> BNF_Def.BNF list ->
+ Proof.context -> Proof.context) ->
+ binding list -> ((string * sort) * typ) list -> Proof.context -> Proof.context
+ val fp_bnf_cmd: (binding list -> typ list list -> BNF_Def.BNF list ->
+ Proof.context -> Proof.context) ->
+ binding list * (string list * string list) -> Proof.context -> Proof.context
+end;
+
+structure BNF_FP_Util : BNF_FP_UTIL =
+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 ());
+
+(*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;
+
+val coN = "co"
+val algN = "alg"
+val IITN = "IITN"
+val iterN = "iter"
+val coiterN = coN ^ iterN
+val uniqueN = "_unique"
+val iter_uniqueN = iterN ^ uniqueN
+val coiter_uniqueN = coiterN ^ uniqueN
+val fldN = "fld"
+val unfN = "unf"
+val fld_coiterN = fldN ^ "_" ^ coiterN
+val map_simpsN = mapN ^ "_simps"
+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 ^ "_simps"
+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 corecN = coN ^ recN
+
+val fld_unfN = fldN ^ "_" ^ unfN
+val unf_fldN = unfN ^ "_" ^ fldN
+fun mk_nchotomyN s = s ^ "_nchotomy"
+fun mk_injectN s = s ^ "_inject"
+fun mk_exhaustN s = s ^ "_exhaust"
+val fld_injectN = mk_injectN fldN
+val fld_exhaustN = mk_exhaustN fldN
+val unf_injectN = mk_injectN unfN
+val unf_exhaustN = mk_exhaustN unfN
+val inductN = "induct"
+val coinductN = coN ^ inductN
+val fld_inductN = fldN ^ "_" ^ inductN
+val fld_induct2N = fld_inductN ^ "2"
+val unf_coinductN = unfN ^ "_" ^ coinductN
+val rel_coinductN = relN ^ "_" ^ coinductN
+val pred_coinductN = predN ^ "_" ^ coinductN
+val uptoN = "upto"
+val unf_coinduct_uptoN = unf_coinductN ^ "_" ^ uptoN
+val rel_coinduct_uptoN = rel_coinductN ^ "_" ^ uptoN
+val pred_coinduct_uptoN = pred_coinductN ^ "_" ^ uptoN
+val hsetN = "Hset"
+val hset_recN = hsetN ^ "_rec"
+val set_inclN = "set_incl"
+val set_set_inclN = "set_set_incl"
+
+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);
+
+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);
+
+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);
+
+(*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);
+
+(* 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_tactics mid mcomp mcong snat bdco bdinf sbd inbd wpull =
+ [mid, mcomp, mcong] @ snat @ [bdco, bdinf] @ sbd @ [inbd, wpull];
+
+fun interleave xs ys = flat (map2 (fn x => fn y => [x, y]) xs ys);
+
+fun fp_sort lhss Ass = Library.sort (Term_Ord.typ_ord o pairself TFree)
+ (subtract (op =) lhss (fold (fold (insert (op =))) Ass [])) @ lhss;
+
+fun mk_fp_bnf timer construct bs sort bnfs deads lives unfold lthy =
+ let
+ (* TODO: assert that none of the deads is a lhs *)
+
+ val name = fold_rev (fn b => fn s => Binding.name_of b ^ s) bs "";
+ fun qualify i bind =
+ let val namei = if i > 0 then name ^ string_of_int i else name;
+ in
+ if member (op =) (#2 (Binding.dest bind)) (namei, true) then bind
+ else Binding.prefix_name namei bind
+ end;
+
+ val Ass = map (map dest_TFree) lives;
+ val Ds = fold (fold Term.add_tfreesT) deads [];
+
+ val timer = time (timer "Construction of BNFs");
+
+ val ((kill_poss, _), (bnfs', (unfold', lthy'))) =
+ normalize_bnfs qualify Ass Ds sort bnfs unfold lthy;
+
+ val Dss = map3 (append oo map o nth) lives kill_poss deads;
+
+ val (bnfs'', lthy'') =
+ fold_map3 (seal_bnf unfold') (map (Binding.suffix_name "BNF") bs) Dss bnfs' lthy';
+
+ val timer = time (timer "Normalization & sealing of BNFs");
+
+ val res = construct bs Dss bnfs'' lthy'';
+
+ val timer = time (timer "FP construction in total");
+ in
+ res
+ end;
+
+fun fp_bnf construct bs eqs lthy =
+ let
+ val timer = time (Timer.startRealTimer ());
+ val (lhss, rhss) = split_list eqs;
+ val sort = fp_sort lhss;
+ val ((bnfs, (Dss, Ass)), (unfold, lthy')) = apfst (apsnd split_list o split_list)
+ (fold_map2 (fn b => bnf_of_typ Smart_Inline (Binding.suffix_name "RAW" b) I sort) bs rhss
+ (empty_unfold, lthy));
+ in
+ mk_fp_bnf timer construct bs sort bnfs Dss Ass unfold 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;
+ val ((bnfs, (Dss, Ass)), (unfold, lthy')) = apfst (apsnd split_list o split_list)
+ (fold_map2 (fn b => fn rawT =>
+ (bnf_of_typ Smart_Inline (Binding.suffix_name "RAW" b) I sort (Syntax.read_typ lthy rawT)))
+ bs raw_bnfs (empty_unfold, lthy));
+ in
+ mk_fp_bnf timer construct bs sort bnfs Dss Ass unfold lthy'
+ end;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Codatatype/Tools/bnf_gfp.ML Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,2784 @@
+(* Title: HOL/Codatatype/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: binding list -> typ list list -> BNF_Def.BNF list -> Proof.context -> Proof.context
+end;
+
+structure BNF_GFP : BNF_GFP =
+struct
+
+open BNF_Def
+open BNF_Util
+open BNF_Tactics
+open BNF_FP_Util
+open BNF_GFP_Util
+open BNF_GFP_Tactics
+
+(*all bnfs have the same lives*)
+fun bnf_gfp bs Dss_insts 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 (fold_rev (fn b => fn s => Binding.name_of b ^ s) bs "");
+
+ fun note thmN thms = snd o Local_Theory.note
+ ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), thms);
+ fun notes thmN thmss = fold2 (fn b => fn thms => snd o Local_Theory.note
+ ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), thms)) bs thmss;
+
+ (* TODO: check if m, n etc are sane *)
+
+ val Dss = map (fn Ds => map TFree (fold Term.add_tfreesT Ds [])) Dss_insts;
+ val deads = distinct (op =) (flat Dss);
+ 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 *)
+ fun mk_FTs Ts = map2 (fn Ds => mk_T_of_bnf Ds Ts) Dss bnfs;
+ val (params, params') = `(map dest_TFree) (deads @ 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_rel_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 witss = map wits_of_bnf bnfs;
+
+ 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 pred_defs = map pred_def_of_bnf bnfs;
+ val rel_congs = map rel_cong_of_bnf bnfs;
+ val rel_converses = map rel_converse_of_bnf bnfs;
+ val rel_defs = map rel_def_of_bnf bnfs;
+ val rel_Grs = map rel_Gr_of_bnf bnfs;
+ val rel_Ids = map rel_Id_of_bnf bnfs;
+ val rel_monos = map rel_mono_of_bnf bnfs;
+ val rel_Os = map rel_O_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");
+
+ (* 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) (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))))
+ (K (mk_map_comp_id_tac map_comp))
+ 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 = HOLogic.mk_Trueprop (HOLogic.mk_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'))
+ 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 in_cong'_thms = map (fn bnf => in_cong_of_bnf bnf OF (replicate m refl)) bnfs;
+
+ val map_arg_cong_thms =
+ let
+ val prems = map2 (fn x => fn y =>
+ HOLogic.mk_Trueprop (HOLogic.mk_eq (x, y))) xFs xFs_copy;
+ val maps = map (fn map => Term.list_comb (map, all_fs)) mapsAsBs;
+ val concls = map3 (fn x => fn y => fn map =>
+ HOLogic.mk_Trueprop (HOLogic.mk_eq (map $ x, map $ 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))) 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
+ HOLogic.mk_Trueprop (HOLogic.mk_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;
+
+ (*transforms defined frees into consts*)
+ 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))) 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))
+ 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
+ HOLogic.mk_Trueprop (HOLogic.mk_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;
+
+ (*transforms defined frees into consts*)
+ 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))],
+ HOLogic.mk_Trueprop (HOLogic.mk_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));
+ 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))
+ 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))
+ 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))
+ 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) (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))))
+ (K (mk_mor_UNIV_tac morE_thms mor_def))
+ 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))
+ end;
+
+ val mor_sum_case_thm =
+ let
+ val maps = map3 (fn s => fn sum_s => fn map =>
+ mk_sum_case (HOLogic.mk_comp (Term.list_comb (map, 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))
+ 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
+ HOLogic.mk_Trueprop (HOLogic.mk_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;
+
+ (*transforms defined frees into consts*)
+ 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
+ HOLogic.mk_Trueprop (HOLogic.mk_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;
+
+ (*transforms defined frees into consts*)
+ 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)))
+ 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)))
+ 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))))
+ 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)))
+ 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))
+ 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))))
+ 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 = HOLogic.mk_Trueprop (HOLogic.mk_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))))
+ 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
+ HOLogic.mk_Trueprop (HOLogic.mk_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;
+
+ (*transforms defined frees into consts*)
+ 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))
+ end;
+
+ val bis_rel_thm =
+ let
+ fun mk_conjunct R s s' b1 b2 rel =
+ list_all_free [b1, b2] (HOLogic.mk_imp
+ (HOLogic.mk_mem (HOLogic.mk_prod (b1, b2), R),
+ HOLogic.mk_mem (HOLogic.mk_prod (s $ b1, s' $ b2),
+ Term.list_comb (rel, 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)
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq (mk_bis As Bs ss B's s's Rs, rhs))))
+ (K (mk_bis_rel_tac m bis_def rel_defs map_comp's map_congs set_natural'ss))
+ 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_rel_thm rel_congs rel_converses));
+
+ 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_rel_thm rel_congs rel_Os))
+ 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_rel_thm rel_Grs mor_image_thms morE_thms coalg_in_thms)
+ 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)
+ 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
+ HOLogic.mk_Trueprop (HOLogic.mk_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;
+
+ (*transforms defined frees into consts*)
+ 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));
+
+ 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))) 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)))
+ 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 true 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;
+
+ (*transforms defined frees into consts*)
+ 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 sbdT_set_def = the (#set_def sbdT_loc_info);
+ val sbdT_Abs_inject = #Abs_inject sbdT_loc_info;
+ val sbdT_Abs_cases = #Abs_cases sbdT_loc_info;
+
+ val Abs_sbdT_inj = mk_Abs_inj_thm sbdT_set_def sbdT_Abs_inject;
+ val Abs_sbdT_bij = mk_Abs_bij_thm lthy sbdT_set_def sbdT_Abs_inject sbdT_Abs_cases;
+
+ 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 = Local_Defs.fold lthy [sbd_def]
+ (@{thm dir_image} OF [Abs_sbdT_inj, sum_Card_order]);
+ val sbd_card_order = Local_Defs.fold 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
+ HOLogic.mk_Trueprop (HOLogic.mk_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;
+
+ (*transforms defined frees into consts*)
+ 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
+ HOLogic.mk_Trueprop (HOLogic.mk_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;
+
+ (*transforms defined frees into consts*)
+ 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
+ HOLogic.mk_Trueprop (HOLogic.mk_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;
+
+ (*transforms defined frees into consts*)
+ 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);
+
+ 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))
+ 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))
+ 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))))
+ 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))
+ 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
+ HOLogic.mk_Trueprop (HOLogic.mk_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;
+
+ (*transforms defined frees into consts*)
+ 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
+ HOLogic.mk_Trueprop (HOLogic.mk_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;
+
+ (*transforms defined frees into consts*)
+ 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
+ HOLogic.mk_Trueprop (HOLogic.mk_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;
+
+ (*transforms defined frees into consts*)
+ 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)));
+
+ 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)));
+
+ 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))) 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)));
+
+ 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)));
+
+ 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)));
+
+ 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_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)));
+
+ 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)));
+
+ 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);
+
+ 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)))
+ 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));
+
+ 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));
+
+ 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_map3 (fn b => fn car_final => fn in_car_final =>
+ typedef false NONE (b, params, NoSyn) car_final NONE
+ (EVERY' [rtac exI, rtac in_car_final] 1)) bs 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 FTs'_setss = mk_setss (passiveBs @ 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 unfTs = map2 (curry (op -->)) Ts FTs;
+ val fldTs = map2 (curry (op -->)) FTs Ts;
+ val coiter_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), coiter_fs), coiter_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" coiter_fTs
+ ||>> mk_Frees "g" coiter_fTs
+ ||>> mk_Frees "s" corec_sTs
+ ||>> mk_Frees "phi" (map (fn T => T --> T --> HOLogic.boolT) Ts);
+
+ fun unf_bind i = Binding.suffix_name ("_" ^ unfN) (nth bs (i - 1));
+ val unf_name = Binding.name_of o unf_bind;
+ val unf_def_bind = rpair [] o Thm.def_binding o unf_bind;
+
+ fun unf_spec i rep str map_FT unfT Jz Jz' =
+ let
+ val lhs = Free (unf_name i, unfT);
+ val rhs = Term.absfree Jz'
+ (Term.list_comb (map_FT, map HOLogic.id_const passiveAs @ Abs_Ts) $
+ (str $ (rep $ Jz)));
+ in
+ HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))
+ end;
+
+ val ((unf_frees, (_, unf_def_frees)), (lthy, lthy_old)) =
+ lthy
+ |> fold_map7 (fn i => fn rep => fn str => fn map => fn unfT => fn Jz => fn Jz' =>
+ Specification.definition
+ (SOME (unf_bind i, NONE, NoSyn), (unf_def_bind i, unf_spec i rep str map unfT Jz Jz')))
+ ks Rep_Ts str_finals map_FTs unfTs Jzs Jzs'
+ |>> apsnd split_list o split_list
+ ||> `Local_Theory.restore;
+
+ (*transforms defined frees into consts*)
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+ fun mk_unfs passive =
+ map (Term.subst_atomic_types (map (Morphism.typ phi) params' ~~ (deads @ passive)) o
+ Morphism.term phi) unf_frees;
+ val unfs = mk_unfs passiveAs;
+ val unf's = mk_unfs passiveBs;
+ val unf_defs = map ((fn thm => thm RS fun_cong) o Morphism.thm phi) unf_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 unfs car_finals str_finals Rep_Ts))
+ (mk_mor_Rep_tac m (mor_def :: unf_defs) Reps Abs_inverses coalg_final_set_thmss
+ map_comp_id_thms map_congL_thms);
+
+ val mor_Abs =
+ Skip_Proof.prove lthy [] []
+ (HOLogic.mk_Trueprop (mk_mor car_finals str_finals UNIVs unfs Abs_Ts))
+ (mk_mor_Abs_tac (mor_def :: unf_defs) Abs_inverses);
+ in
+ (mor_Rep, mor_Abs)
+ end;
+
+ val timer = time (timer "unf definitions & thms");
+
+ fun coiter_bind i = Binding.suffix_name ("_" ^ coN ^ iterN) (nth bs (i - 1));
+ val coiter_name = Binding.name_of o coiter_bind;
+ val coiter_def_bind = rpair [] o Thm.def_binding o coiter_bind;
+
+ fun coiter_spec i T AT abs f z z' =
+ let
+ val coiterT = Library.foldr (op -->) (sTs, AT --> T);
+
+ val lhs = Term.list_comb (Free (coiter_name i, coiterT), ss);
+ val rhs = Term.absfree z' (abs $ (f $ z));
+ in
+ HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))
+ end;
+
+ val ((coiter_frees, (_, coiter_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 (coiter_bind i, NONE, NoSyn), (coiter_def_bind i, coiter_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;
+
+ (*transforms defined frees into consts*)
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+ val coiters = map (fst o dest_Const o Morphism.term phi) coiter_frees;
+ fun mk_coiter Ts ss i = Term.list_comb (Const (nth coiters (i - 1), Library.foldr (op -->)
+ (map fastype_of ss, domain_type (fastype_of (nth ss (i - 1))) --> nth Ts (i - 1))), ss);
+ val coiter_defs = map ((fn thm => thm RS fun_cong) o Morphism.thm phi) coiter_def_frees;
+
+ val mor_coiter_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 unfs (map (mk_coiter Ts ss) ks))))
+ (K (mk_mor_coiter_tac m mor_UNIV_thm unf_defs coiter_defs Abs_inverses' morEs'
+ map_comp_id_thms map_congs))
+ end;
+ val coiter_thms = map (fn thm => (thm OF [mor_coiter_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 unfs 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)))
+ 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 unfs coiter_fs,
+ mk_mor Bs ss UNIVs unfs coiter_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 coiter_fs coiter_fs_copy zs));
+
+ val unique_mor = Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all (Bs @ ss @ coiter_fs @ coiter_fs_copy @ zs)
+ (Logic.list_implies (prems, unique)))
+ (K (mk_unique_mor_tac raw_coind_thms bis_image2_thm));
+ in
+ map (fn thm => conjI RSN (2, thm RS mp)) (split_conj_thm unique_mor)
+ end;
+
+ val (coiter_unique_mor_thms, coiter_unique_mor_thm) =
+ let
+ val prem = HOLogic.mk_Trueprop (mk_mor active_UNIVs ss UNIVs unfs coiter_fs);
+ fun mk_fun_eq f i = HOLogic.mk_eq (f, mk_coiter Ts ss i);
+ val unique = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+ (map2 mk_fun_eq coiter_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 @ coiter_fs) (Logic.mk_implies (prem, unique)))
+ (K (mk_coiter_unique_mor_tac raw_coind_thms bis_thm mor_thm coiter_defs));
+ in
+ `split_conj_thm unique_mor
+ end;
+
+ val (coiter_unique_thms, coiter_unique_thm) = `split_conj_thm (split_conj_prems n
+ (mor_UNIV_thm RS @{thm ssubst[of _ _ "%x. x"]} RS coiter_unique_mor_thm));
+
+ val coiter_unf_thms = map (fn thm => mor_id_thm RS thm RS sym) coiter_unique_mor_thms;
+
+ val coiter_o_unf_thms =
+ let
+ val mor = mor_comp_thm OF [mor_str_thm, mor_coiter_thm];
+ in
+ map2 (fn unique => fn coiter_fld =>
+ trans OF [mor RS unique, coiter_fld]) coiter_unique_mor_thms coiter_unf_thms
+ end;
+
+ val timer = time (timer "coiter definitions & thms");
+
+ val map_unfs = map2 (fn Ds => fn bnf =>
+ Term.list_comb (mk_map_of_bnf Ds (passiveAs @ Ts) (passiveAs @ FTs) bnf,
+ map HOLogic.id_const passiveAs @ unfs)) Dss bnfs;
+
+ fun fld_bind i = Binding.suffix_name ("_" ^ fldN) (nth bs (i - 1));
+ val fld_name = Binding.name_of o fld_bind;
+ val fld_def_bind = rpair [] o Thm.def_binding o fld_bind;
+
+ fun fld_spec i fldT =
+ let
+ val lhs = Free (fld_name i, fldT);
+ val rhs = mk_coiter Ts map_unfs i;
+ in
+ HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))
+ end;
+
+ val ((fld_frees, (_, fld_def_frees)), (lthy, lthy_old)) =
+ lthy
+ |> fold_map2 (fn i => fn fldT =>
+ Specification.definition
+ (SOME (fld_bind i, NONE, NoSyn), (fld_def_bind i, fld_spec i fldT))) ks fldTs
+ |>> apsnd split_list o split_list
+ ||> `Local_Theory.restore;
+
+ (*transforms defined frees into consts*)
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+ fun mk_flds params =
+ map (Term.subst_atomic_types (map (Morphism.typ phi) params' ~~ params) o Morphism.term phi)
+ fld_frees;
+ val flds = mk_flds params';
+ val fld_defs = map (Morphism.thm phi) fld_def_frees;
+
+ val fld_o_unf_thms = map2 (Local_Defs.fold lthy o single) fld_defs coiter_o_unf_thms;
+
+ val unf_o_fld_thms =
+ let
+ fun mk_goal unf fld FT =
+ HOLogic.mk_Trueprop (HOLogic.mk_eq (HOLogic.mk_comp (unf, fld), HOLogic.id_const FT));
+ val goals = map3 mk_goal unfs flds FTs;
+ in
+ map5 (fn goal => fn fld_def => fn coiter => fn map_comp_id => fn map_congL =>
+ Skip_Proof.prove lthy [] [] goal
+ (mk_unf_o_fld_tac fld_def coiter map_comp_id map_congL coiter_o_unf_thms))
+ goals fld_defs coiter_thms map_comp_id_thms map_congL_thms
+ end;
+
+ val unf_fld_thms = map (fn thm => thm RS @{thm pointfree_idE}) unf_o_fld_thms;
+ val fld_unf_thms = map (fn thm => thm RS @{thm pointfree_idE}) fld_o_unf_thms;
+
+ val bij_unf_thms =
+ map2 (fn thm1 => fn thm2 => @{thm o_bij} OF [thm1, thm2]) fld_o_unf_thms unf_o_fld_thms;
+ val inj_unf_thms = map (fn thm => thm RS @{thm bij_is_inj}) bij_unf_thms;
+ val surj_unf_thms = map (fn thm => thm RS @{thm bij_is_surj}) bij_unf_thms;
+ val unf_nchotomy_thms = map (fn thm => thm RS @{thm surjD}) surj_unf_thms;
+ val unf_inject_thms = map (fn thm => thm RS @{thm inj_eq}) inj_unf_thms;
+ val unf_exhaust_thms = map (fn thm => thm RS @{thm exE}) unf_nchotomy_thms;
+
+ val bij_fld_thms =
+ map2 (fn thm1 => fn thm2 => @{thm o_bij} OF [thm1, thm2]) unf_o_fld_thms fld_o_unf_thms;
+ val inj_fld_thms = map (fn thm => thm RS @{thm bij_is_inj}) bij_fld_thms;
+ val surj_fld_thms = map (fn thm => thm RS @{thm bij_is_surj}) bij_fld_thms;
+ val fld_nchotomy_thms = map (fn thm => thm RS @{thm surjD}) surj_fld_thms;
+ val fld_inject_thms = map (fn thm => thm RS @{thm inj_eq}) inj_fld_thms;
+ val fld_exhaust_thms = map (fn thm => thm RS @{thm exE}) fld_nchotomy_thms;
+
+ val fld_coiter_thms = map3 (fn unf_inject => fn coiter => fn unf_fld =>
+ iffD1 OF [unf_inject, trans OF [coiter, unf_fld RS sym]])
+ unf_inject_thms coiter_thms unf_fld_thms;
+
+ val timer = time (timer "fld definitions & thms");
+
+ val corec_Inl_sum_thms =
+ let
+ val mor = mor_comp_thm OF [mor_sum_case_thm, mor_coiter_thm];
+ in
+ map2 (fn unique => fn coiter_unf =>
+ trans OF [mor RS unique, coiter_unf]) coiter_unique_mor_thms coiter_unf_thms
+ end;
+
+ fun corec_bind i = Binding.suffix_name ("_" ^ coN ^ recN) (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 unf => fn sum_s => fn map => mk_sum_case
+ (HOLogic.mk_comp (Term.list_comb (map, passive_ids @ corec_Inls), unf)) sum_s)
+ unfs corec_ss corec_maps;
+
+ val lhs = Term.list_comb (Free (corec_name i, corecT), corec_ss);
+ val rhs = HOLogic.mk_comp (mk_coiter Ts maps i, Inr_const T AT);
+ in
+ HOLogic.mk_Trueprop (HOLogic.mk_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;
+
+ (*transforms defined frees into consts*)
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+ val corecs = map (fst o dest_Const o Morphism.term phi) corec_frees;
+ fun mk_corec ss i = Term.list_comb (Const (nth corecs (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 corec_thms =
+ let
+ fun mk_goal i corec_s corec_map unf z =
+ let
+ val lhs = unf $ (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) (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)))
+ end;
+ val goals = map5 mk_goal ks corec_ss corec_maps_rev unfs zs;
+ in
+ map3 (fn goal => fn coiter => fn map_cong =>
+ Skip_Proof.prove lthy [] [] goal
+ (mk_corec_tac m corec_defs coiter map_cong corec_Inl_sum_thms))
+ goals coiter_thms map_congs
+ end;
+
+ val timer = time (timer "corec definitions & thms");
+
+ val (unf_coinduct_thm, coinduct_params, rel_coinduct_thm, pred_coinduct_thm,
+ unf_coinduct_upto_thm, rel_coinduct_upto_thm, pred_coinduct_upto_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_rels 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 rels = map (Term.subst_atomic_types ((activeAs ~~ Ts) @ (activeBs ~~ Ts))) relsAsBs;
+
+ fun mk_concl phi z1 z2 = HOLogic.mk_imp (phi $ z1 $ z2, HOLogic.mk_eq (z1, z2));
+ val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj
+ (map3 mk_concl phis Jzs1 Jzs2));
+
+ fun mk_rel_prem upto_eq phi unf rel Jz Jz_copy =
+ let
+ val concl = HOLogic.mk_mem (HOLogic.mk_tuple [unf $ Jz, unf $ Jz_copy],
+ Term.list_comb (rel, mk_Ids upto_eq @ phi_rels upto_eq));
+ in
+ HOLogic.mk_Trueprop
+ (list_all_free [Jz, Jz_copy] (HOLogic.mk_imp (phi $ Jz $ Jz_copy, concl)))
+ end;
+
+ val rel_prems = map5 (mk_rel_prem false) phis unfs rels Jzs Jzs_copy;
+ val rel_upto_prems = map5 (mk_rel_prem true) phis unfs rels Jzs Jzs_copy;
+
+ val rel_coinduct_goal = fold_rev Logic.all frees (Logic.list_implies (rel_prems, concl));
+ val coinduct_params = rev (Term.add_tfrees rel_coinduct_goal []);
+
+ val rel_coinduct = Local_Defs.unfold lthy @{thms diag_UNIV}
+ (Skip_Proof.prove lthy [] [] rel_coinduct_goal
+ (K (mk_rel_coinduct_tac ks raw_coind_thm bis_rel_thm)));
+
+ fun mk_unf_prem upto_eq phi unf 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, unf $ 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_unf_prems upto_eq =
+ map7 (mk_unf_prem upto_eq) phis unfs map_FT_nths prodFT_setss Jzs Jzs_copy FJzs
+
+ val unf_prems = mk_unf_prems false;
+ val unf_upto_prems = mk_unf_prems true;
+
+ val unf_coinduct_goal = fold_rev Logic.all frees (Logic.list_implies (unf_prems, concl));
+ val unf_coinduct = Skip_Proof.prove lthy [] [] unf_coinduct_goal
+ (K (mk_unf_coinduct_tac m ks raw_coind_thm bis_def));
+
+ 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 rel_coinduct_upto = singleton (Proof_Context.export names_lthy lthy)
+ (Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all zs (Logic.list_implies (rel_upto_prems, concl)))
+ (K (mk_rel_coinduct_upto_tac m cTs cts rel_coinduct rel_monos rel_Ids)));
+
+ val unf_coinduct_upto = singleton (Proof_Context.export names_lthy lthy)
+ (Skip_Proof.prove lthy [] []
+ (fold_rev Logic.all zs (Logic.list_implies (unf_upto_prems, concl)))
+ (K (mk_unf_coinduct_upto_tac ks cTs cts unf_coinduct bis_def
+ (tcoalg_thm RS bis_diag_thm))));
+
+ val pred_coinduct = rel_coinduct
+ |> Local_Defs.unfold lthy @{thms Id_def'}
+ |> Local_Defs.fold lthy pred_defs;
+ val pred_coinduct_upto = rel_coinduct_upto
+ |> Local_Defs.unfold lthy @{thms Id_def'}
+ |> Local_Defs.fold lthy pred_defs;
+ in
+ (unf_coinduct, rev (Term.add_tfrees unf_coinduct_goal []), rel_coinduct, pred_coinduct,
+ unf_coinduct_upto, rel_coinduct_upto, pred_coinduct_upto)
+ 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 (fn T => fn U => T --> U --> HOLogic.boolT) 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, fs'_copy)), (gs, gs')), us),
+ (Jys, Jys')), (Jys_copy, Jys'_copy)), set_induct_phiss), JRs), Jphis),
+ B1s), B2s), AXs), Xs), f1s), f2s), p1s), p2s), ps), (ys, ys')), 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 "phi" (map (fn T => map (fn U => T --> U --> HOLogic.boolT) Ts) passiveAs)
+ ||>> mk_Frees "R" JRTs
+ ||>> mk_Frees "phi" JphiTs
+ ||>> 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 "p" pTs
+ ||>> 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' unfs mk_maps =
+ mk_coiter Ts' (map2 (fn unf => fn Fmap =>
+ HOLogic.mk_comp (mk_Fmap mk_const fs Ts Fmap, unf)) unfs (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' unfs mk_mapsAB) ks;
+ val fs_copy_maps = map (mk_map_id Ts fs_copy Ts' unfs mk_mapsAB) ks;
+ val gs_maps = map (mk_map_id Ts' gs Ts'' unf's mk_mapsBC) ks;
+ val fgs_maps =
+ map (mk_map_id Ts (map2 (curry HOLogic.mk_comp) gs fs) Ts'' unfs mk_mapsAC) ks;
+ val Xunfs = mk_unfs 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 unf''s = mk_unfs passiveCs;
+ val f1s_maps = map (mk_map_id Ts f1s YTs unfs mk_mapsAY) ks;
+ val f2s_maps = map (mk_map_id Ts' f2s YTs unf's mk_mapsBY) ks;
+ val pid_maps = map (mk_map_id XTs ps Ts'' Xunfs 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 unf unf' = fold_rev Logic.all fs
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq (HOLogic.mk_comp (unf', fs_map),
+ HOLogic.mk_comp (Term.list_comb (map, fs @ fs_maps), unf))));
+ val goals = map4 mk_goal fs_maps map_FTFT's unfs unf's;
+ val cTs = map (SOME o certifyT lthy) FTs';
+ val maps = map5 (fn goal => fn cT => fn coiter => fn map_comp' => fn map_cong =>
+ Skip_Proof.prove lthy [] [] goal
+ (K (mk_map_tac m n cT coiter map_comp' map_cong)))
+ goals cTs coiter_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 coiter_unique_thm)))
+ end;
+
+ val (map_unique_thms, map_unique_thm) =
+ let
+ fun mk_prem u map unf unf' =
+ HOLogic.mk_Trueprop (HOLogic.mk_eq (HOLogic.mk_comp (unf', u),
+ HOLogic.mk_comp (Term.list_comb (map, fs @ us), unf)));
+ val prems = map4 mk_prem us map_FTFT's unfs unf'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)))
+ (mk_map_unique_tac coiter_unique_thm map_comps);
+ in
+ `split_conj_thm unique
+ 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");
+
+ fun mk_set_Ts T = passiveAs @ replicate n (HOLogic.mk_setT T);
+ 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;
+
+ val setss_by_bnf = map (fn i => map2 (mk_hset unfs i) ls passiveAs) ks;
+ val setss_by_bnf' = map (fn i => map2 (mk_hset unf'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 unf z set =
+ relate (set $ z, mk_union (pas_set $ (unf $ z),
+ Library.foldl1 mk_union
+ (map2 (fn X => mk_UNION (X $ (unf $ 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 unfs 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)))
+ 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))))
+ 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 unfs nat i j T) ks) ls passiveAs;
+ val colss' = map2 (fn j => fn T =>
+ map (fn i => mk_hset_rec unf's nat i j T) ks) ls passiveBs;
+ val Xcolss = map2 (fn j => fn T =>
+ map (fn i => mk_hset_rec Xunfs 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)))
+ 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))))
+ 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) unf_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)))
+ 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 $)) unfs Jzs) (map2 (curry (op $)) unf'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_coiter 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 unf_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)
+ 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 unfs fstsTsTs');
+ val mor_snd = HOLogic.mk_Trueprop
+ (mk_mor thePulls (map2 (curry HOLogic.mk_comp) p2id_Fmaps pickF_ss)
+ UNIV's unf's sndsTsTs');
+ val mor_pick = HOLogic.mk_Trueprop
+ (mk_mor thePulls (map2 (curry HOLogic.mk_comp) pid_Fmaps pickF_ss)
+ UNIV''s unf''s (map2 (curry HOLogic.mk_comp) pid_maps picks));
+
+ val goal_fst = fold_rev Logic.all (AXs @ B1s @ B2s @ f1s @ f2s @ p1s @ p2s)
+ (Logic.mk_implies (wpull_prem, mor_fst));
+ val goal_snd = fold_rev Logic.all (AXs @ B1s @ B2s @ f1s @ f2s @ p1s @ p2s)
+ (Logic.mk_implies (wpull_prem, mor_snd));
+ val goal_pick = fold_rev Logic.all (AXs @ B1s @ B2s @ f1s @ f2s @ p1s @ p2s @ ps)
+ (Logic.mk_implies (wpull_prem, mor_pick));
+ in
+ (Skip_Proof.prove lthy [] [] goal_fst (mk_mor_thePull_fst_tac m mor_def map_wpull_thms
+ map_comp's pickWP_assms_tacs),
+ Skip_Proof.prove lthy [] [] goal_snd (mk_mor_thePull_snd_tac m mor_def map_wpull_thms
+ map_comp's pickWP_assms_tacs),
+ Skip_Proof.prove lthy [] [] goal_pick (mk_mor_thePull_pick_tac mor_def coiter_thms
+ map_comp's))
+ 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 coiter_thms set_natural'ss map_wpull_thms
+ pickWP_assms_tacs)))
+ 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) coiter_unique_thms coiter_unf_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 tacss = map9 mk_tactics 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;
+
+ val fld_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 flds (i - m)) (i - m)) (nth support (i - m))
+ and mk_wit support fld 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 => fld $ t))
+ |> minimize_wits
+ end;
+ in
+ map3 (fn fld => fn i => map close_wit o minimize_wits o maps (mk_wit witss fld i))
+ flds (0 upto n - 1) witss
+ end;
+
+ val wit_tac = mk_wit_tac n unf_fld_thms (flat set_simp_thmss) (maps wit_thms_of_bnf bnfs);
+
+ val (Jbnfs, lthy) =
+ fold_map6 (fn tacs => fn b => fn map => fn sets => fn T => fn wits =>
+ add_bnf Dont_Inline user_policy I tacs wit_tac (SOME deads)
+ ((((b, fold_rev Term.absfree fs' map), sets), absdummy T bd), wits))
+ tacss bs fs_maps setss_by_bnf Ts fld_witss lthy;
+
+ val fold_maps = Local_Defs.fold lthy (map (fn bnf =>
+ mk_unabs_def m (map_def_of_bnf bnf RS @{thm meta_eq_to_obj_eq})) Jbnfs);
+
+ val fold_sets = Local_Defs.fold 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, set_set_incl_thmsss, set_induct_thms) =
+ let
+ fun tinst_of unf =
+ map (SOME o certify lthy) (unf :: remove (op =) unf unfs);
+ fun tinst_of' unf = case tinst_of unf 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 unf => map (singleton (Proof_Context.export names_lthy lthy) o
+ fold_sets o Drule.instantiate' [] (tinst_of' unf) o
+ Thm.instantiate (Tinst, []) o Drule.zero_var_indexes))
+ unfs set_incl_hset_thmss;
+
+ val tinst = interleave (map (SOME o certify lthy) unfs) (replicate n NONE)
+ val set_minimal_thms =
+ map (fold_sets o Drule.instantiate' [] tinst o Thm.instantiate (Tinst, []) o
+ Drule.zero_var_indexes)
+ hset_minimal_thms;
+
+ val set_set_incl_thmsss =
+ map2 (fn unf => map (map (singleton (Proof_Context.export names_lthy lthy) o
+ fold_sets o Drule.instantiate' [] (NONE :: tinst_of' unf) o
+ Thm.instantiate (Tinst, []) o Drule.zero_var_indexes)))
+ unfs 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')
+ |> fold_sets |> Local_Defs.unfold 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;
+
+ 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 preds = map2 (fn Ds => mk_pred_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs;
+ val Jpreds = map (mk_pred_of_bnf deads passiveAs passiveBs) Jbnfs;
+
+ val JrelRs = map (fn Jrel => Term.list_comb (Jrel, JRs)) Jrels;
+ val relRs = map (fn rel => Term.list_comb (rel, JRs @ JrelRs)) rels;
+ val Jpredphis = map (fn Jrel => Term.list_comb (Jrel, Jphis)) Jpreds;
+ val predphis = map (fn rel => Term.list_comb (rel, Jphis @ Jpredphis)) preds;
+
+ val in_rels = map in_rel_of_bnf bnfs;
+ val in_Jrels = map in_rel_of_bnf Jbnfs;
+ val Jpred_defs =
+ map (Drule.abs_def o (fn thm => thm RS @{thm eq_reflection}) o pred_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 Jrel_unfold_thms =
+ let
+ fun mk_goal Jz Jz' unf unf' JrelR relR = fold_rev Logic.all (Jz :: Jz' :: JRs)
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (HOLogic.mk_mem (HOLogic.mk_prod (Jz, Jz'), JrelR),
+ HOLogic.mk_mem (HOLogic.mk_prod (unf $ Jz, unf' $ Jz'), relR))));
+ val goals = map6 mk_goal Jzs Jz's unfs unf's JrelRs relRs;
+ in
+ map12 (fn i => fn goal => fn in_rel => fn map_comp => fn map_cong =>
+ fn map_simp => fn set_simps => fn unf_inject => fn unf_fld =>
+ fn set_naturals => fn set_incls => fn set_set_inclss =>
+ Skip_Proof.prove lthy [] [] goal
+ (K (mk_rel_unfold_tac in_Jrels i in_rel map_comp map_cong map_simp set_simps
+ unf_inject unf_fld set_naturals set_incls set_set_inclss)))
+ ks goals in_rels map_comp's map_congs folded_map_simp_thms folded_set_simp_thmss'
+ unf_inject_thms unf_fld_thms set_natural'ss set_incl_thmss set_set_incl_thmsss
+ end;
+
+ val Jpred_unfold_thms =
+ let
+ fun mk_goal Jz Jz' unf unf' Jpredphi predphi = fold_rev Logic.all (Jz :: Jz' :: Jphis)
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (Jpredphi $ Jz $ Jz', predphi $ (unf $ Jz) $ (unf' $ Jz'))));
+ val goals = map6 mk_goal Jzs Jz's unfs unf's Jpredphis predphis;
+ in
+ map3 (fn goal => fn pred_def => fn Jrel_unfold =>
+ Skip_Proof.prove lthy [] [] goal (mk_pred_unfold_tac pred_def Jpred_defs Jrel_unfold))
+ goals pred_defs Jrel_unfold_thms
+ end;
+
+ val timer = time (timer "additional properties");
+
+ val ls' = if m = 1 then [0] else ls;
+ in
+ lthy
+ |> note map_uniqueN [fold_maps map_unique_thm]
+ |> notes map_simpsN (map single folded_map_simp_thms)
+ |> fold2 (fn i => notes (mk_set_simpsN i) o map single) ls' folded_set_simp_thmss
+ |> notes set_inclN set_incl_thmss
+ |> notes set_set_inclN (map flat set_set_incl_thmsss) (* nicer names? *)
+ |> fold2 (fn i => note (mk_set_inductN i) o single) ls' set_induct_thms
+ |> notes rel_unfoldN (map single Jrel_unfold_thms)
+ |> notes pred_unfoldN (map single Jpred_unfold_thms)
+ end;
+ in
+ lthy
+ |> notes coiterN (map single coiter_thms)
+ |> notes coiter_uniqueN (map single coiter_unique_thms)
+ |> notes corecN (map single corec_thms)
+ |> notes unf_fldN (map single unf_fld_thms)
+ |> notes fld_unfN (map single fld_unf_thms)
+ |> notes unf_injectN (map single unf_inject_thms)
+ |> notes unf_exhaustN (map single unf_exhaust_thms)
+ |> notes fld_injectN (map single fld_inject_thms)
+ |> notes fld_exhaustN (map single fld_exhaust_thms)
+ |> notes fld_coiterN (map single fld_coiter_thms)
+ |> note unf_coinductN [unf_coinduct_thm]
+ |> note rel_coinductN [rel_coinduct_thm]
+ |> note pred_coinductN [pred_coinduct_thm]
+ |> note unf_coinduct_uptoN [unf_coinduct_upto_thm]
+ |> note rel_coinduct_uptoN [rel_coinduct_upto_thm]
+ |> note pred_coinduct_uptoN [pred_coinduct_upto_thm]
+ end;
+
+val _ =
+ Outer_Syntax.local_theory @{command_spec "bnf_codata"} "greatest fixed points for BNF equations"
+ (Parse.and_list1
+ ((Parse.binding --| Parse.$$$ ":") -- (Parse.typ --| Parse.$$$ "=" -- Parse.typ)) >>
+ (fp_bnf_cmd bnf_gfp o apsnd split_list o split_list));
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Codatatype/Tools/bnf_gfp_tactics.ML Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,1536 @@
+(* Title: HOL/Codatatype/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_rel_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_coiter_unique_mor_tac: thm list -> thm -> thm -> thm list -> 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_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_coiter_tac: int -> thm -> thm list -> thm list -> thm list -> thm list -> thm list ->
+ thm list -> 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_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_rel_coinduct_tac: 'a list -> thm -> thm -> tactic
+ val mk_rel_coinduct_upto_tac: int -> ctyp option list -> cterm option list -> thm -> thm list ->
+ thm list -> tactic
+ val mk_rel_unfold_tac: thm list -> int -> thm -> thm -> thm -> thm -> thm list -> thm -> thm ->
+ thm list -> thm list -> thm list 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_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_unf_coinduct_tac: int -> int list -> thm -> thm -> tactic
+ val mk_unf_coinduct_upto_tac: int list -> ctyp option list -> cterm option list -> thm -> thm ->
+ thm -> tactic
+ val mk_unf_o_fld_tac: thm -> thm -> thm -> thm -> thm list ->
+ {prems: 'a, context: Proof.context} -> tactic
+ val mk_unique_mor_tac: thm list -> thm -> tactic
+ val mk_wit_tac: int -> 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_GFP_Util
+
+fun mk_coalg_set_tac coalg_def =
+ dtac (coalg_def RS @{thm subst[of _ _ "\<lambda>x. x"]}) 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 @{thm 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, @{thm set_mp}, equalityD2, rec_Suc,
+ UnI2, mk_UnN n i] @
+ [etac @{thm UN_I}, atac]) 1;
+
+fun mk_set_incl_hin_tac incls =
+ if null incls then rtac @{thm 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 @{thm nat_induct}),
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn thm => EVERY'
+ [rtac @{thm ord_eq_le_trans}, rtac thm, rtac @{thm empty_subsetI}]) rec_0s,
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn rec_Suc => EVERY'
+ [rtac @{thm 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 @{thm 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 @{thm 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 @{thm trans[OF fun_cong[OF 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 @{thm 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_rel_tac m bis_def rel_defs map_comps map_congs set_naturalss =
+ let
+ val n = length rel_defs;
+ val thms = ((1 upto n) ~~ map_comps ~~ map_congs ~~ set_naturalss ~~ rel_defs);
+
+ fun mk_if_tac ((((i, map_comp), map_cong), set_naturals), rel_def) =
+ EVERY' [rtac allI, rtac allI, rtac impI, dtac (mk_conjunctN n i),
+ etac allE, etac allE, etac impE, atac, etac bexE, etac conjE,
+ rtac (rel_def RS equalityD2 RS @{thm 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 @{thm ord_eq_le_trans}, rtac thm, rtac subset_trans,
+ etac @{thm image_mono}, rtac @{thm image_subsetI}, etac @{thm diagI}]
+ else EVERY' [rtac @{thm ord_eq_le_trans}, rtac trans, rtac thm,
+ rtac @{thm trans[OF fun_cong[OF 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), rel_def) =
+ EVERY' [dtac (mk_conjunctN n i), rtac allI, rtac allI, rtac impI,
+ etac allE, etac allE, etac impE, atac,
+ dtac (rel_def RS equalityD1 RS @{thm 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 @{thm Pair_eq[THEN subst, of "%x. x"]},
+ 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 @{thm 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 @{thm ord_eq_le_trans}, rtac thm, rtac @{thm image_subsetI},
+ rtac @{thm diag_fst}, etac @{thm set_mp}, atac]
+ else EVERY' [rtac @{thm ord_eq_le_trans}, rtac trans, rtac thm,
+ rtac @{thm trans[OF fun_cong[OF 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_rel rel_congs rel_converses =
+ EVERY' [stac bis_rel, dtac (bis_rel RS @{thm subst[of _ _ "%x. x"]}),
+ REPEAT_DETERM o etac conjE, rtac conjI,
+ CONJ_WRAP' (K (EVERY' [rtac @{thm converse_shift}, etac subset_trans,
+ rtac equalityD2, rtac @{thm converse_Times}])) rel_congs,
+ CONJ_WRAP' (fn (rel_cong, rel_converse) =>
+ EVERY' [rtac allI, rtac allI, rtac impI, rtac @{thm set_mp[OF equalityD2]},
+ rtac (rel_cong RS trans),
+ REPEAT_DETERM_N m o rtac @{thm diag_converse},
+ REPEAT_DETERM_N (length rel_congs) o rtac refl,
+ rtac rel_converse,
+ REPEAT_DETERM o etac allE,
+ rtac @{thm converseI}, etac mp, etac @{thm converseD}]) (rel_congs ~~ rel_converses)] 1;
+
+fun mk_bis_O_tac m bis_rel rel_congs rel_Os =
+ EVERY' [stac bis_rel, REPEAT_DETERM o dtac (bis_rel RS @{thm subst[of _ _ "%x. x"]}),
+ REPEAT_DETERM o etac conjE, rtac conjI,
+ CONJ_WRAP' (K (EVERY' [etac @{thm relcomp_subset_Sigma}, atac])) rel_congs,
+ CONJ_WRAP' (fn (rel_cong, rel_O) =>
+ EVERY' [rtac allI, rtac allI, rtac impI, rtac @{thm set_mp[OF equalityD2]},
+ rtac (rel_cong RS trans),
+ REPEAT_DETERM_N m o rtac @{thm diag_Comp},
+ REPEAT_DETERM_N (length rel_congs) o rtac refl,
+ rtac rel_O,
+ etac @{thm relcompE},
+ REPEAT_DETERM o dtac @{thm Pair_eq[THEN subst, of "%x. x"]},
+ etac conjE, hyp_subst_tac,
+ REPEAT_DETERM o etac allE, rtac @{thm relcompI},
+ etac mp, atac, etac mp, atac]) (rel_congs ~~ rel_Os)] 1;
+
+fun mk_bis_Gr_tac bis_rel rel_Grs mor_images morEs coalg_ins
+ {context = ctxt, prems = _} =
+ Local_Defs.unfold_tac ctxt (bis_rel :: @{thm diag_Gr} :: rel_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
+ Local_Defs.unfold_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 ssubst[of _ _ "%x. x", OF equiv_def]},
+
+ rtac conjI, rtac @{thm ssubst[of _ _ "%x. x", OF refl_on_def]},
+ rtac conjI, rtac lsbis_incl, rtac ballI, rtac @{thm set_mp},
+ rtac incl_lsbis, rtac bis_diag, atac, etac @{thm diagI},
+
+ rtac conjI, rtac @{thm ssubst[of _ _ "%x. x", OF sym_def]},
+ rtac allI, rtac allI, rtac impI, rtac @{thm set_mp},
+ rtac incl_lsbis, rtac bis_converse, rtac sbis_lsbis, etac @{thm converseI},
+
+ rtac @{thm ssubst[of _ _ "%x. x", OF trans_def]},
+ rtac allI, rtac allI, rtac allI, rtac impI, rtac impI, rtac @{thm 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 @{thm ord_eq_le_trans}),
+ etac ((trans OF [@{thm image_id} RS fun_cong, @{thm id_apply}]) RS
+ @{thm ord_eq_le_trans})]) passive_sets),
+ CONJ_WRAP' (fn (i, thm) => EVERY' [rtac (thm RS @{thm 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 @{thm 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 (Local_Defs.unfold_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
+ Local_Defs.unfold_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 @{thm ordIso_ordLeq_trans}, rtac @{thm card_of_Field_ordIso},
+ rtac @{thm Card_order_cpow}, rtac @{thm 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 @{thm ordIso_ordLeq_trans},
+ rtac @{thm clists_Cinfinite},
+ if n = 1 then rtac sbd_Cinfinite else rtac (sbd_Cinfinite RS @{thm Cinfinite_csum1}),
+ rtac @{thm 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 @{thm ordIso_ordLeq_trans}, rtac @{thm card_of_Func_Ffunc},
+ rtac @{thm 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 @{thm ordIso_ordLeq_trans}, rtac @{thm cexp_cprod},
+ rtac @{thm csum_Cnotzero2}, rtac @{thm ctwo_Cnotzero},
+ rtac @{thm 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 @{thm 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 @{thm 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 @{thm 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 @{thm 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_UnN n i), dtac (def RS @{thm subst[of _ _ "%x. x"]}),
+ 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 @{thm ord_eq_le_trans}, rtac subset_trans,
+ rtac @{thm 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 @{thm set_mp}),
+ REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE],
+ dtac @{thm Pair_eq[THEN subst, of "%x. x"]},
+ etac conjE, hyp_subst_tac,
+ dtac (isNode_def RS @{thm subst[of _ _ "%x. x"]}),
+ REPEAT_DETERM o eresolve_tac [exE, conjE],
+ rtac (equalityD2 RS @{thm 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 @{thm 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 @{thm set_mp}),
+ REPEAT_DETERM o eresolve_tac [CollectE, exE], etac conjE,
+ dtac conjunct2, dtac @{thm Pair_eq[THEN subst, of "%x. x"]}, etac conjE,
+ hyp_subst_tac, dtac (isNode_def RS @{thm subst[of _ _ "%x. x"]}),
+ 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 @{thm 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 (Local_Defs.unfold_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 @{thm nat_induct}),
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn Lev_0 =>
+ EVERY' (map rtac [@{thm 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 @{thm 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 @{thm 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 @{thm nat_induct}),
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn Lev_0 =>
+ EVERY' [rtac impI, dtac (Lev_0 RS equalityD1 RS @{thm 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 @{thm 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 @{thm nat_induct}),
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn Lev_0 =>
+ EVERY' [rtac impI, etac conjE, dtac (Lev_0 RS equalityD1 RS @{thm 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 @{thm 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_UnN 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 (@{thm sum_case_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 @{thm 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 @{thm set_mp}), etac @{thm singletonE}, etac ssubst,
+ rtac (rv_Nil RS arg_cong RS @{thm ssubst[of _ _ "%x. x"]}),
+ rtac (mk_sum_casesN n i RS @{thm ssubst[of _ _ "%x. x"]}),
+ 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 @{thm 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 @{thm ssubst[of _ _ "%x. x"]}),
+ rtac (mk_sum_casesN n i RS arg_cong RS trans RS @{thm ssubst[of _ _ "%x. x"]}),
+ etac (from_to_sbd RS arg_cong), REPEAT_DETERM o etac allE,
+ dtac (mk_conjunctN n i), etac mp, etac conjI, etac @{thm 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 @{thm 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 @{thm 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 @{thm set_mp}), rtac (mk_UnN n i''),
+ rtac CollectI, REPEAT_DETERM o rtac exI, rtac conjI, rtac refl,
+ etac conjI, rtac (Lev_0'' RS equalityD2 RS @{thm 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 @{thm 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 @{thm set_mp}),
+ rtac @{thm ssubst_mem[OF append_Cons]}, rtac (mk_UnN 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 @{thm 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 @{thm 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 @{thm 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 @{thm list.inject[THEN subst, of "%x. x"]} 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 @{thm 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 @{thm 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 @{thm list.inject[THEN subst, of "%x. x"]} 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 @{thm subst[of _ _ "%x. x"]})),
+ 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 @{thm 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 @{thm 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 (@{thm sum_case_cong} RS trans) THEN' rtac (mk_sum_casesN n i)),
+ EVERY' (map2 (fn set_natural => fn set_rv_Lev =>
+ EVERY' [rtac conjI, rtac @{thm ord_eq_le_trans}, rtac (set_natural RS trans),
+ rtac @{thm trans[OF fun_cong[OF 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 (@{thm sum_case_cong} RS trans) THEN' rtac (mk_sum_casesN n i)),
+ EVERY' (map2 (fn set_natural => fn set_rv_Lev =>
+ EVERY' [rtac conjI, rtac @{thm ord_eq_le_trans}, rtac (set_natural RS trans),
+ rtac @{thm trans[OF fun_cong[OF 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 @{thm 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 @{thm ord_eq_le_trans}, rtac (set_natural RS trans),
+ rtac @{thm trans[OF fun_cong[OF 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 @{thm 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 @{thm 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 @{thm sum_case_cong}) RS trans),
+ rtac (@{thm list.size(3)} RS arg_cong RS trans RS equalityD2 RS @{thm 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 (@{thm sum_case_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 ssubst, rtac @{thm Pair_eq}, 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 @{thm list.inject[THEN subst, of "%x. x"]}, etac conjE,
+ if i' = i'' then EVERY' [dtac (mk_InN_inject n i'),
+ dtac (Thm.permute_prems 0 2 (to_sbd_inj RS @{thm subst[of _ _ "%x. x"]})),
+ 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 @{thm set_mp}), rtac (mk_UnN 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 @{thm set_mp}),
+ CONJ_WRAP_GEN' (etac (Thm.permute_prems 1 1 UnE)) (fn i'' =>
+ EVERY' [REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE],
+ dtac @{thm list.inject[THEN subst, of "%x. x"]}, etac conjE,
+ if i' = i'' then EVERY' [dtac (mk_InN_inject n i'),
+ dtac (Thm.permute_prems 0 2 (to_sbd_inj RS @{thm subst[of _ _ "%x. x"]})),
+ atac, atac, hyp_subst_tac, atac]
+ else etac (mk_InN_not_InM i' i'' RS notE)])
+ (rev ks),
+ rtac (Lev_Suc RS equalityD2 RS @{thm set_mp}), rtac (mk_UnN 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 @{thm sum_case_cong} THEN' rtac (mk_sum_casesN n i' RS trans),
+ SELECT_GOAL (Local_Defs.unfold_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 @{thm 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 @{thm ord_eq_le_trans}),
+ rtac @{thm ord_eq_le_trans[OF trans[OF fun_cong[OF 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 @{thm ord_eq_le_trans}),
+ rtac @{thm image_subsetI}, rtac ssubst, rtac @{thm proj_in_iff},
+ rtac equiv_LSBIS, etac @{thm 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 = _} =
+ Local_Defs.unfold_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 @{thm 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 = _} =
+ Local_Defs.unfold_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_coiter_tac m mor_UNIV unf_defs coiter_defs Abs_inverses morEs map_comp_ids map_congs =
+ EVERY' [rtac @{thm ssubst[of _ _ "%x. x"]}, rtac mor_UNIV,
+ CONJ_WRAP' (fn ((Abs_inverse, morE), ((unf_def, coiter_def), (map_comp_id, map_cong))) =>
+ EVERY' [rtac ext, rtac (o_apply RS trans RS sym), rtac (unf_def RS trans),
+ rtac (coiter_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)) coiter_defs)])
+ ((Abs_inverses ~~ morEs) ~~ ((unf_defs ~~ coiter_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 @{thm subst[of _ _ "%x. x"]}),
+ 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 @{thm 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 @{thm 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_coiter_unique_mor_tac raw_coinds bis mor coiter_defs =
+ CONJ_WRAP' (fn (raw_coind, coiter_def) =>
+ EVERY' [rtac ext, etac (bis RS raw_coind RS @{thm set_mp} RS @{thm IdD}), rtac mor,
+ rtac @{thm image2_eqI}, rtac refl, rtac (coiter_def RS arg_cong RS trans),
+ rtac (o_apply RS sym), rtac UNIV_I]) (raw_coinds ~~ coiter_defs) 1;
+
+fun mk_unf_o_fld_tac fld_def coiter map_comp_id map_congL coiter_o_unfs
+ {context = ctxt, prems = _} =
+ Local_Defs.unfold_tac ctxt [fld_def] THEN EVERY' [rtac ext, rtac trans, rtac o_apply,
+ rtac trans, rtac coiter, 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}])) coiter_o_unfs),
+ rtac sym, rtac @{thm id_apply}] 1;
+
+fun mk_corec_tac m corec_defs coiter map_cong corec_Inls {context = ctxt, prems = _} =
+ Local_Defs.unfold_tac ctxt corec_defs THEN EVERY' [rtac trans, rtac (o_apply RS arg_cong),
+ rtac trans, rtac coiter, 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_rel_coinduct_tac ks raw_coind bis_rel =
+ EVERY' [rtac rev_mp, rtac raw_coind, rtac ssubst, rtac bis_rel, 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 @{thm set_mp},
+ rtac CollectI, etac @{thm prod_caseI}])) ks] 1;
+
+fun mk_rel_coinduct_upto_tac m cTs cts rel_coinduct rel_monos rel_Ids =
+ EVERY' [rtac rev_mp, rtac (Drule.instantiate' cTs cts rel_coinduct),
+ EVERY' (map2 (fn rel_mono => fn rel_Id =>
+ EVERY' [REPEAT_DETERM o resolve_tac [allI, impI], REPEAT_DETERM o etac allE,
+ etac disjE, etac mp, atac, hyp_subst_tac, rtac (rel_mono RS @{thm set_mp}),
+ REPEAT_DETERM_N m o rtac @{thm subset_refl},
+ REPEAT_DETERM_N (length rel_monos) o rtac @{thm Id_subset},
+ rtac (rel_Id RS equalityD2 RS @{thm set_mp}), rtac @{thm IdI}])
+ rel_monos rel_Ids),
+ rtac impI, REPEAT_DETERM o etac conjE,
+ CONJ_WRAP' (K (rtac impI THEN' etac mp THEN' etac disjI1)) rel_Ids] 1;
+
+fun mk_unf_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 @{thm 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 @{thm set_mp},
+ rtac CollectI, etac (refl RSN (2, @{thm subst_Pair}))]) ks] 1
+ end;
+
+fun mk_unf_coinduct_upto_tac ks cTs cts unf_coinduct bis_def bis_diag =
+ EVERY' [rtac rev_mp, rtac (Drule.instantiate' cTs cts unf_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 @{thm 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 coiter map_comp' map_cong =
+ EVERY' [rtac ext, rtac (o_apply RS trans RS sym), rtac (o_apply RS trans RS sym),
+ rtac (coiter 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_UnN 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 coiter_unique coiter_unf =
+ EVERY' [rtac (coiter_unique RS trans), EVERY' (map (fn thm => rtac (thm RS sym)) maps),
+ rtac coiter_unf] 1;
+
+fun mk_map_comp_tac m n maps map_comps map_congs coiter_unique =
+ EVERY' [rtac coiter_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;
+
+val o_apply_trans_sym = o_apply RS trans RS sym;
+val fst_convol_fun_cong_sym = @{thm fst_convol} RS fun_cong RS sym;
+val snd_convol_fun_cong_sym = @{thm snd_convol} RS fun_cong RS sym;
+
+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 @{thm set_rev_mp}, rtac @{thm 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 coiter_unique map_comps {context = ctxt, prems = _} =
+ rtac coiter_unique 1 THEN
+ Local_Defs.unfold_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 @{thm nat_induct}),
+ REPEAT_DETERM o rtac allI, SELECT_GOAL (Local_Defs.unfold_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 (Local_Defs.unfold_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 @{thm nat_induct}),
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn rec_0 => EVERY' (map rtac [@{thm 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 @{thm 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 [@{thm ordIso_ordLeq_trans}, @{thm card_of_ordIso_subst}, hset_def,
+ ctrans, @{thm UNION_Cinfinite_bound}, @{thm 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 @{thm ord_eq_le_trans},
+ arg_cong RS sym RS @{thm ord_eq_le_trans},
+ mor_hset OF [tcoalg, mor_Rep, UNIV_I] RS @{thm 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 @{thm 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 @{thm set_mp}, rtac equalityD2, rtac @{thm sym[OF proj_image]}, rtac imageE,
+ rtac @{thm 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 @{thm set_mp}),
+ REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE], hyp_subst_tac,
+ rtac (carT_def RS equalityD2 RS @{thm 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 = _} =
+ Local_Defs.unfold_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 (Local_Defs.unfold_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 @{thm subset_UNIV}),
+ CONJ_WRAP' (fn set_natural =>
+ EVERY' [rtac @{thm ord_eq_le_trans}, rtac trans, rtac set_natural,
+ rtac @{thm trans[OF fun_cong[OF 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
+ Local_Defs.unfold_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 (Local_Defs.unfold_tac ctxt @{thms o_apply prod.cases}),
+ rtac (map_comp RS trans),
+ SELECT_GOAL (Local_Defs.unfold_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 coiters map_comps {context = ctxt, prems = _} =
+ Local_Defs.unfold_tac ctxt [mor_def, @{thm thePull_def}] THEN rtac conjI 1 THEN
+ CONJ_WRAP' (K (rtac ballI THEN' rtac UNIV_I)) coiters 1 THEN
+ CONJ_WRAP' (fn (coiter, map_comp) =>
+ EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, @{thm prod_caseE}, conjE],
+ hyp_subst_tac,
+ SELECT_GOAL (Local_Defs.unfold_tac ctxt (coiter :: map_comp :: @{thms comp_def id_def})),
+ rtac refl])
+ (coiters ~~ map_comps) 1;
+
+fun mk_pick_col_tac m j cts rec_0s rec_Sucs coiters 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 @{thm nat_induct}),
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn thm => EVERY'
+ [rtac impI, rtac @{thm ord_eq_le_trans}, rtac thm, rtac @{thm empty_subsetI}]) rec_0s,
+ REPEAT_DETERM o rtac allI,
+ CONJ_WRAP' (fn (rec_Suc, ((coiter, 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 @{thm ord_eq_le_trans}, rtac rec_Suc,
+ rtac @{thm Un_least},
+ SELECT_GOAL (Local_Defs.unfold_tac ctxt [coiter, nth set_naturals (j - 1),
+ @{thm prod.cases}]),
+ etac @{thm ord_eq_le_trans[OF trans [OF fun_cong[OF 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 (Local_Defs.unfold_tac ctxt [coiter, set_natural, @{thm prod.cases}]),
+ etac imageE, hyp_subst_tac, REPEAT_DETERM o etac allE,
+ dtac (mk_conjunctN n i), etac mp, etac @{thm set_mp}, atac])
+ (ks ~~ rev (drop m set_naturals))])
+ (rec_Sucs ~~ ((coiters ~~ 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 ssubst[of _ _ "%x. x", OF wpull_def]}, 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 @{thm 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 unf_flds set_simp wit {context = ctxt, prems = _} =
+ REPEAT_DETERM (atac 1 ORELSE
+ EVERY' [dtac @{thm set_rev_mp}, rtac equalityD1, resolve_tac set_simp,
+ K (Local_Defs.unfold_tac ctxt unf_flds),
+ 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 @{thm set_rev_mp}, rtac equalityD1, resolve_tac set_simp,
+ K (Local_Defs.unfold_tac ctxt unf_flds), REPEAT_DETERM_N n o etac UnE]))))] 1);
+
+fun mk_rel_unfold_tac in_Jrels i in_rel map_comp map_cong map_simp set_simps unf_inject unf_fld
+ 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_Jrel = nth in_Jrels (i - 1);
+ val if_tac =
+ EVERY' [dtac (in_Jrel RS iffD1), REPEAT_DETERM o eresolve_tac [exE, conjE, CollectE],
+ rtac (in_rel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
+ EVERY' (map2 (fn set_natural => fn set_incl =>
+ EVERY' [rtac conjI, rtac @{thm ord_eq_le_trans}, rtac set_natural,
+ rtac @{thm ord_eq_le_trans}, rtac @{thm trans[OF fun_cong[OF image_id] id_apply]},
+ etac (set_incl RS @{thm subset_trans})])
+ passive_set_naturals set_incls),
+ CONJ_WRAP' (fn (in_Jrel, (set_natural, set_set_incls)) =>
+ EVERY' [rtac @{thm ord_eq_le_trans}, rtac set_natural, rtac @{thm image_subsetI},
+ rtac (in_Jrel 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_Jrels ~~ (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 (unf_inject RS iffD2), atac])
+ @{thms fst_conv snd_conv}];
+ val only_if_tac =
+ EVERY' [dtac (in_rel RS iffD1), REPEAT_DETERM o eresolve_tac [exE, conjE, CollectE],
+ rtac (in_Jrel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
+ CONJ_WRAP' (fn (set_simp, passive_set_natural) =>
+ EVERY' [rtac @{thm ord_eq_le_trans}, rtac set_simp, rtac @{thm Un_least},
+ rtac @{thm ord_eq_le_trans}, rtac box_equals, rtac passive_set_natural,
+ rtac (unf_fld RS sym RS arg_cong), rtac @{thm trans[OF fun_cong[OF image_id] id_apply]},
+ atac,
+ CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_least}))
+ (fn (active_set_natural, in_Jrel) => EVERY' [rtac @{thm ord_eq_le_trans},
+ rtac @{thm UN_cong[OF _ refl]}, rtac @{thm box_equals[OF _ _ refl]},
+ rtac active_set_natural, rtac (unf_fld RS sym RS arg_cong), rtac @{thm UN_least},
+ dtac @{thm set_rev_mp}, etac @{thm image_mono}, etac imageE,
+ dtac @{thm ssubst_mem[OF pair_collapse]}, dtac (in_Jrel 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_Jrels))])
+ (set_simps ~~ passive_set_naturals),
+ rtac conjI,
+ REPEAT_DETERM_N 2 o EVERY'[rtac (unf_inject RS iffD1), rtac trans, rtac map_simp,
+ rtac box_equals, rtac map_comp, rtac (unf_fld 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_Jrel => EVERY' [rtac trans, rtac o_apply, dtac @{thm set_rev_mp}, atac,
+ dtac @{thm ssubst_mem[OF pair_collapse]}, dtac (in_Jrel RS iffD1), dtac @{thm someI_ex},
+ REPEAT_DETERM o etac conjE, atac]) in_Jrels),
+ 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/Codatatype/Tools/bnf_gfp_util.ML Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,261 @@
+(* Title: HOL/Codatatype/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_If: term -> 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 Inl_const: typ -> typ -> term
+ val Inr_const: typ -> typ -> term
+
+ val mk_Inl: term -> typ -> term
+ val mk_Inr: term -> typ -> term
+
+ val mk_InN: typ list -> term -> int -> term
+
+ val mk_sumTN: typ list -> typ
+ val mk_sum_case: term -> term -> term
+ val mk_sum_caseN: term list -> term
+
+ val mk_disjIN: int -> int -> thm
+ val mk_specN: int -> thm -> thm
+ val mk_sum_casesN: int -> int -> thm
+ val mk_sumEN: int -> 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_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_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_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_sumTN Ts = Library.foldr1 (mk_sumT) Ts;
+
+fun Inl_const LT RT = Const (@{const_name Inl}, LT --> mk_sumT (LT, RT));
+fun mk_Inl t RT = Inl_const (fastype_of t) RT $ t;
+
+fun Inr_const LT RT = Const (@{const_name Inr}, RT --> mk_sumT (LT, RT));
+fun mk_Inr t LT = Inr_const LT (fastype_of t) $ t;
+
+fun mk_InN [_] t 1 = t
+ | mk_InN (_ :: Ts) t 1 = mk_Inl t (mk_sumTN Ts)
+ | mk_InN (LT :: Ts) t m = mk_Inr (mk_InN Ts t (m - 1)) LT
+ | mk_InN Ts t _ = raise (TYPE ("mk_InN", Ts, [t]));
+
+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;
+
+fun mk_sum_caseN [f] = f
+ | mk_sum_caseN (f :: fs) = mk_sum_case f (mk_sum_caseN fs);
+
+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_sum_casesN 1 1 = @{thm refl}
+ | mk_sum_casesN _ 1 = @{thm sum.cases(1)}
+ | mk_sum_casesN 2 2 = @{thm sum.cases(2)}
+ | mk_sum_casesN n m = trans OF [@{thm sum_case_step(2)}, mk_sum_casesN (n - 1) (m - 1)];
+
+local
+ fun mk_sumEN' 1 = @{thm obj_sum_step}
+ | mk_sumEN' n = mk_sumEN' (n - 1) RSN (2, @{thm obj_sum_step});
+in
+ fun mk_sumEN 1 = @{thm obj_sum_base}
+ | mk_sumEN 2 = @{thm sumE}
+ | mk_sumEN n = (mk_sumEN' (n - 2) RSN (2, @{thm obj_sumE})) OF replicate n (impI RS allI);
+end;
+
+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/Codatatype/Tools/bnf_lfp.ML Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,1763 @@
+(* Title: HOL/Codatatype/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: binding list -> typ list list -> BNF_Def.BNF list -> Proof.context -> Proof.context
+end;
+
+structure BNF_LFP : BNF_LFP =
+struct
+
+open BNF_Def
+open BNF_Util
+open BNF_Tactics
+open BNF_FP_Util
+open BNF_LFP_Util
+open BNF_LFP_Tactics
+
+(*all bnfs have the same lives*)
+fun bnf_lfp bs Dss_insts 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 (fold_rev (fn b => fn s => Binding.name_of b ^ s) bs "");
+
+ fun note thmN thms = snd o Local_Theory.note
+ ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), thms);
+ fun notes thmN thmss = fold2 (fn b => fn thms => snd o Local_Theory.note
+ ((Binding.qualify true (Binding.name_of b) (Binding.name thmN), []), thms)) bs thmss;
+
+ (* TODO: check if m, n etc are sane *)
+
+ val Dss = map (fn Ds => map TFree (fold Term.add_tfreesT Ds [])) Dss_insts;
+ val deads = distinct (op =) (flat Dss);
+ 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 *)
+ fun mk_FTs Ts = map2 (fn Ds => mk_T_of_bnf Ds Ts) Dss bnfs;
+ val (params, params') = `(map dest_TFree) (deads @ 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);
+
+ 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)) (m upto m + n - 1);
+ val reachable = fixpoint (op =) enrich [];
+ val _ = if map snd reachable = (m upto m + n - 1) then ()
+ else error "The datatype could not be generated because nonemptiness could not be proved";
+
+ 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) (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))))
+ (K (mk_map_comp_id_tac map_comp))
+ 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 = HOLogic.mk_Trueprop (HOLogic.mk_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'))
+ 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
+ HOLogic.mk_Trueprop (HOLogic.mk_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;
+
+ (*transforms defined frees into consts*)
+ 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))) 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))
+ 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)))
+ 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
+ HOLogic.mk_Trueprop (HOLogic.mk_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;
+
+ (*transforms defined frees into consts*)
+ 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],
+ HOLogic.mk_Trueprop (HOLogic.mk_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));
+ 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))
+ 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))
+ 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))
+ 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))
+ 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))
+ end;
+
+ val mor_convol_thm =
+ let
+ val maps = map3 (fn s => fn prod_s => fn map =>
+ mk_convol (HOLogic.mk_comp (s, Term.list_comb (map, 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))
+ 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) (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))))
+ (K (mk_mor_UNIV_tac m morE_thms mor_def))
+ 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 = HOLogic.mk_Trueprop (HOLogic.mk_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))
+ 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));
+
+ 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));
+
+ 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));
+ 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))
+ 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));
+
+ 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)))
+ (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)));
+
+ 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)));
+ 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
+ HOLogic.mk_Trueprop (HOLogic.mk_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;
+
+ (*transforms defined frees into consts*)
+ 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));
+
+ 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));
+
+ 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));
+
+ 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));
+ 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 true 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 =
+ mk_Abs_inverse_thm (the (#set_def IIT_loc_info)) (#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 "phi" (replicate n (initT --> HOLogic.boolT));
+
+ 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
+ HOLogic.mk_Trueprop (HOLogic.mk_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;
+
+ (*transforms defined frees into consts*)
+ 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));
+
+ 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)
+
+ 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))
+ 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);
+
+ 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));
+ 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))
+ end;
+
+ val timer = time (timer "Initiality definition & thms");
+
+ val ((T_names, (T_glob_infos, T_loc_infos)), lthy) =
+ lthy
+ |> fold_map2 (fn b => fn car_init => typedef true NONE (b, params, NoSyn) car_init NONE
+ (EVERY' [rtac ssubst, rtac @{thm ex_in_conv}, resolve_tac alg_not_empty_thms,
+ rtac alg_init_thm] 1)) bs 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 set_defs = map (the o #set_def) T_loc_infos;
+ 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;
+
+ val T_subset1s = map mk_T_subset1 set_defs;
+ val T_subset2s = map mk_T_subset2 set_defs;
+
+ 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));
+
+ 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
+ (map2 (curry op RS) T_subset1s 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 iterT = 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 (((((((((((Izs, (Izs1, Izs1')), (Izs2, Izs2')), (xFs, xFs')), yFs), (AFss, AFss')),
+ (iter_f, iter_f')), fs), phis), phi2s), rec_ss), names_lthy) = names_lthy
+ |> mk_Frees "z" Ts
+ ||>> 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") iterT
+ ||>> mk_Frees "f" fTs
+ ||>> mk_Frees "phi" (map (fn T => T --> HOLogic.boolT) Ts)
+ ||>> mk_Frees "phi" (map2 (fn T => fn U => T --> U --> HOLogic.boolT) Ts Ts')
+ ||>> mk_Frees "s" rec_sTs;
+
+ fun fld_bind i = Binding.suffix_name ("_" ^ fldN) (nth bs (i - 1));
+ val fld_name = Binding.name_of o fld_bind;
+ val fld_def_bind = rpair [] o Thm.def_binding o fld_bind;
+
+ fun fld_spec i abs str map_FT_init x x' =
+ let
+ val fldT = nth FTs (i - 1) --> nth Ts (i - 1);
+
+ val lhs = Free (fld_name i, fldT);
+ val rhs = Term.absfree x' (abs $ (str $
+ (Term.list_comb (map_FT_init, map HOLogic.id_const passiveAs @ Rep_Ts) $ x)));
+ in
+ HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))
+ end;
+
+ val ((fld_frees, (_, fld_def_frees)), (lthy, lthy_old)) =
+ lthy
+ |> fold_map6 (fn i => fn abs => fn str => fn map => fn x => fn x' =>
+ Specification.definition
+ (SOME (fld_bind i, NONE, NoSyn), (fld_def_bind i, fld_spec i abs str map x x')))
+ ks Abs_Ts str_inits map_FT_inits xFs xFs'
+ |>> apsnd split_list o split_list
+ ||> `Local_Theory.restore;
+
+ (*transforms defined frees into consts*)
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+ fun mk_flds passive =
+ map (Term.subst_atomic_types (map (Morphism.typ phi) params' ~~ (deads @ passive)) o
+ Morphism.term phi) fld_frees;
+ val flds = mk_flds passiveAs;
+ val fld's = mk_flds passiveBs;
+ val fld_defs = map (Morphism.thm phi) fld_def_frees;
+
+ val (mor_Rep_thm, mor_Abs_thm) =
+ let
+ val copy = alg_init_thm RS copy_alg_thm;
+ fun mk_bij inj subset1 subset2 Rep cases = @{thm bij_betwI'} OF
+ [inj, Rep RS subset2, subset1 RS cases];
+ val bijs = map5 mk_bij Rep_injects T_subset1s T_subset2s Reps Rep_casess;
+ val mor_Rep =
+ Skip_Proof.prove lthy [] []
+ (HOLogic.mk_Trueprop (mk_mor UNIVs flds car_inits str_inits Rep_Ts))
+ (mk_mor_Rep_tac fld_defs copy bijs inver_Abss inver_Reps);
+
+ 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 flds Abs_Ts))
+ (K (mk_mor_Abs_tac inv inver_Abss inver_Reps));
+ in
+ (mor_Rep, mor_Abs)
+ end;
+
+ val timer = time (timer "fld definitions & thms");
+
+ val iter_fun = Term.absfree iter_f'
+ (mk_mor UNIVs flds active_UNIVs ss (map (mk_nthN n iter_f) ks));
+ val iter = HOLogic.choice_const iterT $ iter_fun;
+
+ fun iter_bind i = Binding.suffix_name ("_" ^ iterN) (nth bs (i - 1));
+ val iter_name = Binding.name_of o iter_bind;
+ val iter_def_bind = rpair [] o Thm.def_binding o iter_bind;
+
+ fun iter_spec i T AT =
+ let
+ val iterT = Library.foldr (op -->) (sTs, T --> AT);
+
+ val lhs = Term.list_comb (Free (iter_name i, iterT), ss);
+ val rhs = mk_nthN n iter i;
+ in
+ HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))
+ end;
+
+ val ((iter_frees, (_, iter_def_frees)), (lthy, lthy_old)) =
+ lthy
+ |> fold_map3 (fn i => fn T => fn AT =>
+ Specification.definition
+ (SOME (iter_bind i, NONE, NoSyn), (iter_def_bind i, iter_spec i T AT)))
+ ks Ts activeAs
+ |>> apsnd split_list o split_list
+ ||> `Local_Theory.restore;
+
+ (*transforms defined frees into consts*)
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+ val iters = map (fst o dest_Const o Morphism.term phi) iter_frees;
+ fun mk_iter Ts ss i = Term.list_comb (Const (nth iters (i - 1), Library.foldr (op -->)
+ (map fastype_of ss, nth Ts (i - 1) --> range_type (fastype_of (nth ss (i - 1))))), ss);
+ val iter_defs = map (Morphism.thm phi) iter_def_frees;
+
+ val mor_iter_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 iterT;
+ val ct = certify lthy iter_fun
+ in
+ singleton (Proof_Context.export names_lthy lthy)
+ (Skip_Proof.prove lthy [] []
+ (HOLogic.mk_Trueprop (mk_mor UNIVs flds active_UNIVs ss (map (mk_iter Ts ss) ks)))
+ (K (mk_mor_iter_tac cT ct iter_defs ex_mor (mor_comp RS mor_cong))))
+ end;
+
+ val iter_thms = map (fn morE => rule_by_tactic lthy
+ ((rtac CollectI THEN' CONJ_WRAP' (K (rtac @{thm subset_UNIV})) (1 upto m + n)) 1)
+ (mor_iter_thm RS morE)) morE_thms;
+
+ val (iter_unique_mor_thms, iter_unique_mor_thm) =
+ let
+ val prem = HOLogic.mk_Trueprop (mk_mor UNIVs flds active_UNIVs ss fs);
+ fun mk_fun_eq f i = HOLogic.mk_eq (f, mk_iter 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_iter_unique_mor_tac type_defs init_unique_mor_thms T_subset2s Reps
+ mor_comp_thm mor_Abs_thm mor_iter_thm));
+ in
+ `split_conj_thm unique_mor
+ end;
+
+ val (iter_unique_thms, iter_unique_thm) =
+ `split_conj_thm (mk_conjIN n RS
+ (mor_UNIV_thm RS @{thm ssubst[of _ _ "%x. x"]} RS iter_unique_mor_thm))
+
+ val iter_fld_thms =
+ map (fn thm => (mor_incl_thm OF replicate n @{thm subset_UNIV}) RS thm RS sym)
+ iter_unique_mor_thms;
+
+ val fld_o_iter_thms =
+ let
+ val mor = mor_comp_thm OF [mor_iter_thm, mor_str_thm];
+ in
+ map2 (fn unique => fn iter_fld =>
+ trans OF [mor RS unique, iter_fld]) iter_unique_mor_thms iter_fld_thms
+ end;
+
+ val timer = time (timer "iter definitions & thms");
+
+ val map_flds = map2 (fn Ds => fn bnf =>
+ Term.list_comb (mk_map_of_bnf Ds (passiveAs @ FTs) (passiveAs @ Ts) bnf,
+ map HOLogic.id_const passiveAs @ flds)) Dss bnfs;
+
+ fun unf_bind i = Binding.suffix_name ("_" ^ unfN) (nth bs (i - 1));
+ val unf_name = Binding.name_of o unf_bind;
+ val unf_def_bind = rpair [] o Thm.def_binding o unf_bind;
+
+ fun unf_spec i FT T =
+ let
+ val unfT = T --> FT;
+
+ val lhs = Free (unf_name i, unfT);
+ val rhs = mk_iter Ts map_flds i;
+ in
+ HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))
+ end;
+
+ val ((unf_frees, (_, unf_def_frees)), (lthy, lthy_old)) =
+ lthy
+ |> fold_map3 (fn i => fn FT => fn T =>
+ Specification.definition
+ (SOME (unf_bind i, NONE, NoSyn), (unf_def_bind i, unf_spec i FT T))) ks FTs Ts
+ |>> apsnd split_list o split_list
+ ||> `Local_Theory.restore;
+
+ (*transforms defined frees into consts*)
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+ fun mk_unfs params =
+ map (Term.subst_atomic_types (map (Morphism.typ phi) params' ~~ params) o Morphism.term phi)
+ unf_frees;
+ val unfs = mk_unfs params';
+ val unf_defs = map (Morphism.thm phi) unf_def_frees;
+
+ val fld_o_unf_thms = map2 (Local_Defs.fold lthy o single) unf_defs fld_o_iter_thms;
+
+ val unf_o_fld_thms =
+ let
+ fun mk_goal unf fld FT =
+ HOLogic.mk_Trueprop (HOLogic.mk_eq (HOLogic.mk_comp (unf, fld), HOLogic.id_const FT));
+ val goals = map3 mk_goal unfs flds FTs;
+ in
+ map5 (fn goal => fn unf_def => fn iter => fn map_comp_id => fn map_congL =>
+ Skip_Proof.prove lthy [] [] goal
+ (K (mk_unf_o_fld_tac unf_def iter map_comp_id map_congL fld_o_iter_thms)))
+ goals unf_defs iter_thms map_comp_id_thms map_congL_thms
+ end;
+
+ val unf_fld_thms = map (fn thm => thm RS @{thm pointfree_idE}) unf_o_fld_thms;
+ val fld_unf_thms = map (fn thm => thm RS @{thm pointfree_idE}) fld_o_unf_thms;
+
+ val bij_unf_thms =
+ map2 (fn thm1 => fn thm2 => @{thm o_bij} OF [thm1, thm2]) fld_o_unf_thms unf_o_fld_thms;
+ val inj_unf_thms = map (fn thm => thm RS @{thm bij_is_inj}) bij_unf_thms;
+ val surj_unf_thms = map (fn thm => thm RS @{thm bij_is_surj}) bij_unf_thms;
+ val unf_nchotomy_thms = map (fn thm => thm RS @{thm surjD}) surj_unf_thms;
+ val unf_inject_thms = map (fn thm => thm RS @{thm inj_eq}) inj_unf_thms;
+ val unf_exhaust_thms = map (fn thm => thm RS @{thm exE}) unf_nchotomy_thms;
+
+ val bij_fld_thms =
+ map2 (fn thm1 => fn thm2 => @{thm o_bij} OF [thm1, thm2]) unf_o_fld_thms fld_o_unf_thms;
+ val inj_fld_thms = map (fn thm => thm RS @{thm bij_is_inj}) bij_fld_thms;
+ val surj_fld_thms = map (fn thm => thm RS @{thm bij_is_surj}) bij_fld_thms;
+ val fld_nchotomy_thms = map (fn thm => thm RS @{thm surjD}) surj_fld_thms;
+ val fld_inject_thms = map (fn thm => thm RS @{thm inj_eq}) inj_fld_thms;
+ val fld_exhaust_thms = map (fn thm => thm RS @{thm exE}) fld_nchotomy_thms;
+
+ val timer = time (timer "unf definitions & thms");
+
+ val fst_rec_pair_thms =
+ let
+ val mor = mor_comp_thm OF [mor_iter_thm, mor_convol_thm];
+ in
+ map2 (fn unique => fn iter_fld =>
+ trans OF [mor RS unique, iter_fld]) iter_unique_mor_thms iter_fld_thms
+ end;
+
+ fun rec_bind i = Binding.suffix_name ("_" ^ 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 fld => fn prod_s => fn map =>
+ mk_convol (HOLogic.mk_comp (fld, Term.list_comb (map, passive_ids @ rec_fsts)), prod_s))
+ flds 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_iter Ts maps i);
+ in
+ HOLogic.mk_Trueprop (HOLogic.mk_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;
+
+ (*transforms defined frees into consts*)
+ val phi = Proof_Context.export_morphism lthy_old lthy;
+ val recs = map (fst o dest_Const o Morphism.term phi) rec_frees;
+ fun mk_rec ss i = Term.list_comb (Const (nth recs (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 rec_thms =
+ let
+ fun mk_goal i rec_s rec_map fld x =
+ let
+ val lhs = mk_rec rec_ss i $ (fld $ x);
+ val rhs = rec_s $ (Term.list_comb (rec_map, passive_ids @ convols) $ x);
+ in
+ fold_rev Logic.all (x :: rec_ss) (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)))
+ end;
+ val goals = map5 mk_goal ks rec_ss rec_maps_rev flds xFs;
+ in
+ map2 (fn goal => fn iter =>
+ Skip_Proof.prove lthy [] [] goal (mk_rec_tac rec_defs iter fst_rec_pair_thms))
+ goals iter_thms
+ end;
+
+ val timer = time (timer "rec definitions & thms");
+
+ val (fld_induct_thm, induct_params) =
+ let
+ fun mk_prem phi fld 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 $ (fld $ x));
+ in
+ Logic.all x (Logic.list_implies (IHs, concl))
+ end;
+
+ val prems = map4 mk_prem phis flds 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_fld_induct_tac m set_natural'ss init_induct_thm morE_thms mor_Abs_thm
+ Rep_inverses Abs_inverses Reps T_subset1s T_subset2s)), rev (Term.add_tfrees goal []))
+ end;
+
+ val cTs = map (SOME o certifyT lthy o TFree) induct_params;
+
+ val weak_fld_induct_thms =
+ let fun insts i = (replicate (i - 1) TrueI) @ (@{thm asm_rl} :: replicate (n - i) TrueI);
+ in map (fn i => (fld_induct_thm OF insts i) RS mk_conjunctN n i) ks end;
+
+ val (fld_induct2_thm, induct2_params) =
+ let
+ fun mk_prem phi fld fld' 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 $ (fld $ x) $ (fld' $ y));
+ in
+ fold_rev Logic.all [x, y] (Logic.list_implies (IHs, concl))
+ end;
+
+ val prems = map7 mk_prem phi2s flds fld'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_fld_induct2_tac cTs cts fld_induct_thm weak_fld_induct_thms)),
+ 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 (fn T => fn U => T --> U --> HOLogic.boolT) 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 "phi" 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_iter_arg fs Ts fld fmap =
+ HOLogic.mk_comp (fld, Term.list_comb (fmap, fs @ map HOLogic.id_const Ts));
+ fun mk_map Ts fs Ts' flds mk_maps =
+ mk_iter Ts (map2 (mk_map_iter_arg fs Ts') flds (mk_maps Ts'));
+ val pmapsABT' = mk_passive_maps passiveAs passiveBs;
+ val fs_maps = map (mk_map Ts fs Ts' fld's pmapsABT') ks;
+ val fs_copy_maps = map (mk_map Ts fs_copy Ts' fld's pmapsABT') ks;
+ val Yflds = mk_flds passiveYs;
+ val f1s_maps = map (mk_map Ts f1s YTs Yflds (mk_passive_maps passiveAs passiveYs)) ks;
+ val f2s_maps = map (mk_map Ts' f2s YTs Yflds (mk_passive_maps passiveBs passiveYs)) ks;
+ val p1s_maps = map (mk_map XTs p1s Ts flds (mk_passive_maps passiveXs passiveAs)) ks;
+ val p2s_maps = map (mk_map XTs p2s Ts' fld's (mk_passive_maps passiveXs passiveBs)) ks;
+
+ val (map_simp_thms, map_thms) =
+ let
+ fun mk_goal fs_map map fld fld' = fold_rev Logic.all fs
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq (HOLogic.mk_comp (fs_map, fld),
+ HOLogic.mk_comp (fld', Term.list_comb (map, fs @ fs_maps)))));
+ val goals = map4 mk_goal fs_maps map_FTFT's flds fld's;
+ val maps = map4 (fn goal => fn iter => fn map_comp_id => fn map_cong =>
+ Skip_Proof.prove lthy [] [] goal
+ (K (mk_map_tac m n iter map_comp_id map_cong)))
+ goals iter_thms map_comp_id_thms map_congs;
+ in
+ map_split (fn thm => (thm RS @{thm pointfreeE}, thm)) maps
+ end;
+
+ val (map_unique_thms, map_unique_thm) =
+ let
+ fun mk_prem u map fld fld' =
+ HOLogic.mk_Trueprop (HOLogic.mk_eq (HOLogic.mk_comp (u, fld),
+ HOLogic.mk_comp (fld', Term.list_comb (map, fs @ us))));
+ val prems = map4 mk_prem us map_FTFT's flds fld'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 iter_unique_mor_thm map_comp_id_thms map_congs));
+ 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_iter Ts cols) ks) colss;
+ val setss_by_bnf = transpose setss_by_range;
+
+ val (set_simp_thmss, set_thmss) =
+ let
+ fun mk_goal sets fld set col map =
+ HOLogic.mk_Trueprop (HOLogic.mk_eq (HOLogic.mk_comp (set, fld),
+ HOLogic.mk_comp (col, Term.list_comb (map, passive_ids @ sets))));
+ val goalss =
+ map3 (fn sets => map4 (mk_goal sets) flds sets) setss_by_range colss map_setss;
+ val setss = map (map2 (fn iter => fn goal =>
+ Skip_Proof.prove lthy [] [] goal (K (mk_set_tac iter)))
+ iter_thms) goalss;
+
+ fun mk_simp_goal pas_set act_sets sets fld z set =
+ Logic.all z (HOLogic.mk_Trueprop (HOLogic.mk_eq (set $ (fld $ 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 flds 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))))
+ set_natural'ss) ls simp_goalss setss;
+ in
+ (set_simpss, setss)
+ 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) fld_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)))
+ 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) fld_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)))
+ 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) fld_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));
+ 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 flds 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) fld_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));
+ 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) fld_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 fld_inject_thms)));
+ 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 tacss = map9 mk_tactics 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;
+
+ val fld_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 flds (i - m)) (i - m)) (nth support (i - m))
+ and mk_wit support fld 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 => fld $ t))
+ |> minimize_wits
+ end;
+ in
+ map3 (fn fld => fn i => map close_wit o minimize_wits o maps (mk_wit witss fld i))
+ flds (0 upto n - 1) witss
+ end;
+
+ fun wit_tac _ = mk_wit_tac n (flat set_simp_thmss) (maps wit_thms_of_bnf bnfs);
+
+ val (Ibnfs, lthy) =
+ fold_map6 (fn tacs => fn b => fn map => fn sets => fn T => fn wits =>
+ add_bnf Dont_Inline user_policy I tacs wit_tac (SOME deads)
+ ((((b, fold_rev Term.absfree fs' map), sets), absdummy T bd), wits))
+ tacss bs fs_maps setss_by_bnf Ts fld_witss lthy;
+
+ val fold_maps = Local_Defs.fold lthy (map (fn bnf =>
+ mk_unabs_def m (map_def_of_bnf bnf RS @{thm meta_eq_to_obj_eq})) Ibnfs);
+
+ val fold_sets = Local_Defs.fold 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 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 preds = map2 (fn Ds => mk_pred_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs;
+ val Ipreds = map (mk_pred_of_bnf deads passiveAs passiveBs) Ibnfs;
+
+ val IrelRs = map (fn Irel => Term.list_comb (Irel, IRs)) Irels;
+ val relRs = map (fn rel => Term.list_comb (rel, IRs @ IrelRs)) rels;
+ val Ipredphis = map (fn Irel => Term.list_comb (Irel, Iphis)) Ipreds;
+ val predphis = map (fn rel => Term.list_comb (rel, Iphis @ Ipredphis)) preds;
+
+ val in_rels = map in_rel_of_bnf bnfs;
+ val in_Irels = map in_rel_of_bnf Ibnfs;
+ val pred_defs = map pred_def_of_bnf bnfs;
+ val Ipred_defs =
+ map (Drule.abs_def o (fn thm => thm RS @{thm eq_reflection}) o pred_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 Irel_unfold_thms =
+ let
+ fun mk_goal xF yF fld fld' IrelR relR = fold_rev Logic.all (xF :: yF :: IRs)
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (HOLogic.mk_mem (HOLogic.mk_prod (fld $ xF, fld' $ yF), IrelR),
+ HOLogic.mk_mem (HOLogic.mk_prod (xF, yF), relR))));
+ val goals = map6 mk_goal xFs yFs flds fld's IrelRs relRs;
+ in
+ map12 (fn i => fn goal => fn in_rel => fn map_comp => fn map_cong =>
+ fn map_simp => fn set_simps => fn fld_inject => fn fld_unf =>
+ fn set_naturals => fn set_incls => fn set_set_inclss =>
+ Skip_Proof.prove lthy [] [] goal
+ (K (mk_rel_unfold_tac in_Irels i in_rel map_comp map_cong map_simp set_simps
+ fld_inject fld_unf set_naturals set_incls set_set_inclss)))
+ ks goals in_rels map_comp's map_congs folded_map_simp_thms folded_set_simp_thmss'
+ fld_inject_thms fld_unf_thms set_natural'ss set_incl_thmss set_set_incl_thmsss
+ end;
+
+ val Ipred_unfold_thms =
+ let
+ fun mk_goal xF yF fld fld' Ipredphi predphi = fold_rev Logic.all (xF :: yF :: Iphis)
+ (HOLogic.mk_Trueprop (HOLogic.mk_eq
+ (Ipredphi $ (fld $ xF) $ (fld' $ yF), predphi $ xF $ yF)));
+ val goals = map6 mk_goal xFs yFs flds fld's Ipredphis predphis;
+ in
+ map3 (fn goal => fn pred_def => fn Irel_unfold =>
+ Skip_Proof.prove lthy [] [] goal (mk_pred_unfold_tac pred_def Ipred_defs Irel_unfold))
+ goals pred_defs Irel_unfold_thms
+ end;
+
+ val timer = time (timer "additional properties");
+
+ val ls' = if m = 1 then [0] else ls
+ in
+ lthy
+ |> note map_uniqueN [fold_maps map_unique_thm]
+ |> notes map_simpsN (map single folded_map_simp_thms)
+ |> fold2 (fn i => notes (mk_set_simpsN i) o map single) ls' folded_set_simp_thmss
+ |> notes set_inclN set_incl_thmss
+ |> notes set_set_inclN (map flat set_set_incl_thmsss)
+ |> notes rel_unfoldN (map single Irel_unfold_thms)
+ |> notes pred_unfoldN (map single Ipred_unfold_thms)
+ end;
+
+ in
+ lthy
+ |> notes iterN (map single iter_thms)
+ |> notes iter_uniqueN (map single iter_unique_thms)
+ |> notes recN (map single rec_thms)
+ |> notes unf_fldN (map single unf_fld_thms)
+ |> notes fld_unfN (map single fld_unf_thms)
+ |> notes unf_injectN (map single unf_inject_thms)
+ |> notes unf_exhaustN (map single unf_exhaust_thms)
+ |> notes fld_injectN (map single fld_inject_thms)
+ |> notes fld_exhaustN (map single fld_exhaust_thms)
+ |> note fld_inductN [fld_induct_thm]
+ |> note fld_induct2N [fld_induct2_thm]
+ end;
+
+val _ =
+ Outer_Syntax.local_theory @{command_spec "bnf_data"} "least fixed points for BNF equations"
+ (Parse.and_list1
+ ((Parse.binding --| Parse.$$$ ":") -- (Parse.typ --| Parse.$$$ "=" -- Parse.typ)) >>
+ (fp_bnf_cmd bnf_lfp o apsnd split_list o split_list));
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Codatatype/Tools/bnf_lfp_tactics.ML Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,835 @@
+(* Title: HOL/Codatatype/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_ex_copy_alg_tac: int -> thm -> thm -> tactic
+ val mk_fld_induct2_tac: ctyp option list -> cterm option list -> thm -> thm list ->
+ {prems: 'a, context: Proof.context} -> tactic
+ val mk_fld_induct_tac: int -> thm list list -> thm -> thm list -> thm -> thm list -> thm list ->
+ thm list -> thm list -> thm list -> 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_iter_unique_mor_tac: thm list -> 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_iter_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_rel_unfold_tac: thm list -> int -> thm -> thm -> thm -> thm -> thm list -> thm ->
+ thm -> thm list -> thm list -> thm list list -> 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_unf_o_fld_tac: thm -> thm -> thm -> thm -> thm 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
+
+fun mk_alg_set_tac alg_def =
+ dtac (alg_def RS @{thm subst[of _ _ "\<lambda>x. x"]}) 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 @{thm 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 @{thm set_mp}, stac @{thm id_apply}, atac]))
+ map_id's THEN'
+ CONJ_WRAP' (fn thm =>
+ (EVERY' [rtac ballI, rtac trans, rtac @{thm 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 @{thm subset_UNIV},
+ (EVERY' [rtac @{thm ord_eq_le_trans}, rtac thm, rtac @{thm image_subsetI},
+ etac bspec, etac @{thm 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 @{thm set_mp}, etac imageI];
+ fun Collect_tac set_natural' =
+ CONJ_WRAP' (fn thm =>
+ FIRST' [rtac @{thm subset_UNIV},
+ (EVERY' [rtac @{thm 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 @{thm subst[of _ _ "\<lambda>x. x"]}) THEN'
+ dtac (alg_def RS @{thm subst[of _ _ "\<lambda>x. x"]}) 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 @{thm 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 @{thm 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 @{thm set_mp},
+ rtac equalityD1, etac @{thm bij_betw_imageE}, rtac imageI, etac thm,
+ REPEAT_DETERM o rtac @{thm 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 @{thm 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 @{thm 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 ssubst, rtac @{thm Pair_eq}, 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 @{thm 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 @{thm 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
+ Local_Defs.unfold_tac ctxt (Abs_inverse :: @{thms fst_conv snd_conv}) 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 @{thm subset_UNIV},
+ EVERY' [rtac @{thm ord_eq_le_trans}, resolve_tac set_natural', rtac @{thm 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 (Local_Defs.unfold_tac ctxt (Abs_inverse :: @{thms fst_conv snd_conv})) 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 @{thm 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 @{thm subset_UNIV},
+ REPEAT_DETERM_N n o (etac @{thm 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 @{thm subset_UNIV},
+ REPEAT_DETERM_N n o (etac @{thm 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 @{thm 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 fld_defs copy bijs inver_Abss inver_Reps {context = ctxt, prems = _} =
+ (K (Local_Defs.unfold_tac ctxt fld_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 @{thm subset_UNIV}, rtac conjI, rtac inver_Rep, rtac inver_Abs])
+ inver_Abss inver_Reps)) 1;
+
+fun mk_mor_iter_tac cT ct iter_defs ex_mor mor =
+ (EVERY' (map stac iter_defs) THEN' EVERY' [rtac rev_mp, rtac ex_mor, rtac impI] THEN'
+ REPEAT_DETERM_N (length iter_defs) o etac exE THEN'
+ rtac (Drule.instantiate' [SOME cT] [SOME ct] @{thm someI}) THEN' etac mor) 1;
+
+fun mk_iter_unique_mor_tac type_defs init_unique_mors subsets Reps mor_comp mor_Abs mor_iter =
+ let
+ fun mk_subset subset Rep = etac subset ORELSE' rtac (Rep RS subset);
+ 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' (map2 mk_subset subsets Reps),
+ rtac mor_comp, rtac mor_Abs, atac,
+ rtac mor_comp, rtac mor_Abs, rtac mor_iter];
+ in
+ CONJ_WRAP' mk_unique type_defs 1
+ end;
+
+fun mk_unf_o_fld_tac unf_def iter map_comp_id map_congL fld_o_iters =
+ EVERY' [stac unf_def, rtac ext, rtac trans, rtac o_apply, rtac trans, rtac iter,
+ 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}])) fld_o_iters),
+ rtac sym, rtac @{thm id_apply}] 1;
+
+fun mk_rec_tac rec_defs iter fst_recs {context = ctxt, prems = _}=
+ Local_Defs.unfold_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 (iter RS @{thm arg_cong[of _ _ snd]}),
+ rtac @{thm snd_convol'}] 1;
+
+fun mk_fld_induct_tac m set_natural'ss init_induct morEs mor_Abs Rep_invs Abs_invs Reps
+ subset1s subset2s =
+ let
+ val n = length set_natural'ss;
+ val ks = 1 upto n;
+
+ fun mk_IH_tac Rep_inv Abs_inv set_natural' subset =
+ DETERM o EVERY' [dtac @{thm meta_mp}, rtac (Rep_inv RS arg_cong RS subst), etac bspec,
+ dtac @{thm set_rev_mp}, rtac equalityD1, rtac set_natural', etac imageE,
+ hyp_subst_tac, rtac (Abs_inv RS ssubst), rtac subset, etac @{thm 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' (map4 mk_IH_tac Rep_invs Abs_invs (drop m set_natural's) subset1s), atac];
+
+ fun mk_induct_tac ((Rep, Rep_inv), subset) =
+ EVERY' [rtac (Rep_inv RS arg_cong RS subst), etac (Rep RS subset 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) ~~ subset2s) THEN'
+ rtac init_induct THEN'
+ DETERM o CONJ_WRAP' mk_closed_tac
+ (ks ~~ (morEs ~~ set_natural'ss))) 1
+ end;
+
+fun mk_fld_induct2_tac cTs cts fld_induct weak_fld_inducts {context = ctxt, prems = _} =
+ let
+ val n = length weak_fld_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 @{thm meta_mp} THEN_ALL_NEW Goal.norm_hhf_tac,
+ REPEAT_DETERM o dtac @{thm meta_spec}, etac (spec RS @{thm meta_mp}), atac],
+ atac];
+ in
+ EVERY' [rtac rev_mp, rtac (Drule.instantiate' cTs cts fld_induct),
+ EVERY' (map2 mk_inner_induct_tac weak_fld_inducts ks), rtac impI,
+ REPEAT_DETERM o eresolve_tac [conjE, allE],
+ CONJ_WRAP' (K atac) ks] 1
+ end;
+
+fun mk_map_tac m n iter map_comp_id map_cong =
+ EVERY' [rtac ext, rtac trans, rtac o_apply, rtac trans, rtac iter, 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, @{thm id_apply}])),
+ rtac sym, rtac o_apply] 1;
+
+fun mk_map_unique_tac m mor_def iter_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, @{thm id_apply}]))];
+ in
+ EVERY' [rtac iter_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 iter = EVERY' [rtac ext, rtac trans, rtac o_apply,
+ rtac trans, rtac iter, 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', @{thm trans[OF fun_cong[OF 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 @{thm 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 @{thm 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 fld_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 @{thm meta_mp}, atac,
+ dtac @{thm 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 @{thm 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 @{thm set_rev_mp}, atac,
+ REPEAT_DETERM o etac conjE, atac]];
+
+ fun mk_wpull wpull map_simp set_simps set_setss fld_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 fld_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 fld_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 @{thm 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 @{thm set_rev_mp}, rtac equalityD1, resolve_tac set_simp,
+ REPEAT_DETERM_N n o etac UnE]))))] 1);
+
+fun mk_rel_unfold_tac in_Irels i in_rel map_comp map_cong map_simp set_simps fld_inject
+ fld_unf 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_Irel = nth in_Irels (i - 1);
+ val le_arg_cong_fld_unf = fld_unf RS arg_cong RS @{thm ord_eq_le_trans};
+ val eq_arg_cong_fld_unf = fld_unf RS arg_cong RS trans;
+ val if_tac =
+ EVERY' [dtac (in_Irel RS iffD1), REPEAT_DETERM o eresolve_tac [exE, conjE, CollectE],
+ rtac (in_rel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
+ EVERY' (map2 (fn set_natural => fn set_incl =>
+ EVERY' [rtac conjI, rtac @{thm ord_eq_le_trans}, rtac set_natural,
+ rtac @{thm ord_eq_le_trans}, rtac @{thm trans[OF fun_cong[OF image_id] id_apply]},
+ rtac (set_incl RS @{thm subset_trans}), etac le_arg_cong_fld_unf])
+ passive_set_naturals set_incls),
+ CONJ_WRAP' (fn (in_Irel, (set_natural, set_set_incls)) =>
+ EVERY' [rtac @{thm ord_eq_le_trans}, rtac set_natural, rtac @{thm image_subsetI},
+ rtac (in_Irel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
+ CONJ_WRAP' (fn thm =>
+ EVERY' (map etac [thm RS @{thm subset_trans}, le_arg_cong_fld_unf]))
+ set_set_incls,
+ rtac conjI, rtac refl, rtac refl])
+ (in_Irels ~~ (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 (fld_inject RS iffD1), rtac trans, rtac sym, rtac map_simp,
+ etac eq_arg_cong_fld_unf])
+ @{thms fst_conv snd_conv}];
+ val only_if_tac =
+ EVERY' [dtac (in_rel RS iffD1), REPEAT_DETERM o eresolve_tac [exE, conjE, CollectE],
+ rtac (in_Irel RS iffD2), rtac exI, rtac conjI, rtac CollectI,
+ CONJ_WRAP' (fn (set_simp, passive_set_natural) =>
+ EVERY' [rtac @{thm ord_eq_le_trans}, rtac set_simp, rtac @{thm Un_least},
+ rtac @{thm ord_eq_le_trans}, rtac @{thm box_equals[OF _ refl]},
+ rtac passive_set_natural, rtac @{thm trans[OF fun_cong[OF image_id] id_apply]},
+ atac,
+ CONJ_WRAP_GEN' (rtac (Thm.permute_prems 0 1 @{thm Un_least}))
+ (fn (active_set_natural, in_Irel) => EVERY' [rtac @{thm ord_eq_le_trans},
+ rtac @{thm UN_cong[OF _ refl]}, rtac active_set_natural, rtac @{thm UN_least},
+ dtac @{thm set_rev_mp}, etac @{thm image_mono}, etac imageE,
+ dtac @{thm ssubst_mem[OF pair_collapse]}, dtac (in_Irel 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_Irels))])
+ (set_simps ~~ passive_set_naturals),
+ rtac conjI,
+ REPEAT_DETERM_N 2 o EVERY'[rtac trans, rtac map_simp, rtac (fld_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_Irel => EVERY' [rtac trans, rtac o_apply, dtac @{thm set_rev_mp}, atac,
+ dtac @{thm ssubst_mem[OF pair_collapse]}, dtac (in_Irel RS iffD1), dtac @{thm someI_ex},
+ REPEAT_DETERM o etac conjE, atac]) in_Irels),
+ 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/Codatatype/Tools/bnf_lfp_util.ML Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,76 @@
+(* Title: Codatatype_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/Codatatype/Tools/bnf_tactics.ML Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,317 @@
+(* Title: HOL/Codatatype/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 select_prem_tac: int -> (int -> tactic) -> int -> int -> tactic
+ val fo_rtac: thm -> Proof.context -> int -> tactic
+ val subst_tac: Proof.context -> thm list -> int -> tactic
+ val substs_tac: Proof.context -> thm list -> int -> 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 -> thm
+ val mk_Abs_inj_thm: thm -> thm -> thm
+ val mk_Abs_inverse_thm: thm -> thm -> thm
+ val mk_T_I: thm -> thm
+ val mk_T_subset1: thm -> thm
+ val mk_T_subset2: thm -> thm
+
+ val mk_Card_order_tac: thm -> tactic
+ 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_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
+ val mk_map_wppull_tac: thm -> thm -> thm -> thm -> thm list -> tactic
+ val mk_map_wpull_tac: thm -> thm list -> thm -> tactic
+ val mk_set_natural': thm -> thm
+ val mk_simple_wit_tac: thm list -> tactic
+
+ val mk_pred_unfold_tac: thm -> thm list -> thm -> {prems: 'a, context: Proof.context} -> tactic
+ val mk_rel_Gr_tac: thm -> thm -> thm -> thm -> thm -> thm -> thm -> thm list ->
+ {prems: thm list, context: Proof.context} -> tactic
+ val mk_rel_Id_tac: int -> thm -> thm -> {prems: 'a, context: Proof.context} -> tactic
+ val mk_rel_O_tac: thm -> thm -> thm -> thm -> thm -> thm list ->
+ {prems: thm list, context: Proof.context} -> tactic
+ val mk_in_rel_tac: thm -> int -> {prems: 'b, context: Proof.context} -> tactic
+ val mk_rel_converse_tac: thm -> tactic
+ val mk_rel_converse_le_tac: thm -> thm -> thm -> thm -> thm list ->
+ {prems: thm list, context: Proof.context} -> tactic
+ val mk_rel_mono_tac: thm -> thm -> {prems: 'a, context: Proof.context} -> tactic
+end;
+
+structure BNF_Tactics : BNF_TACTICS =
+struct
+
+open BNF_Util
+
+val set_mp = @{thm set_mp};
+
+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);
+
+(*unlike "unfold_tac", succeeds when the RHS contains schematic variables not in the LHS*)
+fun subst_tac ctxt = EqSubst.eqsubst_tac ctxt [0];
+fun substs_tac ctxt = REPEAT_DETERM oo subst_tac ctxt;
+
+
+
+(* Theorems for typedefs with UNIV as representing set *)
+
+(*equivalent to UNIV_I for the representing set of the particular type T*)
+fun mk_T_subset1 def = set_mp OF [def RS meta_eq_to_obj_eq RS equalityD2];
+fun mk_T_subset2 def = set_mp OF [def RS meta_eq_to_obj_eq RS equalityD1];
+fun mk_T_I def = UNIV_I RS mk_T_subset1 def;
+
+fun mk_Abs_inverse_thm def inv = mk_T_I def RS inv;
+fun mk_Abs_inj_thm def inj = inj OF (replicate 2 (mk_T_I def));
+fun mk_Abs_bij_thm ctxt def inj surj = rule_by_tactic ctxt ((rtac surj THEN' etac exI) 1)
+ (mk_Abs_inj_thm def inj 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 mk_Card_order_tac card_order =
+ (rtac conjE THEN'
+ rtac @{thm card_order_on_Card_order} THEN'
+ rtac card_order THEN'
+ atac) 1;
+
+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 @{thm 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_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_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_simple_wit_tac wit_thms = ALLGOALS (atac ORELSE' eresolve_tac (@{thm emptyE} :: wit_thms));
+
+fun mk_map_comp_id_tac map_comp =
+ (rtac trans THEN' rtac map_comp) 1 THEN
+ REPEAT_DETERM (stac @{thm o_id} 1) THEN
+ rtac refl 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;
+
+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_rel_Gr_tac rel_def map_id map_cong map_wpull in_cong map_id' map_comp set_naturals
+ {context = ctxt, prems = _} =
+ let
+ val n = length set_naturals;
+ in
+ if null set_naturals then
+ Local_Defs.unfold_tac ctxt [rel_def] THEN EVERY' [rtac @{thm Gr_UNIV_id}, rtac map_id] 1
+ else Local_Defs.unfold_tac ctxt [rel_def, @{thm Gr_def}] THEN
+ EVERY' [rtac equalityI, rtac subsetI,
+ REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE, @{thm relcompE}, @{thm converseE}],
+ REPEAT_DETERM o dtac @{thm Pair_eq[THEN subst, of "%x. x"]},
+ REPEAT_DETERM o etac conjE, hyp_subst_tac,
+ rtac CollectI, rtac exI, rtac conjI, stac @{thm Pair_eq}, 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 @{thm Pair_eq[THEN subst, of "%x. x"]},
+ REPEAT_DETERM o etac conjE, hyp_subst_tac,
+ rtac allE, rtac subst, rtac @{thm wpull_def}, rtac map_wpull,
+ REPEAT_DETERM_N n o rtac @{thm wpull_Gr}, etac allE, etac impE, rtac conjI, atac,
+ rtac conjI, REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac CollectI,
+ CONJ_WRAP' (fn thm => EVERY' [rtac (thm RS @{thm ord_eq_le_trans}),
+ rtac @{thm image_mono}, atac]) set_naturals,
+ rtac sym, rtac map_id', REPEAT_DETERM o eresolve_tac [bexE, conjE],
+ rtac @{thm relcompI}, rtac @{thm converseI},
+ REPEAT_DETERM_N 2 o EVERY' [rtac CollectI, rtac exI,
+ rtac conjI, stac @{thm Pair_eq}, rtac conjI, rtac refl, etac sym,
+ etac @{thm set_rev_mp}, rtac equalityD1, rtac in_cong,
+ REPEAT_DETERM_N n o rtac @{thm Gr_def}]] 1
+ end;
+
+fun mk_rel_Id_tac n rel_Gr map_id {context = ctxt, prems = _} =
+ Local_Defs.unfold_tac ctxt [rel_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 @{thm subset_UNIV}, rtac subsetI, rtac CollectI,
+ CONJ_WRAP' (K (rtac @{thm subset_UNIV})) (1 upto n), rtac refl] 1);
+
+fun mk_rel_mono_tac rel_def in_mono {context = ctxt, prems = _} =
+ Local_Defs.unfold_tac ctxt [rel_def] 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_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_rel_converse_le_tac rel_def rel_Id map_cong map_comp set_naturals
+ {context = ctxt, prems = _} =
+ let
+ val n = length set_naturals;
+ in
+ if null set_naturals then
+ Local_Defs.unfold_tac ctxt [rel_Id] THEN rtac equalityD2 1 THEN rtac @{thm converse_Id} 1
+ else Local_Defs.unfold_tac ctxt [rel_def, @{thm Gr_def}] THEN
+ EVERY' [rtac @{thm subrelI},
+ REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE, @{thm relcompE}, @{thm converseE}],
+ REPEAT_DETERM o dtac @{thm Pair_eq[THEN subst, of "%x. x"]},
+ 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, stac @{thm Pair_eq}, 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_rel_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_rel_O_tac rel_def rel_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 Local_Defs.unfold_tac ctxt [rel_Id] THEN rtac (@{thm Id_O_R} RS sym) 1
+ else Local_Defs.unfold_tac ctxt [rel_def, @{thm Gr_def}] THEN
+ EVERY' [rtac equalityI, rtac @{thm subrelI},
+ REPEAT_DETERM o eresolve_tac [CollectE, exE, conjE, @{thm relcompE}, @{thm converseE}],
+ REPEAT_DETERM o dtac @{thm Pair_eq[THEN subst, of "%x. x"]},
+ REPEAT_DETERM o etac conjE, hyp_subst_tac,
+ rtac @{thm relcompI}, rtac @{thm relcompI}, rtac @{thm converseI},
+ rtac CollectI, rtac exI, rtac conjI, stac @{thm Pair_eq}, 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, stac @{thm Pair_eq}, 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, stac @{thm Pair_eq}, 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, stac @{thm Pair_eq}, 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 @{thm Pair_eq[THEN subst, of "%x. x"]},
+ 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, stac @{thm Pair_eq}, 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_rel_tac rel_def m {context = ctxt, prems = _} =
+ let
+ val ls' = replicate (Int.max (1, m)) ();
+ in
+ Local_Defs.unfold_tac ctxt (rel_def ::
+ @{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;
+
+fun mk_pred_unfold_tac pred_def pred_defs rel_unfold {context = ctxt, prems = _} =
+ Local_Defs.unfold_tac ctxt (@{thm mem_Collect_eq_split} :: pred_def :: pred_defs) THEN
+ rtac rel_unfold 1;
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Codatatype/Tools/bnf_util.ML Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,508 @@
+(* Title: HOL/Codatatype/Tools/bnf_util.ML
+ Author: Dmitriy Traytel, TU Muenchen
+ Copyright 2012
+
+General library functions.
+*)
+
+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 transpose: 'a list list -> 'a list 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_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_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 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_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
+
+ (*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 mk_nthI: int -> int -> thm
+ val mk_nth_conv: int -> int -> thm
+ val mk_ordLeq_csum: int -> int -> thm -> thm
+ val mk_UnN: int -> int -> thm
+
+ val ctrans: thm
+ val o_apply: thm
+ val mk_sym: thm -> thm
+ val mk_trans: thm -> thm -> thm
+ val mk_unabs_def: int -> thm -> thm
+
+ val mk_permute: ''a list -> ''a list -> 'b list -> 'b list
+ val find_indices: ''a list -> ''a list -> int list
+
+ 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;
+
+(*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 mk_TFrees n = apfst (map TFree) o Variable.invent_types (replicate n (HOLogic.typeS));
+fun mk_TFreess ns = apfst (map (map TFree)) o
+ fold_map Variable.invent_types (map (fn n => replicate n (HOLogic.typeS)) ns);
+
+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 names => Variable.variant_fixes names ctxt) oo mk_names;
+fun mk_Frees x Ts ctxt = mk_fresh_names ctxt (length Ts) x
+ |>> (fn names => map2 (curry Free) names Ts);
+fun mk_Freess x Tss ctxt =
+ fold_map2 (fn name => fn Ts => fn ctxt =>
+ mk_fresh_names ctxt (length Ts) name) (mk_names (length Tss) x) Tss ctxt
+ |>> (fn namess => map2 (map2 (curry Free)) namess Tss);
+fun mk_Frees' x Ts ctxt = mk_fresh_names ctxt (length Ts) x
+ |>> (fn names => `(map Free) (names ~~ Ts));
+fun mk_Freess' x Tss ctxt =
+ fold_map2 (fn name => fn Ts => fn ctxt =>
+ mk_fresh_names ctxt (length Ts) name) (mk_names (length Tss) x) Tss ctxt
+ |>> (fn namess => map_split (`(map Free) o (op ~~)) (namess ~~ Tss));
+
+
+(** Types **)
+
+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 **)
+
+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;
+
+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};
+
+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 mk_conjIN 1 = @{thm TrueE[OF TrueI]}
+ | mk_conjIN n = mk_conjIN (n - 1) RSN (2, conjI);
+
+fun mk_ordLeq_csum 1 1 thm = thm
+ | mk_ordLeq_csum _ 1 thm = @{thm ordLeq_transitive} OF [thm, @{thm ordLeq_csum1}]
+ | mk_ordLeq_csum 2 2 thm = @{thm ordLeq_transitive} OF [thm, @{thm ordLeq_csum2}]
+ | mk_ordLeq_csum n m thm = @{thm ordLeq_transitive} OF
+ [mk_ordLeq_csum (n - 1) (m - 1) thm, @{thm ordLeq_csum2[OF Card_order_csum]}];
+
+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_UnN' 0 = @{thm UnI2}
+ | mk_UnN' m = mk_UnN' (m - 1) RS @{thm UnI1};
+in
+ fun mk_UnN 1 1 = @{thm TrueE[OF TrueI]}
+ | mk_UnN n 1 = Library.foldr1 (op RS o swap) (replicate (n - 1) @{thm UnI1})
+ | mk_UnN n m = mk_UnN' (n - m)
+end;
+
+fun transpose [] = []
+ | transpose ([] :: xss) = transpose xss
+ | transpose xss = map hd xss :: transpose (map tl xss);
+
+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]]};
+
+
+end;
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Ordinals_and_Cardinals/Cardinal_Arithmetic.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,878 @@
+(* Title: HOL/Ordinals_and_Cardinals/Cardinal_Arithmetic.thy
+ Author: Dmitriy Traytel, TU Muenchen
+ Copyright 2012
+
+Cardinal arithmetic.
+*)
+
+header {* Cardinal Arithmetic *}
+
+theory Cardinal_Arithmetic
+imports Cardinal_Order_Relation_Base
+begin
+
+text {*
+ The following collection of lemmas should be seen as an user interface to the HOL Theory
+ of cardinals. It is not expected to be complete in any sense, since its
+ development was driven by demand arising from the development of the (co)datatype package.
+*}
+
+(*library candidate*)
+lemma dir_image: "\<lbrakk>\<And>x y. (f x = f y) = (x = y); Card_order r\<rbrakk> \<Longrightarrow> r =o dir_image r f"
+by (rule dir_image_ordIso) (auto simp add: inj_on_def card_order_on_def)
+
+(*should supersede a weaker lemma from the library*)
+lemma dir_image_Field: "Field (dir_image r f) = f ` Field r"
+unfolding dir_image_def Field_def Range_def Domain_def by fastforce
+
+lemma card_order_dir_image:
+ assumes bij: "bij f" and co: "card_order r"
+ shows "card_order (dir_image r f)"
+proof -
+ from assms have "Field (dir_image r f) = UNIV"
+ using card_order_on_Card_order[of UNIV r] unfolding bij_def dir_image_Field by auto
+ moreover from bij have "\<And>x y. (f x = f y) = (x = y)" unfolding bij_def inj_on_def by auto
+ with co have "Card_order (dir_image r f)"
+ using card_order_on_Card_order[of UNIV r] Card_order_ordIso2[OF _ dir_image] by blast
+ ultimately show ?thesis by auto
+qed
+
+(*library candidate*)
+lemma ordIso_refl: "Card_order r \<Longrightarrow> r =o r"
+by (rule card_order_on_ordIso)
+
+(*library candidate*)
+lemma ordLeq_refl: "Card_order r \<Longrightarrow> r \<le>o r"
+by (rule ordIso_imp_ordLeq, rule card_order_on_ordIso)
+
+(*library candidate*)
+lemma card_of_ordIso_subst: "A = B \<Longrightarrow> |A| =o |B|"
+by (simp only: ordIso_refl card_of_Card_order)
+
+(*library candidate*)
+lemma card_of_Times_Plus_distrib:
+ "|A <*> (B <+> C)| =o |A <*> B <+> A <*> C|" (is "|?RHS| =o |?LHS|")
+proof -
+ let ?f = "\<lambda>(a, bc). case bc of Inl b \<Rightarrow> Inl (a, b) | Inr c \<Rightarrow> Inr (a, c)"
+ have "bij_betw ?f ?RHS ?LHS" unfolding bij_betw_def inj_on_def by force
+ thus ?thesis using card_of_ordIso by blast
+qed
+
+(*library candidate*)
+lemma Field_card_order: "card_order r \<Longrightarrow> Field r = UNIV"
+using card_order_on_Card_order[of UNIV r] by simp
+
+subsection {* Zero *}
+
+definition czero where
+ "czero = card_of {}"
+
+lemma czero_ordIso:
+ "czero =o czero"
+using card_of_empty_ordIso by (simp add: czero_def)
+
+lemma card_of_ordIso_czero_iff_empty:
+ "|A| =o (czero :: 'a rel) \<longleftrightarrow> A = ({} :: 'a set)"
+unfolding czero_def by (rule iffI[OF card_of_empty2]) (auto simp: card_of_refl)
+
+(* A "not czero" Cardinal predicate *)
+abbreviation Cnotzero where
+ "Cnotzero (r :: 'a rel) \<equiv> \<not>(r =o (czero :: 'a rel)) \<and> Card_order r"
+
+(*helper*)
+lemma Cnotzero_imp_not_empty: "Cnotzero r \<Longrightarrow> Field r \<noteq> {}"
+by (metis Card_order_iff_ordIso_card_of czero_def)
+
+lemma czeroI:
+ "\<lbrakk>Card_order r; Field r = {}\<rbrakk> \<Longrightarrow> r =o czero"
+using Cnotzero_imp_not_empty ordIso_transitive[OF _ czero_ordIso] by blast
+
+lemma czeroE:
+ "r =o czero \<Longrightarrow> Field r = {}"
+unfolding czero_def
+by (drule card_of_cong) (simp only: Field_card_of card_of_empty2)
+
+lemma Cnotzero_mono:
+ "\<lbrakk>Cnotzero r; Card_order q; r \<le>o q\<rbrakk> \<Longrightarrow> Cnotzero q"
+apply (rule ccontr)
+apply auto
+apply (drule czeroE)
+apply (erule notE)
+apply (erule czeroI)
+apply (drule card_of_mono2)
+apply (simp only: card_of_empty3)
+done
+
+subsection {* Infinite cardinals *}
+
+definition cinfinite where
+ "cinfinite r = infinite (Field r)"
+
+abbreviation Cinfinite where
+ "Cinfinite r \<equiv> cinfinite r \<and> Card_order r"
+
+lemma natLeq_ordLeq_cinfinite:
+ assumes inf: "Cinfinite r"
+ shows "natLeq \<le>o r"
+proof -
+ from inf have "natLeq \<le>o |Field r|" by (simp add: cinfinite_def infinite_iff_natLeq_ordLeq)
+ also from inf have "|Field r| =o r" by (simp add: card_of_unique ordIso_symmetric)
+ finally show ?thesis .
+qed
+
+lemma cinfinite_not_czero: "cinfinite r \<Longrightarrow> \<not> (r =o (czero :: 'a rel))"
+unfolding cinfinite_def by (metis czeroE finite.emptyI)
+
+lemma Cinfinite_Cnotzero: "Cinfinite r \<Longrightarrow> Cnotzero r"
+by (metis cinfinite_not_czero)
+
+lemma Cinfinite_cong: "\<lbrakk>r1 =o r2; Cinfinite r1\<rbrakk> \<Longrightarrow> Cinfinite r2"
+by (metis Card_order_ordIso2 card_of_mono2 card_of_ordLeq_infinite cinfinite_def ordIso_iff_ordLeq)
+
+lemma cinfinite_mono: "\<lbrakk>r1 \<le>o r2; cinfinite r1\<rbrakk> \<Longrightarrow> cinfinite r2"
+by (metis card_of_mono2 card_of_ordLeq_infinite cinfinite_def)
+
+
+subsection {* Binary sum *}
+
+definition csum (infixr "+c" 65) where
+ "r1 +c r2 \<equiv> |Field r1 <+> Field r2|"
+
+lemma Card_order_csum:
+ "Card_order (r1 +c r2)"
+unfolding csum_def by (simp add: card_of_Card_order)
+
+lemma csum_Cnotzero1:
+ "Cnotzero r1 \<Longrightarrow> Cnotzero (r1 +c r2)"
+unfolding csum_def
+by (metis Cnotzero_imp_not_empty Field_card_of Plus_eq_empty_conv card_of_card_order_on czeroE)
+
+lemma csum_Cnotzero2:
+ "Cnotzero r2 \<Longrightarrow> Cnotzero (r1 +c r2)"
+unfolding csum_def
+by (metis Cnotzero_imp_not_empty Field_card_of Plus_eq_empty_conv card_of_card_order_on czeroE)
+
+lemma card_order_csum:
+ assumes "card_order r1" "card_order r2"
+ shows "card_order (r1 +c r2)"
+proof -
+ have "Field r1 = UNIV" "Field r2 = UNIV" using assms card_order_on_Card_order by auto
+ thus ?thesis unfolding csum_def by (auto simp: card_of_card_order_on)
+qed
+
+lemma cinfinite_csum:
+ "cinfinite r1 \<or> cinfinite r2 \<Longrightarrow> cinfinite (r1 +c r2)"
+unfolding cinfinite_def csum_def by (auto simp: Field_card_of)
+
+lemma Cinfinite_csum:
+ "Cinfinite r1 \<or> Cinfinite r2 \<Longrightarrow> Cinfinite (r1 +c r2)"
+unfolding cinfinite_def csum_def by (metis Field_card_of card_of_Card_order finite_Plus_iff)
+
+lemma Cinfinite_csum_strong:
+ "\<lbrakk>Cinfinite r1; Cinfinite r2\<rbrakk> \<Longrightarrow> Cinfinite (r1 +c r2)"
+by (metis Cinfinite_csum)
+
+lemma Cinfinite_csum1:
+ "Cinfinite r1 \<Longrightarrow> Cinfinite (r1 +c r2)"
+unfolding cinfinite_def csum_def by (metis Field_card_of card_of_Card_order finite_Plus_iff)
+
+lemma csum_cong: "\<lbrakk>p1 =o r1; p2 =o r2\<rbrakk> \<Longrightarrow> p1 +c p2 =o r1 +c r2"
+by (simp only: csum_def ordIso_Plus_cong)
+
+lemma csum_cong1: "p1 =o r1 \<Longrightarrow> p1 +c q =o r1 +c q"
+by (simp only: csum_def ordIso_Plus_cong1)
+
+lemma csum_cong2: "p2 =o r2 \<Longrightarrow> q +c p2 =o q +c r2"
+by (simp only: csum_def ordIso_Plus_cong2)
+
+lemma csum_mono: "\<lbrakk>p1 \<le>o r1; p2 \<le>o r2\<rbrakk> \<Longrightarrow> p1 +c p2 \<le>o r1 +c r2"
+by (simp only: csum_def ordLeq_Plus_mono)
+
+lemma csum_mono1: "p1 \<le>o r1 \<Longrightarrow> p1 +c q \<le>o r1 +c q"
+by (simp only: csum_def ordLeq_Plus_mono1)
+
+lemma csum_mono2: "p2 \<le>o r2 \<Longrightarrow> q +c p2 \<le>o q +c r2"
+by (simp only: csum_def ordLeq_Plus_mono2)
+
+lemma ordLeq_csum1: "Card_order p1 \<Longrightarrow> p1 \<le>o p1 +c p2"
+by (simp only: csum_def Card_order_Plus1)
+
+lemma ordLeq_csum2: "Card_order p2 \<Longrightarrow> p2 \<le>o p1 +c p2"
+by (simp only: csum_def Card_order_Plus2)
+
+lemma csum_com: "p1 +c p2 =o p2 +c p1"
+by (simp only: csum_def card_of_Plus_commute)
+
+lemma csum_assoc: "(p1 +c p2) +c p3 =o p1 +c p2 +c p3"
+by (simp only: csum_def Field_card_of card_of_Plus_assoc)
+
+lemma Plus_csum: "|A <+> B| =o |A| +c |B|"
+by (simp only: csum_def Field_card_of card_of_refl)
+
+lemma Un_csum: "|A \<union> B| \<le>o |A| +c |B|"
+using ordLeq_ordIso_trans[OF card_of_Un_Plus_ordLeq Plus_csum] by blast
+
+
+subsection {* One *}
+
+definition cone where
+ "cone = card_of {()}"
+
+lemma Card_order_cone: "Card_order cone"
+unfolding cone_def by (rule card_of_Card_order)
+
+lemma single_cone:
+ "|{x}| =o cone"
+proof -
+ let ?f = "\<lambda>x. ()"
+ have "bij_betw ?f {x} {()}" unfolding bij_betw_def by auto
+ thus ?thesis unfolding cone_def using card_of_ordIso by blast
+qed
+
+lemma cone_not_czero: "\<not> (cone =o czero)"
+unfolding czero_def cone_def by (metis empty_not_insert card_of_empty3[of "{()}"] ordIso_iff_ordLeq)
+
+lemma cone_Cnotzero: "Cnotzero cone"
+by (simp add: cone_not_czero Card_order_cone)
+
+lemma cone_ordLeq_Cnotzero: "Cnotzero r \<Longrightarrow> cone \<le>o r"
+unfolding cone_def by (metis Card_order_singl_ordLeq czeroI)
+
+
+subsection{* Two *}
+
+definition ctwo where
+ "ctwo = |UNIV :: bool set|"
+
+lemma Card_order_ctwo: "Card_order ctwo"
+unfolding ctwo_def by (rule card_of_Card_order)
+
+lemma cone_ordLeq_ctwo: "cone \<le>o ctwo"
+unfolding cone_def ctwo_def card_of_ordLeq[symmetric] by auto
+
+lemma ctwo_not_czero: "\<not> (ctwo =o czero)"
+using card_of_empty3[of "UNIV :: bool set"] ordIso_iff_ordLeq
+unfolding czero_def ctwo_def by (metis UNIV_not_empty)
+
+lemma ctwo_Cnotzero: "Cnotzero ctwo"
+by (simp add: ctwo_not_czero Card_order_ctwo)
+
+
+subsection {* Family sum *}
+
+definition Csum where
+ "Csum r rs \<equiv> |SIGMA i : Field r. Field (rs i)|"
+
+(* Similar setup to the one for SIGMA from theory Big_Operators: *)
+syntax "_Csum" ::
+ "pttrn => ('a * 'a) set => 'b * 'b set => (('a * 'b) * ('a * 'b)) set"
+ ("(3CSUM _:_. _)" [0, 51, 10] 10)
+
+translations
+ "CSUM i:r. rs" == "CONST Csum r (%i. rs)"
+
+lemma SIGMA_CSUM: "|SIGMA i : I. As i| = (CSUM i : |I|. |As i| )"
+by (auto simp: Csum_def Field_card_of)
+
+(* NB: Always, under the cardinal operator,
+operations on sets are reduced automatically to operations on cardinals.
+This should make cardinal reasoning more direct and natural. *)
+
+
+subsection {* Product *}
+
+definition cprod (infixr "*c" 80) where
+ "r1 *c r2 = |Field r1 <*> Field r2|"
+
+lemma Times_cprod: "|A \<times> B| =o |A| *c |B|"
+by (simp only: cprod_def Field_card_of card_of_refl)
+
+lemma card_order_cprod:
+ assumes "card_order r1" "card_order r2"
+ shows "card_order (r1 *c r2)"
+proof -
+ have "Field r1 = UNIV" "Field r2 = UNIV" using assms card_order_on_Card_order by auto
+ thus ?thesis by (auto simp: cprod_def card_of_card_order_on)
+qed
+
+lemma Card_order_cprod: "Card_order (r1 *c r2)"
+by (simp only: cprod_def Field_card_of card_of_card_order_on)
+
+lemma cprod_cong2: "p2 =o r2 \<Longrightarrow> q *c p2 =o q *c r2"
+by (simp only: cprod_def ordIso_Times_cong2)
+
+lemma cprod_mono1: "p1 \<le>o r1 \<Longrightarrow> p1 *c q \<le>o r1 *c q"
+by (simp only: cprod_def ordLeq_Times_mono1)
+
+lemma cprod_mono2: "p2 \<le>o r2 \<Longrightarrow> q *c p2 \<le>o q *c r2"
+by (simp only: cprod_def ordLeq_Times_mono2)
+
+lemma ordLeq_cprod1: "\<lbrakk>Card_order p1; Cnotzero p2\<rbrakk> \<Longrightarrow> p1 \<le>o p1 *c p2"
+unfolding cprod_def by (metis Card_order_Times1 czeroI)
+
+lemma ordLeq_cprod2: "\<lbrakk>Cnotzero p1; Card_order p2\<rbrakk> \<Longrightarrow> p2 \<le>o p1 *c p2"
+unfolding cprod_def by (metis Card_order_Times2 czeroI)
+
+lemma cinfinite_cprod: "\<lbrakk>cinfinite r1; cinfinite r2\<rbrakk> \<Longrightarrow> cinfinite (r1 *c r2)"
+by (simp add: cinfinite_def cprod_def Field_card_of infinite_cartesian_product)
+
+lemma cinfinite_cprod2: "\<lbrakk>Cnotzero r1; Cinfinite r2\<rbrakk> \<Longrightarrow> cinfinite (r1 *c r2)"
+by (metis cinfinite_mono ordLeq_cprod2)
+
+lemma Cinfinite_cprod2: "\<lbrakk>Cnotzero r1; Cinfinite r2\<rbrakk> \<Longrightarrow> Cinfinite (r1 *c r2)"
+by (blast intro: cinfinite_cprod2 Card_order_cprod)
+
+lemma cprod_com: "p1 *c p2 =o p2 *c p1"
+by (simp only: cprod_def card_of_Times_commute)
+
+lemma card_of_Csum_Times:
+ "\<forall>i \<in> I. |A i| \<le>o |B| \<Longrightarrow> (CSUM i : |I|. |A i| ) \<le>o |I| *c |B|"
+by (simp only: Csum_def cprod_def Field_card_of card_of_Sigma_Times)
+
+lemma card_of_Csum_Times':
+ assumes "Card_order r" "\<forall>i \<in> I. |A i| \<le>o r"
+ shows "(CSUM i : |I|. |A i| ) \<le>o |I| *c r"
+proof -
+ from assms(1) have *: "r =o |Field r|" by (simp add: card_of_unique)
+ with assms(2) have "\<forall>i \<in> I. |A i| \<le>o |Field r|" by (blast intro: ordLeq_ordIso_trans)
+ hence "(CSUM i : |I|. |A i| ) \<le>o |I| *c |Field r|" by (simp only: card_of_Csum_Times)
+ also from * have "|I| *c |Field r| \<le>o |I| *c r"
+ by (simp only: Field_card_of card_of_refl cprod_def ordIso_imp_ordLeq)
+ finally show ?thesis .
+qed
+
+lemma cprod_csum_distrib1: "r1 *c r2 +c r1 *c r3 =o r1 *c (r2 +c r3)"
+unfolding csum_def cprod_def by (simp add: Field_card_of card_of_Times_Plus_distrib ordIso_symmetric)
+
+lemma csum_absorb2': "\<lbrakk>Card_order r2; r1 \<le>o r2; cinfinite r1 \<or> cinfinite r2\<rbrakk> \<Longrightarrow> r1 +c r2 =o r2"
+unfolding csum_def by (metis Card_order_Plus_infinite cinfinite_def cinfinite_mono)
+
+lemma csum_absorb1':
+ assumes card: "Card_order r2"
+ and r12: "r1 \<le>o r2" and cr12: "cinfinite r1 \<or> cinfinite r2"
+ shows "r2 +c r1 =o r2"
+by (rule ordIso_transitive, rule csum_com, rule csum_absorb2', (simp only: assms)+)
+
+lemma csum_absorb1: "\<lbrakk>Cinfinite r2; r1 \<le>o r2\<rbrakk> \<Longrightarrow> r2 +c r1 =o r2"
+by (rule csum_absorb1') auto
+
+lemma cprod_infinite1': "\<lbrakk>Cinfinite r; Cnotzero p; p \<le>o r\<rbrakk> \<Longrightarrow> r *c p =o r"
+unfolding cinfinite_def cprod_def
+by (rule Card_order_Times_infinite[THEN conjunct1]) (blast intro: czeroI)+
+
+lemma cprod_infinite: "Cinfinite r \<Longrightarrow> r *c r =o r"
+using cprod_infinite1' Cinfinite_Cnotzero ordLeq_refl by blast
+
+
+subsection {* Exponentiation *}
+
+definition cexp (infixr "^c" 80) where
+ "r1 ^c r2 \<equiv> |Func (Field r2) (Field r1)|"
+
+definition ccexp (infixr "^^c" 80) where
+ "r1 ^^c r2 \<equiv> |Pfunc (Field r2) (Field r1)|"
+
+lemma cexp_ordLeq_ccexp: "r1 ^c r2 \<le>o r1 ^^c r2"
+unfolding cexp_def ccexp_def by (rule card_of_mono1) (rule Func_Pfunc)
+
+lemma card_order_ccexp:
+ assumes "card_order r1" "card_order r2"
+ shows "card_order (r1 ^^c r2)"
+proof -
+ have "Field r1 = UNIV" "Field r2 = UNIV" using assms card_order_on_Card_order by auto
+ thus ?thesis unfolding ccexp_def Pfunc_def
+ by (auto simp: card_of_card_order_on split: option.split)
+qed
+
+lemma Card_order_cexp: "Card_order (r1 ^c r2)"
+unfolding cexp_def by (rule card_of_Card_order)
+
+lemma cexp_mono':
+ assumes 1: "p1 \<le>o r1" and 2: "p2 \<le>o r2"
+ and n1: "Field p1 \<noteq> {} \<or> cone \<le>o r1 ^c r2"
+ and n2: "Field p2 = {} \<Longrightarrow> Field r2 = {}"
+ shows "p1 ^c p2 \<le>o r1 ^c r2"
+proof(cases "Field p1 = {}")
+ case True
+ hence "|Field (p1 ^c p2)| \<le>o cone"
+ unfolding czero_def cone_def cexp_def Field_card_of
+ by (cases "Field p2 = {}", auto intro: card_of_ordLeqI2 simp: Func_empty)
+ (metis Func_is_emp card_of_empty ex_in_conv)
+ hence "p1 ^c p2 \<le>o cone" by (simp add: Field_card_of cexp_def)
+ thus ?thesis using True n1 ordLeq_transitive by auto
+next
+ case False
+ have 1: "|Field p1| \<le>o |Field r1|" and 2: "|Field p2| \<le>o |Field r2|"
+ using 1 2 by (auto simp: card_of_mono2)
+ obtain f1 where f1: "f1 ` Field r1 = Field p1"
+ using 1 unfolding card_of_ordLeq2[OF False, symmetric] by auto
+ obtain f2 where f2: "inj_on f2 (Field p2)" "f2 ` Field p2 \<subseteq> Field r2"
+ using 2 unfolding card_of_ordLeq[symmetric] by blast
+ have 0: "Func_map (Field p2) f1 f2 ` (Field (r1 ^c r2)) = Field (p1 ^c p2)"
+ unfolding cexp_def Field_card_of using Func_map_surj[OF f1 f2 n2, symmetric] .
+ have 00: "Field (p1 ^c p2) \<noteq> {}" unfolding cexp_def Field_card_of Func_is_emp
+ using False by simp
+ show ?thesis
+ using 0 card_of_ordLeq2[OF 00] unfolding cexp_def Field_card_of by blast
+qed
+
+lemma cexp_mono:
+ assumes 1: "p1 \<le>o r1" and 2: "p2 \<le>o r2"
+ and n1: "Cnotzero p1 \<or> cone \<le>o r1 ^c r2"
+ and n2: "p2 =o czero \<Longrightarrow> r2 =o czero" and card: "Card_order p2"
+ shows "p1 ^c p2 \<le>o r1 ^c r2"
+proof (rule cexp_mono'[OF 1 2])
+ show "Field p1 \<noteq> {} \<or> cone \<le>o r1 ^c r2"
+ proof (cases "Cnotzero p1")
+ case True show ?thesis using Cnotzero_imp_not_empty[OF True] by (rule disjI1)
+ next
+ case False with n1 show ?thesis by blast
+ qed
+qed (rule czeroI[OF card, THEN n2, THEN czeroE])
+
+lemma cexp_mono1:
+ assumes 1: "p1 \<le>o r1"
+ and n1: "Cnotzero p1 \<or> cone \<le>o r1 ^c q" and q: "Card_order q"
+ shows "p1 ^c q \<le>o r1 ^c q"
+using ordLeq_refl[OF q] by (rule cexp_mono[OF 1 _ n1]) (auto simp: q)
+
+lemma cexp_mono1_Cnotzero: "\<lbrakk>p1 \<le>o r1; Cnotzero p1; Card_order q\<rbrakk> \<Longrightarrow> p1 ^c q \<le>o r1 ^c q"
+by (simp add: cexp_mono1)
+
+lemma cexp_mono1_cone_ordLeq: "\<lbrakk>p1 \<le>o r1; cone \<le>o r1 ^c q; Card_order q\<rbrakk> \<Longrightarrow> p1 ^c q \<le>o r1 ^c q"
+using assms by (simp add: cexp_mono1)
+
+lemma cexp_mono2':
+ assumes 2: "p2 \<le>o r2" and q: "Card_order q"
+ and n1: "Field q \<noteq> {} \<or> cone \<le>o q ^c r2"
+ and n2: "Field p2 = {} \<Longrightarrow> Field r2 = {}"
+ shows "q ^c p2 \<le>o q ^c r2"
+using ordLeq_refl[OF q] by (rule cexp_mono'[OF _ 2 n1 n2]) auto
+
+lemma cexp_mono2:
+ assumes 2: "p2 \<le>o r2" and q: "Card_order q"
+ and n1: "Cnotzero q \<or> cone \<le>o q ^c r2"
+ and n2: "p2 =o czero \<Longrightarrow> r2 =o czero" and card: "Card_order p2"
+ shows "q ^c p2 \<le>o q ^c r2"
+using ordLeq_refl[OF q] by (rule cexp_mono[OF _ 2 n1 n2 card]) auto
+
+lemma cexp_mono2_Cnotzero:
+ assumes "p2 \<le>o r2" "Cnotzero q" and n2: "Cnotzero p2"
+ shows "q ^c p2 \<le>o q ^c r2"
+proof (rule cexp_mono)
+ assume *: "p2 =o czero"
+ have False using n2 czeroI czeroE[OF *] by blast
+ thus "r2 =o czero" by blast
+qed (auto simp add: assms ordLeq_refl)
+
+lemma cexp_cong:
+ assumes 1: "p1 =o r1" and 2: "p2 =o r2"
+ and p1: "Cnotzero p1 \<or> cone \<le>o r1 ^c r2" and Cr: "Card_order r2"
+ and r1: "Cnotzero r1 \<or> cone \<le>o p1 ^c p2" and Cp: "Card_order p2"
+ shows "p1 ^c p2 =o r1 ^c r2"
+proof -
+ obtain f where "bij_betw f (Field p2) (Field r2)"
+ using 2 card_of_ordIso[of "Field p2" "Field r2"] card_of_cong by auto
+ hence 0: "Field p2 = {} \<longleftrightarrow> Field r2 = {}" unfolding bij_betw_def by auto
+ have r: "p2 =o czero \<Longrightarrow> r2 =o czero"
+ and p: "r2 =o czero \<Longrightarrow> p2 =o czero"
+ using 0 Cr Cp czeroE czeroI by auto
+ show ?thesis using 0 1 2 unfolding ordIso_iff_ordLeq
+ using r p cexp_mono[OF _ _ p1 _ Cp] cexp_mono[OF _ _ r1 _ Cr]
+ by blast
+qed
+
+lemma cexp_cong1:
+ assumes 1: "p1 =o r1" and q: "Card_order q"
+ and p1: "Cnotzero p1 \<or> cone \<le>o r1 ^c q"
+ and r1: "Cnotzero r1 \<or> cone \<le>o p1 ^c q"
+ shows "p1 ^c q =o r1 ^c q"
+by (rule cexp_cong[OF 1 _ p1 q r1 q]) (rule ordIso_refl[OF q])
+
+lemma cexp_cong1_Cnotzero:
+ assumes "p1 =o r1" "Card_order q" "Cnotzero p1" "Cnotzero r1"
+ shows "p1 ^c q =o r1 ^c q"
+by (rule cexp_cong1, auto simp add: assms)
+
+lemma cexp_cong2:
+ assumes 2: "p2 =o r2" and q: "Card_order q"
+ and p: "Card_order p2" and r: "Card_order r2"
+ shows "Cnotzero q \<or> (cone \<le>o q ^c p2 \<and> cone \<le>o q ^c r2) \<Longrightarrow>
+ q ^c p2 =o q ^c r2"
+by (rule cexp_cong[OF _ 2]) (auto simp only: ordIso_refl q p r)
+
+lemma cexp_cong2_Cnotzero:
+ assumes 2: "p2 =o r2" and q: "Cnotzero q"
+ and p: "Card_order p2"
+ shows "q ^c p2 =o q ^c r2"
+by (rule cexp_cong[OF _ 2]) (auto simp only: ordIso_refl Card_order_ordIso2[OF p 2] q p)
+
+lemma cexp_czero: "r ^c czero =o cone"
+unfolding cexp_def czero_def Field_card_of Func_empty by (rule single_cone)
+
+lemma cexp_cone:
+ assumes "Card_order r"
+ shows "r ^c cone =o r"
+proof -
+ have "r ^c cone =o |Field r|"
+ unfolding cexp_def cone_def Field_card_of Func_empty
+ card_of_ordIso[symmetric] bij_betw_def Func_def inj_on_def image_def
+ by (rule exI[of _ "\<lambda>f. case f () of Some a \<Rightarrow> a"]) auto
+ also have "|Field r| =o r" by (rule card_of_Field_ordIso[OF assms])
+ finally show ?thesis .
+qed
+
+lemma cexp_cprod:
+ assumes r1: "Cnotzero r1"
+ shows "(r1 ^c r2) ^c r3 =o r1 ^c (r2 *c r3)" (is "?L =o ?R")
+proof -
+ have "?L =o r1 ^c (r3 *c r2)"
+ unfolding cprod_def cexp_def Field_card_of
+ using card_of_Func_Times by(rule ordIso_symmetric)
+ also have "r1 ^c (r3 *c r2) =o ?R"
+ apply(rule cexp_cong2) using cprod_com r1 by (auto simp: Card_order_cprod)
+ finally show ?thesis .
+qed
+
+lemma cexp_cprod_ordLeq:
+ assumes r1: "Cnotzero r1" and r2: "Cinfinite r2"
+ and r3: "Cnotzero r3" "r3 \<le>o r2"
+ shows "(r1 ^c r2) ^c r3 =o r1 ^c r2" (is "?L =o ?R")
+proof-
+ have "?L =o r1 ^c (r2 *c r3)" using cexp_cprod[OF r1] .
+ also have "r1 ^c (r2 *c r3) =o ?R"
+ apply(rule cexp_cong2)
+ apply(rule cprod_infinite1'[OF r2 r3]) using r1 r2 by (fastforce simp: Card_order_cprod)+
+ finally show ?thesis .
+qed
+
+lemma Cnotzero_UNIV: "Cnotzero |UNIV|"
+by (auto simp: card_of_Card_order card_of_ordIso_czero_iff_empty)
+
+lemma Pow_cexp_ctwo:
+ "|Pow A| =o ctwo ^c |A|"
+unfolding ctwo_def cexp_def Field_card_of by (rule card_of_Pow_Func)
+
+lemma ordLess_ctwo_cexp:
+ assumes "Card_order r"
+ shows "r <o ctwo ^c r"
+proof -
+ have "r <o |Pow (Field r)|" using assms by (rule Card_order_Pow)
+ also have "|Pow (Field r)| =o ctwo ^c r"
+ unfolding ctwo_def cexp_def Field_card_of by (rule card_of_Pow_Func)
+ finally show ?thesis .
+qed
+
+lemma ordLeq_cexp1:
+ assumes "Cnotzero r" "Card_order q"
+ shows "q \<le>o q ^c r"
+proof (cases "q =o (czero :: 'a rel)")
+ case True thus ?thesis by (simp only: card_of_empty cexp_def czero_def ordIso_ordLeq_trans)
+next
+ case False
+ thus ?thesis
+ apply -
+ apply (rule ordIso_ordLeq_trans)
+ apply (rule ordIso_symmetric)
+ apply (rule cexp_cone)
+ apply (rule assms(2))
+ apply (rule cexp_mono2)
+ apply (rule cone_ordLeq_Cnotzero)
+ apply (rule assms(1))
+ apply (rule assms(2))
+ apply (rule disjI1)
+ apply (rule conjI)
+ apply (rule notI)
+ apply (erule notE)
+ apply (rule ordIso_transitive)
+ apply assumption
+ apply (rule czero_ordIso)
+ apply (rule assms(2))
+ apply (rule notE)
+ apply (rule cone_not_czero)
+ apply assumption
+ apply (rule Card_order_cone)
+ done
+qed
+
+lemma ordLeq_cexp2:
+ assumes "ctwo \<le>o q" "Card_order r"
+ shows "r \<le>o q ^c r"
+proof (cases "r =o (czero :: 'a rel)")
+ case True thus ?thesis by (simp only: card_of_empty cexp_def czero_def ordIso_ordLeq_trans)
+next
+ case False thus ?thesis
+ apply -
+ apply (rule ordLess_imp_ordLeq)
+ apply (rule ordLess_ordLeq_trans)
+ apply (rule ordLess_ctwo_cexp)
+ apply (rule assms(2))
+ apply (rule cexp_mono1)
+ apply (rule assms(1))
+ apply (rule disjI1)
+ apply (rule ctwo_Cnotzero)
+ apply (rule assms(2))
+ done
+qed
+
+lemma Cnotzero_cexp:
+ assumes "Cnotzero q" "Card_order r"
+ shows "Cnotzero (q ^c r)"
+proof (cases "r =o czero")
+ case False thus ?thesis
+ apply -
+ apply (rule Cnotzero_mono)
+ apply (rule assms(1))
+ apply (rule Card_order_cexp)
+ apply (rule ordLeq_cexp1)
+ apply (rule conjI)
+ apply (rule notI)
+ apply (erule notE)
+ apply (rule ordIso_transitive)
+ apply assumption
+ apply (rule czero_ordIso)
+ apply (rule assms(2))
+ apply (rule conjunct2)
+ apply (rule assms(1))
+ done
+next
+ case True thus ?thesis
+ apply -
+ apply (rule Cnotzero_mono)
+ apply (rule cone_Cnotzero)
+ apply (rule Card_order_cexp)
+ apply (rule ordIso_imp_ordLeq)
+ apply (rule ordIso_symmetric)
+ apply (rule ordIso_transitive)
+ apply (rule cexp_cong2)
+ apply assumption
+ apply (rule conjunct2)
+ apply (rule assms(1))
+ apply (rule assms(2))
+ apply (simp only: card_of_Card_order czero_def)
+ apply (rule disjI1)
+ apply (rule assms(1))
+ apply (rule cexp_czero)
+ done
+qed
+
+lemma Cinfinite_ctwo_cexp:
+ "Cinfinite r \<Longrightarrow> Cinfinite (ctwo ^c r)"
+unfolding ctwo_def cexp_def cinfinite_def Field_card_of
+by (rule conjI, rule infinite_Func, auto, rule card_of_card_order_on)
+
+lemma cinfinite_cexp: "\<lbrakk>ctwo \<le>o q; Cinfinite r\<rbrakk> \<Longrightarrow> cinfinite (q ^c r)"
+by (metis assms cinfinite_mono ordLeq_cexp2)
+
+lemma cinfinite_ccexp: "\<lbrakk>ctwo \<le>o q; Cinfinite r\<rbrakk> \<Longrightarrow> cinfinite (q ^^c r)"
+by (rule cinfinite_mono[OF cexp_ordLeq_ccexp cinfinite_cexp])
+
+lemma Cinfinite_cexp:
+ "\<lbrakk>ctwo \<le>o q; Cinfinite r\<rbrakk> \<Longrightarrow> Cinfinite (q ^c r)"
+by (simp add: cinfinite_cexp Card_order_cexp)
+
+lemma ctwo_ordLess_natLeq:
+ "ctwo <o natLeq"
+unfolding ctwo_def using finite_iff_ordLess_natLeq finite_UNIV by fast
+
+lemma ctwo_ordLess_Cinfinite: "Cinfinite r \<Longrightarrow> ctwo <o r"
+by (metis ctwo_ordLess_natLeq natLeq_ordLeq_cinfinite ordLess_ordLeq_trans)
+
+lemma ctwo_ordLeq_Cinfinite:
+ assumes "Cinfinite r"
+ shows "ctwo \<le>o r"
+by (rule ordLess_imp_ordLeq[OF ctwo_ordLess_Cinfinite[OF assms]])
+
+lemma Cinfinite_ordLess_cexp:
+ assumes r: "Cinfinite r"
+ shows "r <o r ^c r"
+proof -
+ have "r <o ctwo ^c r" using r by (simp only: ordLess_ctwo_cexp)
+ also have "ctwo ^c r \<le>o r ^c r"
+ by (rule cexp_mono1[OF ctwo_ordLeq_Cinfinite]) (auto simp: r ctwo_not_czero Card_order_ctwo)
+ finally show ?thesis .
+qed
+
+lemma infinite_ordLeq_cexp:
+ assumes "Cinfinite r"
+ shows "r \<le>o r ^c r"
+by (rule ordLess_imp_ordLeq[OF Cinfinite_ordLess_cexp[OF assms]])
+
+lemma cone_ordLeq_iff_Field:
+ assumes "cone \<le>o r"
+ shows "Field r \<noteq> {}"
+proof (rule ccontr)
+ assume "\<not> Field r \<noteq> {}"
+ hence "Field r = {}" by simp
+ thus False using card_of_empty3
+ card_of_mono2[OF assms] Cnotzero_imp_not_empty[OF cone_Cnotzero] by auto
+qed
+
+lemma cone_ordLeq_cexp: "cone \<le>o r1 \<Longrightarrow> cone \<le>o r1 ^c r2"
+by (simp add: cexp_def cone_def Func_non_emp card_of_singl_ordLeq cone_ordLeq_iff_Field)
+
+lemma Card_order_czero: "Card_order czero"
+by (simp only: card_of_Card_order czero_def)
+
+lemma cexp_mono2'':
+ assumes 2: "p2 \<le>o r2"
+ and n1: "Cnotzero q"
+ and n2: "Card_order p2"
+ shows "q ^c p2 \<le>o q ^c r2"
+proof (cases "p2 =o (czero :: 'a rel)")
+ case True
+ hence "q ^c p2 =o q ^c (czero :: 'a rel)" using n1 n2 cexp_cong2 Card_order_czero by blast
+ also have "q ^c (czero :: 'a rel) =o cone" using cexp_czero by blast
+ also have "cone \<le>o q ^c r2" using cone_ordLeq_cexp cone_ordLeq_Cnotzero n1 by blast
+ finally show ?thesis .
+next
+ case False thus ?thesis using assms cexp_mono2' czeroI by metis
+qed
+
+lemma Un_Cinfinite_bound: "\<lbrakk>|A| \<le>o r; |B| \<le>o r; Cinfinite r\<rbrakk> \<Longrightarrow> |A \<union> B| \<le>o r"
+by (auto simp add: cinfinite_def card_of_Un_ordLeq_infinite_Field)
+
+lemma UNION_Cinfinite_bound: "\<lbrakk>|I| \<le>o r; \<forall>i \<in> I. |A i| \<le>o r; Cinfinite r\<rbrakk> \<Longrightarrow> |\<Union>i \<in> I. A i| \<le>o r"
+by (auto simp add: card_of_UNION_ordLeq_infinite_Field cinfinite_def)
+
+lemma csum_cinfinite_bound:
+ assumes "p \<le>o r" "q \<le>o r" "Card_order p" "Card_order q" "Cinfinite r"
+ shows "p +c q \<le>o r"
+proof -
+ from assms(1-4) have "|Field p| \<le>o r" "|Field q| \<le>o r"
+ unfolding card_order_on_def using card_of_least ordLeq_transitive by blast+
+ with assms show ?thesis unfolding cinfinite_def csum_def
+ by (blast intro: card_of_Plus_ordLeq_infinite_Field)
+qed
+
+lemma csum_cexp: "\<lbrakk>Cinfinite r1; Cinfinite r2; Card_order q; ctwo \<le>o q\<rbrakk> \<Longrightarrow>
+ q ^c r1 +c q ^c r2 \<le>o q ^c (r1 +c r2)"
+apply (rule csum_cinfinite_bound)
+apply (rule cexp_mono2)
+apply (rule ordLeq_csum1)
+apply (erule conjunct2)
+apply assumption
+apply (rule disjI2)
+apply (rule ordLeq_transitive)
+apply (rule cone_ordLeq_ctwo)
+apply (rule ordLeq_transitive)
+apply assumption
+apply (rule ordLeq_cexp1)
+apply (rule Cinfinite_Cnotzero)
+apply (rule Cinfinite_csum)
+apply (rule disjI1)
+apply assumption
+apply assumption
+apply (rule notE)
+apply (rule cinfinite_not_czero[of r1])
+apply (erule conjunct1)
+apply assumption
+apply (erule conjunct2)
+apply (rule cexp_mono2)
+apply (rule ordLeq_csum2)
+apply (erule conjunct2)
+apply assumption
+apply (rule disjI2)
+apply (rule ordLeq_transitive)
+apply (rule cone_ordLeq_ctwo)
+apply (rule ordLeq_transitive)
+apply assumption
+apply (rule ordLeq_cexp1)
+apply (rule Cinfinite_Cnotzero)
+apply (rule Cinfinite_csum)
+apply (rule disjI1)
+apply assumption
+apply assumption
+apply (rule notE)
+apply (rule cinfinite_not_czero[of r2])
+apply (erule conjunct1)
+apply assumption
+apply (erule conjunct2)
+apply (rule Card_order_cexp)
+apply (rule Card_order_cexp)
+apply (rule Cinfinite_cexp)
+apply assumption
+apply (rule Cinfinite_csum)
+by (rule disjI1)
+
+lemma csum_cexp': "\<lbrakk>Cinfinite r; Card_order q; ctwo \<le>o q\<rbrakk> \<Longrightarrow> q +c r \<le>o q ^c r"
+apply (rule csum_cinfinite_bound)
+ apply (metis Cinfinite_Cnotzero ordLeq_cexp1)
+ apply (metis ordLeq_cexp2)
+ apply blast+
+by (metis Cinfinite_cexp)
+
+lemma cprod_cinfinite_bound:
+ assumes "p \<le>o r" "q \<le>o r" "Card_order p" "Card_order q" "Cinfinite r"
+ shows "p *c q \<le>o r"
+proof -
+ from assms(1-4) have "|Field p| \<le>o r" "|Field q| \<le>o r"
+ unfolding card_order_on_def using card_of_least ordLeq_transitive by blast+
+ with assms show ?thesis unfolding cinfinite_def cprod_def
+ by (blast intro: card_of_Times_ordLeq_infinite_Field)
+qed
+
+lemma cprod_csum_cexp:
+ "r1 *c r2 \<le>o (r1 +c r2) ^c ctwo"
+unfolding cprod_def csum_def cexp_def ctwo_def Field_card_of
+proof -
+ let ?f = "\<lambda>(a, b). %x. if x then Some (Inl a) else Some (Inr b)"
+ have "inj_on ?f (Field r1 \<times> Field r2)" (is "inj_on _ ?LHS")
+ by (auto simp: inj_on_def fun_eq_iff split: bool.split)
+ moreover
+ have "?f ` ?LHS \<subseteq> Func (UNIV :: bool set) (Field r1 <+> Field r2)" (is "_ \<subseteq> ?RHS")
+ by (auto simp: Func_def)
+ ultimately show "|?LHS| \<le>o |?RHS|" using card_of_ordLeq by blast
+qed
+
+lemma card_of_Sigma_ordLeq_Cinfinite:
+ "\<lbrakk>Cinfinite r; |I| \<le>o r; \<forall>i \<in> I. |A i| \<le>o r\<rbrakk> \<Longrightarrow> |SIGMA i : I. A i| \<le>o r"
+unfolding cinfinite_def by (blast intro: card_of_Sigma_ordLeq_infinite_Field)
+
+
+(* cardSuc *)
+
+lemma Cinfinite_cardSuc: "Cinfinite r \<Longrightarrow> Cinfinite (cardSuc r)"
+by (simp add: cinfinite_def cardSuc_Card_order cardSuc_finite)
+
+lemma cardSuc_UNION_Cinfinite:
+ assumes "Cinfinite r" "relChain (cardSuc r) As" "B \<le> (UN i : Field (cardSuc r). As i)" "|B| <=o r"
+ shows "EX i : Field (cardSuc r). B \<le> As i"
+using cardSuc_UNION assms unfolding cinfinite_def by blast
+
+subsection {* Powerset *}
+
+definition cpow where "cpow r = |Pow (Field r)|"
+
+lemma card_order_cpow: "card_order r \<Longrightarrow> card_order (cpow r)"
+by (simp only: cpow_def Field_card_order Pow_UNIV card_of_card_order_on)
+
+lemma cpow_greater_eq: "Card_order r \<Longrightarrow> r \<le>o cpow r"
+by (rule ordLess_imp_ordLeq) (simp only: cpow_def Card_order_Pow)
+
+lemma Card_order_cpow: "Card_order (cpow r)"
+unfolding cpow_def by (rule card_of_Card_order)
+
+lemma Cinfinite_cpow: "Cinfinite r \<Longrightarrow> Cinfinite (cpow r)"
+unfolding cpow_def cinfinite_def by (metis Field_card_of card_of_Card_order infinite_Pow)
+
+lemma cardSuc_ordLeq_cpow: "Card_order r \<Longrightarrow> cardSuc r \<le>o cpow r"
+unfolding cpow_def by (metis Card_order_Pow cardSuc_ordLess_ordLeq card_of_Card_order)
+
+lemma cpow_cexp_ctwo: "cpow r =o ctwo ^c r"
+unfolding cpow_def ctwo_def cexp_def Field_card_of by (rule card_of_Pow_Func)
+
+subsection {* Lists *}
+
+definition clists where "clists r = |lists (Field r)|"
+
+lemma clists_Cinfinite: "Cinfinite r \<Longrightarrow> clists r =o r"
+unfolding cinfinite_def clists_def by (blast intro: Card_order_lists_infinite)
+
+lemma Card_order_clists: "Card_order (clists r)"
+unfolding clists_def by (rule card_of_Card_order)
+
+lemma Cnotzero_clists: "Cnotzero (clists r)"
+by (simp add: clists_def card_of_ordIso_czero_iff_empty lists_not_empty) (rule card_of_Card_order)
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Ordinals_and_Cardinals/Cardinal_Order_Relation.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,1097 @@
+(* Title: HOL/Ordinals_and_Cardinals/Cardinal_Order_Relation.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Cardinal-order relations.
+*)
+
+header {* Cardinal-Order Relations *}
+
+theory Cardinal_Order_Relation
+imports Cardinal_Order_Relation_Base Constructions_on_Wellorders
+begin
+
+declare
+ card_order_on_well_order_on[simp]
+ card_of_card_order_on[simp]
+ card_of_well_order_on[simp]
+ Field_card_of[simp]
+ card_of_Card_order[simp]
+ card_of_Well_order[simp]
+ card_of_least[simp]
+ card_of_unique[simp]
+ card_of_mono1[simp]
+ card_of_mono2[simp]
+ card_of_cong[simp]
+ card_of_Field_ordLess[simp]
+ card_of_Field_ordIso[simp]
+ card_of_underS[simp]
+ ordLess_Field[simp]
+ card_of_empty[simp]
+ card_of_empty1[simp]
+ card_of_image[simp]
+ card_of_singl_ordLeq[simp]
+ Card_order_singl_ordLeq[simp]
+ card_of_Pow[simp]
+ Card_order_Pow[simp]
+ card_of_set_type[simp]
+ card_of_Plus1[simp]
+ Card_order_Plus1[simp]
+ card_of_Plus2[simp]
+ Card_order_Plus2[simp]
+ card_of_Plus_mono1[simp]
+ card_of_Plus_mono2[simp]
+ card_of_Plus_mono[simp]
+ card_of_Plus_cong2[simp]
+ card_of_Plus_cong[simp]
+ card_of_Un1[simp]
+ card_of_diff[simp]
+ card_of_Un_Plus_ordLeq[simp]
+ card_of_Times1[simp]
+ card_of_Times2[simp]
+ card_of_Times3[simp]
+ card_of_Times_mono1[simp]
+ card_of_Times_mono2[simp]
+ card_of_Times_cong1[simp]
+ card_of_Times_cong2[simp]
+ card_of_ordIso_finite[simp]
+ finite_ordLess_infinite2[simp]
+ card_of_Times_same_infinite[simp]
+ card_of_Times_infinite_simps[simp]
+ card_of_Plus_infinite1[simp]
+ card_of_Plus_infinite2[simp]
+ card_of_Plus_ordLess_infinite[simp]
+ card_of_Plus_ordLess_infinite_Field[simp]
+ card_of_lists_infinite[simp]
+ infinite_cartesian_product[simp]
+ cardSuc_Card_order[simp]
+ cardSuc_greater[simp]
+ cardSuc_ordLeq[simp]
+ cardSuc_ordLeq_ordLess[simp]
+ cardSuc_mono_ordLeq[simp]
+ cardSuc_invar_ordIso[simp]
+ card_of_cardSuc_finite[simp]
+ cardSuc_finite[simp]
+ card_of_Plus_ordLeq_infinite_Field[simp]
+ curr_in[intro, simp]
+ Func_empty[simp]
+ Func_map_empty[simp]
+ Func_is_emp[simp]
+
+
+subsection {* Cardinal of a set *}
+
+lemma card_of_inj_rel: assumes INJ: "!! x y y'. \<lbrakk>(x,y) : R; (x,y') : R\<rbrakk> \<Longrightarrow> y = y'"
+shows "|{y. EX x. (x,y) : R}| <=o |{x. EX y. (x,y) : R}|"
+proof-
+ let ?Y = "{y. EX x. (x,y) : R}" let ?X = "{x. EX y. (x,y) : R}"
+ let ?f = "% y. SOME x. (x,y) : R"
+ have "?f ` ?Y <= ?X" using someI by force (* FIXME: takes a bit long *)
+ moreover have "inj_on ?f ?Y"
+ unfolding inj_on_def proof(auto)
+ fix y1 x1 y2 x2
+ assume *: "(x1, y1) \<in> R" "(x2, y2) \<in> R" and **: "?f y1 = ?f y2"
+ hence "(?f y1,y1) : R" using someI[of "% x. (x,y1) : R"] by auto
+ moreover have "(?f y2,y2) : R" using * someI[of "% x. (x,y2) : R"] by auto
+ ultimately show "y1 = y2" using ** INJ by auto
+ qed
+ ultimately show "|?Y| <=o |?X|" using card_of_ordLeq by blast
+qed
+
+lemma card_of_unique2: "\<lbrakk>card_order_on B r; bij_betw f A B\<rbrakk> \<Longrightarrow> r =o |A|"
+using card_of_ordIso card_of_unique ordIso_equivalence by blast
+
+lemma internalize_card_of_ordLess:
+"( |A| <o r) = (\<exists>B < Field r. |A| =o |B| \<and> |B| <o r)"
+proof
+ assume "|A| <o r"
+ then obtain p where 1: "Field p < Field r \<and> |A| =o p \<and> p <o r"
+ using internalize_ordLess[of "|A|" r] by blast
+ hence "Card_order p" using card_of_Card_order Card_order_ordIso2 by blast
+ hence "|Field p| =o p" using card_of_Field_ordIso by blast
+ hence "|A| =o |Field p| \<and> |Field p| <o r"
+ using 1 ordIso_equivalence ordIso_ordLess_trans by blast
+ thus "\<exists>B < Field r. |A| =o |B| \<and> |B| <o r" using 1 by blast
+next
+ assume "\<exists>B < Field r. |A| =o |B| \<and> |B| <o r"
+ thus "|A| <o r" using ordIso_ordLess_trans by blast
+qed
+
+lemma internalize_card_of_ordLess2:
+"( |A| <o |C| ) = (\<exists>B < C. |A| =o |B| \<and> |B| <o |C| )"
+using internalize_card_of_ordLess[of "A" "|C|"] Field_card_of[of C] by auto
+
+lemma Card_order_omax:
+assumes "finite R" and "R \<noteq> {}" and "\<forall>r\<in>R. Card_order r"
+shows "Card_order (omax R)"
+proof-
+ have "\<forall>r\<in>R. Well_order r"
+ using assms unfolding card_order_on_def by simp
+ thus ?thesis using assms apply - apply(drule omax_in) by auto
+qed
+
+lemma Card_order_omax2:
+assumes "finite I" and "I \<noteq> {}"
+shows "Card_order (omax {|A i| | i. i \<in> I})"
+proof-
+ let ?R = "{|A i| | i. i \<in> I}"
+ have "finite ?R" and "?R \<noteq> {}" using assms by auto
+ moreover have "\<forall>r\<in>?R. Card_order r"
+ using card_of_Card_order by auto
+ ultimately show ?thesis by(rule Card_order_omax)
+qed
+
+
+subsection {* Cardinals versus set operations on arbitrary sets *}
+
+lemma subset_ordLeq_strict:
+assumes "A \<le> B" and "|A| <o |B|"
+shows "A < B"
+proof-
+ {assume "\<not>(A < B)"
+ hence "A = B" using assms(1) by blast
+ hence False using assms(2) not_ordLess_ordIso card_of_refl by blast
+ }
+ thus ?thesis by blast
+qed
+
+corollary subset_ordLeq_diff:
+assumes "A \<le> B" and "|A| <o |B|"
+shows "B - A \<noteq> {}"
+using assms subset_ordLeq_strict by blast
+
+lemma card_of_empty4:
+"|{}::'b set| <o |A::'a set| = (A \<noteq> {})"
+proof(intro iffI notI)
+ assume *: "|{}::'b set| <o |A|" and "A = {}"
+ hence "|A| =o |{}::'b set|"
+ using card_of_ordIso unfolding bij_betw_def inj_on_def by blast
+ hence "|{}::'b set| =o |A|" using ordIso_symmetric by blast
+ with * show False using not_ordLess_ordIso[of "|{}::'b set|" "|A|"] by blast
+next
+ assume "A \<noteq> {}"
+ hence "(\<not> (\<exists>f. inj_on f A \<and> f ` A \<subseteq> {}))"
+ unfolding inj_on_def by blast
+ thus "| {} | <o | A |"
+ using card_of_ordLess by blast
+qed
+
+lemma card_of_empty5:
+"|A| <o |B| \<Longrightarrow> B \<noteq> {}"
+using card_of_empty not_ordLess_ordLeq by blast
+
+lemma Well_order_card_of_empty:
+"Well_order r \<Longrightarrow> |{}| \<le>o r" by simp
+
+lemma card_of_UNIV[simp]:
+"|A :: 'a set| \<le>o |UNIV :: 'a set|"
+using card_of_mono1[of A] by simp
+
+lemma card_of_UNIV2[simp]:
+"Card_order r \<Longrightarrow> (r :: 'a rel) \<le>o |UNIV :: 'a set|"
+using card_of_UNIV[of "Field r"] card_of_Field_ordIso
+ ordIso_symmetric ordIso_ordLeq_trans by blast
+
+lemma card_of_Pow_mono[simp]:
+assumes "|A| \<le>o |B|"
+shows "|Pow A| \<le>o |Pow B|"
+proof-
+ obtain f where "inj_on f A \<and> f ` A \<le> B"
+ using assms card_of_ordLeq[of A B] by auto
+ hence "inj_on (image f) (Pow A) \<and> (image f) ` (Pow A) \<le> (Pow B)"
+ by (auto simp add: inj_on_image_Pow image_Pow_mono)
+ thus ?thesis using card_of_ordLeq[of "Pow A"] by metis
+qed
+
+lemma ordIso_Pow_mono[simp]:
+assumes "r \<le>o r'"
+shows "|Pow(Field r)| \<le>o |Pow(Field r')|"
+using assms card_of_mono2 card_of_Pow_mono by blast
+
+lemma card_of_Pow_cong[simp]:
+assumes "|A| =o |B|"
+shows "|Pow A| =o |Pow B|"
+proof-
+ obtain f where "bij_betw f A B"
+ using assms card_of_ordIso[of A B] by auto
+ hence "bij_betw (image f) (Pow A) (Pow B)"
+ by (auto simp add: bij_betw_image_Pow)
+ thus ?thesis using card_of_ordIso[of "Pow A"] by auto
+qed
+
+lemma ordIso_Pow_cong[simp]:
+assumes "r =o r'"
+shows "|Pow(Field r)| =o |Pow(Field r')|"
+using assms card_of_cong card_of_Pow_cong by blast
+
+corollary Card_order_Plus_empty1:
+"Card_order r \<Longrightarrow> r =o |(Field r) <+> {}|"
+using card_of_Plus_empty1 card_of_Field_ordIso ordIso_equivalence by blast
+
+corollary Card_order_Plus_empty2:
+"Card_order r \<Longrightarrow> r =o |{} <+> (Field r)|"
+using card_of_Plus_empty2 card_of_Field_ordIso ordIso_equivalence by blast
+
+lemma Card_order_Un1:
+shows "Card_order r \<Longrightarrow> |Field r| \<le>o |(Field r) \<union> B| "
+using card_of_Un1 card_of_Field_ordIso ordIso_symmetric ordIso_ordLeq_trans by auto
+
+lemma card_of_Un2[simp]:
+shows "|A| \<le>o |B \<union> A|"
+using inj_on_id[of A] card_of_ordLeq[of A _] by fastforce
+
+lemma Card_order_Un2:
+shows "Card_order r \<Longrightarrow> |Field r| \<le>o |A \<union> (Field r)| "
+using card_of_Un2 card_of_Field_ordIso ordIso_symmetric ordIso_ordLeq_trans by auto
+
+lemma Un_Plus_bij_betw:
+assumes "A Int B = {}"
+shows "\<exists>f. bij_betw f (A \<union> B) (A <+> B)"
+proof-
+ let ?f = "\<lambda> c. if c \<in> A then Inl c else Inr c"
+ have "bij_betw ?f (A \<union> B) (A <+> B)"
+ using assms by(unfold bij_betw_def inj_on_def, auto)
+ thus ?thesis by blast
+qed
+
+lemma card_of_Un_Plus_ordIso:
+assumes "A Int B = {}"
+shows "|A \<union> B| =o |A <+> B|"
+using assms card_of_ordIso[of "A \<union> B"] Un_Plus_bij_betw[of A B] by auto
+
+lemma card_of_Un_Plus_ordIso1:
+"|A \<union> B| =o |A <+> (B - A)|"
+using card_of_Un_Plus_ordIso[of A "B - A"] by auto
+
+lemma card_of_Un_Plus_ordIso2:
+"|A \<union> B| =o |(A - B) <+> B|"
+using card_of_Un_Plus_ordIso[of "A - B" B] by auto
+
+lemma card_of_Times_singl1: "|A| =o |A \<times> {b}|"
+proof-
+ have "bij_betw fst (A \<times> {b}) A" unfolding bij_betw_def inj_on_def by force
+ thus ?thesis using card_of_ordIso ordIso_symmetric by blast
+qed
+
+corollary Card_order_Times_singl1:
+"Card_order r \<Longrightarrow> r =o |(Field r) \<times> {b}|"
+using card_of_Times_singl1[of _ b] card_of_Field_ordIso ordIso_equivalence by blast
+
+lemma card_of_Times_singl2: "|A| =o |{b} \<times> A|"
+proof-
+ have "bij_betw snd ({b} \<times> A) A" unfolding bij_betw_def inj_on_def by force
+ thus ?thesis using card_of_ordIso ordIso_symmetric by blast
+qed
+
+corollary Card_order_Times_singl2:
+"Card_order r \<Longrightarrow> r =o |{a} \<times> (Field r)|"
+using card_of_Times_singl2[of _ a] card_of_Field_ordIso ordIso_equivalence by blast
+
+lemma card_of_Times_assoc: "|(A \<times> B) \<times> C| =o |A \<times> B \<times> C|"
+proof -
+ let ?f = "\<lambda>((a,b),c). (a,(b,c))"
+ have "A \<times> B \<times> C \<subseteq> ?f ` ((A \<times> B) \<times> C)"
+ proof
+ fix x assume "x \<in> A \<times> B \<times> C"
+ then obtain a b c where *: "a \<in> A" "b \<in> B" "c \<in> C" "x = (a, b, c)" by blast
+ let ?x = "((a, b), c)"
+ from * have "?x \<in> (A \<times> B) \<times> C" "x = ?f ?x" by auto
+ thus "x \<in> ?f ` ((A \<times> B) \<times> C)" by blast
+ qed
+ hence "bij_betw ?f ((A \<times> B) \<times> C) (A \<times> B \<times> C)"
+ unfolding bij_betw_def inj_on_def by auto
+ thus ?thesis using card_of_ordIso by blast
+qed
+
+corollary Card_order_Times3:
+"Card_order r \<Longrightarrow> |Field r| \<le>o |(Field r) \<times> (Field r)|"
+using card_of_Times3 card_of_Field_ordIso
+ ordIso_ordLeq_trans ordIso_symmetric by blast
+
+lemma card_of_Times_mono[simp]:
+assumes "|A| \<le>o |B|" and "|C| \<le>o |D|"
+shows "|A \<times> C| \<le>o |B \<times> D|"
+using assms card_of_Times_mono1[of A B C] card_of_Times_mono2[of C D B]
+ ordLeq_transitive[of "|A \<times> C|"] by blast
+
+corollary ordLeq_Times_mono:
+assumes "r \<le>o r'" and "p \<le>o p'"
+shows "|(Field r) \<times> (Field p)| \<le>o |(Field r') \<times> (Field p')|"
+using assms card_of_mono2[of r r'] card_of_mono2[of p p'] card_of_Times_mono by blast
+
+corollary ordIso_Times_cong1[simp]:
+assumes "r =o r'"
+shows "|(Field r) \<times> C| =o |(Field r') \<times> C|"
+using assms card_of_cong card_of_Times_cong1 by blast
+
+lemma card_of_Times_cong[simp]:
+assumes "|A| =o |B|" and "|C| =o |D|"
+shows "|A \<times> C| =o |B \<times> D|"
+using assms
+by (auto simp add: ordIso_iff_ordLeq)
+
+corollary ordIso_Times_cong:
+assumes "r =o r'" and "p =o p'"
+shows "|(Field r) \<times> (Field p)| =o |(Field r') \<times> (Field p')|"
+using assms card_of_cong[of r r'] card_of_cong[of p p'] card_of_Times_cong by blast
+
+lemma card_of_Sigma_mono2:
+assumes "inj_on f (I::'i set)" and "f ` I \<le> (J::'j set)"
+shows "|SIGMA i : I. (A::'j \<Rightarrow> 'a set) (f i)| \<le>o |SIGMA j : J. A j|"
+proof-
+ let ?LEFT = "SIGMA i : I. A (f i)"
+ let ?RIGHT = "SIGMA j : J. A j"
+ obtain u where u_def: "u = (\<lambda>(i::'i,a::'a). (f i,a))" by blast
+ have "inj_on u ?LEFT \<and> u `?LEFT \<le> ?RIGHT"
+ using assms unfolding u_def inj_on_def by auto
+ thus ?thesis using card_of_ordLeq by blast
+qed
+
+lemma card_of_Sigma_mono:
+assumes INJ: "inj_on f I" and IM: "f ` I \<le> J" and
+ LEQ: "\<forall>j \<in> J. |A j| \<le>o |B j|"
+shows "|SIGMA i : I. A (f i)| \<le>o |SIGMA j : J. B j|"
+proof-
+ have "\<forall>i \<in> I. |A(f i)| \<le>o |B(f i)|"
+ using IM LEQ by blast
+ hence "|SIGMA i : I. A (f i)| \<le>o |SIGMA i : I. B (f i)|"
+ using card_of_Sigma_mono1[of I] by metis
+ moreover have "|SIGMA i : I. B (f i)| \<le>o |SIGMA j : J. B j|"
+ using INJ IM card_of_Sigma_mono2 by blast
+ ultimately show ?thesis using ordLeq_transitive by blast
+qed
+
+
+lemma ordLeq_Sigma_mono1:
+assumes "\<forall>i \<in> I. p i \<le>o r i"
+shows "|SIGMA i : I. Field(p i)| \<le>o |SIGMA i : I. Field(r i)|"
+using assms by (auto simp add: card_of_Sigma_mono1)
+
+
+lemma ordLeq_Sigma_mono:
+assumes "inj_on f I" and "f ` I \<le> J" and
+ "\<forall>j \<in> J. p j \<le>o r j"
+shows "|SIGMA i : I. Field(p(f i))| \<le>o |SIGMA j : J. Field(r j)|"
+using assms card_of_mono2 card_of_Sigma_mono
+ [of f I J "\<lambda> i. Field(p i)" "\<lambda> j. Field(r j)"] by metis
+
+
+lemma card_of_Sigma_cong1:
+assumes "\<forall>i \<in> I. |A i| =o |B i|"
+shows "|SIGMA i : I. A i| =o |SIGMA i : I. B i|"
+using assms by (auto simp add: card_of_Sigma_mono1 ordIso_iff_ordLeq)
+
+
+lemma card_of_Sigma_cong2:
+assumes "bij_betw f (I::'i set) (J::'j set)"
+shows "|SIGMA i : I. (A::'j \<Rightarrow> 'a set) (f i)| =o |SIGMA j : J. A j|"
+proof-
+ let ?LEFT = "SIGMA i : I. A (f i)"
+ let ?RIGHT = "SIGMA j : J. A j"
+ obtain u where u_def: "u = (\<lambda>(i::'i,a::'a). (f i,a))" by blast
+ have "bij_betw u ?LEFT ?RIGHT"
+ using assms unfolding u_def bij_betw_def inj_on_def by auto
+ thus ?thesis using card_of_ordIso by blast
+qed
+
+lemma card_of_Sigma_cong:
+assumes BIJ: "bij_betw f I J" and
+ ISO: "\<forall>j \<in> J. |A j| =o |B j|"
+shows "|SIGMA i : I. A (f i)| =o |SIGMA j : J. B j|"
+proof-
+ have "\<forall>i \<in> I. |A(f i)| =o |B(f i)|"
+ using ISO BIJ unfolding bij_betw_def by blast
+ hence "|SIGMA i : I. A (f i)| =o |SIGMA i : I. B (f i)|"
+ using card_of_Sigma_cong1 by metis
+ moreover have "|SIGMA i : I. B (f i)| =o |SIGMA j : J. B j|"
+ using BIJ card_of_Sigma_cong2 by blast
+ ultimately show ?thesis using ordIso_transitive by blast
+qed
+
+lemma ordIso_Sigma_cong1:
+assumes "\<forall>i \<in> I. p i =o r i"
+shows "|SIGMA i : I. Field(p i)| =o |SIGMA i : I. Field(r i)|"
+using assms by (auto simp add: card_of_Sigma_cong1)
+
+lemma ordLeq_Sigma_cong:
+assumes "bij_betw f I J" and
+ "\<forall>j \<in> J. p j =o r j"
+shows "|SIGMA i : I. Field(p(f i))| =o |SIGMA j : J. Field(r j)|"
+using assms card_of_cong card_of_Sigma_cong
+ [of f I J "\<lambda> j. Field(p j)" "\<lambda> j. Field(r j)"] by blast
+
+corollary ordLeq_Sigma_Times:
+"\<forall>i \<in> I. p i \<le>o r \<Longrightarrow> |SIGMA i : I. Field (p i)| \<le>o |I \<times> (Field r)|"
+by (auto simp add: card_of_Sigma_Times)
+
+lemma card_of_UNION_Sigma2:
+assumes
+"!! i j. \<lbrakk>{i,j} <= I; i ~= j\<rbrakk> \<Longrightarrow> A i Int A j = {}"
+shows
+"|\<Union>i\<in>I. A i| =o |Sigma I A|"
+proof-
+ let ?L = "\<Union>i\<in>I. A i" let ?R = "Sigma I A"
+ have "|?L| <=o |?R|" using card_of_UNION_Sigma .
+ moreover have "|?R| <=o |?L|"
+ proof-
+ have "inj_on snd ?R"
+ unfolding inj_on_def using assms by auto
+ moreover have "snd ` ?R <= ?L" by auto
+ ultimately show ?thesis using card_of_ordLeq by blast
+ qed
+ ultimately show ?thesis by(simp add: ordIso_iff_ordLeq)
+qed
+
+corollary Plus_into_Times:
+assumes A2: "a1 \<noteq> a2 \<and> {a1,a2} \<le> A" and
+ B2: "b1 \<noteq> b2 \<and> {b1,b2} \<le> B"
+shows "\<exists>f. inj_on f (A <+> B) \<and> f ` (A <+> B) \<le> A \<times> B"
+using assms by (auto simp add: card_of_Plus_Times card_of_ordLeq)
+
+corollary Plus_into_Times_types:
+assumes A2: "(a1::'a) \<noteq> a2" and B2: "(b1::'b) \<noteq> b2"
+shows "\<exists>(f::'a + 'b \<Rightarrow> 'a * 'b). inj f"
+using assms Plus_into_Times[of a1 a2 UNIV b1 b2 UNIV]
+by auto
+
+corollary Times_same_infinite_bij_betw:
+assumes "infinite A"
+shows "\<exists>f. bij_betw f (A \<times> A) A"
+using assms by (auto simp add: card_of_ordIso)
+
+corollary Times_same_infinite_bij_betw_types:
+assumes INF: "infinite(UNIV::'a set)"
+shows "\<exists>(f::('a * 'a) => 'a). bij f"
+using assms Times_same_infinite_bij_betw[of "UNIV::'a set"]
+by auto
+
+corollary Times_infinite_bij_betw:
+assumes INF: "infinite A" and NE: "B \<noteq> {}" and INJ: "inj_on g B \<and> g ` B \<le> A"
+shows "(\<exists>f. bij_betw f (A \<times> B) A) \<and> (\<exists>h. bij_betw h (B \<times> A) A)"
+proof-
+ have "|B| \<le>o |A|" using INJ card_of_ordLeq by blast
+ thus ?thesis using INF NE
+ by (auto simp add: card_of_ordIso card_of_Times_infinite)
+qed
+
+corollary Times_infinite_bij_betw_types:
+assumes INF: "infinite(UNIV::'a set)" and
+ BIJ: "inj(g::'b \<Rightarrow> 'a)"
+shows "(\<exists>(f::('b * 'a) => 'a). bij f) \<and> (\<exists>(h::('a * 'b) => 'a). bij h)"
+using assms Times_infinite_bij_betw[of "UNIV::'a set" "UNIV::'b set" g]
+by auto
+
+lemma card_of_Times_ordLeq_infinite:
+"\<lbrakk>infinite C; |A| \<le>o |C|; |B| \<le>o |C|\<rbrakk>
+ \<Longrightarrow> |A <*> B| \<le>o |C|"
+by(simp add: card_of_Sigma_ordLeq_infinite)
+
+corollary Plus_infinite_bij_betw:
+assumes INF: "infinite A" and INJ: "inj_on g B \<and> g ` B \<le> A"
+shows "(\<exists>f. bij_betw f (A <+> B) A) \<and> (\<exists>h. bij_betw h (B <+> A) A)"
+proof-
+ have "|B| \<le>o |A|" using INJ card_of_ordLeq by blast
+ thus ?thesis using INF
+ by (auto simp add: card_of_ordIso)
+qed
+
+corollary Plus_infinite_bij_betw_types:
+assumes INF: "infinite(UNIV::'a set)" and
+ BIJ: "inj(g::'b \<Rightarrow> 'a)"
+shows "(\<exists>(f::('b + 'a) => 'a). bij f) \<and> (\<exists>(h::('a + 'b) => 'a). bij h)"
+using assms Plus_infinite_bij_betw[of "UNIV::'a set" g "UNIV::'b set"]
+by auto
+
+lemma card_of_Un_infinite_simps[simp]:
+"\<lbrakk>infinite A; |B| \<le>o |A| \<rbrakk> \<Longrightarrow> |A \<union> B| =o |A|"
+"\<lbrakk>infinite A; |B| \<le>o |A| \<rbrakk> \<Longrightarrow> |B \<union> A| =o |A|"
+using card_of_Un_infinite by auto
+
+corollary Card_order_Un_infinite:
+assumes INF: "infinite(Field r)" and CARD: "Card_order r" and
+ LEQ: "p \<le>o r"
+shows "| (Field r) \<union> (Field p) | =o r \<and> | (Field p) \<union> (Field r) | =o r"
+proof-
+ have "| Field r \<union> Field p | =o | Field r | \<and>
+ | Field p \<union> Field r | =o | Field r |"
+ using assms by (auto simp add: card_of_Un_infinite)
+ thus ?thesis
+ using assms card_of_Field_ordIso[of r]
+ ordIso_transitive[of "|Field r \<union> Field p|"]
+ ordIso_transitive[of _ "|Field r|"] by blast
+qed
+
+corollary subset_ordLeq_diff_infinite:
+assumes INF: "infinite B" and SUB: "A \<le> B" and LESS: "|A| <o |B|"
+shows "infinite (B - A)"
+using assms card_of_Un_diff_infinite card_of_ordIso_finite by blast
+
+lemma card_of_Times_ordLess_infinite[simp]:
+assumes INF: "infinite C" and
+ LESS1: "|A| <o |C|" and LESS2: "|B| <o |C|"
+shows "|A \<times> B| <o |C|"
+proof(cases "A = {} \<or> B = {}")
+ assume Case1: "A = {} \<or> B = {}"
+ hence "A \<times> B = {}" by blast
+ moreover have "C \<noteq> {}" using
+ LESS1 card_of_empty5 by blast
+ ultimately show ?thesis by(auto simp add: card_of_empty4)
+next
+ assume Case2: "\<not>(A = {} \<or> B = {})"
+ {assume *: "|C| \<le>o |A \<times> B|"
+ hence "infinite (A \<times> B)" using INF card_of_ordLeq_finite by blast
+ hence 1: "infinite A \<or> infinite B" using finite_cartesian_product by blast
+ {assume Case21: "|A| \<le>o |B|"
+ hence "infinite B" using 1 card_of_ordLeq_finite by blast
+ hence "|A \<times> B| =o |B|" using Case2 Case21
+ by (auto simp add: card_of_Times_infinite)
+ hence False using LESS2 not_ordLess_ordLeq * ordLeq_ordIso_trans by blast
+ }
+ moreover
+ {assume Case22: "|B| \<le>o |A|"
+ hence "infinite A" using 1 card_of_ordLeq_finite by blast
+ hence "|A \<times> B| =o |A|" using Case2 Case22
+ by (auto simp add: card_of_Times_infinite)
+ hence False using LESS1 not_ordLess_ordLeq * ordLeq_ordIso_trans by blast
+ }
+ ultimately have False using ordLeq_total card_of_Well_order[of A]
+ card_of_Well_order[of B] by blast
+ }
+ thus ?thesis using ordLess_or_ordLeq[of "|A \<times> B|" "|C|"]
+ card_of_Well_order[of "A \<times> B"] card_of_Well_order[of "C"] by auto
+qed
+
+lemma card_of_Times_ordLess_infinite_Field[simp]:
+assumes INF: "infinite (Field r)" and r: "Card_order r" and
+ LESS1: "|A| <o r" and LESS2: "|B| <o r"
+shows "|A \<times> B| <o r"
+proof-
+ let ?C = "Field r"
+ have 1: "r =o |?C| \<and> |?C| =o r" using r card_of_Field_ordIso
+ ordIso_symmetric by blast
+ hence "|A| <o |?C|" "|B| <o |?C|"
+ using LESS1 LESS2 ordLess_ordIso_trans by blast+
+ hence "|A <*> B| <o |?C|" using INF
+ card_of_Times_ordLess_infinite by blast
+ thus ?thesis using 1 ordLess_ordIso_trans by blast
+qed
+
+lemma card_of_Un_ordLess_infinite[simp]:
+assumes INF: "infinite C" and
+ LESS1: "|A| <o |C|" and LESS2: "|B| <o |C|"
+shows "|A \<union> B| <o |C|"
+using assms card_of_Plus_ordLess_infinite card_of_Un_Plus_ordLeq
+ ordLeq_ordLess_trans by blast
+
+lemma card_of_Un_ordLess_infinite_Field[simp]:
+assumes INF: "infinite (Field r)" and r: "Card_order r" and
+ LESS1: "|A| <o r" and LESS2: "|B| <o r"
+shows "|A Un B| <o r"
+proof-
+ let ?C = "Field r"
+ have 1: "r =o |?C| \<and> |?C| =o r" using r card_of_Field_ordIso
+ ordIso_symmetric by blast
+ hence "|A| <o |?C|" "|B| <o |?C|"
+ using LESS1 LESS2 ordLess_ordIso_trans by blast+
+ hence "|A Un B| <o |?C|" using INF
+ card_of_Un_ordLess_infinite by blast
+ thus ?thesis using 1 ordLess_ordIso_trans by blast
+qed
+
+lemma card_of_Un_singl_ordLess_infinite1:
+assumes "infinite B" and "|A| <o |B|"
+shows "|{a} Un A| <o |B|"
+proof-
+ have "|{a}| <o |B|" using assms by auto
+ thus ?thesis using assms card_of_Un_ordLess_infinite[of B] by fastforce
+qed
+
+lemma card_of_Un_singl_ordLess_infinite:
+assumes "infinite B"
+shows "( |A| <o |B| ) = ( |{a} Un A| <o |B| )"
+using assms card_of_Un_singl_ordLess_infinite1[of B A]
+proof(auto)
+ assume "|insert a A| <o |B|"
+ moreover have "|A| <=o |insert a A|" using card_of_mono1[of A] by blast
+ ultimately show "|A| <o |B|" using ordLeq_ordLess_trans by blast
+qed
+
+
+subsection {* Cardinals versus lists *}
+
+lemma Card_order_lists: "Card_order r \<Longrightarrow> r \<le>o |lists(Field r) |"
+using card_of_lists card_of_Field_ordIso ordIso_ordLeq_trans ordIso_symmetric by blast
+
+lemma Union_set_lists:
+"Union(set ` (lists A)) = A"
+unfolding lists_def2 proof(auto)
+ fix a assume "a \<in> A"
+ hence "set [a] \<le> A \<and> a \<in> set [a]" by auto
+ thus "\<exists>l. set l \<le> A \<and> a \<in> set l" by blast
+qed
+
+lemma inj_on_map_lists:
+assumes "inj_on f A"
+shows "inj_on (map f) (lists A)"
+using assms Union_set_lists[of A] inj_on_mapI[of f "lists A"] by auto
+
+lemma map_lists_mono:
+assumes "f ` A \<le> B"
+shows "(map f) ` (lists A) \<le> lists B"
+using assms unfolding lists_def2 by (auto, blast) (* lethal combination of methods :) *)
+
+lemma map_lists_surjective:
+assumes "f ` A = B"
+shows "(map f) ` (lists A) = lists B"
+using assms unfolding lists_def2
+proof (auto, blast)
+ fix l' assume *: "set l' \<le> f ` A"
+ have "set l' \<le> f ` A \<longrightarrow> l' \<in> map f ` {l. set l \<le> A}"
+ proof(induct l', auto)
+ fix l a
+ assume "a \<in> A" and "set l \<le> A" and
+ IH: "f ` (set l) \<le> f ` A"
+ hence "set (a # l) \<le> A" by auto
+ hence "map f (a # l) \<in> map f ` {l. set l \<le> A}" by blast
+ thus "f a # map f l \<in> map f ` {l. set l \<le> A}" by auto
+ qed
+ thus "l' \<in> map f ` {l. set l \<le> A}" using * by auto
+qed
+
+lemma bij_betw_map_lists:
+assumes "bij_betw f A B"
+shows "bij_betw (map f) (lists A) (lists B)"
+using assms unfolding bij_betw_def
+by(auto simp add: inj_on_map_lists map_lists_surjective)
+
+lemma card_of_lists_mono[simp]:
+assumes "|A| \<le>o |B|"
+shows "|lists A| \<le>o |lists B|"
+proof-
+ obtain f where "inj_on f A \<and> f ` A \<le> B"
+ using assms card_of_ordLeq[of A B] by auto
+ hence "inj_on (map f) (lists A) \<and> (map f) ` (lists A) \<le> (lists B)"
+ by (auto simp add: inj_on_map_lists map_lists_mono)
+ thus ?thesis using card_of_ordLeq[of "lists A"] by metis
+qed
+
+lemma ordIso_lists_mono:
+assumes "r \<le>o r'"
+shows "|lists(Field r)| \<le>o |lists(Field r')|"
+using assms card_of_mono2 card_of_lists_mono by blast
+
+lemma card_of_lists_cong[simp]:
+assumes "|A| =o |B|"
+shows "|lists A| =o |lists B|"
+proof-
+ obtain f where "bij_betw f A B"
+ using assms card_of_ordIso[of A B] by auto
+ hence "bij_betw (map f) (lists A) (lists B)"
+ by (auto simp add: bij_betw_map_lists)
+ thus ?thesis using card_of_ordIso[of "lists A"] by auto
+qed
+
+lemma ordIso_lists_cong:
+assumes "r =o r'"
+shows "|lists(Field r)| =o |lists(Field r')|"
+using assms card_of_cong card_of_lists_cong by blast
+
+corollary lists_infinite_bij_betw:
+assumes "infinite A"
+shows "\<exists>f. bij_betw f (lists A) A"
+using assms card_of_lists_infinite card_of_ordIso by blast
+
+corollary lists_infinite_bij_betw_types:
+assumes "infinite(UNIV :: 'a set)"
+shows "\<exists>(f::'a list \<Rightarrow> 'a). bij f"
+using assms assms lists_infinite_bij_betw[of "UNIV::'a set"]
+using lists_UNIV by auto
+
+
+subsection {* Cardinals versus the set-of-finite-sets operator *}
+
+definition Fpow :: "'a set \<Rightarrow> 'a set set"
+where "Fpow A \<equiv> {X. X \<le> A \<and> finite X}"
+
+lemma Fpow_mono: "A \<le> B \<Longrightarrow> Fpow A \<le> Fpow B"
+unfolding Fpow_def by auto
+
+lemma empty_in_Fpow: "{} \<in> Fpow A"
+unfolding Fpow_def by auto
+
+lemma Fpow_not_empty: "Fpow A \<noteq> {}"
+using empty_in_Fpow by blast
+
+lemma Fpow_subset_Pow: "Fpow A \<le> Pow A"
+unfolding Fpow_def by auto
+
+lemma card_of_Fpow[simp]: "|A| \<le>o |Fpow A|"
+proof-
+ let ?h = "\<lambda> a. {a}"
+ have "inj_on ?h A \<and> ?h ` A \<le> Fpow A"
+ unfolding inj_on_def Fpow_def by auto
+ thus ?thesis using card_of_ordLeq by metis
+qed
+
+lemma Card_order_Fpow: "Card_order r \<Longrightarrow> r \<le>o |Fpow(Field r) |"
+using card_of_Fpow card_of_Field_ordIso ordIso_ordLeq_trans ordIso_symmetric by blast
+
+lemma Fpow_Pow_finite: "Fpow A = Pow A Int {A. finite A}"
+unfolding Fpow_def Pow_def by blast
+
+lemma inj_on_image_Fpow:
+assumes "inj_on f A"
+shows "inj_on (image f) (Fpow A)"
+using assms Fpow_subset_Pow[of A] subset_inj_on[of "image f" "Pow A"]
+ inj_on_image_Pow by blast
+
+lemma image_Fpow_mono:
+assumes "f ` A \<le> B"
+shows "(image f) ` (Fpow A) \<le> Fpow B"
+using assms by(unfold Fpow_def, auto)
+
+lemma image_Fpow_surjective:
+assumes "f ` A = B"
+shows "(image f) ` (Fpow A) = Fpow B"
+using assms proof(unfold Fpow_def, auto)
+ fix Y assume *: "Y \<le> f ` A" and **: "finite Y"
+ hence "\<forall>b \<in> Y. \<exists>a. a \<in> A \<and> f a = b" by auto
+ with bchoice[of Y "\<lambda>b a. a \<in> A \<and> f a = b"]
+ obtain g where 1: "\<forall>b \<in> Y. g b \<in> A \<and> f(g b) = b" by blast
+ obtain X where X_def: "X = g ` Y" by blast
+ have "f ` X = Y \<and> X \<le> A \<and> finite X"
+ by(unfold X_def, force simp add: ** 1)
+ thus "Y \<in> (image f) ` {X. X \<le> A \<and> finite X}" by auto
+qed
+
+lemma bij_betw_image_Fpow:
+assumes "bij_betw f A B"
+shows "bij_betw (image f) (Fpow A) (Fpow B)"
+using assms unfolding bij_betw_def
+by (auto simp add: inj_on_image_Fpow image_Fpow_surjective)
+
+lemma card_of_Fpow_mono[simp]:
+assumes "|A| \<le>o |B|"
+shows "|Fpow A| \<le>o |Fpow B|"
+proof-
+ obtain f where "inj_on f A \<and> f ` A \<le> B"
+ using assms card_of_ordLeq[of A B] by auto
+ hence "inj_on (image f) (Fpow A) \<and> (image f) ` (Fpow A) \<le> (Fpow B)"
+ by (auto simp add: inj_on_image_Fpow image_Fpow_mono)
+ thus ?thesis using card_of_ordLeq[of "Fpow A"] by auto
+qed
+
+lemma ordIso_Fpow_mono:
+assumes "r \<le>o r'"
+shows "|Fpow(Field r)| \<le>o |Fpow(Field r')|"
+using assms card_of_mono2 card_of_Fpow_mono by blast
+
+lemma card_of_Fpow_cong[simp]:
+assumes "|A| =o |B|"
+shows "|Fpow A| =o |Fpow B|"
+proof-
+ obtain f where "bij_betw f A B"
+ using assms card_of_ordIso[of A B] by auto
+ hence "bij_betw (image f) (Fpow A) (Fpow B)"
+ by (auto simp add: bij_betw_image_Fpow)
+ thus ?thesis using card_of_ordIso[of "Fpow A"] by auto
+qed
+
+lemma ordIso_Fpow_cong:
+assumes "r =o r'"
+shows "|Fpow(Field r)| =o |Fpow(Field r')|"
+using assms card_of_cong card_of_Fpow_cong by blast
+
+lemma card_of_Fpow_lists: "|Fpow A| \<le>o |lists A|"
+proof-
+ have "set ` (lists A) = Fpow A"
+ unfolding lists_def2 Fpow_def using finite_list finite_set by blast
+ thus ?thesis using card_of_ordLeq2[of "Fpow A"] Fpow_not_empty[of A] by blast
+qed
+
+lemma card_of_Fpow_infinite[simp]:
+assumes "infinite A"
+shows "|Fpow A| =o |A|"
+using assms card_of_Fpow_lists card_of_lists_infinite card_of_Fpow
+ ordLeq_ordIso_trans ordIso_iff_ordLeq by blast
+
+corollary Fpow_infinite_bij_betw:
+assumes "infinite A"
+shows "\<exists>f. bij_betw f (Fpow A) A"
+using assms card_of_Fpow_infinite card_of_ordIso by blast
+
+
+subsection {* The cardinal $\omega$ and the finite cardinals *}
+
+subsubsection {* First as well-orders *}
+
+lemma Field_natLess: "Field natLess = (UNIV::nat set)"
+by(unfold Field_def, auto)
+
+lemma natLeq_ofilter_less: "ofilter natLeq {0 ..< n}"
+by(auto simp add: natLeq_wo_rel wo_rel.ofilter_def,
+ simp add: Field_natLeq, unfold rel.under_def, auto)
+
+lemma natLeq_ofilter_leq: "ofilter natLeq {0 .. n}"
+by(auto simp add: natLeq_wo_rel wo_rel.ofilter_def,
+ simp add: Field_natLeq, unfold rel.under_def, auto)
+
+lemma natLeq_ofilter_iff:
+"ofilter natLeq A = (A = UNIV \<or> (\<exists>n. A = {0 ..< n}))"
+proof(rule iffI)
+ assume "ofilter natLeq A"
+ hence "\<forall>m n. n \<in> A \<and> m \<le> n \<longrightarrow> m \<in> A"
+ by(auto simp add: natLeq_wo_rel wo_rel.ofilter_def rel.under_def)
+ thus "A = UNIV \<or> (\<exists>n. A = {0 ..< n})" using closed_nat_set_iff by blast
+next
+ assume "A = UNIV \<or> (\<exists>n. A = {0 ..< n})"
+ thus "ofilter natLeq A"
+ by(auto simp add: natLeq_ofilter_less natLeq_UNIV_ofilter)
+qed
+
+lemma natLeq_under_leq: "under natLeq n = {0 .. n}"
+unfolding rel.under_def by auto
+
+corollary natLeq_on_ofilter:
+"ofilter(natLeq_on n) {0 ..< n}"
+by (auto simp add: natLeq_on_ofilter_less_eq)
+
+lemma natLeq_on_ofilter_less:
+"n < m \<Longrightarrow> ofilter (natLeq_on m) {0 .. n}"
+by(auto simp add: natLeq_on_wo_rel wo_rel.ofilter_def,
+ simp add: Field_natLeq_on, unfold rel.under_def, auto)
+
+lemma natLeq_on_ordLess_natLeq: "natLeq_on n <o natLeq"
+using Field_natLeq Field_natLeq_on[of n] nat_infinite
+ finite_ordLess_infinite[of "natLeq_on n" natLeq]
+ natLeq_Well_order natLeq_on_Well_order[of n] by auto
+
+lemma natLeq_on_injective:
+"natLeq_on m = natLeq_on n \<Longrightarrow> m = n"
+using Field_natLeq_on[of m] Field_natLeq_on[of n]
+ atLeastLessThan_injective[of m n] by auto
+
+lemma natLeq_on_injective_ordIso:
+"(natLeq_on m =o natLeq_on n) = (m = n)"
+proof(auto simp add: natLeq_on_Well_order ordIso_reflexive)
+ assume "natLeq_on m =o natLeq_on n"
+ then obtain f where "bij_betw f {0..<m} {0..<n}"
+ using Field_natLeq_on assms unfolding ordIso_def iso_def[abs_def] by auto
+ thus "m = n" using atLeastLessThan_injective2 by blast
+qed
+
+
+subsubsection {* Then as cardinals *}
+
+lemma ordIso_natLeq_infinite1:
+"|A| =o natLeq \<Longrightarrow> infinite A"
+using ordIso_symmetric ordIso_imp_ordLeq infinite_iff_natLeq_ordLeq by blast
+
+lemma ordIso_natLeq_infinite2:
+"natLeq =o |A| \<Longrightarrow> infinite A"
+using ordIso_imp_ordLeq infinite_iff_natLeq_ordLeq by blast
+
+lemma ordLeq_natLeq_on_imp_finite:
+assumes "|A| \<le>o natLeq_on n"
+shows "finite A"
+proof-
+ have "|A| \<le>o |{0 ..< n}|"
+ using assms card_of_less ordIso_symmetric ordLeq_ordIso_trans by blast
+ thus ?thesis by (auto simp add: card_of_ordLeq_finite)
+qed
+
+
+subsubsection {* "Backwards compatibility" with the numeric cardinal operator for finite sets *}
+
+lemma finite_card_of_iff_card:
+assumes FIN: "finite A" and FIN': "finite B"
+shows "( |A| =o |B| ) = (card A = card B)"
+using assms card_of_ordIso[of A B] bij_betw_iff_card[of A B] by blast
+
+lemma finite_card_of_iff_card3:
+assumes FIN: "finite A" and FIN': "finite B"
+shows "( |A| <o |B| ) = (card A < card B)"
+proof-
+ have "( |A| <o |B| ) = (~ ( |B| \<le>o |A| ))" by simp
+ also have "... = (~ (card B \<le> card A))"
+ using assms by(simp add: finite_card_of_iff_card2)
+ also have "... = (card A < card B)" by auto
+ finally show ?thesis .
+qed
+
+lemma card_Field_natLeq_on:
+"card(Field(natLeq_on n)) = n"
+using Field_natLeq_on card_atLeastLessThan by auto
+
+
+subsection {* The successor of a cardinal *}
+
+lemma embed_implies_ordIso_Restr:
+assumes WELL: "Well_order r" and WELL': "Well_order r'" and EMB: "embed r' r f"
+shows "r' =o Restr r (f ` (Field r'))"
+using assms embed_implies_iso_Restr Well_order_Restr unfolding ordIso_def by blast
+
+lemma cardSuc_Well_order[simp]:
+"Card_order r \<Longrightarrow> Well_order(cardSuc r)"
+using cardSuc_Card_order unfolding card_order_on_def by blast
+
+lemma Field_cardSuc_not_empty:
+assumes "Card_order r"
+shows "Field (cardSuc r) \<noteq> {}"
+proof
+ assume "Field(cardSuc r) = {}"
+ hence "|Field(cardSuc r)| \<le>o r" using assms Card_order_empty[of r] by auto
+ hence "cardSuc r \<le>o r" using assms card_of_Field_ordIso
+ cardSuc_Card_order ordIso_symmetric ordIso_ordLeq_trans by blast
+ thus False using cardSuc_greater not_ordLess_ordLeq assms by blast
+qed
+
+lemma cardSuc_mono_ordLess[simp]:
+assumes CARD: "Card_order r" and CARD': "Card_order r'"
+shows "(cardSuc r <o cardSuc r') = (r <o r')"
+proof-
+ have 0: "Well_order r \<and> Well_order r' \<and> Well_order(cardSuc r) \<and> Well_order(cardSuc r')"
+ using assms by auto
+ thus ?thesis
+ using not_ordLeq_iff_ordLess not_ordLeq_iff_ordLess[of r r']
+ using cardSuc_mono_ordLeq[of r' r] assms by blast
+qed
+
+lemma card_of_Plus_ordLeq_infinite[simp]:
+assumes C: "infinite C" and A: "|A| \<le>o |C|" and B: "|B| \<le>o |C|"
+shows "|A <+> B| \<le>o |C|"
+proof-
+ let ?r = "cardSuc |C|"
+ have "Card_order ?r \<and> infinite (Field ?r)" using assms by simp
+ moreover have "|A| <o ?r" and "|B| <o ?r" using A B by auto
+ ultimately have "|A <+> B| <o ?r"
+ using card_of_Plus_ordLess_infinite_Field by blast
+ thus ?thesis using C by simp
+qed
+
+lemma card_of_Un_ordLeq_infinite[simp]:
+assumes C: "infinite C" and A: "|A| \<le>o |C|" and B: "|B| \<le>o |C|"
+shows "|A Un B| \<le>o |C|"
+using assms card_of_Plus_ordLeq_infinite card_of_Un_Plus_ordLeq
+ordLeq_transitive by metis
+
+
+subsection {* Others *}
+
+lemma under_mono[simp]:
+assumes "Well_order r" and "(i,j) \<in> r"
+shows "under r i \<subseteq> under r j"
+using assms unfolding rel.under_def order_on_defs
+trans_def by blast
+
+lemma underS_under:
+assumes "i \<in> Field r"
+shows "underS r i = under r i - {i}"
+using assms unfolding rel.underS_def rel.under_def by auto
+
+lemma relChain_under:
+assumes "Well_order r"
+shows "relChain r (\<lambda> i. under r i)"
+using assms unfolding relChain_def by auto
+
+lemma infinite_card_of_diff_singl:
+assumes "infinite A"
+shows "|A - {a}| =o |A|"
+by (metis assms card_of_infinite_diff_finitte finite.emptyI finite_insert)
+
+lemma card_of_vimage:
+assumes "B \<subseteq> range f"
+shows "|B| \<le>o |f -` B|"
+apply(rule surj_imp_ordLeq[of _ f])
+using assms by (metis Int_absorb2 image_vimage_eq order_refl)
+
+lemma surj_card_of_vimage:
+assumes "surj f"
+shows "|B| \<le>o |f -` B|"
+by (metis assms card_of_vimage subset_UNIV)
+
+(* bounded powerset *)
+definition Bpow where
+"Bpow r A \<equiv> {X . X \<subseteq> A \<and> |X| \<le>o r}"
+
+lemma Bpow_empty[simp]:
+assumes "Card_order r"
+shows "Bpow r {} = {{}}"
+using assms unfolding Bpow_def by auto
+
+lemma singl_in_Bpow:
+assumes rc: "Card_order r"
+and r: "Field r \<noteq> {}" and a: "a \<in> A"
+shows "{a} \<in> Bpow r A"
+proof-
+ have "|{a}| \<le>o r" using r rc by auto
+ thus ?thesis unfolding Bpow_def using a by auto
+qed
+
+lemma ordLeq_card_Bpow:
+assumes rc: "Card_order r" and r: "Field r \<noteq> {}"
+shows "|A| \<le>o |Bpow r A|"
+proof-
+ have "inj_on (\<lambda> a. {a}) A" unfolding inj_on_def by auto
+ moreover have "(\<lambda> a. {a}) ` A \<subseteq> Bpow r A"
+ using singl_in_Bpow[OF assms] by auto
+ ultimately show ?thesis unfolding card_of_ordLeq[symmetric] by blast
+qed
+
+lemma infinite_Bpow:
+assumes rc: "Card_order r" and r: "Field r \<noteq> {}"
+and A: "infinite A"
+shows "infinite (Bpow r A)"
+using ordLeq_card_Bpow[OF rc r]
+by (metis A card_of_ordLeq_infinite)
+
+lemma Bpow_ordLeq_Func_Field:
+assumes rc: "Card_order r" and r: "Field r \<noteq> {}" and A: "infinite A"
+shows "|Bpow r A| \<le>o |Func (Field r) A|"
+proof-
+ let ?F = "\<lambda> f. {x | x a. f a = Some x}"
+ {fix X assume "X \<in> Bpow r A - {{}}"
+ hence XA: "X \<subseteq> A" and "|X| \<le>o r"
+ and X: "X \<noteq> {}" unfolding Bpow_def by auto
+ hence "|X| \<le>o |Field r|" by (metis Field_card_of card_of_mono2)
+ then obtain F where 1: "X = F ` (Field r)"
+ using card_of_ordLeq2[OF X] by metis
+ def f \<equiv> "\<lambda> i. if i \<in> Field r then Some (F i) else None"
+ have "\<exists> f \<in> Func (Field r) A. X = ?F f"
+ apply (intro bexI[of _ f]) using 1 XA unfolding Func_def f_def by auto
+ }
+ hence "Bpow r A - {{}} \<subseteq> ?F ` (Func (Field r) A)" by auto
+ hence "|Bpow r A - {{}}| \<le>o |Func (Field r) A|"
+ by (rule surj_imp_ordLeq)
+ moreover
+ {have 2: "infinite (Bpow r A)" using infinite_Bpow[OF rc r A] .
+ have "|Bpow r A| =o |Bpow r A - {{}}|"
+ using card_of_infinite_diff_finitte
+ by (metis Pow_empty 2 finite_Pow_iff infinite_imp_nonempty ordIso_symmetric)
+ }
+ ultimately show ?thesis by (metis ordIso_ordLeq_trans)
+qed
+
+lemma Func_emp2[simp]: "A \<noteq> {} \<Longrightarrow> Func A {} = {}" by auto
+
+lemma empty_in_Func[simp]:
+"B \<noteq> {} \<Longrightarrow> empty \<in> Func {} B"
+unfolding Func_def by auto
+
+lemma Func_mono[simp]:
+assumes "B1 \<subseteq> B2"
+shows "Func A B1 \<subseteq> Func A B2"
+using assms unfolding Func_def by force
+
+lemma Pfunc_mono[simp]:
+assumes "A1 \<subseteq> A2" and "B1 \<subseteq> B2"
+shows "Pfunc A B1 \<subseteq> Pfunc A B2"
+using assms in_mono unfolding Pfunc_def apply safe
+apply(case_tac "x a", auto)
+by (metis in_mono option.simps(5))
+
+lemma card_of_Func_UNIV_UNIV:
+"|Func (UNIV::'a set) (UNIV::'b set)| =o |UNIV::('a \<Rightarrow> 'b) set|"
+using card_of_Func_UNIV[of "UNIV::'b set"] by auto
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Ordinals_and_Cardinals/Cardinal_Order_Relation_Base.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,2579 @@
+(* Title: HOL/Ordinals_and_Cardinals/Cardinal_Order_Relation_Base.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Cardinal-order relations (base).
+*)
+
+header {* Cardinal-Order Relations (Base) *}
+
+theory Cardinal_Order_Relation_Base
+imports Constructions_on_Wellorders_Base
+begin
+
+
+text{* In this section, we define cardinal-order relations to be minim well-orders
+on their field. Then we define the cardinal of a set to be {\em some} cardinal-order
+relation on that set, which will be unique up to order isomorphism. Then we study
+the connection between cardinals and:
+\begin{itemize}
+\item standard set-theoretic constructions: products,
+sums, unions, lists, powersets, set-of finite sets operator;
+\item finiteness and infiniteness (in particular, with the numeric cardinal operator
+for finite sets, @{text "card"}, from the theory @{text "Finite_Sets.thy"}).
+\end{itemize}
+%
+On the way, we define the canonical $\omega$ cardinal and finite cardinals. We also
+define (again, up to order isomorphism) the successor of a cardinal, and show that
+any cardinal admits a successor.
+
+Main results of this section are the existence of cardinal relations and the
+facts that, in the presence of infiniteness,
+most of the standard set-theoretic constructions (except for the powerset)
+{\em do not increase cardinality}. In particular, e.g., the set of words/lists over
+any infinite set has the same cardinality (hence, is in bijection) with that set.
+*}
+
+
+subsection {* Cardinal orders *}
+
+
+text{* A cardinal order in our setting shall be a well-order {\em minim} w.r.t. the
+order-embedding relation, @{text "\<le>o"} (which is the same as being {\em minimal} w.r.t. the
+strict order-embedding relation, @{text "<o"}), among all the well-orders on its field. *}
+
+definition card_order_on :: "'a set \<Rightarrow> 'a rel \<Rightarrow> bool"
+where
+"card_order_on A r \<equiv> well_order_on A r \<and> (\<forall>r'. well_order_on A r' \<longrightarrow> r \<le>o r')"
+
+
+abbreviation "Card_order r \<equiv> card_order_on (Field r) r"
+abbreviation "card_order r \<equiv> card_order_on UNIV r"
+
+
+lemma card_order_on_well_order_on:
+assumes "card_order_on A r"
+shows "well_order_on A r"
+using assms unfolding card_order_on_def by simp
+
+
+lemma card_order_on_Card_order:
+"card_order_on A r \<Longrightarrow> A = Field r \<and> Card_order r"
+unfolding card_order_on_def using rel.well_order_on_Field by blast
+
+
+text{* The existence of a cardinal relation on any given set (which will mean
+that any set has a cardinal) follows from two facts:
+\begin{itemize}
+\item Zermelo's theorem (proved in @{text "Zorn.thy"} as theorem @{text "well_order_on"}),
+which states that on any given set there exists a well-order;
+\item The well-founded-ness of @{text "<o"}, ensuring that then there exists a minimal
+such well-order, i.e., a cardinal order.
+\end{itemize}
+*}
+
+
+theorem card_order_on: "\<exists>r. card_order_on A r"
+proof-
+ obtain R where R_def: "R = {r. well_order_on A r}" by blast
+ have 1: "R \<noteq> {} \<and> (\<forall>r \<in> R. Well_order r)"
+ using well_order_on[of A] R_def rel.well_order_on_Well_order by blast
+ hence "\<exists>r \<in> R. \<forall>r' \<in> R. r \<le>o r'"
+ using exists_minim_Well_order[of R] by auto
+ thus ?thesis using R_def unfolding card_order_on_def by auto
+qed
+
+
+lemma card_order_on_ordIso:
+assumes CO: "card_order_on A r" and CO': "card_order_on A r'"
+shows "r =o r'"
+using assms unfolding card_order_on_def
+using ordIso_iff_ordLeq by blast
+
+
+lemma Card_order_ordIso:
+assumes CO: "Card_order r" and ISO: "r' =o r"
+shows "Card_order r'"
+using ISO unfolding ordIso_def
+proof(unfold card_order_on_def, auto)
+ fix p' assume "well_order_on (Field r') p'"
+ hence 0: "Well_order p' \<and> Field p' = Field r'"
+ using rel.well_order_on_Well_order by blast
+ obtain f where 1: "iso r' r f" and 2: "Well_order r \<and> Well_order r'"
+ using ISO unfolding ordIso_def by auto
+ hence 3: "inj_on f (Field r') \<and> f ` (Field r') = Field r"
+ by (auto simp add: iso_iff embed_inj_on)
+ let ?p = "dir_image p' f"
+ have 4: "p' =o ?p \<and> Well_order ?p"
+ using 0 2 3 by (auto simp add: dir_image_ordIso Well_order_dir_image)
+ moreover have "Field ?p = Field r"
+ using 0 3 by (auto simp add: dir_image_Field2 order_on_defs)
+ ultimately have "well_order_on (Field r) ?p" by auto
+ hence "r \<le>o ?p" using CO unfolding card_order_on_def by auto
+ thus "r' \<le>o p'"
+ using ISO 4 ordLeq_ordIso_trans ordIso_ordLeq_trans ordIso_symmetric by blast
+qed
+
+
+lemma Card_order_ordIso2:
+assumes CO: "Card_order r" and ISO: "r =o r'"
+shows "Card_order r'"
+using assms Card_order_ordIso ordIso_symmetric by blast
+
+
+subsection {* Cardinal of a set *}
+
+
+text{* We define the cardinal of set to be {\em some} cardinal order on that set.
+We shall prove that this notion is unique up to order isomorphism, meaning
+that order isomorphism shall be the true identity of cardinals. *}
+
+
+definition card_of :: "'a set \<Rightarrow> 'a rel" ("|_|" )
+where "card_of A = (SOME r. card_order_on A r)"
+
+
+lemma card_of_card_order_on: "card_order_on A |A|"
+unfolding card_of_def by (auto simp add: card_order_on someI_ex)
+
+
+lemma card_of_well_order_on: "well_order_on A |A|"
+using card_of_card_order_on card_order_on_def by blast
+
+
+lemma Field_card_of: "Field |A| = A"
+using card_of_card_order_on[of A] unfolding card_order_on_def
+using rel.well_order_on_Field by blast
+
+
+lemma card_of_Card_order: "Card_order |A|"
+by (simp only: card_of_card_order_on Field_card_of)
+
+
+corollary ordIso_card_of_imp_Card_order:
+"r =o |A| \<Longrightarrow> Card_order r"
+using card_of_Card_order Card_order_ordIso by blast
+
+
+lemma card_of_Well_order: "Well_order |A|"
+using card_of_Card_order unfolding card_order_on_def by auto
+
+
+lemma card_of_refl: "|A| =o |A|"
+using card_of_Well_order ordIso_reflexive by blast
+
+
+lemma card_of_least: "well_order_on A r \<Longrightarrow> |A| \<le>o r"
+using card_of_card_order_on unfolding card_order_on_def by blast
+
+
+lemma card_of_ordIso:
+"(\<exists>f. bij_betw f A B) = ( |A| =o |B| )"
+proof(auto)
+ fix f assume *: "bij_betw f A B"
+ then obtain r where "well_order_on B r \<and> |A| =o r"
+ using Well_order_iso_copy card_of_well_order_on by blast
+ hence "|B| \<le>o |A|" using card_of_least
+ ordLeq_ordIso_trans ordIso_symmetric by blast
+ moreover
+ {let ?g = "inv_into A f"
+ have "bij_betw ?g B A" using * bij_betw_inv_into by blast
+ then obtain r where "well_order_on A r \<and> |B| =o r"
+ using Well_order_iso_copy card_of_well_order_on by blast
+ hence "|A| \<le>o |B|" using card_of_least
+ ordLeq_ordIso_trans ordIso_symmetric by blast
+ }
+ ultimately show "|A| =o |B|" using ordIso_iff_ordLeq by blast
+next
+ assume "|A| =o |B|"
+ then obtain f where "iso ( |A| ) ( |B| ) f"
+ unfolding ordIso_def by auto
+ hence "bij_betw f A B" unfolding iso_def Field_card_of by simp
+ thus "\<exists>f. bij_betw f A B" by auto
+qed
+
+
+lemma card_of_ordLeq:
+"(\<exists>f. inj_on f A \<and> f ` A \<le> B) = ( |A| \<le>o |B| )"
+proof(auto)
+ fix f assume *: "inj_on f A" and **: "f ` A \<le> B"
+ {assume "|B| <o |A|"
+ hence "|B| \<le>o |A|" using ordLeq_iff_ordLess_or_ordIso by blast
+ then obtain g where "embed ( |B| ) ( |A| ) g"
+ unfolding ordLeq_def by auto
+ hence 1: "inj_on g B \<and> g ` B \<le> A" using embed_inj_on[of "|B|" "|A|" "g"]
+ card_of_Well_order[of "B"] Field_card_of[of "B"] Field_card_of[of "A"]
+ embed_Field[of "|B|" "|A|" g] by auto
+ obtain h where "bij_betw h A B"
+ using * ** 1 Cantor_Bernstein[of f] by fastforce
+ hence "|A| =o |B|" using card_of_ordIso by blast
+ hence "|A| \<le>o |B|" using ordIso_iff_ordLeq by auto
+ }
+ thus "|A| \<le>o |B|" using ordLess_or_ordLeq[of "|B|" "|A|"]
+ by (auto simp: card_of_Well_order)
+next
+ assume *: "|A| \<le>o |B|"
+ obtain f where "embed ( |A| ) ( |B| ) f"
+ using * unfolding ordLeq_def by auto
+ hence "inj_on f A \<and> f ` A \<le> B" using embed_inj_on[of "|A|" "|B|" f]
+ card_of_Well_order[of "A"] Field_card_of[of "A"] Field_card_of[of "B"]
+ embed_Field[of "|A|" "|B|" f] by auto
+ thus "\<exists>f. inj_on f A \<and> f ` A \<le> B" by auto
+qed
+
+
+lemma card_of_ordLeq2:
+"A \<noteq> {} \<Longrightarrow> (\<exists>g. g ` B = A) = ( |A| \<le>o |B| )"
+using card_of_ordLeq[of A B] inj_on_iff_surj[of A B] by auto
+
+
+lemma card_of_ordLess:
+"(\<not>(\<exists>f. inj_on f A \<and> f ` A \<le> B)) = ( |B| <o |A| )"
+proof-
+ have "(\<not>(\<exists>f. inj_on f A \<and> f ` A \<le> B)) = (\<not> |A| \<le>o |B| )"
+ using card_of_ordLeq by blast
+ also have "\<dots> = ( |B| <o |A| )"
+ using card_of_Well_order[of A] card_of_Well_order[of B]
+ not_ordLeq_iff_ordLess by blast
+ finally show ?thesis .
+qed
+
+
+lemma card_of_ordLess2:
+"B \<noteq> {} \<Longrightarrow> (\<not>(\<exists>f. f ` A = B)) = ( |A| <o |B| )"
+using card_of_ordLess[of B A] inj_on_iff_surj[of B A] by auto
+
+
+lemma card_of_ordIsoI:
+assumes "bij_betw f A B"
+shows "|A| =o |B|"
+using assms unfolding card_of_ordIso[symmetric] by auto
+
+
+lemma card_of_ordLeqI:
+assumes "inj_on f A" and "\<And> a. a \<in> A \<Longrightarrow> f a \<in> B"
+shows "|A| \<le>o |B|"
+using assms unfolding card_of_ordLeq[symmetric] by auto
+
+
+lemma card_of_unique:
+"card_order_on A r \<Longrightarrow> r =o |A|"
+by (simp only: card_order_on_ordIso card_of_card_order_on)
+
+
+lemma card_of_mono1:
+"A \<le> B \<Longrightarrow> |A| \<le>o |B|"
+using inj_on_id[of A] card_of_ordLeq[of A B] by fastforce
+
+
+lemma card_of_mono2:
+assumes "r \<le>o r'"
+shows "|Field r| \<le>o |Field r'|"
+proof-
+ obtain f where
+ 1: "well_order_on (Field r) r \<and> well_order_on (Field r) r \<and> embed r r' f"
+ using assms unfolding ordLeq_def
+ by (auto simp add: rel.well_order_on_Well_order)
+ hence "inj_on f (Field r) \<and> f ` (Field r) \<le> Field r'"
+ by (auto simp add: embed_inj_on embed_Field)
+ thus "|Field r| \<le>o |Field r'|" using card_of_ordLeq by blast
+qed
+
+
+lemma card_of_cong: "r =o r' \<Longrightarrow> |Field r| =o |Field r'|"
+by (simp add: ordIso_iff_ordLeq card_of_mono2)
+
+
+lemma card_of_Field_ordLess: "Well_order r \<Longrightarrow> |Field r| \<le>o r"
+using card_of_least card_of_well_order_on rel.well_order_on_Well_order by blast
+
+
+lemma card_of_Field_ordIso:
+assumes "Card_order r"
+shows "|Field r| =o r"
+proof-
+ have "card_order_on (Field r) r"
+ using assms card_order_on_Card_order by blast
+ moreover have "card_order_on (Field r) |Field r|"
+ using card_of_card_order_on by blast
+ ultimately show ?thesis using card_order_on_ordIso by blast
+qed
+
+
+lemma Card_order_iff_ordIso_card_of:
+"Card_order r = (r =o |Field r| )"
+using ordIso_card_of_imp_Card_order card_of_Field_ordIso ordIso_symmetric by blast
+
+
+lemma Card_order_iff_ordLeq_card_of:
+"Card_order r = (r \<le>o |Field r| )"
+proof-
+ have "Card_order r = (r =o |Field r| )"
+ unfolding Card_order_iff_ordIso_card_of by simp
+ also have "... = (r \<le>o |Field r| \<and> |Field r| \<le>o r)"
+ unfolding ordIso_iff_ordLeq by simp
+ also have "... = (r \<le>o |Field r| )"
+ using card_of_Field_ordLess
+ by (auto simp: card_of_Field_ordLess ordLeq_Well_order_simp)
+ finally show ?thesis .
+qed
+
+
+lemma Card_order_iff_Restr_underS:
+assumes "Well_order r"
+shows "Card_order r = (\<forall>a \<in> Field r. Restr r (rel.underS r a) <o |Field r| )"
+using assms unfolding Card_order_iff_ordLeq_card_of
+using ordLeq_iff_ordLess_Restr card_of_Well_order by blast
+
+
+lemma card_of_underS:
+assumes r: "Card_order r" and a: "a : Field r"
+shows "|rel.underS r a| <o r"
+proof-
+ let ?A = "rel.underS r a" let ?r' = "Restr r ?A"
+ have 1: "Well_order r"
+ using r unfolding card_order_on_def by simp
+ have "Well_order ?r'" using 1 Well_order_Restr by auto
+ moreover have "card_order_on (Field ?r') |Field ?r'|"
+ using card_of_card_order_on .
+ ultimately have "|Field ?r'| \<le>o ?r'"
+ unfolding card_order_on_def by simp
+ moreover have "Field ?r' = ?A"
+ using 1 wo_rel.underS_ofilter Field_Restr_ofilter
+ unfolding wo_rel_def by fastforce
+ ultimately have "|?A| \<le>o ?r'" by simp
+ also have "?r' <o |Field r|"
+ using 1 a r Card_order_iff_Restr_underS by blast
+ also have "|Field r| =o r"
+ using r ordIso_symmetric unfolding Card_order_iff_ordIso_card_of by auto
+ finally show ?thesis .
+qed
+
+
+lemma ordLess_Field:
+assumes "r <o r'"
+shows "|Field r| <o r'"
+proof-
+ have "well_order_on (Field r) r" using assms unfolding ordLess_def
+ by (auto simp add: rel.well_order_on_Well_order)
+ hence "|Field r| \<le>o r" using card_of_least by blast
+ thus ?thesis using assms ordLeq_ordLess_trans by blast
+qed
+
+
+lemma internalize_card_of_ordLeq:
+"( |A| \<le>o r) = (\<exists>B \<le> Field r. |A| =o |B| \<and> |B| \<le>o r)"
+proof
+ assume "|A| \<le>o r"
+ then obtain p where 1: "Field p \<le> Field r \<and> |A| =o p \<and> p \<le>o r"
+ using internalize_ordLeq[of "|A|" r] by blast
+ hence "Card_order p" using card_of_Card_order Card_order_ordIso2 by blast
+ hence "|Field p| =o p" using card_of_Field_ordIso by blast
+ hence "|A| =o |Field p| \<and> |Field p| \<le>o r"
+ using 1 ordIso_equivalence ordIso_ordLeq_trans by blast
+ thus "\<exists>B \<le> Field r. |A| =o |B| \<and> |B| \<le>o r" using 1 by blast
+next
+ assume "\<exists>B \<le> Field r. |A| =o |B| \<and> |B| \<le>o r"
+ thus "|A| \<le>o r" using ordIso_ordLeq_trans by blast
+qed
+
+
+lemma internalize_card_of_ordLeq2:
+"( |A| \<le>o |C| ) = (\<exists>B \<le> C. |A| =o |B| \<and> |B| \<le>o |C| )"
+using internalize_card_of_ordLeq[of "A" "|C|"] Field_card_of[of C] by auto
+
+
+
+subsection {* Cardinals versus set operations on arbitrary sets *}
+
+
+text{* Here we embark in a long journey of simple results showing
+that the standard set-theoretic operations are well-behaved w.r.t. the notion of
+cardinal -- essentially, this means that they preserve the ``cardinal identity"
+@{text "=o"} and are monotonic w.r.t. @{text "\<le>o"}.
+*}
+
+
+lemma card_of_empty: "|{}| \<le>o |A|"
+using card_of_ordLeq inj_on_id by blast
+
+
+lemma card_of_empty1:
+assumes "Well_order r \<or> Card_order r"
+shows "|{}| \<le>o r"
+proof-
+ have "Well_order r" using assms unfolding card_order_on_def by auto
+ hence "|Field r| <=o r"
+ using assms card_of_Field_ordLess by blast
+ moreover have "|{}| \<le>o |Field r|" by (simp add: card_of_empty)
+ ultimately show ?thesis using ordLeq_transitive by blast
+qed
+
+
+corollary Card_order_empty:
+"Card_order r \<Longrightarrow> |{}| \<le>o r" by (simp add: card_of_empty1)
+
+
+lemma card_of_empty2:
+assumes LEQ: "|A| =o |{}|"
+shows "A = {}"
+using assms card_of_ordIso[of A] bij_betw_empty2 by blast
+
+
+lemma card_of_empty3:
+assumes LEQ: "|A| \<le>o |{}|"
+shows "A = {}"
+using assms
+by (simp add: ordIso_iff_ordLeq card_of_empty1 card_of_empty2
+ ordLeq_Well_order_simp)
+
+
+lemma card_of_empty_ordIso:
+"|{}::'a set| =o |{}::'b set|"
+using card_of_ordIso unfolding bij_betw_def inj_on_def by blast
+
+
+lemma card_of_image:
+"|f ` A| <=o |A|"
+proof(cases "A = {}", simp add: card_of_empty)
+ assume "A ~= {}"
+ hence "f ` A ~= {}" by auto
+ thus "|f ` A| \<le>o |A|"
+ using card_of_ordLeq2[of "f ` A" A] by auto
+qed
+
+
+lemma surj_imp_ordLeq:
+assumes "B <= f ` A"
+shows "|B| <=o |A|"
+proof-
+ have "|B| <=o |f ` A|" using assms card_of_mono1 by auto
+ thus ?thesis using card_of_image ordLeq_transitive by blast
+qed
+
+
+lemma card_of_ordLeqI2:
+assumes "B \<subseteq> f ` A"
+shows "|B| \<le>o |A|"
+using assms by (metis surj_imp_ordLeq)
+
+
+lemma card_of_singl_ordLeq:
+assumes "A \<noteq> {}"
+shows "|{b}| \<le>o |A|"
+proof-
+ obtain a where *: "a \<in> A" using assms by auto
+ let ?h = "\<lambda> b'::'b. if b' = b then a else undefined"
+ have "inj_on ?h {b} \<and> ?h ` {b} \<le> A"
+ using * unfolding inj_on_def by auto
+ thus ?thesis using card_of_ordLeq by blast
+qed
+
+
+corollary Card_order_singl_ordLeq:
+"\<lbrakk>Card_order r; Field r \<noteq> {}\<rbrakk> \<Longrightarrow> |{b}| \<le>o r"
+using card_of_singl_ordLeq[of "Field r" b]
+ card_of_Field_ordIso[of r] ordLeq_ordIso_trans by blast
+
+
+lemma card_of_Pow: "|A| <o |Pow A|"
+using card_of_ordLess2[of "Pow A" A] Cantors_paradox[of A]
+ Pow_not_empty[of A] by auto
+
+
+lemma infinite_Pow:
+assumes "infinite A"
+shows "infinite (Pow A)"
+proof-
+ have "|A| \<le>o |Pow A|" by (metis card_of_Pow ordLess_imp_ordLeq)
+ thus ?thesis by (metis assms finite_Pow_iff)
+qed
+
+
+corollary Card_order_Pow:
+"Card_order r \<Longrightarrow> r <o |Pow(Field r)|"
+using card_of_Pow card_of_Field_ordIso ordIso_ordLess_trans ordIso_symmetric by blast
+
+
+corollary card_of_set_type: "|UNIV::'a set| <o |UNIV::'a set set|"
+using card_of_Pow[of "UNIV::'a set"] by simp
+
+
+lemma card_of_Plus1: "|A| \<le>o |A <+> B|"
+proof-
+ have "Inl ` A \<le> A <+> B" by auto
+ thus ?thesis using inj_Inl[of A] card_of_ordLeq by blast
+qed
+
+
+corollary Card_order_Plus1:
+"Card_order r \<Longrightarrow> r \<le>o |(Field r) <+> B|"
+using card_of_Plus1 card_of_Field_ordIso ordIso_ordLeq_trans ordIso_symmetric by blast
+
+
+lemma card_of_Plus2: "|B| \<le>o |A <+> B|"
+proof-
+ have "Inr ` B \<le> A <+> B" by auto
+ thus ?thesis using inj_Inr[of B] card_of_ordLeq by blast
+qed
+
+
+corollary Card_order_Plus2:
+"Card_order r \<Longrightarrow> r \<le>o |A <+> (Field r)|"
+using card_of_Plus2 card_of_Field_ordIso ordIso_ordLeq_trans ordIso_symmetric by blast
+
+
+lemma card_of_Plus_empty1: "|A| =o |A <+> {}|"
+proof-
+ have "bij_betw Inl A (A <+> {})" unfolding bij_betw_def inj_on_def by auto
+ thus ?thesis using card_of_ordIso by auto
+qed
+
+
+lemma card_of_Plus_empty2: "|A| =o |{} <+> A|"
+proof-
+ have "bij_betw Inr A ({} <+> A)" unfolding bij_betw_def inj_on_def by auto
+ thus ?thesis using card_of_ordIso by auto
+qed
+
+
+lemma card_of_Plus_commute: "|A <+> B| =o |B <+> A|"
+proof-
+ let ?f = "\<lambda>(c::'a + 'b). case c of Inl a \<Rightarrow> Inr a
+ | Inr b \<Rightarrow> Inl b"
+ have "bij_betw ?f (A <+> B) (B <+> A)"
+ unfolding bij_betw_def inj_on_def by force
+ thus ?thesis using card_of_ordIso by blast
+qed
+
+
+lemma card_of_Plus_assoc:
+fixes A :: "'a set" and B :: "'b set" and C :: "'c set"
+shows "|(A <+> B) <+> C| =o |A <+> B <+> C|"
+proof -
+ def f \<equiv> "\<lambda>(k::('a + 'b) + 'c).
+ case k of Inl ab \<Rightarrow> (case ab of Inl a \<Rightarrow> Inl a
+ |Inr b \<Rightarrow> Inr (Inl b))
+ |Inr c \<Rightarrow> Inr (Inr c)"
+ have "A <+> B <+> C \<subseteq> f ` ((A <+> B) <+> C)"
+ proof
+ fix x assume x: "x \<in> A <+> B <+> C"
+ show "x \<in> f ` ((A <+> B) <+> C)"
+ proof(cases x)
+ case (Inl a)
+ hence "a \<in> A" "x = f (Inl (Inl a))"
+ using x unfolding f_def by auto
+ thus ?thesis by auto
+ next
+ case (Inr bc) note 1 = Inr show ?thesis
+ proof(cases bc)
+ case (Inl b)
+ hence "b \<in> B" "x = f (Inl (Inr b))"
+ using x 1 unfolding f_def by auto
+ thus ?thesis by auto
+ next
+ case (Inr c)
+ hence "c \<in> C" "x = f (Inr c)"
+ using x 1 unfolding f_def by auto
+ thus ?thesis by auto
+ qed
+ qed
+ qed
+ hence "bij_betw f ((A <+> B) <+> C) (A <+> B <+> C)"
+ unfolding bij_betw_def inj_on_def f_def by auto
+ thus ?thesis using card_of_ordIso by blast
+qed
+
+
+lemma card_of_Plus_mono1:
+assumes "|A| \<le>o |B|"
+shows "|A <+> C| \<le>o |B <+> C|"
+proof-
+ obtain f where 1: "inj_on f A \<and> f ` A \<le> B"
+ using assms card_of_ordLeq[of A] by fastforce
+ obtain g where g_def:
+ "g = (\<lambda>d. case d of Inl a \<Rightarrow> Inl(f a) | Inr (c::'c) \<Rightarrow> Inr c)" by blast
+ have "inj_on g (A <+> C) \<and> g ` (A <+> C) \<le> (B <+> C)"
+ proof-
+ {fix d1 and d2 assume "d1 \<in> A <+> C \<and> d2 \<in> A <+> C" and
+ "g d1 = g d2"
+ hence "d1 = d2" using 1 unfolding inj_on_def
+ by(case_tac d1, case_tac d2, auto simp add: g_def)
+ }
+ moreover
+ {fix d assume "d \<in> A <+> C"
+ hence "g d \<in> B <+> C" using 1
+ by(case_tac d, auto simp add: g_def)
+ }
+ ultimately show ?thesis unfolding inj_on_def by auto
+ qed
+ thus ?thesis using card_of_ordLeq by metis
+qed
+
+
+corollary ordLeq_Plus_mono1:
+assumes "r \<le>o r'"
+shows "|(Field r) <+> C| \<le>o |(Field r') <+> C|"
+using assms card_of_mono2 card_of_Plus_mono1 by blast
+
+
+lemma card_of_Plus_mono2:
+assumes "|A| \<le>o |B|"
+shows "|C <+> A| \<le>o |C <+> B|"
+using assms card_of_Plus_mono1[of A B C]
+ card_of_Plus_commute[of C A] card_of_Plus_commute[of B C]
+ ordIso_ordLeq_trans[of "|C <+> A|"] ordLeq_ordIso_trans[of "|C <+> A|"]
+by blast
+
+
+corollary ordLeq_Plus_mono2:
+assumes "r \<le>o r'"
+shows "|A <+> (Field r)| \<le>o |A <+> (Field r')|"
+using assms card_of_mono2 card_of_Plus_mono2 by blast
+
+
+lemma card_of_Plus_mono:
+assumes "|A| \<le>o |B|" and "|C| \<le>o |D|"
+shows "|A <+> C| \<le>o |B <+> D|"
+using assms card_of_Plus_mono1[of A B C] card_of_Plus_mono2[of C D B]
+ ordLeq_transitive[of "|A <+> C|"] by blast
+
+
+corollary ordLeq_Plus_mono:
+assumes "r \<le>o r'" and "p \<le>o p'"
+shows "|(Field r) <+> (Field p)| \<le>o |(Field r') <+> (Field p')|"
+using assms card_of_mono2[of r r'] card_of_mono2[of p p'] card_of_Plus_mono by blast
+
+
+lemma card_of_Plus_cong1:
+assumes "|A| =o |B|"
+shows "|A <+> C| =o |B <+> C|"
+using assms by (simp add: ordIso_iff_ordLeq card_of_Plus_mono1)
+
+
+corollary ordIso_Plus_cong1:
+assumes "r =o r'"
+shows "|(Field r) <+> C| =o |(Field r') <+> C|"
+using assms card_of_cong card_of_Plus_cong1 by blast
+
+
+lemma card_of_Plus_cong2:
+assumes "|A| =o |B|"
+shows "|C <+> A| =o |C <+> B|"
+using assms by (simp add: ordIso_iff_ordLeq card_of_Plus_mono2)
+
+
+corollary ordIso_Plus_cong2:
+assumes "r =o r'"
+shows "|A <+> (Field r)| =o |A <+> (Field r')|"
+using assms card_of_cong card_of_Plus_cong2 by blast
+
+
+lemma card_of_Plus_cong:
+assumes "|A| =o |B|" and "|C| =o |D|"
+shows "|A <+> C| =o |B <+> D|"
+using assms by (simp add: ordIso_iff_ordLeq card_of_Plus_mono)
+
+
+corollary ordIso_Plus_cong:
+assumes "r =o r'" and "p =o p'"
+shows "|(Field r) <+> (Field p)| =o |(Field r') <+> (Field p')|"
+using assms card_of_cong[of r r'] card_of_cong[of p p'] card_of_Plus_cong by blast
+
+
+lemma card_of_Un1:
+shows "|A| \<le>o |A \<union> B| "
+using inj_on_id[of A] card_of_ordLeq[of A _] by fastforce
+
+
+lemma card_of_diff:
+shows "|A - B| \<le>o |A|"
+using inj_on_id[of "A - B"] card_of_ordLeq[of "A - B" _] by fastforce
+
+
+lemma card_of_Un_Plus_ordLeq:
+"|A \<union> B| \<le>o |A <+> B|"
+proof-
+ let ?f = "\<lambda> c. if c \<in> A then Inl c else Inr c"
+ have "inj_on ?f (A \<union> B) \<and> ?f ` (A \<union> B) \<le> A <+> B"
+ unfolding inj_on_def by auto
+ thus ?thesis using card_of_ordLeq by blast
+qed
+
+
+lemma card_of_Times1:
+assumes "A \<noteq> {}"
+shows "|B| \<le>o |B \<times> A|"
+proof(cases "B = {}", simp add: card_of_empty)
+ assume *: "B \<noteq> {}"
+ have "fst `(B \<times> A) = B" unfolding image_def using assms by auto
+ thus ?thesis using inj_on_iff_surj[of B "B \<times> A"]
+ card_of_ordLeq[of B "B \<times> A"] * by blast
+qed
+
+
+corollary Card_order_Times1:
+"\<lbrakk>Card_order r; B \<noteq> {}\<rbrakk> \<Longrightarrow> r \<le>o |(Field r) \<times> B|"
+using card_of_Times1[of B] card_of_Field_ordIso
+ ordIso_ordLeq_trans ordIso_symmetric by blast
+
+
+lemma card_of_Times_commute: "|A \<times> B| =o |B \<times> A|"
+proof-
+ let ?f = "\<lambda>(a::'a,b::'b). (b,a)"
+ have "bij_betw ?f (A \<times> B) (B \<times> A)"
+ unfolding bij_betw_def inj_on_def by auto
+ thus ?thesis using card_of_ordIso by blast
+qed
+
+
+lemma card_of_Times2:
+assumes "A \<noteq> {}" shows "|B| \<le>o |A \<times> B|"
+using assms card_of_Times1[of A B] card_of_Times_commute[of B A]
+ ordLeq_ordIso_trans by blast
+
+
+corollary Card_order_Times2:
+"\<lbrakk>Card_order r; A \<noteq> {}\<rbrakk> \<Longrightarrow> r \<le>o |A \<times> (Field r)|"
+using card_of_Times2[of A] card_of_Field_ordIso
+ ordIso_ordLeq_trans ordIso_symmetric by blast
+
+
+lemma card_of_Times3: "|A| \<le>o |A \<times> A|"
+using card_of_Times1[of A]
+by(cases "A = {}", simp add: card_of_empty, blast)
+
+
+lemma card_of_Plus_Times_bool: "|A <+> A| =o |A \<times> (UNIV::bool set)|"
+proof-
+ let ?f = "\<lambda>c::'a + 'a. case c of Inl a \<Rightarrow> (a,True)
+ |Inr a \<Rightarrow> (a,False)"
+ have "bij_betw ?f (A <+> A) (A \<times> (UNIV::bool set))"
+ proof-
+ {fix c1 and c2 assume "?f c1 = ?f c2"
+ hence "c1 = c2"
+ by(case_tac "c1", case_tac "c2", auto, case_tac "c2", auto)
+ }
+ moreover
+ {fix c assume "c \<in> A <+> A"
+ hence "?f c \<in> A \<times> (UNIV::bool set)"
+ by(case_tac c, auto)
+ }
+ moreover
+ {fix a bl assume *: "(a,bl) \<in> A \<times> (UNIV::bool set)"
+ have "(a,bl) \<in> ?f ` ( A <+> A)"
+ proof(cases bl)
+ assume bl hence "?f(Inl a) = (a,bl)" by auto
+ thus ?thesis using * by force
+ next
+ assume "\<not> bl" hence "?f(Inr a) = (a,bl)" by auto
+ thus ?thesis using * by force
+ qed
+ }
+ ultimately show ?thesis unfolding bij_betw_def inj_on_def by auto
+ qed
+ thus ?thesis using card_of_ordIso by blast
+qed
+
+
+lemma card_of_Times_mono1:
+assumes "|A| \<le>o |B|"
+shows "|A \<times> C| \<le>o |B \<times> C|"
+proof-
+ obtain f where 1: "inj_on f A \<and> f ` A \<le> B"
+ using assms card_of_ordLeq[of A] by fastforce
+ obtain g where g_def:
+ "g = (\<lambda>(a,c::'c). (f a,c))" by blast
+ have "inj_on g (A \<times> C) \<and> g ` (A \<times> C) \<le> (B \<times> C)"
+ using 1 unfolding inj_on_def using g_def by auto
+ thus ?thesis using card_of_ordLeq by metis
+qed
+
+
+corollary ordLeq_Times_mono1:
+assumes "r \<le>o r'"
+shows "|(Field r) \<times> C| \<le>o |(Field r') \<times> C|"
+using assms card_of_mono2 card_of_Times_mono1 by blast
+
+
+lemma card_of_Times_mono2:
+assumes "|A| \<le>o |B|"
+shows "|C \<times> A| \<le>o |C \<times> B|"
+using assms card_of_Times_mono1[of A B C]
+ card_of_Times_commute[of C A] card_of_Times_commute[of B C]
+ ordIso_ordLeq_trans[of "|C \<times> A|"] ordLeq_ordIso_trans[of "|C \<times> A|"]
+by blast
+
+
+corollary ordLeq_Times_mono2:
+assumes "r \<le>o r'"
+shows "|A \<times> (Field r)| \<le>o |A \<times> (Field r')|"
+using assms card_of_mono2 card_of_Times_mono2 by blast
+
+
+lemma card_of_Times_cong1:
+assumes "|A| =o |B|"
+shows "|A \<times> C| =o |B \<times> C|"
+using assms by (simp add: ordIso_iff_ordLeq card_of_Times_mono1)
+
+
+lemma card_of_Times_cong2:
+assumes "|A| =o |B|"
+shows "|C \<times> A| =o |C \<times> B|"
+using assms by (simp add: ordIso_iff_ordLeq card_of_Times_mono2)
+
+
+corollary ordIso_Times_cong2:
+assumes "r =o r'"
+shows "|A \<times> (Field r)| =o |A \<times> (Field r')|"
+using assms card_of_cong card_of_Times_cong2 by blast
+
+
+lemma card_of_Sigma_mono1:
+assumes "\<forall>i \<in> I. |A i| \<le>o |B i|"
+shows "|SIGMA i : I. A i| \<le>o |SIGMA i : I. B i|"
+proof-
+ have "\<forall>i. i \<in> I \<longrightarrow> (\<exists>f. inj_on f (A i) \<and> f ` (A i) \<le> B i)"
+ using assms by (auto simp add: card_of_ordLeq)
+ with choice[of "\<lambda> i f. i \<in> I \<longrightarrow> inj_on f (A i) \<and> f ` (A i) \<le> B i"]
+ obtain F where 1: "\<forall>i \<in> I. inj_on (F i) (A i) \<and> (F i) ` (A i) \<le> B i" by fastforce
+ obtain g where g_def: "g = (\<lambda>(i,a::'b). (i,F i a))" by blast
+ have "inj_on g (Sigma I A) \<and> g ` (Sigma I A) \<le> (Sigma I B)"
+ using 1 unfolding inj_on_def using g_def by force
+ thus ?thesis using card_of_ordLeq by metis
+qed
+
+
+corollary card_of_Sigma_Times:
+"\<forall>i \<in> I. |A i| \<le>o |B| \<Longrightarrow> |SIGMA i : I. A i| \<le>o |I \<times> B|"
+using card_of_Sigma_mono1[of I A "\<lambda>i. B"] .
+
+
+lemma card_of_UNION_Sigma:
+"|\<Union>i \<in> I. A i| \<le>o |SIGMA i : I. A i|"
+using Ex_inj_on_UNION_Sigma[of I A] card_of_ordLeq by metis
+
+
+lemma card_of_bool:
+assumes "a1 \<noteq> a2"
+shows "|UNIV::bool set| =o |{a1,a2}|"
+proof-
+ let ?f = "\<lambda> bl. case bl of True \<Rightarrow> a1 | False \<Rightarrow> a2"
+ have "bij_betw ?f UNIV {a1,a2}"
+ proof-
+ {fix bl1 and bl2 assume "?f bl1 = ?f bl2"
+ hence "bl1 = bl2" using assms by (case_tac bl1, case_tac bl2, auto)
+ }
+ moreover
+ {fix bl have "?f bl \<in> {a1,a2}" by (case_tac bl, auto)
+ }
+ moreover
+ {fix a assume *: "a \<in> {a1,a2}"
+ have "a \<in> ?f ` UNIV"
+ proof(cases "a = a1")
+ assume "a = a1"
+ hence "?f True = a" by auto thus ?thesis by blast
+ next
+ assume "a \<noteq> a1" hence "a = a2" using * by auto
+ hence "?f False = a" by auto thus ?thesis by blast
+ qed
+ }
+ ultimately show ?thesis unfolding bij_betw_def inj_on_def
+ by (metis image_subsetI order_eq_iff subsetI)
+ qed
+ thus ?thesis using card_of_ordIso by blast
+qed
+
+
+lemma card_of_Plus_Times_aux:
+assumes A2: "a1 \<noteq> a2 \<and> {a1,a2} \<le> A" and
+ LEQ: "|A| \<le>o |B|"
+shows "|A <+> B| \<le>o |A \<times> B|"
+proof-
+ have 1: "|UNIV::bool set| \<le>o |A|"
+ using A2 card_of_mono1[of "{a1,a2}"] card_of_bool[of a1 a2]
+ ordIso_ordLeq_trans[of "|UNIV::bool set|"] by metis
+ (* *)
+ have "|A <+> B| \<le>o |B <+> B|"
+ using LEQ card_of_Plus_mono1 by blast
+ moreover have "|B <+> B| =o |B \<times> (UNIV::bool set)|"
+ using card_of_Plus_Times_bool by blast
+ moreover have "|B \<times> (UNIV::bool set)| \<le>o |B \<times> A|"
+ using 1 by (simp add: card_of_Times_mono2)
+ moreover have " |B \<times> A| =o |A \<times> B|"
+ using card_of_Times_commute by blast
+ ultimately show "|A <+> B| \<le>o |A \<times> B|"
+ using ordLeq_ordIso_trans[of "|A <+> B|" "|B <+> B|" "|B \<times> (UNIV::bool set)|"]
+ ordLeq_transitive[of "|A <+> B|" "|B \<times> (UNIV::bool set)|" "|B \<times> A|"]
+ ordLeq_ordIso_trans[of "|A <+> B|" "|B \<times> A|" "|A \<times> B|"]
+ by blast
+qed
+
+
+lemma card_of_Plus_Times:
+assumes A2: "a1 \<noteq> a2 \<and> {a1,a2} \<le> A" and
+ B2: "b1 \<noteq> b2 \<and> {b1,b2} \<le> B"
+shows "|A <+> B| \<le>o |A \<times> B|"
+proof-
+ {assume "|A| \<le>o |B|"
+ hence ?thesis using assms by (auto simp add: card_of_Plus_Times_aux)
+ }
+ moreover
+ {assume "|B| \<le>o |A|"
+ hence "|B <+> A| \<le>o |B \<times> A|"
+ using assms by (auto simp add: card_of_Plus_Times_aux)
+ hence ?thesis
+ using card_of_Plus_commute card_of_Times_commute
+ ordIso_ordLeq_trans ordLeq_ordIso_trans by metis
+ }
+ ultimately show ?thesis
+ using card_of_Well_order[of A] card_of_Well_order[of B]
+ ordLeq_total[of "|A|"] by metis
+qed
+
+
+lemma card_of_ordLeq_finite:
+assumes "|A| \<le>o |B|" and "finite B"
+shows "finite A"
+using assms unfolding ordLeq_def
+using embed_inj_on[of "|A|" "|B|"] embed_Field[of "|A|" "|B|"]
+ Field_card_of[of "A"] Field_card_of[of "B"] inj_on_finite[of _ "A" "B"] by fastforce
+
+
+lemma card_of_ordLeq_infinite:
+assumes "|A| \<le>o |B|" and "infinite A"
+shows "infinite B"
+using assms card_of_ordLeq_finite by auto
+
+
+lemma card_of_ordIso_finite:
+assumes "|A| =o |B|"
+shows "finite A = finite B"
+using assms unfolding ordIso_def iso_def[abs_def]
+by (auto simp: bij_betw_finite Field_card_of)
+
+
+lemma card_of_ordIso_finite_Field:
+assumes "Card_order r" and "r =o |A|"
+shows "finite(Field r) = finite A"
+using assms card_of_Field_ordIso card_of_ordIso_finite ordIso_equivalence by blast
+
+
+subsection {* Cardinals versus set operations involving infinite sets *}
+
+
+text{* Here we show that, for infinite sets, most set-theoretic constructions
+do not increase the cardinality. The cornerstone for this is
+theorem @{text "Card_order_Times_same_infinite"}, which states that self-product
+does not increase cardinality -- the proof of this fact adapts a standard
+set-theoretic argument, as presented, e.g., in the proof of theorem 1.5.11
+at page 47 in \cite{card-book}. Then everything else follows fairly easily. *}
+
+
+lemma infinite_iff_card_of_nat:
+"infinite A = ( |UNIV::nat set| \<le>o |A| )"
+by (auto simp add: infinite_iff_countable_subset card_of_ordLeq)
+
+
+lemma finite_iff_cardOf_nat:
+"finite A = ( |A| <o |UNIV :: nat set| )"
+using infinite_iff_card_of_nat[of A]
+not_ordLeq_iff_ordLess[of "|A|" "|UNIV :: nat set|"]
+by (fastforce simp: card_of_Well_order)
+
+lemma finite_ordLess_infinite2:
+assumes "finite A" and "infinite B"
+shows "|A| <o |B|"
+using assms
+finite_ordLess_infinite[of "|A|" "|B|"]
+card_of_Well_order[of A] card_of_Well_order[of B]
+Field_card_of[of A] Field_card_of[of B] by auto
+
+
+text{* The next two results correspond to the ZF fact that all infinite cardinals are
+limit ordinals: *}
+
+lemma Card_order_infinite_not_under:
+assumes CARD: "Card_order r" and INF: "infinite (Field r)"
+shows "\<not> (\<exists>a. Field r = rel.under r a)"
+proof(auto)
+ have 0: "Well_order r \<and> wo_rel r \<and> Refl r"
+ using CARD unfolding wo_rel_def card_order_on_def order_on_defs by auto
+ fix a assume *: "Field r = rel.under r a"
+ show False
+ proof(cases "a \<in> Field r")
+ assume Case1: "a \<notin> Field r"
+ hence "rel.under r a = {}" unfolding Field_def rel.under_def by auto
+ thus False using INF * by auto
+ next
+ let ?r' = "Restr r (rel.underS r a)"
+ assume Case2: "a \<in> Field r"
+ hence 1: "rel.under r a = rel.underS r a \<union> {a} \<and> a \<notin> rel.underS r a"
+ using 0 rel.Refl_under_underS rel.underS_notIn by fastforce
+ have 2: "wo_rel.ofilter r (rel.underS r a) \<and> rel.underS r a < Field r"
+ using 0 wo_rel.underS_ofilter * 1 Case2 by auto
+ hence "?r' <o r" using 0 using ofilter_ordLess by blast
+ moreover
+ have "Field ?r' = rel.underS r a \<and> Well_order ?r'"
+ using 2 0 Field_Restr_ofilter[of r] Well_order_Restr[of r] by blast
+ ultimately have "|rel.underS r a| <o r" using ordLess_Field[of ?r'] by auto
+ moreover have "|rel.under r a| =o r" using * CARD card_of_Field_ordIso[of r] by auto
+ ultimately have "|rel.underS r a| <o |rel.under r a|"
+ using ordIso_symmetric ordLess_ordIso_trans by blast
+ moreover
+ {have "\<exists>f. bij_betw f (rel.under r a) (rel.underS r a)"
+ using infinite_imp_bij_betw[of "Field r" a] INF * 1 by auto
+ hence "|rel.under r a| =o |rel.underS r a|" using card_of_ordIso by blast
+ }
+ ultimately show False using not_ordLess_ordIso ordIso_symmetric by blast
+ qed
+qed
+
+
+lemma infinite_Card_order_limit:
+assumes r: "Card_order r" and "infinite (Field r)"
+and a: "a : Field r"
+shows "EX b : Field r. a \<noteq> b \<and> (a,b) : r"
+proof-
+ have "Field r \<noteq> rel.under r a"
+ using assms Card_order_infinite_not_under by blast
+ moreover have "rel.under r a \<le> Field r"
+ using rel.under_Field .
+ ultimately have "rel.under r a < Field r" by blast
+ then obtain b where 1: "b : Field r \<and> ~ (b,a) : r"
+ unfolding rel.under_def by blast
+ moreover have ba: "b \<noteq> a"
+ using 1 r unfolding card_order_on_def well_order_on_def
+ linear_order_on_def partial_order_on_def preorder_on_def refl_on_def by auto
+ ultimately have "(a,b) : r"
+ using a r unfolding card_order_on_def well_order_on_def linear_order_on_def
+ total_on_def by blast
+ thus ?thesis using 1 ba by auto
+qed
+
+
+theorem Card_order_Times_same_infinite:
+assumes CO: "Card_order r" and INF: "infinite(Field r)"
+shows "|Field r \<times> Field r| \<le>o r"
+proof-
+ obtain phi where phi_def:
+ "phi = (\<lambda>r::'a rel. Card_order r \<and> infinite(Field r) \<and>
+ \<not> |Field r \<times> Field r| \<le>o r )" by blast
+ have temp1: "\<forall>r. phi r \<longrightarrow> Well_order r"
+ unfolding phi_def card_order_on_def by auto
+ have Ft: "\<not>(\<exists>r. phi r)"
+ proof
+ assume "\<exists>r. phi r"
+ hence "{r. phi r} \<noteq> {} \<and> {r. phi r} \<le> {r. Well_order r}"
+ using temp1 by auto
+ then obtain r where 1: "phi r" and 2: "\<forall>r'. phi r' \<longrightarrow> r \<le>o r'" and
+ 3: "Card_order r \<and> Well_order r"
+ using exists_minim_Well_order[of "{r. phi r}"] temp1 phi_def by blast
+ let ?A = "Field r" let ?r' = "bsqr r"
+ have 4: "Well_order ?r' \<and> Field ?r' = ?A \<times> ?A \<and> |?A| =o r"
+ using 3 bsqr_Well_order Field_bsqr card_of_Field_ordIso by blast
+ have 5: "Card_order |?A \<times> ?A| \<and> Well_order |?A \<times> ?A|"
+ using card_of_Card_order card_of_Well_order by blast
+ (* *)
+ have "r <o |?A \<times> ?A|"
+ using 1 3 5 ordLess_or_ordLeq unfolding phi_def by blast
+ moreover have "|?A \<times> ?A| \<le>o ?r'"
+ using card_of_least[of "?A \<times> ?A"] 4 by auto
+ ultimately have "r <o ?r'" using ordLess_ordLeq_trans by auto
+ then obtain f where 6: "embed r ?r' f" and 7: "\<not> bij_betw f ?A (?A \<times> ?A)"
+ unfolding ordLess_def embedS_def[abs_def]
+ by (auto simp add: Field_bsqr)
+ let ?B = "f ` ?A"
+ have "|?A| =o |?B|"
+ using 3 6 embed_inj_on inj_on_imp_bij_betw card_of_ordIso by blast
+ hence 8: "r =o |?B|" using 4 ordIso_transitive ordIso_symmetric by blast
+ (* *)
+ have "wo_rel.ofilter ?r' ?B"
+ using 6 embed_Field_ofilter 3 4 by blast
+ hence "wo_rel.ofilter ?r' ?B \<and> ?B \<noteq> ?A \<times> ?A \<and> ?B \<noteq> Field ?r'"
+ using 7 unfolding bij_betw_def using 6 3 embed_inj_on 4 by auto
+ hence temp2: "wo_rel.ofilter ?r' ?B \<and> ?B < ?A \<times> ?A"
+ using 4 wo_rel_def[of ?r'] wo_rel.ofilter_def[of ?r' ?B] by blast
+ have "\<not> (\<exists>a. Field r = rel.under r a)"
+ using 1 unfolding phi_def using Card_order_infinite_not_under[of r] by auto
+ then obtain A1 where temp3: "wo_rel.ofilter r A1 \<and> A1 < ?A" and 9: "?B \<le> A1 \<times> A1"
+ using temp2 3 bsqr_ofilter[of r ?B] by blast
+ hence "|?B| \<le>o |A1 \<times> A1|" using card_of_mono1 by blast
+ hence 10: "r \<le>o |A1 \<times> A1|" using 8 ordIso_ordLeq_trans by blast
+ let ?r1 = "Restr r A1"
+ have "?r1 <o r" using temp3 ofilter_ordLess 3 by blast
+ moreover
+ {have "well_order_on A1 ?r1" using 3 temp3 well_order_on_Restr by blast
+ hence "|A1| \<le>o ?r1" using 3 Well_order_Restr card_of_least by blast
+ }
+ ultimately have 11: "|A1| <o r" using ordLeq_ordLess_trans by blast
+ (* *)
+ have "infinite (Field r)" using 1 unfolding phi_def by simp
+ hence "infinite ?B" using 8 3 card_of_ordIso_finite_Field[of r ?B] by blast
+ hence "infinite A1" using 9 infinite_super finite_cartesian_product by blast
+ moreover have temp4: "Field |A1| = A1 \<and> Well_order |A1| \<and> Card_order |A1|"
+ using card_of_Card_order[of A1] card_of_Well_order[of A1]
+ by (simp add: Field_card_of)
+ moreover have "\<not> r \<le>o | A1 |"
+ using temp4 11 3 using not_ordLeq_iff_ordLess by blast
+ ultimately have "infinite(Field |A1| ) \<and> Card_order |A1| \<and> \<not> r \<le>o | A1 |"
+ by (simp add: card_of_card_order_on)
+ hence "|Field |A1| \<times> Field |A1| | \<le>o |A1|"
+ using 2 unfolding phi_def by blast
+ hence "|A1 \<times> A1 | \<le>o |A1|" using temp4 by auto
+ hence "r \<le>o |A1|" using 10 ordLeq_transitive by blast
+ thus False using 11 not_ordLess_ordLeq by auto
+ qed
+ thus ?thesis using assms unfolding phi_def by blast
+qed
+
+
+corollary card_of_Times_same_infinite:
+assumes "infinite A"
+shows "|A \<times> A| =o |A|"
+proof-
+ let ?r = "|A|"
+ have "Field ?r = A \<and> Card_order ?r"
+ using Field_card_of card_of_Card_order[of A] by fastforce
+ hence "|A \<times> A| \<le>o |A|"
+ using Card_order_Times_same_infinite[of ?r] assms by auto
+ thus ?thesis using card_of_Times3 ordIso_iff_ordLeq by blast
+qed
+
+
+lemma card_of_Times_infinite:
+assumes INF: "infinite A" and NE: "B \<noteq> {}" and LEQ: "|B| \<le>o |A|"
+shows "|A \<times> B| =o |A| \<and> |B \<times> A| =o |A|"
+proof-
+ have "|A| \<le>o |A \<times> B| \<and> |A| \<le>o |B \<times> A|"
+ using assms by (simp add: card_of_Times1 card_of_Times2)
+ moreover
+ {have "|A \<times> B| \<le>o |A \<times> A| \<and> |B \<times> A| \<le>o |A \<times> A|"
+ using LEQ card_of_Times_mono1 card_of_Times_mono2 by blast
+ moreover have "|A \<times> A| =o |A|" using INF card_of_Times_same_infinite by blast
+ ultimately have "|A \<times> B| \<le>o |A| \<and> |B \<times> A| \<le>o |A|"
+ using ordLeq_ordIso_trans[of "|A \<times> B|"] ordLeq_ordIso_trans[of "|B \<times> A|"] by auto
+ }
+ ultimately show ?thesis by (simp add: ordIso_iff_ordLeq)
+qed
+
+
+corollary card_of_Times_infinite_simps:
+"\<lbrakk>infinite A; B \<noteq> {}; |B| \<le>o |A|\<rbrakk> \<Longrightarrow> |A \<times> B| =o |A|"
+"\<lbrakk>infinite A; B \<noteq> {}; |B| \<le>o |A|\<rbrakk> \<Longrightarrow> |A| =o |A \<times> B|"
+"\<lbrakk>infinite A; B \<noteq> {}; |B| \<le>o |A|\<rbrakk> \<Longrightarrow> |B \<times> A| =o |A|"
+"\<lbrakk>infinite A; B \<noteq> {}; |B| \<le>o |A|\<rbrakk> \<Longrightarrow> |A| =o |B \<times> A|"
+by (auto simp add: card_of_Times_infinite ordIso_symmetric)
+
+
+corollary Card_order_Times_infinite:
+assumes INF: "infinite(Field r)" and CARD: "Card_order r" and
+ NE: "Field p \<noteq> {}" and LEQ: "p \<le>o r"
+shows "| (Field r) \<times> (Field p) | =o r \<and> | (Field p) \<times> (Field r) | =o r"
+proof-
+ have "|Field r \<times> Field p| =o |Field r| \<and> |Field p \<times> Field r| =o |Field r|"
+ using assms by (simp add: card_of_Times_infinite card_of_mono2)
+ thus ?thesis
+ using assms card_of_Field_ordIso[of r]
+ ordIso_transitive[of "|Field r \<times> Field p|"]
+ ordIso_transitive[of _ "|Field r|"] by blast
+qed
+
+
+lemma card_of_Sigma_ordLeq_infinite:
+assumes INF: "infinite B" and
+ LEQ_I: "|I| \<le>o |B|" and LEQ: "\<forall>i \<in> I. |A i| \<le>o |B|"
+shows "|SIGMA i : I. A i| \<le>o |B|"
+proof(cases "I = {}", simp add: card_of_empty)
+ assume *: "I \<noteq> {}"
+ have "|SIGMA i : I. A i| \<le>o |I \<times> B|"
+ using LEQ card_of_Sigma_Times by blast
+ moreover have "|I \<times> B| =o |B|"
+ using INF * LEQ_I by (auto simp add: card_of_Times_infinite)
+ ultimately show ?thesis using ordLeq_ordIso_trans by blast
+qed
+
+
+lemma card_of_Sigma_ordLeq_infinite_Field:
+assumes INF: "infinite (Field r)" and r: "Card_order r" and
+ LEQ_I: "|I| \<le>o r" and LEQ: "\<forall>i \<in> I. |A i| \<le>o r"
+shows "|SIGMA i : I. A i| \<le>o r"
+proof-
+ let ?B = "Field r"
+ have 1: "r =o |?B| \<and> |?B| =o r" using r card_of_Field_ordIso
+ ordIso_symmetric by blast
+ hence "|I| \<le>o |?B|" "\<forall>i \<in> I. |A i| \<le>o |?B|"
+ using LEQ_I LEQ ordLeq_ordIso_trans by blast+
+ hence "|SIGMA i : I. A i| \<le>o |?B|" using INF LEQ
+ card_of_Sigma_ordLeq_infinite by blast
+ thus ?thesis using 1 ordLeq_ordIso_trans by blast
+qed
+
+
+lemma card_of_Times_ordLeq_infinite_Field:
+"\<lbrakk>infinite (Field r); |A| \<le>o r; |B| \<le>o r; Card_order r\<rbrakk>
+ \<Longrightarrow> |A <*> B| \<le>o r"
+by(simp add: card_of_Sigma_ordLeq_infinite_Field)
+
+
+lemma card_of_UNION_ordLeq_infinite:
+assumes INF: "infinite B" and
+ LEQ_I: "|I| \<le>o |B|" and LEQ: "\<forall>i \<in> I. |A i| \<le>o |B|"
+shows "|\<Union> i \<in> I. A i| \<le>o |B|"
+proof(cases "I = {}", simp add: card_of_empty)
+ assume *: "I \<noteq> {}"
+ have "|\<Union> i \<in> I. A i| \<le>o |SIGMA i : I. A i|"
+ using card_of_UNION_Sigma by blast
+ moreover have "|SIGMA i : I. A i| \<le>o |B|"
+ using assms card_of_Sigma_ordLeq_infinite by blast
+ ultimately show ?thesis using ordLeq_transitive by blast
+qed
+
+
+corollary card_of_UNION_ordLeq_infinite_Field:
+assumes INF: "infinite (Field r)" and r: "Card_order r" and
+ LEQ_I: "|I| \<le>o r" and LEQ: "\<forall>i \<in> I. |A i| \<le>o r"
+shows "|\<Union> i \<in> I. A i| \<le>o r"
+proof-
+ let ?B = "Field r"
+ have 1: "r =o |?B| \<and> |?B| =o r" using r card_of_Field_ordIso
+ ordIso_symmetric by blast
+ hence "|I| \<le>o |?B|" "\<forall>i \<in> I. |A i| \<le>o |?B|"
+ using LEQ_I LEQ ordLeq_ordIso_trans by blast+
+ hence "|\<Union> i \<in> I. A i| \<le>o |?B|" using INF LEQ
+ card_of_UNION_ordLeq_infinite by blast
+ thus ?thesis using 1 ordLeq_ordIso_trans by blast
+qed
+
+
+lemma card_of_Plus_infinite1:
+assumes INF: "infinite A" and LEQ: "|B| \<le>o |A|"
+shows "|A <+> B| =o |A|"
+proof(cases "B = {}", simp add: card_of_Plus_empty1 card_of_Plus_empty2 ordIso_symmetric)
+ let ?Inl = "Inl::'a \<Rightarrow> 'a + 'b" let ?Inr = "Inr::'b \<Rightarrow> 'a + 'b"
+ assume *: "B \<noteq> {}"
+ then obtain b1 where 1: "b1 \<in> B" by blast
+ show ?thesis
+ proof(cases "B = {b1}")
+ assume Case1: "B = {b1}"
+ have 2: "bij_betw ?Inl A ((?Inl ` A))"
+ unfolding bij_betw_def inj_on_def by auto
+ hence 3: "infinite (?Inl ` A)"
+ using INF bij_betw_finite[of ?Inl A] by blast
+ let ?A' = "?Inl ` A \<union> {?Inr b1}"
+ obtain g where "bij_betw g (?Inl ` A) ?A'"
+ using 3 infinite_imp_bij_betw2[of "?Inl ` A"] by auto
+ moreover have "?A' = A <+> B" using Case1 by blast
+ ultimately have "bij_betw g (?Inl ` A) (A <+> B)" by simp
+ hence "bij_betw (g o ?Inl) A (A <+> B)"
+ using 2 by (auto simp add: bij_betw_trans)
+ thus ?thesis using card_of_ordIso ordIso_symmetric by blast
+ next
+ assume Case2: "B \<noteq> {b1}"
+ with * 1 obtain b2 where 3: "b1 \<noteq> b2 \<and> {b1,b2} \<le> B" by fastforce
+ obtain f where "inj_on f B \<and> f ` B \<le> A"
+ using LEQ card_of_ordLeq[of B] by fastforce
+ with 3 have "f b1 \<noteq> f b2 \<and> {f b1, f b2} \<le> A"
+ unfolding inj_on_def by auto
+ with 3 have "|A <+> B| \<le>o |A \<times> B|"
+ by (auto simp add: card_of_Plus_Times)
+ moreover have "|A \<times> B| =o |A|"
+ using assms * by (simp add: card_of_Times_infinite_simps)
+ ultimately have "|A <+> B| \<le>o |A|" using ordLeq_ordIso_trans by metis
+ thus ?thesis using card_of_Plus1 ordIso_iff_ordLeq by blast
+ qed
+qed
+
+
+lemma card_of_Plus_infinite2:
+assumes INF: "infinite A" and LEQ: "|B| \<le>o |A|"
+shows "|B <+> A| =o |A|"
+using assms card_of_Plus_commute card_of_Plus_infinite1
+ordIso_equivalence by blast
+
+
+lemma card_of_Plus_infinite:
+assumes INF: "infinite A" and LEQ: "|B| \<le>o |A|"
+shows "|A <+> B| =o |A| \<and> |B <+> A| =o |A|"
+using assms by (auto simp: card_of_Plus_infinite1 card_of_Plus_infinite2)
+
+
+corollary Card_order_Plus_infinite:
+assumes INF: "infinite(Field r)" and CARD: "Card_order r" and
+ LEQ: "p \<le>o r"
+shows "| (Field r) <+> (Field p) | =o r \<and> | (Field p) <+> (Field r) | =o r"
+proof-
+ have "| Field r <+> Field p | =o | Field r | \<and>
+ | Field p <+> Field r | =o | Field r |"
+ using assms by (simp add: card_of_Plus_infinite card_of_mono2)
+ thus ?thesis
+ using assms card_of_Field_ordIso[of r]
+ ordIso_transitive[of "|Field r <+> Field p|"]
+ ordIso_transitive[of _ "|Field r|"] by blast
+qed
+
+
+lemma card_of_Un_infinite:
+assumes INF: "infinite A" and LEQ: "|B| \<le>o |A|"
+shows "|A \<union> B| =o |A| \<and> |B \<union> A| =o |A|"
+proof-
+ have "|A \<union> B| \<le>o |A <+> B|" by (rule card_of_Un_Plus_ordLeq)
+ moreover have "|A <+> B| =o |A|"
+ using assms by (metis card_of_Plus_infinite)
+ ultimately have "|A \<union> B| \<le>o |A|" using ordLeq_ordIso_trans by blast
+ hence "|A \<union> B| =o |A|" using card_of_Un1 ordIso_iff_ordLeq by blast
+ thus ?thesis using Un_commute[of B A] by auto
+qed
+
+
+lemma card_of_Un_diff_infinite:
+assumes INF: "infinite A" and LESS: "|B| <o |A|"
+shows "|A - B| =o |A|"
+proof-
+ obtain C where C_def: "C = A - B" by blast
+ have "|A \<union> B| =o |A|"
+ using assms ordLeq_iff_ordLess_or_ordIso card_of_Un_infinite by blast
+ moreover have "C \<union> B = A \<union> B" unfolding C_def by auto
+ ultimately have 1: "|C \<union> B| =o |A|" by auto
+ (* *)
+ {assume *: "|C| \<le>o |B|"
+ moreover
+ {assume **: "finite B"
+ hence "finite C"
+ using card_of_ordLeq_finite * by blast
+ hence False using ** INF card_of_ordIso_finite 1 by blast
+ }
+ hence "infinite B" by auto
+ ultimately have False
+ using card_of_Un_infinite 1 ordIso_equivalence(1,3) LESS not_ordLess_ordIso by metis
+ }
+ hence 2: "|B| \<le>o |C|" using card_of_Well_order ordLeq_total by blast
+ {assume *: "finite C"
+ hence "finite B" using card_of_ordLeq_finite 2 by blast
+ hence False using * INF card_of_ordIso_finite 1 by blast
+ }
+ hence "infinite C" by auto
+ hence "|C| =o |A|"
+ using card_of_Un_infinite 1 2 ordIso_equivalence(1,3) by metis
+ thus ?thesis unfolding C_def .
+qed
+
+
+lemma card_of_Plus_ordLess_infinite:
+assumes INF: "infinite C" and
+ LESS1: "|A| <o |C|" and LESS2: "|B| <o |C|"
+shows "|A <+> B| <o |C|"
+proof(cases "A = {} \<or> B = {}")
+ assume Case1: "A = {} \<or> B = {}"
+ hence "|A| =o |A <+> B| \<or> |B| =o |A <+> B|"
+ using card_of_Plus_empty1 card_of_Plus_empty2 by blast
+ hence "|A <+> B| =o |A| \<or> |A <+> B| =o |B|"
+ using ordIso_symmetric[of "|A|"] ordIso_symmetric[of "|B|"] by blast
+ thus ?thesis using LESS1 LESS2
+ ordIso_ordLess_trans[of "|A <+> B|" "|A|"]
+ ordIso_ordLess_trans[of "|A <+> B|" "|B|"] by blast
+next
+ assume Case2: "\<not>(A = {} \<or> B = {})"
+ {assume *: "|C| \<le>o |A <+> B|"
+ hence "infinite (A <+> B)" using INF card_of_ordLeq_finite by blast
+ hence 1: "infinite A \<or> infinite B" using finite_Plus by blast
+ {assume Case21: "|A| \<le>o |B|"
+ hence "infinite B" using 1 card_of_ordLeq_finite by blast
+ hence "|A <+> B| =o |B|" using Case2 Case21
+ by (auto simp add: card_of_Plus_infinite)
+ hence False using LESS2 not_ordLess_ordLeq * ordLeq_ordIso_trans by blast
+ }
+ moreover
+ {assume Case22: "|B| \<le>o |A|"
+ hence "infinite A" using 1 card_of_ordLeq_finite by blast
+ hence "|A <+> B| =o |A|" using Case2 Case22
+ by (auto simp add: card_of_Plus_infinite)
+ hence False using LESS1 not_ordLess_ordLeq * ordLeq_ordIso_trans by blast
+ }
+ ultimately have False using ordLeq_total card_of_Well_order[of A]
+ card_of_Well_order[of B] by blast
+ }
+ thus ?thesis using ordLess_or_ordLeq[of "|A <+> B|" "|C|"]
+ card_of_Well_order[of "A <+> B"] card_of_Well_order[of "C"] by auto
+qed
+
+
+lemma card_of_Plus_ordLess_infinite_Field:
+assumes INF: "infinite (Field r)" and r: "Card_order r" and
+ LESS1: "|A| <o r" and LESS2: "|B| <o r"
+shows "|A <+> B| <o r"
+proof-
+ let ?C = "Field r"
+ have 1: "r =o |?C| \<and> |?C| =o r" using r card_of_Field_ordIso
+ ordIso_symmetric by blast
+ hence "|A| <o |?C|" "|B| <o |?C|"
+ using LESS1 LESS2 ordLess_ordIso_trans by blast+
+ hence "|A <+> B| <o |?C|" using INF
+ card_of_Plus_ordLess_infinite by blast
+ thus ?thesis using 1 ordLess_ordIso_trans by blast
+qed
+
+
+lemma infinite_card_of_insert:
+assumes "infinite A"
+shows "|insert a A| =o |A|"
+proof-
+ have iA: "insert a A = A \<union> {a}" by simp
+ show ?thesis
+ using infinite_imp_bij_betw2[OF assms] unfolding iA
+ by (metis bij_betw_inv card_of_ordIso)
+qed
+
+
+subsection {* Cardinals versus lists *}
+
+
+text{* The next is an auxiliary operator, which shall be used for inductive
+proofs of facts concerning the cardinality of @{text "List"} : *}
+
+definition nlists :: "'a set \<Rightarrow> nat \<Rightarrow> 'a list set"
+where "nlists A n \<equiv> {l. set l \<le> A \<and> length l = n}"
+
+
+lemma lists_def2: "lists A = {l. set l \<le> A}"
+using in_listsI by blast
+
+
+lemma lists_UNION_nlists: "lists A = (\<Union> n. nlists A n)"
+unfolding lists_def2 nlists_def by blast
+
+
+lemma card_of_lists: "|A| \<le>o |lists A|"
+proof-
+ let ?h = "\<lambda> a. [a]"
+ have "inj_on ?h A \<and> ?h ` A \<le> lists A"
+ unfolding inj_on_def lists_def2 by auto
+ thus ?thesis by (metis card_of_ordLeq)
+qed
+
+
+lemma nlists_0: "nlists A 0 = {[]}"
+unfolding nlists_def by auto
+
+
+lemma nlists_not_empty:
+assumes "A \<noteq> {}"
+shows "nlists A n \<noteq> {}"
+proof(induct n, simp add: nlists_0)
+ fix n assume "nlists A n \<noteq> {}"
+ then obtain a and l where "a \<in> A \<and> l \<in> nlists A n" using assms by auto
+ hence "a # l \<in> nlists A (Suc n)" unfolding nlists_def by auto
+ thus "nlists A (Suc n) \<noteq> {}" by auto
+qed
+
+
+lemma Nil_in_lists: "[] \<in> lists A"
+unfolding lists_def2 by auto
+
+
+lemma lists_not_empty: "lists A \<noteq> {}"
+using Nil_in_lists by blast
+
+
+lemma card_of_nlists_Succ: "|nlists A (Suc n)| =o |A \<times> (nlists A n)|"
+proof-
+ let ?B = "A \<times> (nlists A n)" let ?h = "\<lambda>(a,l). a # l"
+ have "inj_on ?h ?B \<and> ?h ` ?B \<le> nlists A (Suc n)"
+ unfolding inj_on_def nlists_def by auto
+ moreover have "nlists A (Suc n) \<le> ?h ` ?B"
+ proof(auto)
+ fix l assume "l \<in> nlists A (Suc n)"
+ hence 1: "length l = Suc n \<and> set l \<le> A" unfolding nlists_def by auto
+ then obtain a and l' where 2: "l = a # l'" by (auto simp: length_Suc_conv)
+ hence "a \<in> A \<and> set l' \<le> A \<and> length l' = n" using 1 by auto
+ thus "l \<in> ?h ` ?B" using 2 unfolding nlists_def by auto
+ qed
+ ultimately have "bij_betw ?h ?B (nlists A (Suc n))"
+ unfolding bij_betw_def by auto
+ thus ?thesis using card_of_ordIso ordIso_symmetric by blast
+qed
+
+
+lemma card_of_nlists_infinite:
+assumes "infinite A"
+shows "|nlists A n| \<le>o |A|"
+proof(induct n)
+ have "A \<noteq> {}" using assms by auto
+ thus "|nlists A 0| \<le>o |A|" by (simp add: nlists_0 card_of_singl_ordLeq)
+next
+ fix n assume IH: "|nlists A n| \<le>o |A|"
+ have "|nlists A (Suc n)| =o |A \<times> (nlists A n)|"
+ using card_of_nlists_Succ by blast
+ moreover
+ {have "nlists A n \<noteq> {}" using assms nlists_not_empty[of A] by blast
+ hence "|A \<times> (nlists A n)| =o |A|"
+ using assms IH by (auto simp add: card_of_Times_infinite)
+ }
+ ultimately show "|nlists A (Suc n)| \<le>o |A|"
+ using ordIso_transitive ordIso_iff_ordLeq by blast
+qed
+
+
+lemma card_of_lists_infinite:
+assumes "infinite A"
+shows "|lists A| =o |A|"
+proof-
+ have "|lists A| \<le>o |A|"
+ using assms
+ by (auto simp add: lists_UNION_nlists card_of_UNION_ordLeq_infinite
+ infinite_iff_card_of_nat card_of_nlists_infinite)
+ thus ?thesis using card_of_lists ordIso_iff_ordLeq by blast
+qed
+
+
+lemma Card_order_lists_infinite:
+assumes "Card_order r" and "infinite(Field r)"
+shows "|lists(Field r)| =o r"
+using assms card_of_lists_infinite card_of_Field_ordIso ordIso_transitive by blast
+
+
+
+subsection {* The cardinal $\omega$ and the finite cardinals *}
+
+
+text{* The cardinal $\omega$, of natural numbers, shall be the standard non-strict
+order relation on
+@{text "nat"}, that we abbreviate by @{text "natLeq"}. The finite cardinals
+shall be the restrictions of these relations to the numbers smaller than
+fixed numbers @{text "n"}, that we abbreviate by @{text "natLeq_on n"}. *}
+
+abbreviation "(natLeq::(nat * nat) set) \<equiv> {(x,y). x \<le> y}"
+abbreviation "(natLess::(nat * nat) set) \<equiv> {(x,y). x < y}"
+
+abbreviation natLeq_on :: "nat \<Rightarrow> (nat * nat) set"
+where "natLeq_on n \<equiv> {(x,y). x < n \<and> y < n \<and> x \<le> y}"
+
+lemma infinite_cartesian_product:
+assumes "infinite A" "infinite B"
+shows "infinite (A \<times> B)"
+proof
+ assume "finite (A \<times> B)"
+ from assms(1) have "A \<noteq> {}" by auto
+ with `finite (A \<times> B)` have "finite B" using finite_cartesian_productD2 by auto
+ with assms(2) show False by simp
+qed
+
+
+
+subsubsection {* First as well-orders *}
+
+
+lemma Field_natLeq: "Field natLeq = (UNIV::nat set)"
+by(unfold Field_def, auto)
+
+
+lemma natLeq_Refl: "Refl natLeq"
+unfolding refl_on_def Field_def by auto
+
+
+lemma natLeq_trans: "trans natLeq"
+unfolding trans_def by auto
+
+
+lemma natLeq_Preorder: "Preorder natLeq"
+unfolding preorder_on_def
+by (auto simp add: natLeq_Refl natLeq_trans)
+
+
+lemma natLeq_antisym: "antisym natLeq"
+unfolding antisym_def by auto
+
+
+lemma natLeq_Partial_order: "Partial_order natLeq"
+unfolding partial_order_on_def
+by (auto simp add: natLeq_Preorder natLeq_antisym)
+
+
+lemma natLeq_Total: "Total natLeq"
+unfolding total_on_def by auto
+
+
+lemma natLeq_Linear_order: "Linear_order natLeq"
+unfolding linear_order_on_def
+by (auto simp add: natLeq_Partial_order natLeq_Total)
+
+
+lemma natLeq_natLess_Id: "natLess = natLeq - Id"
+by auto
+
+
+lemma natLeq_Well_order: "Well_order natLeq"
+unfolding well_order_on_def
+using natLeq_Linear_order wf_less natLeq_natLess_Id by auto
+
+
+corollary natLeq_well_order_on: "well_order_on UNIV natLeq"
+using natLeq_Well_order Field_natLeq by auto
+
+
+lemma natLeq_wo_rel: "wo_rel natLeq"
+unfolding wo_rel_def using natLeq_Well_order .
+
+
+lemma natLeq_UNIV_ofilter: "wo_rel.ofilter natLeq UNIV"
+using natLeq_wo_rel Field_natLeq wo_rel.Field_ofilter[of natLeq] by auto
+
+
+lemma closed_nat_set_iff:
+assumes "\<forall>(m::nat) n. n \<in> A \<and> m \<le> n \<longrightarrow> m \<in> A"
+shows "A = UNIV \<or> (\<exists>n. A = {0 ..< n})"
+proof-
+ {assume "A \<noteq> UNIV" hence "\<exists>n. n \<notin> A" by blast
+ moreover obtain n where n_def: "n = (LEAST n. n \<notin> A)" by blast
+ ultimately have 1: "n \<notin> A \<and> (\<forall>m. m < n \<longrightarrow> m \<in> A)"
+ using LeastI_ex[of "\<lambda> n. n \<notin> A"] n_def Least_le[of "\<lambda> n. n \<notin> A"] by fastforce
+ have "A = {0 ..< n}"
+ proof(auto simp add: 1)
+ fix m assume *: "m \<in> A"
+ {assume "n \<le> m" with assms * have "n \<in> A" by blast
+ hence False using 1 by auto
+ }
+ thus "m < n" by fastforce
+ qed
+ hence "\<exists>n. A = {0 ..< n}" by blast
+ }
+ thus ?thesis by blast
+qed
+
+
+lemma Field_natLeq_on: "Field (natLeq_on n) = {0 ..< n}"
+unfolding Field_def by auto
+
+
+lemma natLeq_underS_less: "rel.underS natLeq n = {0 ..< n}"
+unfolding rel.underS_def by auto
+
+
+lemma Restr_natLeq: "Restr natLeq {0 ..< n} = natLeq_on n"
+by auto
+
+
+lemma Restr_natLeq2:
+"Restr natLeq (rel.underS natLeq n) = natLeq_on n"
+by (auto simp add: Restr_natLeq natLeq_underS_less)
+
+
+lemma natLeq_on_Well_order: "Well_order(natLeq_on n)"
+using Restr_natLeq[of n] natLeq_Well_order
+ Well_order_Restr[of natLeq "{0..<n}"] by auto
+
+
+corollary natLeq_on_well_order_on: "well_order_on {0 ..< n} (natLeq_on n)"
+using natLeq_on_Well_order Field_natLeq_on by auto
+
+
+lemma natLeq_on_wo_rel: "wo_rel(natLeq_on n)"
+unfolding wo_rel_def using natLeq_on_Well_order .
+
+
+lemma natLeq_on_ofilter_less_eq:
+"n \<le> m \<Longrightarrow> wo_rel.ofilter (natLeq_on m) {0 ..< n}"
+by (auto simp add: natLeq_on_wo_rel wo_rel.ofilter_def,
+ simp add: Field_natLeq_on, unfold rel.under_def, auto)
+
+
+lemma natLeq_on_ofilter_iff:
+"wo_rel.ofilter (natLeq_on m) A = (\<exists>n \<le> m. A = {0 ..< n})"
+proof(rule iffI)
+ assume *: "wo_rel.ofilter (natLeq_on m) A"
+ hence 1: "A \<le> {0..<m}"
+ by (auto simp add: natLeq_on_wo_rel wo_rel.ofilter_def rel.under_def Field_natLeq_on)
+ hence "\<forall>n1 n2. n2 \<in> A \<and> n1 \<le> n2 \<longrightarrow> n1 \<in> A"
+ using * by(fastforce simp add: natLeq_on_wo_rel wo_rel.ofilter_def rel.under_def)
+ hence "A = UNIV \<or> (\<exists>n. A = {0 ..< n})" using closed_nat_set_iff by blast
+ thus "\<exists>n \<le> m. A = {0 ..< n}" using 1 atLeastLessThan_less_eq by blast
+next
+ assume "(\<exists>n\<le>m. A = {0 ..< n})"
+ thus "wo_rel.ofilter (natLeq_on m) A" by (auto simp add: natLeq_on_ofilter_less_eq)
+qed
+
+
+
+subsubsection {* Then as cardinals *}
+
+
+lemma natLeq_Card_order: "Card_order natLeq"
+proof(auto simp add: natLeq_Well_order
+ Card_order_iff_Restr_underS Restr_natLeq2, simp add: Field_natLeq)
+ fix n have "finite(Field (natLeq_on n))"
+ unfolding Field_natLeq_on by auto
+ moreover have "infinite(UNIV::nat set)" by auto
+ ultimately show "natLeq_on n <o |UNIV::nat set|"
+ using finite_ordLess_infinite[of "natLeq_on n" "|UNIV::nat set|"]
+ Field_card_of[of "UNIV::nat set"]
+ card_of_Well_order[of "UNIV::nat set"] natLeq_on_Well_order[of n] by auto
+qed
+
+
+corollary card_of_Field_natLeq:
+"|Field natLeq| =o natLeq"
+using Field_natLeq natLeq_Card_order Card_order_iff_ordIso_card_of[of natLeq]
+ ordIso_symmetric[of natLeq] by blast
+
+
+corollary card_of_nat:
+"|UNIV::nat set| =o natLeq"
+using Field_natLeq card_of_Field_natLeq by auto
+
+
+corollary infinite_iff_natLeq_ordLeq:
+"infinite A = ( natLeq \<le>o |A| )"
+using infinite_iff_card_of_nat[of A] card_of_nat
+ ordIso_ordLeq_trans ordLeq_ordIso_trans ordIso_symmetric by blast
+
+
+corollary finite_iff_ordLess_natLeq:
+"finite A = ( |A| <o natLeq)"
+using infinite_iff_natLeq_ordLeq not_ordLeq_iff_ordLess
+ card_of_Well_order natLeq_Well_order by blast
+
+
+lemma ordIso_natLeq_on_imp_finite:
+"|A| =o natLeq_on n \<Longrightarrow> finite A"
+unfolding ordIso_def iso_def[abs_def]
+by (auto simp: Field_natLeq_on bij_betw_finite Field_card_of)
+
+
+lemma natLeq_on_Card_order: "Card_order (natLeq_on n)"
+proof(unfold card_order_on_def,
+ auto simp add: natLeq_on_Well_order, simp add: Field_natLeq_on)
+ fix r assume "well_order_on {0..<n} r"
+ thus "natLeq_on n \<le>o r"
+ using finite_atLeastLessThan natLeq_on_well_order_on
+ finite_well_order_on_ordIso ordIso_iff_ordLeq by blast
+qed
+
+
+corollary card_of_Field_natLeq_on:
+"|Field (natLeq_on n)| =o natLeq_on n"
+using Field_natLeq_on natLeq_on_Card_order
+ Card_order_iff_ordIso_card_of[of "natLeq_on n"]
+ ordIso_symmetric[of "natLeq_on n"] by blast
+
+
+corollary card_of_less:
+"|{0 ..< n}| =o natLeq_on n"
+using Field_natLeq_on card_of_Field_natLeq_on by auto
+
+
+lemma natLeq_on_ordLeq_less_eq:
+"((natLeq_on m) \<le>o (natLeq_on n)) = (m \<le> n)"
+proof
+ assume "natLeq_on m \<le>o natLeq_on n"
+ then obtain f where "inj_on f {0..<m} \<and> f ` {0..<m} \<le> {0..<n}"
+ using Field_natLeq_on[of m] Field_natLeq_on[of n]
+ unfolding ordLeq_def using embed_inj_on[of "natLeq_on m" "natLeq_on n"]
+ embed_Field[of "natLeq_on m" "natLeq_on n"] using natLeq_on_Well_order[of m] by fastforce
+ thus "m \<le> n" using atLeastLessThan_less_eq2 by blast
+next
+ assume "m \<le> n"
+ hence "inj_on id {0..<m} \<and> id ` {0..<m} \<le> {0..<n}" unfolding inj_on_def by auto
+ hence "|{0..<m}| \<le>o |{0..<n}|" using card_of_ordLeq by blast
+ thus "natLeq_on m \<le>o natLeq_on n"
+ using card_of_less ordIso_ordLeq_trans ordLeq_ordIso_trans ordIso_symmetric by blast
+qed
+
+
+lemma natLeq_on_ordLeq_less:
+"((natLeq_on m) <o (natLeq_on n)) = (m < n)"
+using not_ordLeq_iff_ordLess[of "natLeq_on m" "natLeq_on n"]
+natLeq_on_Well_order natLeq_on_ordLeq_less_eq by auto
+
+
+
+subsubsection {* "Backwards compatibility" with the numeric cardinal operator for finite sets *}
+
+
+lemma finite_card_of_iff_card2:
+assumes FIN: "finite A" and FIN': "finite B"
+shows "( |A| \<le>o |B| ) = (card A \<le> card B)"
+using assms card_of_ordLeq[of A B] inj_on_iff_card_le[of A B] by blast
+
+
+lemma finite_imp_card_of_natLeq_on:
+assumes "finite A"
+shows "|A| =o natLeq_on (card A)"
+proof-
+ obtain h where "bij_betw h A {0 ..< card A}"
+ using assms ex_bij_betw_finite_nat by blast
+ thus ?thesis using card_of_ordIso card_of_less ordIso_equivalence by blast
+qed
+
+
+lemma finite_iff_card_of_natLeq_on:
+"finite A = (\<exists>n. |A| =o natLeq_on n)"
+using finite_imp_card_of_natLeq_on[of A]
+by(auto simp add: ordIso_natLeq_on_imp_finite)
+
+
+
+subsection {* The successor of a cardinal *}
+
+
+text{* First we define @{text "isCardSuc r r'"}, the notion of @{text "r'"}
+being a successor cardinal of @{text "r"}. Although the definition does
+not require @{text "r"} to be a cardinal, only this case will be meaningful. *}
+
+
+definition isCardSuc :: "'a rel \<Rightarrow> 'a set rel \<Rightarrow> bool"
+where
+"isCardSuc r r' \<equiv>
+ Card_order r' \<and> r <o r' \<and>
+ (\<forall>(r''::'a set rel). Card_order r'' \<and> r <o r'' \<longrightarrow> r' \<le>o r'')"
+
+
+text{* Now we introduce the cardinal-successor operator @{text "cardSuc"},
+by picking {\em some} cardinal-order relation fulfilling @{text "isCardSuc"}.
+Again, the picked item shall be proved unique up to order-isomorphism. *}
+
+
+definition cardSuc :: "'a rel \<Rightarrow> 'a set rel"
+where
+"cardSuc r \<equiv> SOME r'. isCardSuc r r'"
+
+
+lemma exists_minim_Card_order:
+"\<lbrakk>R \<noteq> {}; \<forall>r \<in> R. Card_order r\<rbrakk> \<Longrightarrow> \<exists>r \<in> R. \<forall>r' \<in> R. r \<le>o r'"
+unfolding card_order_on_def using exists_minim_Well_order by blast
+
+
+lemma exists_isCardSuc:
+assumes "Card_order r"
+shows "\<exists>r'. isCardSuc r r'"
+proof-
+ let ?R = "{(r'::'a set rel). Card_order r' \<and> r <o r'}"
+ have "|Pow(Field r)| \<in> ?R \<and> (\<forall>r \<in> ?R. Card_order r)" using assms
+ by (simp add: card_of_Card_order Card_order_Pow)
+ then obtain r where "r \<in> ?R \<and> (\<forall>r' \<in> ?R. r \<le>o r')"
+ using exists_minim_Card_order[of ?R] by blast
+ thus ?thesis unfolding isCardSuc_def by auto
+qed
+
+
+lemma cardSuc_isCardSuc:
+assumes "Card_order r"
+shows "isCardSuc r (cardSuc r)"
+unfolding cardSuc_def using assms
+by (simp add: exists_isCardSuc someI_ex)
+
+
+lemma cardSuc_Card_order:
+"Card_order r \<Longrightarrow> Card_order(cardSuc r)"
+using cardSuc_isCardSuc unfolding isCardSuc_def by blast
+
+
+lemma cardSuc_greater:
+"Card_order r \<Longrightarrow> r <o cardSuc r"
+using cardSuc_isCardSuc unfolding isCardSuc_def by blast
+
+
+lemma cardSuc_ordLeq:
+"Card_order r \<Longrightarrow> r \<le>o cardSuc r"
+using cardSuc_greater ordLeq_iff_ordLess_or_ordIso by blast
+
+
+text{* The minimality property of @{text "cardSuc"} originally present in its definition
+is local to the type @{text "'a set rel"}, i.e., that of @{text "cardSuc r"}: *}
+
+lemma cardSuc_least_aux:
+"\<lbrakk>Card_order (r::'a rel); Card_order (r'::'a set rel); r <o r'\<rbrakk> \<Longrightarrow> cardSuc r \<le>o r'"
+using cardSuc_isCardSuc unfolding isCardSuc_def by blast
+
+
+text{* But from this we can infer general minimality: *}
+
+lemma cardSuc_least:
+assumes CARD: "Card_order r" and CARD': "Card_order r'" and LESS: "r <o r'"
+shows "cardSuc r \<le>o r'"
+proof-
+ let ?p = "cardSuc r"
+ have 0: "Well_order ?p \<and> Well_order r'"
+ using assms cardSuc_Card_order unfolding card_order_on_def by blast
+ {assume "r' <o ?p"
+ then obtain r'' where 1: "Field r'' < Field ?p" and 2: "r' =o r'' \<and> r'' <o ?p"
+ using internalize_ordLess[of r' ?p] by blast
+ (* *)
+ have "Card_order r''" using CARD' Card_order_ordIso2 2 by blast
+ moreover have "r <o r''" using LESS 2 ordLess_ordIso_trans by blast
+ ultimately have "?p \<le>o r''" using cardSuc_least_aux CARD by blast
+ hence False using 2 not_ordLess_ordLeq by blast
+ }
+ thus ?thesis using 0 ordLess_or_ordLeq by blast
+qed
+
+
+lemma cardSuc_ordLess_ordLeq:
+assumes CARD: "Card_order r" and CARD': "Card_order r'"
+shows "(r <o r') = (cardSuc r \<le>o r')"
+proof(auto simp add: assms cardSuc_least)
+ assume "cardSuc r \<le>o r'"
+ thus "r <o r'" using assms cardSuc_greater ordLess_ordLeq_trans by blast
+qed
+
+
+lemma cardSuc_ordLeq_ordLess:
+assumes CARD: "Card_order r" and CARD': "Card_order r'"
+shows "(r' <o cardSuc r) = (r' \<le>o r)"
+proof-
+ have "Well_order r \<and> Well_order r'"
+ using assms unfolding card_order_on_def by auto
+ moreover have "Well_order(cardSuc r)"
+ using assms cardSuc_Card_order card_order_on_def by blast
+ ultimately show ?thesis
+ using assms cardSuc_ordLess_ordLeq[of r r']
+ not_ordLeq_iff_ordLess[of r r'] not_ordLeq_iff_ordLess[of r' "cardSuc r"] by blast
+qed
+
+
+lemma cardSuc_mono_ordLeq:
+assumes CARD: "Card_order r" and CARD': "Card_order r'"
+shows "(cardSuc r \<le>o cardSuc r') = (r \<le>o r')"
+using assms cardSuc_ordLeq_ordLess cardSuc_ordLess_ordLeq cardSuc_Card_order by blast
+
+
+lemma cardSuc_invar_ordIso:
+assumes CARD: "Card_order r" and CARD': "Card_order r'"
+shows "(cardSuc r =o cardSuc r') = (r =o r')"
+proof-
+ have 0: "Well_order r \<and> Well_order r' \<and> Well_order(cardSuc r) \<and> Well_order(cardSuc r')"
+ using assms by (simp add: card_order_on_well_order_on cardSuc_Card_order)
+ thus ?thesis
+ using ordIso_iff_ordLeq[of r r'] ordIso_iff_ordLeq
+ using cardSuc_mono_ordLeq[of r r'] cardSuc_mono_ordLeq[of r' r] assms by blast
+qed
+
+
+lemma cardSuc_natLeq_on_Suc:
+"cardSuc(natLeq_on n) =o natLeq_on(Suc n)"
+proof-
+ obtain r r' p where r_def: "r = natLeq_on n" and
+ r'_def: "r' = cardSuc(natLeq_on n)" and
+ p_def: "p = natLeq_on(Suc n)" by blast
+ (* Preliminary facts: *)
+ have CARD: "Card_order r \<and> Card_order r' \<and> Card_order p" unfolding r_def r'_def p_def
+ using cardSuc_ordLess_ordLeq natLeq_on_Card_order cardSuc_Card_order by blast
+ hence WELL: "Well_order r \<and> Well_order r' \<and> Well_order p"
+ unfolding card_order_on_def by force
+ have FIELD: "Field r = {0..<n} \<and> Field p = {0..<(Suc n)}"
+ unfolding r_def p_def Field_natLeq_on by simp
+ hence FIN: "finite (Field r)" by force
+ have "r <o r'" using CARD unfolding r_def r'_def using cardSuc_greater by blast
+ hence "|Field r| <o r'" using CARD card_of_Field_ordIso ordIso_ordLess_trans by blast
+ hence LESS: "|Field r| <o |Field r'|"
+ using CARD card_of_Field_ordIso ordLess_ordIso_trans ordIso_symmetric by blast
+ (* Main proof: *)
+ have "r' \<le>o p" using CARD unfolding r_def r'_def p_def
+ using natLeq_on_ordLeq_less cardSuc_ordLess_ordLeq by blast
+ moreover have "p \<le>o r'"
+ proof-
+ {assume "r' <o p"
+ then obtain f where 0: "embedS r' p f" unfolding ordLess_def by force
+ let ?q = "Restr p (f ` Field r')"
+ have 1: "embed r' p f" using 0 unfolding embedS_def by force
+ hence 2: "f ` Field r' < {0..<(Suc n)}"
+ using WELL FIELD 0 by (auto simp add: embedS_iff)
+ have "wo_rel.ofilter p (f ` Field r')" using embed_Field_ofilter 1 WELL by blast
+ then obtain m where "m \<le> Suc n" and 3: "f ` (Field r') = {0..<m}"
+ unfolding p_def by (auto simp add: natLeq_on_ofilter_iff)
+ hence 4: "m \<le> n" using 2 by force
+ (* *)
+ have "bij_betw f (Field r') (f ` (Field r'))"
+ using 1 WELL embed_inj_on unfolding bij_betw_def by force
+ moreover have "finite(f ` (Field r'))" using 3 finite_atLeastLessThan[of 0 m] by force
+ ultimately have 5: "finite (Field r') \<and> card(Field r') = card (f ` (Field r'))"
+ using bij_betw_same_card bij_betw_finite by metis
+ hence "card(Field r') \<le> card(Field r)" using 3 4 FIELD by force
+ hence "|Field r'| \<le>o |Field r|" using FIN 5 finite_card_of_iff_card2 by blast
+ hence False using LESS not_ordLess_ordLeq by auto
+ }
+ thus ?thesis using WELL CARD by (fastforce simp: not_ordLess_iff_ordLeq)
+ qed
+ ultimately show ?thesis using ordIso_iff_ordLeq unfolding r'_def p_def by blast
+qed
+
+
+lemma card_of_cardSuc_finite:
+"finite(Field(cardSuc |A| )) = finite A"
+proof
+ assume *: "finite (Field (cardSuc |A| ))"
+ have 0: "|Field(cardSuc |A| )| =o cardSuc |A|"
+ using card_of_Card_order cardSuc_Card_order card_of_Field_ordIso by blast
+ hence "|A| \<le>o |Field(cardSuc |A| )|"
+ using card_of_Card_order[of A] cardSuc_ordLeq[of "|A|"] ordIso_symmetric
+ ordLeq_ordIso_trans by blast
+ thus "finite A" using * card_of_ordLeq_finite by blast
+next
+ assume "finite A"
+ then obtain n where "|A| =o natLeq_on n" using finite_iff_card_of_natLeq_on by blast
+ hence "cardSuc |A| =o cardSuc(natLeq_on n)"
+ using card_of_Card_order cardSuc_invar_ordIso natLeq_on_Card_order by blast
+ hence "cardSuc |A| =o natLeq_on(Suc n)"
+ using cardSuc_natLeq_on_Suc ordIso_transitive by blast
+ hence "cardSuc |A| =o |{0..<(Suc n)}|" using card_of_less ordIso_equivalence by blast
+ moreover have "|Field (cardSuc |A| ) | =o cardSuc |A|"
+ using card_of_Field_ordIso cardSuc_Card_order card_of_Card_order by blast
+ ultimately have "|Field (cardSuc |A| ) | =o |{0..<(Suc n)}|"
+ using ordIso_equivalence by blast
+ thus "finite (Field (cardSuc |A| ))"
+ using card_of_ordIso_finite finite_atLeastLessThan by blast
+qed
+
+
+lemma cardSuc_finite:
+assumes "Card_order r"
+shows "finite (Field (cardSuc r)) = finite (Field r)"
+proof-
+ let ?A = "Field r"
+ have "|?A| =o r" using assms by (simp add: card_of_Field_ordIso)
+ hence "cardSuc |?A| =o cardSuc r" using assms
+ by (simp add: card_of_Card_order cardSuc_invar_ordIso)
+ moreover have "|Field (cardSuc |?A| ) | =o cardSuc |?A|"
+ by (simp add: card_of_card_order_on Field_card_of card_of_Field_ordIso cardSuc_Card_order)
+ moreover
+ {have "|Field (cardSuc r) | =o cardSuc r"
+ using assms by (simp add: card_of_Field_ordIso cardSuc_Card_order)
+ hence "cardSuc r =o |Field (cardSuc r) |"
+ using ordIso_symmetric by blast
+ }
+ ultimately have "|Field (cardSuc |?A| ) | =o |Field (cardSuc r) |"
+ using ordIso_transitive by blast
+ hence "finite (Field (cardSuc |?A| )) = finite (Field (cardSuc r))"
+ using card_of_ordIso_finite by blast
+ thus ?thesis by (simp only: card_of_cardSuc_finite)
+qed
+
+
+lemma card_of_Plus_ordLeq_infinite_Field:
+assumes r: "infinite (Field r)" and A: "|A| \<le>o r" and B: "|B| \<le>o r"
+and c: "Card_order r"
+shows "|A <+> B| \<le>o r"
+proof-
+ let ?r' = "cardSuc r"
+ have "Card_order ?r' \<and> infinite (Field ?r')" using assms
+ by (simp add: cardSuc_Card_order cardSuc_finite)
+ moreover have "|A| <o ?r'" and "|B| <o ?r'" using A B c
+ by (auto simp: card_of_card_order_on Field_card_of cardSuc_ordLeq_ordLess)
+ ultimately have "|A <+> B| <o ?r'"
+ using card_of_Plus_ordLess_infinite_Field by blast
+ thus ?thesis using c r
+ by (simp add: card_of_card_order_on Field_card_of cardSuc_ordLeq_ordLess)
+qed
+
+
+lemma card_of_Un_ordLeq_infinite_Field:
+assumes C: "infinite (Field r)" and A: "|A| \<le>o r" and B: "|B| \<le>o r"
+and "Card_order r"
+shows "|A Un B| \<le>o r"
+using assms card_of_Plus_ordLeq_infinite_Field card_of_Un_Plus_ordLeq
+ordLeq_transitive by blast
+
+
+
+subsection {* Regular cardinals *}
+
+
+definition cofinal where
+"cofinal A r \<equiv>
+ ALL a : Field r. EX b : A. a \<noteq> b \<and> (a,b) : r"
+
+
+definition regular where
+"regular r \<equiv>
+ ALL K. K \<le> Field r \<and> cofinal K r \<longrightarrow> |K| =o r"
+
+
+definition relChain where
+"relChain r As \<equiv>
+ ALL i j. (i,j) \<in> r \<longrightarrow> As i \<le> As j"
+
+lemma regular_UNION:
+assumes r: "Card_order r" "regular r"
+and As: "relChain r As"
+and Bsub: "B \<le> (UN i : Field r. As i)"
+and cardB: "|B| <o r"
+shows "EX i : Field r. B \<le> As i"
+proof-
+ let ?phi = "%b j. j : Field r \<and> b : As j"
+ have "ALL b : B. EX j. ?phi b j" using Bsub by blast
+ then obtain f where f: "!! b. b : B \<Longrightarrow> ?phi b (f b)"
+ using bchoice[of B ?phi] by blast
+ let ?K = "f ` B"
+ {assume 1: "!! i. i : Field r \<Longrightarrow> ~ B \<le> As i"
+ have 2: "cofinal ?K r"
+ unfolding cofinal_def proof auto
+ fix i assume i: "i : Field r"
+ with 1 obtain b where b: "b : B \<and> b \<notin> As i" by blast
+ hence "i \<noteq> f b \<and> ~ (f b,i) : r"
+ using As f unfolding relChain_def by auto
+ hence "i \<noteq> f b \<and> (i, f b) : r" using r
+ unfolding card_order_on_def well_order_on_def linear_order_on_def
+ total_on_def using i f b by auto
+ with b show "\<exists>b\<in>B. i \<noteq> f b \<and> (i, f b) \<in> r" by blast
+ qed
+ moreover have "?K \<le> Field r" using f by blast
+ ultimately have "|?K| =o r" using 2 r unfolding regular_def by blast
+ moreover
+ {
+ have "|?K| <=o |B|" using card_of_image .
+ hence "|?K| <o r" using cardB ordLeq_ordLess_trans by blast
+ }
+ ultimately have False using not_ordLess_ordIso by blast
+ }
+ thus ?thesis by blast
+qed
+
+
+lemma infinite_cardSuc_regular:
+assumes r_inf: "infinite (Field r)" and r_card: "Card_order r"
+shows "regular (cardSuc r)"
+proof-
+ let ?r' = "cardSuc r"
+ have r': "Card_order ?r'"
+ "!! p. Card_order p \<longrightarrow> (p \<le>o r) = (p <o ?r')"
+ using r_card by (auto simp: cardSuc_Card_order cardSuc_ordLeq_ordLess)
+ show ?thesis
+ unfolding regular_def proof auto
+ fix K assume 1: "K \<le> Field ?r'" and 2: "cofinal K ?r'"
+ hence "|K| \<le>o |Field ?r'|" by (simp only: card_of_mono1)
+ also have 22: "|Field ?r'| =o ?r'"
+ using r' by (simp add: card_of_Field_ordIso[of ?r'])
+ finally have "|K| \<le>o ?r'" .
+ moreover
+ {let ?L = "UN j : K. rel.underS ?r' j"
+ let ?J = "Field r"
+ have rJ: "r =o |?J|"
+ using r_card card_of_Field_ordIso ordIso_symmetric by blast
+ assume "|K| <o ?r'"
+ hence "|K| <=o r" using r' card_of_Card_order[of K] by blast
+ hence "|K| \<le>o |?J|" using rJ ordLeq_ordIso_trans by blast
+ moreover
+ {have "ALL j : K. |rel.underS ?r' j| <o ?r'"
+ using r' 1 by (auto simp: card_of_underS)
+ hence "ALL j : K. |rel.underS ?r' j| \<le>o r"
+ using r' card_of_Card_order by blast
+ hence "ALL j : K. |rel.underS ?r' j| \<le>o |?J|"
+ using rJ ordLeq_ordIso_trans by blast
+ }
+ ultimately have "|?L| \<le>o |?J|"
+ using r_inf card_of_UNION_ordLeq_infinite by blast
+ hence "|?L| \<le>o r" using rJ ordIso_symmetric ordLeq_ordIso_trans by blast
+ hence "|?L| <o ?r'" using r' card_of_Card_order by blast
+ moreover
+ {
+ have "Field ?r' \<le> ?L"
+ using 2 unfolding rel.underS_def cofinal_def by auto
+ hence "|Field ?r'| \<le>o |?L|" by (simp add: card_of_mono1)
+ hence "?r' \<le>o |?L|"
+ using 22 ordIso_ordLeq_trans ordIso_symmetric by blast
+ }
+ ultimately have "|?L| <o |?L|" using ordLess_ordLeq_trans by blast
+ hence False using ordLess_irreflexive by blast
+ }
+ ultimately show "|K| =o ?r'"
+ unfolding ordLeq_iff_ordLess_or_ordIso by blast
+ qed
+qed
+
+lemma cardSuc_UNION:
+assumes r: "Card_order r" and "infinite (Field r)"
+and As: "relChain (cardSuc r) As"
+and Bsub: "B \<le> (UN i : Field (cardSuc r). As i)"
+and cardB: "|B| <=o r"
+shows "EX i : Field (cardSuc r). B \<le> As i"
+proof-
+ let ?r' = "cardSuc r"
+ have "Card_order ?r' \<and> |B| <o ?r'"
+ using r cardB cardSuc_ordLeq_ordLess cardSuc_Card_order
+ card_of_Card_order by blast
+ moreover have "regular ?r'"
+ using assms by(simp add: infinite_cardSuc_regular)
+ ultimately show ?thesis
+ using As Bsub cardB regular_UNION by blast
+qed
+
+
+subsection {* Others *}
+
+(* FIXME: finitte ~> finite? *)
+lemma card_of_infinite_diff_finitte:
+assumes "infinite A" and "finite B"
+shows "|A - B| =o |A|"
+by (metis assms card_of_Un_diff_infinite finite_ordLess_infinite2)
+
+(* function space *)
+definition Func where
+"Func A B \<equiv>
+ {f. (\<forall> a. f a \<noteq> None \<longleftrightarrow> a \<in> A) \<and> (\<forall> a \<in> A. case f a of Some b \<Rightarrow> b \<in> B |None \<Rightarrow> True)}"
+
+lemma Func_empty:
+"Func {} B = {empty}"
+unfolding Func_def by auto
+
+lemma Func_elim:
+assumes "g \<in> Func A B" and "a \<in> A"
+shows "\<exists> b. b \<in> B \<and> g a = Some b"
+using assms unfolding Func_def by (cases "g a") force+
+
+definition curr where
+"curr A f \<equiv> \<lambda> a. if a \<in> A then Some (\<lambda> b. f (a,b)) else None"
+
+lemma curr_in:
+assumes f: "f \<in> Func (A <*> B) C"
+shows "curr A f \<in> Func A (Func B C)"
+using assms unfolding curr_def Func_def by auto
+
+lemma curr_inj:
+assumes "f1 \<in> Func (A <*> B) C" and "f2 \<in> Func (A <*> B) C"
+shows "curr A f1 = curr A f2 \<longleftrightarrow> f1 = f2"
+proof safe
+ assume c: "curr A f1 = curr A f2"
+ show "f1 = f2"
+ proof (rule ext, clarify)
+ fix a b show "f1 (a, b) = f2 (a, b)"
+ proof (cases "(a,b) \<in> A <*> B")
+ case False
+ thus ?thesis using assms unfolding Func_def
+ apply(cases "f1 (a,b)") apply(cases "f2 (a,b)", fastforce, fastforce)
+ apply(cases "f2 (a,b)") by auto
+ next
+ case True hence a: "a \<in> A" and b: "b \<in> B" by auto
+ thus ?thesis
+ using c unfolding curr_def fun_eq_iff
+ apply(elim allE[of _ a]) apply simp unfolding fun_eq_iff by auto
+ qed
+ qed
+qed
+
+lemma curr_surj:
+assumes "g \<in> Func A (Func B C)"
+shows "\<exists> f \<in> Func (A <*> B) C. curr A f = g"
+proof
+ let ?f = "\<lambda> ab. case g (fst ab) of None \<Rightarrow> None | Some g1 \<Rightarrow> g1 (snd ab)"
+ show "curr A ?f = g"
+ proof (rule ext)
+ fix a show "curr A ?f a = g a"
+ proof (cases "a \<in> A")
+ case False
+ hence "g a = None" using assms unfolding Func_def by auto
+ thus ?thesis unfolding curr_def using False by simp
+ next
+ case True
+ obtain g1 where "g1 \<in> Func B C" and "g a = Some g1"
+ using assms using Func_elim[OF assms True] by blast
+ thus ?thesis using True unfolding curr_def by auto
+ qed
+ qed
+ show "?f \<in> Func (A <*> B) C"
+ unfolding Func_def mem_Collect_eq proof(intro conjI allI ballI)
+ fix ab show "?f ab \<noteq> None \<longleftrightarrow> ab \<in> A \<times> B"
+ proof(cases "g (fst ab)")
+ case None
+ hence "fst ab \<notin> A" using assms unfolding Func_def by force
+ thus ?thesis using None by auto
+ next
+ case (Some g1)
+ hence fst: "fst ab \<in> A" and g1: "g1 \<in> Func B C"
+ using assms unfolding Func_def[of A] by force+
+ hence "?f ab \<noteq> None \<longleftrightarrow> g1 (snd ab) \<noteq> None" using Some by auto
+ also have "... \<longleftrightarrow> snd ab \<in> B" using g1 unfolding Func_def by auto
+ also have "... \<longleftrightarrow> ab \<in> A \<times> B" using fst by (cases ab, auto)
+ finally show ?thesis .
+ qed
+ next
+ fix ab assume ab: "ab \<in> A \<times> B"
+ hence "fst ab \<in> A" and "snd ab \<in> B" by(cases ab, auto)
+ then obtain g1 where "g1 \<in> Func B C" and "g (fst ab) = Some g1"
+ using assms using Func_elim[OF assms] by blast
+ thus "case ?f ab of Some c \<Rightarrow> c \<in> C |None \<Rightarrow> True"
+ unfolding Func_def by auto
+ qed
+qed
+
+(* FIXME: betwe ~> betw? *)
+lemma bij_betwe_curr:
+"bij_betw (curr A) (Func (A <*> B) C) (Func A (Func B C))"
+unfolding bij_betw_def inj_on_def image_def
+using curr_in curr_inj curr_surj by blast
+
+lemma card_of_Func_Times:
+"|Func (A <*> B) C| =o |Func A (Func B C)|"
+unfolding card_of_ordIso[symmetric]
+using bij_betwe_curr by blast
+
+definition Func_map where
+"Func_map B2 f1 f2 g b2 \<equiv>
+ if b2 \<in> B2 then case g (f2 b2) of None \<Rightarrow> None | Some a1 \<Rightarrow> Some (f1 a1)
+ else None"
+
+lemma Func_map:
+assumes g: "g \<in> Func A2 A1" and f1: "f1 ` A1 \<subseteq> B1" and f2: "f2 ` B2 \<subseteq> A2"
+shows "Func_map B2 f1 f2 g \<in> Func B2 B1"
+unfolding Func_def mem_Collect_eq proof(intro conjI allI ballI)
+ fix b2 show "Func_map B2 f1 f2 g b2 \<noteq> None \<longleftrightarrow> b2 \<in> B2"
+ proof(cases "b2 \<in> B2")
+ case True
+ hence "f2 b2 \<in> A2" using f2 by auto
+ then obtain a1 where "g (f2 b2) = Some a1" and "a1 \<in> A1"
+ using g unfolding Func_def by(cases "g (f2 b2)", fastforce+)
+ thus ?thesis unfolding Func_map_def using True by auto
+ qed(unfold Func_map_def, auto)
+next
+ fix b2 assume b2: "b2 \<in> B2"
+ hence "f2 b2 \<in> A2" using f2 by auto
+ then obtain a1 where "g (f2 b2) = Some a1" and "a1 \<in> A1"
+ using g unfolding Func_def by(cases "g (f2 b2)", fastforce+)
+ thus "case Func_map B2 f1 f2 g b2 of None \<Rightarrow> True | Some b1 \<Rightarrow> b1 \<in> B1"
+ unfolding Func_map_def using b2 f1 by auto
+qed
+
+lemma Func_map_empty:
+"Func_map B2 f1 f2 empty = empty"
+unfolding Func_map_def[abs_def] by (rule ext, auto)
+
+lemma Func_non_emp:
+assumes "B \<noteq> {}"
+shows "Func A B \<noteq> {}"
+proof-
+ obtain b where b: "b \<in> B" using assms by auto
+ hence "(\<lambda> a. if a \<in> A then Some b else None) \<in> Func A B"
+ unfolding Func_def by auto
+ thus ?thesis by blast
+qed
+
+lemma Func_is_emp:
+"Func A B = {} \<longleftrightarrow> A \<noteq> {} \<and> B = {}" (is "?L \<longleftrightarrow> ?R")
+proof
+ assume L: ?L
+ moreover {assume "A = {}" hence False using L Func_empty by auto}
+ moreover {assume "B \<noteq> {}" hence False using L Func_non_emp by metis}
+ ultimately show ?R by blast
+next
+ assume R: ?R
+ moreover
+ {fix f assume "f \<in> Func A B"
+ moreover obtain a where "a \<in> A" using R by blast
+ ultimately obtain b where "b \<in> B" unfolding Func_def by(cases "f a", force+)
+ with R have False by auto
+ }
+ thus ?L by blast
+qed
+
+lemma Func_map_surj:
+assumes B1: "f1 ` A1 = B1" and A2: "inj_on f2 B2" "f2 ` B2 \<subseteq> A2"
+and B2A2: "B2 = {} \<Longrightarrow> A2 = {}"
+shows "Func B2 B1 = Func_map B2 f1 f2 ` Func A2 A1"
+proof(cases "B2 = {}")
+ case True
+ thus ?thesis using B2A2 by (auto simp: Func_empty Func_map_empty)
+next
+ case False note B2 = False
+ show ?thesis
+proof safe
+ fix h assume h: "h \<in> Func B2 B1"
+ def j1 \<equiv> "inv_into A1 f1"
+ have "\<forall> a2 \<in> f2 ` B2. \<exists> b2. b2 \<in> B2 \<and> f2 b2 = a2" by blast
+ then obtain k where k: "\<forall> a2 \<in> f2 ` B2. k a2 \<in> B2 \<and> f2 (k a2) = a2" by metis
+ {fix b2 assume b2: "b2 \<in> B2"
+ hence "f2 (k (f2 b2)) = f2 b2" using k A2(2) by auto
+ moreover have "k (f2 b2) \<in> B2" using b2 A2(2) k by auto
+ ultimately have "k (f2 b2) = b2" using b2 A2(1) unfolding inj_on_def by blast
+ } note kk = this
+ obtain b22 where b22: "b22 \<in> B2" using B2 by auto
+ def j2 \<equiv> "\<lambda> a2. if a2 \<in> f2 ` B2 then k a2 else b22"
+ have j2A2: "j2 ` A2 \<subseteq> B2" unfolding j2_def using k b22 by auto
+ have j2: "\<And> b2. b2 \<in> B2 \<Longrightarrow> j2 (f2 b2) = b2"
+ using kk unfolding j2_def by auto
+ def g \<equiv> "Func_map A2 j1 j2 h"
+ have "Func_map B2 f1 f2 g = h"
+ proof (rule ext)
+ fix b2 show "Func_map B2 f1 f2 g b2 = h b2"
+ proof(cases "b2 \<in> B2")
+ case True
+ show ?thesis
+ proof (cases "h b2")
+ case (Some b1)
+ hence b1: "b1 \<in> f1 ` A1" using True h unfolding B1 Func_def by auto
+ show ?thesis
+ using Some True A2 f_inv_into_f[OF b1]
+ unfolding g_def Func_map_def j1_def j2[OF True] by auto
+ qed(insert A2 True j2[OF True], unfold g_def Func_map_def, auto)
+ qed(insert h, unfold Func_def Func_map_def, auto)
+ qed
+ moreover have "g \<in> Func A2 A1" unfolding g_def apply(rule Func_map[OF h])
+ using inv_into_into j2A2 B1 A2 inv_into_into
+ unfolding j1_def image_def by(force, force)
+ ultimately show "h \<in> Func_map B2 f1 f2 ` Func A2 A1"
+ unfolding Func_map_def[abs_def] unfolding image_def by auto
+qed(insert B1 Func_map[OF _ _ A2(2)], auto)
+qed
+
+(* partial-function space: *)
+definition Pfunc where
+"Pfunc A B \<equiv>
+ {f. (\<forall>a. f a \<noteq> None \<longrightarrow> a \<in> A) \<and>
+ (\<forall>a. case f a of None \<Rightarrow> True | Some b \<Rightarrow> b \<in> B)}"
+
+lemma Func_Pfunc:
+"Func A B \<subseteq> Pfunc A B"
+unfolding Func_def Pfunc_def by auto
+
+lemma Pfunc_Func:
+"Pfunc A B = (\<Union> A' \<in> Pow A. Func A' B)"
+proof safe
+ fix f assume f: "f \<in> Pfunc A B"
+ show "f \<in> (\<Union>A'\<in>Pow A. Func A' B)"
+ proof (intro UN_I)
+ let ?A' = "{a. f a \<noteq> None}"
+ show "?A' \<in> Pow A" using f unfolding Pow_def Pfunc_def by auto
+ show "f \<in> Func ?A' B" using f unfolding Func_def Pfunc_def by auto
+ qed
+next
+ fix f A' assume "f \<in> Func A' B" and "A' \<subseteq> A"
+ thus "f \<in> Pfunc A B" unfolding Func_def Pfunc_def by auto
+qed
+
+lemma card_of_Pow_Func:
+"|Pow A| =o |Func A (UNIV::bool set)|"
+proof-
+ def F \<equiv> "\<lambda> A' a. if a \<in> A then (if a \<in> A' then Some True else Some False)
+ else None"
+ have "bij_betw F (Pow A) (Func A (UNIV::bool set))"
+ unfolding bij_betw_def inj_on_def proof (intro ballI impI conjI)
+ fix A1 A2 assume A1: "A1 \<in> Pow A" and A2: "A2 \<in> Pow A" and eq: "F A1 = F A2"
+ show "A1 = A2"
+ proof-
+ {fix a
+ have "a \<in> A1 \<longleftrightarrow> F A1 a = Some True" using A1 unfolding F_def Pow_def by auto
+ also have "... \<longleftrightarrow> F A2 a = Some True" unfolding eq ..
+ also have "... \<longleftrightarrow> a \<in> A2" using A2 unfolding F_def Pow_def by auto
+ finally have "a \<in> A1 \<longleftrightarrow> a \<in> A2" .
+ }
+ thus ?thesis by auto
+ qed
+ next
+ show "F ` Pow A = Func A UNIV"
+ proof safe
+ fix f assume f: "f \<in> Func A (UNIV::bool set)"
+ show "f \<in> F ` Pow A" unfolding image_def mem_Collect_eq proof(intro bexI)
+ let ?A1 = "{a \<in> A. f a = Some True}"
+ show "f = F ?A1" unfolding F_def apply(rule ext)
+ using f unfolding Func_def mem_Collect_eq by (auto,force)
+ qed auto
+ qed(unfold Func_def mem_Collect_eq F_def, auto)
+ qed
+ thus ?thesis unfolding card_of_ordIso[symmetric] by blast
+qed
+
+lemma card_of_Func_mono:
+fixes A1 A2 :: "'a set" and B :: "'b set"
+assumes A12: "A1 \<subseteq> A2" and B: "B \<noteq> {}"
+shows "|Func A1 B| \<le>o |Func A2 B|"
+proof-
+ obtain bb where bb: "bb \<in> B" using B by auto
+ def F \<equiv> "\<lambda> (f1::'a \<Rightarrow> 'b option) a. if a \<in> A2 then (if a \<in> A1 then f1 a else Some bb)
+ else None"
+ show ?thesis unfolding card_of_ordLeq[symmetric] proof(intro exI[of _ F] conjI)
+ show "inj_on F (Func A1 B)" unfolding inj_on_def proof safe
+ fix f g assume f: "f \<in> Func A1 B" and g: "g \<in> Func A1 B" and eq: "F f = F g"
+ show "f = g"
+ proof(rule ext)
+ fix a show "f a = g a"
+ proof(cases "a \<in> A1")
+ case True
+ thus ?thesis using eq A12 unfolding F_def fun_eq_iff
+ by (elim allE[of _ a]) auto
+ qed(insert f g, unfold Func_def, fastforce)
+ qed
+ qed
+ qed(insert bb, unfold Func_def F_def, force)
+qed
+
+lemma card_of_Pfunc_Pow_Func:
+assumes "B \<noteq> {}"
+shows "|Pfunc A B| \<le>o |Pow A <*> Func A B|"
+proof-
+ have "|Pfunc A B| =o |\<Union> A' \<in> Pow A. Func A' B|" (is "_ =o ?K")
+ unfolding Pfunc_Func by(rule card_of_refl)
+ also have "?K \<le>o |Sigma (Pow A) (\<lambda> A'. Func A' B)|" using card_of_UNION_Sigma .
+ also have "|Sigma (Pow A) (\<lambda> A'. Func A' B)| \<le>o |Pow A <*> Func A B|"
+ apply(rule card_of_Sigma_mono1) using card_of_Func_mono[OF _ assms] by auto
+ finally show ?thesis .
+qed
+
+lemma ordLeq_Func:
+assumes "{b1,b2} \<subseteq> B" "b1 \<noteq> b2"
+shows "|A| \<le>o |Func A B|"
+unfolding card_of_ordLeq[symmetric] proof(intro exI conjI)
+ let ?F = "\<lambda> aa a. if a \<in> A then (if a = aa then Some b1 else Some b2)
+ else None"
+ show "inj_on ?F A" using assms unfolding inj_on_def fun_eq_iff by auto
+ show "?F ` A \<subseteq> Func A B" using assms unfolding Func_def apply auto
+ by (metis option.simps(3))
+qed
+
+lemma infinite_Func:
+assumes A: "infinite A" and B: "{b1,b2} \<subseteq> B" "b1 \<noteq> b2"
+shows "infinite (Func A B)"
+using ordLeq_Func[OF B] by (metis A card_of_ordLeq_finite)
+
+(* alternative function space avoiding the option type, with undefined instead of None *)
+definition Ffunc where
+"Ffunc A B = {f . (\<forall> a \<in> A. f a \<in> B) \<and> (\<forall> a. a \<notin> A \<longrightarrow> f a = undefined)}"
+
+lemma card_of_Func_Ffunc:
+"|Ffunc A B| =o |Func A B|"
+unfolding card_of_ordIso[symmetric] proof
+ let ?F = "\<lambda> f a. if a \<in> A then Some (f a) else None"
+ show "bij_betw ?F (Ffunc A B) (Func A B)"
+ unfolding bij_betw_def unfolding inj_on_def proof(intro conjI ballI impI)
+ fix f g assume f: "f \<in> Ffunc A B" and g: "g \<in> Ffunc A B" and eq: "?F f = ?F g"
+ show "f = g"
+ proof(rule ext)
+ fix a
+ show "f a = g a"
+ proof(cases "a \<in> A")
+ case True
+ have "Some (f a) = ?F f a" using True by auto
+ also have "... = ?F g a" using eq unfolding fun_eq_iff by(rule allE)
+ also have "... = Some (g a)" using True by auto
+ finally have "Some (f a) = Some (g a)" .
+ thus ?thesis by simp
+ qed(insert f g, unfold Ffunc_def, auto)
+ qed
+ next
+ show "?F ` Ffunc A B = Func A B"
+ proof safe
+ fix f assume f: "f \<in> Func A B"
+ def g \<equiv> "\<lambda> a. case f a of Some b \<Rightarrow> b | None \<Rightarrow> undefined"
+ have "g \<in> Ffunc A B"
+ using f unfolding g_def Func_def Ffunc_def by force+
+ moreover have "f = ?F g"
+ proof(rule ext)
+ fix a show "f a = ?F g a"
+ using f unfolding Func_def g_def by (cases "a \<in> A") force+
+ qed
+ ultimately show "f \<in> ?F ` (Ffunc A B)" by blast
+ qed(unfold Ffunc_def Func_def, auto)
+ qed
+qed
+
+lemma card_of_Func_UNIV:
+"|Func (UNIV::'a set) (B::'b set)| =o |{f::'a \<Rightarrow> 'b. range f \<subseteq> B}|"
+apply(rule ordIso_symmetric) proof(intro card_of_ordIsoI)
+ let ?F = "\<lambda> f (a::'a). Some ((f a)::'b)"
+ show "bij_betw ?F {f. range f \<subseteq> B} (Func UNIV B)"
+ unfolding bij_betw_def inj_on_def proof safe
+ fix h :: "'a \<Rightarrow> 'b option" assume h: "h \<in> Func UNIV B"
+ hence "\<forall> a. \<exists> b. h a = Some b" unfolding Func_def by auto
+ then obtain f where f: "\<forall> a. h a = Some (f a)" by metis
+ hence "range f \<subseteq> B" using h unfolding Func_def by auto
+ thus "h \<in> (\<lambda>f a. Some (f a)) ` {f. range f \<subseteq> B}" using f unfolding image_def by auto
+ qed(unfold Func_def fun_eq_iff, auto)
+qed
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Ordinals_and_Cardinals/Constructions_on_Wellorders.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,795 @@
+(* Title: HOL/Ordinals_and_Cardinals/Constructions_on_Wellorders.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Constructions on wellorders.
+*)
+
+header {* Constructions on Wellorders *}
+
+theory Constructions_on_Wellorders
+imports Constructions_on_Wellorders_Base Wellorder_Embedding
+begin
+
+declare
+ ordLeq_Well_order_simp[simp]
+ ordLess_Well_order_simp[simp]
+ ordIso_Well_order_simp[simp]
+ not_ordLeq_iff_ordLess[simp]
+ not_ordLess_iff_ordLeq[simp]
+
+
+subsection {* Restriction to a set *}
+
+lemma Restr_incr2:
+"r <= r' \<Longrightarrow> Restr r A <= Restr r' A"
+by blast
+
+lemma Restr_incr:
+"\<lbrakk>r \<le> r'; A \<le> A'\<rbrakk> \<Longrightarrow> Restr r A \<le> Restr r' A'"
+by blast
+
+lemma Restr_Int:
+"Restr (Restr r A) B = Restr r (A Int B)"
+by blast
+
+lemma Restr_iff: "(a,b) : Restr r A = (a : A \<and> b : A \<and> (a,b) : r)"
+by (auto simp add: Field_def)
+
+lemma Restr_subset1: "Restr r A \<le> r"
+by auto
+
+lemma Restr_subset2: "Restr r A \<le> A \<times> A"
+by auto
+
+lemma wf_Restr:
+"wf r \<Longrightarrow> wf(Restr r A)"
+using wf_subset Restr_subset by blast
+
+lemma Restr_incr1:
+"A \<le> B \<Longrightarrow> Restr r A \<le> Restr r B"
+by blast
+
+
+subsection {* Order filters versus restrictions and embeddings *}
+
+lemma ofilter_Restr:
+assumes WELL: "Well_order r" and
+ OFA: "ofilter r A" and OFB: "ofilter r B" and SUB: "A \<le> B"
+shows "ofilter (Restr r B) A"
+proof-
+ let ?rB = "Restr r B"
+ have Well: "wo_rel r" unfolding wo_rel_def using WELL .
+ hence Refl: "Refl r" by (auto simp add: wo_rel.REFL)
+ hence Field: "Field ?rB = Field r Int B"
+ using Refl_Field_Restr by blast
+ have WellB: "wo_rel ?rB \<and> Well_order ?rB" using WELL
+ by (auto simp add: Well_order_Restr wo_rel_def)
+ (* Main proof *)
+ show ?thesis
+ proof(auto simp add: WellB wo_rel.ofilter_def)
+ fix a assume "a \<in> A"
+ hence "a \<in> Field r \<and> a \<in> B" using assms Well
+ by (auto simp add: wo_rel.ofilter_def)
+ with Field show "a \<in> Field(Restr r B)" by auto
+ next
+ fix a b assume *: "a \<in> A" and "b \<in> under (Restr r B) a"
+ hence "b \<in> under r a"
+ using WELL OFB SUB ofilter_Restr_under[of r B a] by auto
+ thus "b \<in> A" using * Well OFA by(auto simp add: wo_rel.ofilter_def)
+ qed
+qed
+
+lemma ofilter_subset_iso:
+assumes WELL: "Well_order r" and
+ OFA: "ofilter r A" and OFB: "ofilter r B"
+shows "(A = B) = iso (Restr r A) (Restr r B) id"
+using assms
+by (auto simp add: ofilter_subset_embedS_iso)
+
+
+subsection {* Ordering the well-orders by existence of embeddings *}
+
+corollary ordLeq_refl_on: "refl_on {r. Well_order r} ordLeq"
+using ordLeq_reflexive unfolding ordLeq_def refl_on_def
+by blast
+
+corollary ordLeq_trans: "trans ordLeq"
+using trans_def[of ordLeq] ordLeq_transitive by blast
+
+corollary ordLeq_preorder_on: "preorder_on {r. Well_order r} ordLeq"
+by(auto simp add: preorder_on_def ordLeq_refl_on ordLeq_trans)
+
+corollary ordIso_refl_on: "refl_on {r. Well_order r} ordIso"
+using ordIso_reflexive unfolding refl_on_def ordIso_def
+by blast
+
+corollary ordIso_trans: "trans ordIso"
+using trans_def[of ordIso] ordIso_transitive by blast
+
+corollary ordIso_sym: "sym ordIso"
+by (auto simp add: sym_def ordIso_symmetric)
+
+corollary ordIso_equiv: "equiv {r. Well_order r} ordIso"
+by (auto simp add: equiv_def ordIso_sym ordIso_refl_on ordIso_trans)
+
+lemma ordLess_irrefl: "irrefl ordLess"
+by(unfold irrefl_def, auto simp add: ordLess_irreflexive)
+
+lemma ordLess_or_ordIso:
+assumes WELL: "Well_order r" and WELL': "Well_order r'"
+shows "r <o r' \<or> r' <o r \<or> r =o r'"
+unfolding ordLess_def ordIso_def
+using assms embedS_or_iso[of r r'] by auto
+
+corollary ordLeq_ordLess_Un_ordIso:
+"ordLeq = ordLess \<union> ordIso"
+by (auto simp add: ordLeq_iff_ordLess_or_ordIso)
+
+lemma not_ordLeq_ordLess:
+"r \<le>o r' \<Longrightarrow> \<not> r' <o r"
+using not_ordLess_ordLeq by blast
+
+lemma ordIso_or_ordLess:
+assumes WELL: "Well_order r" and WELL': "Well_order r'"
+shows "r =o r' \<or> r <o r' \<or> r' <o r"
+using assms ordLess_or_ordLeq ordLeq_iff_ordLess_or_ordIso by blast
+
+lemmas ord_trans = ordIso_transitive ordLeq_transitive ordLess_transitive
+ ordIso_ordLeq_trans ordLeq_ordIso_trans
+ ordIso_ordLess_trans ordLess_ordIso_trans
+ ordLess_ordLeq_trans ordLeq_ordLess_trans
+
+lemma ofilter_ordLeq:
+assumes "Well_order r" and "ofilter r A"
+shows "Restr r A \<le>o r"
+proof-
+ have "A \<le> Field r" using assms by (auto simp add: wo_rel_def wo_rel.ofilter_def)
+ thus ?thesis using assms
+ by (simp add: ofilter_subset_ordLeq wo_rel.Field_ofilter
+ wo_rel_def Restr_Field)
+qed
+
+corollary under_Restr_ordLeq:
+"Well_order r \<Longrightarrow> Restr r (under r a) \<le>o r"
+by (auto simp add: ofilter_ordLeq wo_rel.under_ofilter wo_rel_def)
+
+
+subsection {* Copy via direct images *}
+
+lemma Id_dir_image: "dir_image Id f \<le> Id"
+unfolding dir_image_def by auto
+
+lemma Un_dir_image:
+"dir_image (r1 \<union> r2) f = (dir_image r1 f) \<union> (dir_image r2 f)"
+unfolding dir_image_def by auto
+
+lemma Int_dir_image:
+assumes "inj_on f (Field r1 \<union> Field r2)"
+shows "dir_image (r1 Int r2) f = (dir_image r1 f) Int (dir_image r2 f)"
+proof
+ show "dir_image (r1 Int r2) f \<le> (dir_image r1 f) Int (dir_image r2 f)"
+ using assms unfolding dir_image_def inj_on_def by auto
+next
+ show "(dir_image r1 f) Int (dir_image r2 f) \<le> dir_image (r1 Int r2) f"
+ proof(clarify)
+ fix a' b'
+ assume "(a',b') \<in> dir_image r1 f" "(a',b') \<in> dir_image r2 f"
+ then obtain a1 b1 a2 b2
+ where 1: "a' = f a1 \<and> b' = f b1 \<and> a' = f a2 \<and> b' = f b2" and
+ 2: "(a1,b1) \<in> r1 \<and> (a2,b2) \<in> r2" and
+ 3: "{a1,b1} \<le> Field r1 \<and> {a2,b2} \<le> Field r2"
+ unfolding dir_image_def Field_def by blast
+ hence "a1 = a2 \<and> b1 = b2" using assms unfolding inj_on_def by auto
+ hence "a' = f a1 \<and> b' = f b1 \<and> (a1,b1) \<in> r1 Int r2 \<and> (a2,b2) \<in> r1 Int r2"
+ using 1 2 by auto
+ thus "(a',b') \<in> dir_image (r1 \<inter> r2) f"
+ unfolding dir_image_def by blast
+ qed
+qed
+
+
+subsection {* Ordinal-like sum of two (disjoint) well-orders *}
+
+text{* This is roughly obtained by ``concatenating" the two well-orders -- thus, all elements
+of the first will be smaller than all elements of the second. This construction
+only makes sense if the fields of the two well-order relations are disjoint. *}
+
+definition Osum :: "'a rel \<Rightarrow> 'a rel \<Rightarrow> 'a rel" (infix "Osum" 60)
+where
+"r Osum r' = r \<union> r' \<union> {(a,a'). a \<in> Field r \<and> a' \<in> Field r'}"
+
+abbreviation Osum2 :: "'a rel \<Rightarrow> 'a rel \<Rightarrow> 'a rel" (infix "\<union>o" 60)
+where "r \<union>o r' \<equiv> r Osum r'"
+
+lemma Field_Osum: "Field(r Osum r') = Field r \<union> Field r'"
+unfolding Osum_def Field_def by blast
+
+lemma Osum_Refl:
+assumes FLD: "Field r Int Field r' = {}" and
+ REFL: "Refl r" and REFL': "Refl r'"
+shows "Refl (r Osum r')"
+using assms (* Need first unfold Field_Osum, only then Osum_def *)
+unfolding refl_on_def Field_Osum unfolding Osum_def by blast
+
+lemma Osum_trans:
+assumes FLD: "Field r Int Field r' = {}" and
+ TRANS: "trans r" and TRANS': "trans r'"
+shows "trans (r Osum r')"
+proof(unfold trans_def, auto)
+ fix x y z assume *: "(x, y) \<in> r \<union>o r'" and **: "(y, z) \<in> r \<union>o r'"
+ show "(x, z) \<in> r \<union>o r'"
+ proof-
+ {assume Case1: "(x,y) \<in> r"
+ hence 1: "x \<in> Field r \<and> y \<in> Field r" unfolding Field_def by auto
+ have ?thesis
+ proof-
+ {assume Case11: "(y,z) \<in> r"
+ hence "(x,z) \<in> r" using Case1 TRANS trans_def[of r] by blast
+ hence ?thesis unfolding Osum_def by auto
+ }
+ moreover
+ {assume Case12: "(y,z) \<in> r'"
+ hence "y \<in> Field r'" unfolding Field_def by auto
+ hence False using FLD 1 by auto
+ }
+ moreover
+ {assume Case13: "z \<in> Field r'"
+ hence ?thesis using 1 unfolding Osum_def by auto
+ }
+ ultimately show ?thesis using ** unfolding Osum_def by blast
+ qed
+ }
+ moreover
+ {assume Case2: "(x,y) \<in> r'"
+ hence 2: "x \<in> Field r' \<and> y \<in> Field r'" unfolding Field_def by auto
+ have ?thesis
+ proof-
+ {assume Case21: "(y,z) \<in> r"
+ hence "y \<in> Field r" unfolding Field_def by auto
+ hence False using FLD 2 by auto
+ }
+ moreover
+ {assume Case22: "(y,z) \<in> r'"
+ hence "(x,z) \<in> r'" using Case2 TRANS' trans_def[of r'] by blast
+ hence ?thesis unfolding Osum_def by auto
+ }
+ moreover
+ {assume Case23: "y \<in> Field r"
+ hence False using FLD 2 by auto
+ }
+ ultimately show ?thesis using ** unfolding Osum_def by blast
+ qed
+ }
+ moreover
+ {assume Case3: "x \<in> Field r \<and> y \<in> Field r'"
+ have ?thesis
+ proof-
+ {assume Case31: "(y,z) \<in> r"
+ hence "y \<in> Field r" unfolding Field_def by auto
+ hence False using FLD Case3 by auto
+ }
+ moreover
+ {assume Case32: "(y,z) \<in> r'"
+ hence "z \<in> Field r'" unfolding Field_def by blast
+ hence ?thesis unfolding Osum_def using Case3 by auto
+ }
+ moreover
+ {assume Case33: "y \<in> Field r"
+ hence False using FLD Case3 by auto
+ }
+ ultimately show ?thesis using ** unfolding Osum_def by blast
+ qed
+ }
+ ultimately show ?thesis using * unfolding Osum_def by blast
+ qed
+qed
+
+lemma Osum_Preorder:
+"\<lbrakk>Field r Int Field r' = {}; Preorder r; Preorder r'\<rbrakk> \<Longrightarrow> Preorder (r Osum r')"
+unfolding preorder_on_def using Osum_Refl Osum_trans by blast
+
+lemma Osum_antisym:
+assumes FLD: "Field r Int Field r' = {}" and
+ AN: "antisym r" and AN': "antisym r'"
+shows "antisym (r Osum r')"
+proof(unfold antisym_def, auto)
+ fix x y assume *: "(x, y) \<in> r \<union>o r'" and **: "(y, x) \<in> r \<union>o r'"
+ show "x = y"
+ proof-
+ {assume Case1: "(x,y) \<in> r"
+ hence 1: "x \<in> Field r \<and> y \<in> Field r" unfolding Field_def by auto
+ have ?thesis
+ proof-
+ have "(y,x) \<in> r \<Longrightarrow> ?thesis"
+ using Case1 AN antisym_def[of r] by blast
+ moreover
+ {assume "(y,x) \<in> r'"
+ hence "y \<in> Field r'" unfolding Field_def by auto
+ hence False using FLD 1 by auto
+ }
+ moreover
+ have "x \<in> Field r' \<Longrightarrow> False" using FLD 1 by auto
+ ultimately show ?thesis using ** unfolding Osum_def by blast
+ qed
+ }
+ moreover
+ {assume Case2: "(x,y) \<in> r'"
+ hence 2: "x \<in> Field r' \<and> y \<in> Field r'" unfolding Field_def by auto
+ have ?thesis
+ proof-
+ {assume "(y,x) \<in> r"
+ hence "y \<in> Field r" unfolding Field_def by auto
+ hence False using FLD 2 by auto
+ }
+ moreover
+ have "(y,x) \<in> r' \<Longrightarrow> ?thesis"
+ using Case2 AN' antisym_def[of r'] by blast
+ moreover
+ {assume "y \<in> Field r"
+ hence False using FLD 2 by auto
+ }
+ ultimately show ?thesis using ** unfolding Osum_def by blast
+ qed
+ }
+ moreover
+ {assume Case3: "x \<in> Field r \<and> y \<in> Field r'"
+ have ?thesis
+ proof-
+ {assume "(y,x) \<in> r"
+ hence "y \<in> Field r" unfolding Field_def by auto
+ hence False using FLD Case3 by auto
+ }
+ moreover
+ {assume Case32: "(y,x) \<in> r'"
+ hence "x \<in> Field r'" unfolding Field_def by blast
+ hence False using FLD Case3 by auto
+ }
+ moreover
+ have "\<not> y \<in> Field r" using FLD Case3 by auto
+ ultimately show ?thesis using ** unfolding Osum_def by blast
+ qed
+ }
+ ultimately show ?thesis using * unfolding Osum_def by blast
+ qed
+qed
+
+lemma Osum_Partial_order:
+"\<lbrakk>Field r Int Field r' = {}; Partial_order r; Partial_order r'\<rbrakk> \<Longrightarrow>
+ Partial_order (r Osum r')"
+unfolding partial_order_on_def using Osum_Preorder Osum_antisym by blast
+
+lemma Osum_Total:
+assumes FLD: "Field r Int Field r' = {}" and
+ TOT: "Total r" and TOT': "Total r'"
+shows "Total (r Osum r')"
+using assms
+unfolding total_on_def Field_Osum unfolding Osum_def by blast
+
+lemma Osum_Linear_order:
+"\<lbrakk>Field r Int Field r' = {}; Linear_order r; Linear_order r'\<rbrakk> \<Longrightarrow>
+ Linear_order (r Osum r')"
+unfolding linear_order_on_def using Osum_Partial_order Osum_Total by blast
+
+lemma Osum_wf:
+assumes FLD: "Field r Int Field r' = {}" and
+ WF: "wf r" and WF': "wf r'"
+shows "wf (r Osum r')"
+unfolding wf_eq_minimal2 unfolding Field_Osum
+proof(intro allI impI, elim conjE)
+ fix A assume *: "A \<subseteq> Field r \<union> Field r'" and **: "A \<noteq> {}"
+ obtain B where B_def: "B = A Int Field r" by blast
+ show "\<exists>a\<in>A. \<forall>a'\<in>A. (a', a) \<notin> r \<union>o r'"
+ proof(cases "B = {}")
+ assume Case1: "B \<noteq> {}"
+ hence "B \<noteq> {} \<and> B \<le> Field r" using B_def by auto
+ then obtain a where 1: "a \<in> B" and 2: "\<forall>a1 \<in> B. (a1,a) \<notin> r"
+ using WF unfolding wf_eq_minimal2 by blast
+ hence 3: "a \<in> Field r \<and> a \<notin> Field r'" using B_def FLD by auto
+ (* *)
+ have "\<forall>a1 \<in> A. (a1,a) \<notin> r Osum r'"
+ proof(intro ballI)
+ fix a1 assume **: "a1 \<in> A"
+ {assume Case11: "a1 \<in> Field r"
+ hence "(a1,a) \<notin> r" using B_def ** 2 by auto
+ moreover
+ have "(a1,a) \<notin> r'" using 3 by (auto simp add: Field_def)
+ ultimately have "(a1,a) \<notin> r Osum r'"
+ using 3 unfolding Osum_def by auto
+ }
+ moreover
+ {assume Case12: "a1 \<notin> Field r"
+ hence "(a1,a) \<notin> r" unfolding Field_def by auto
+ moreover
+ have "(a1,a) \<notin> r'" using 3 unfolding Field_def by auto
+ ultimately have "(a1,a) \<notin> r Osum r'"
+ using 3 unfolding Osum_def by auto
+ }
+ ultimately show "(a1,a) \<notin> r Osum r'" by blast
+ qed
+ thus ?thesis using 1 B_def by auto
+ next
+ assume Case2: "B = {}"
+ hence 1: "A \<noteq> {} \<and> A \<le> Field r'" using * ** B_def by auto
+ then obtain a' where 2: "a' \<in> A" and 3: "\<forall>a1' \<in> A. (a1',a') \<notin> r'"
+ using WF' unfolding wf_eq_minimal2 by blast
+ hence 4: "a' \<in> Field r' \<and> a' \<notin> Field r" using 1 FLD by blast
+ (* *)
+ have "\<forall>a1' \<in> A. (a1',a') \<notin> r Osum r'"
+ proof(unfold Osum_def, auto simp add: 3)
+ fix a1' assume "(a1', a') \<in> r"
+ thus False using 4 unfolding Field_def by blast
+ next
+ fix a1' assume "a1' \<in> A" and "a1' \<in> Field r"
+ thus False using Case2 B_def by auto
+ qed
+ thus ?thesis using 2 by blast
+ qed
+qed
+
+lemma Osum_minus_Id:
+assumes TOT: "Total r" and TOT': "Total r'" and
+ NID: "\<not> (r \<le> Id)" and NID': "\<not> (r' \<le> Id)"
+shows "(r Osum r') - Id \<le> (r - Id) Osum (r' - Id)"
+proof-
+ {fix a a' assume *: "(a,a') \<in> (r Osum r')" and **: "a \<noteq> a'"
+ have "(a,a') \<in> (r - Id) Osum (r' - Id)"
+ proof-
+ {assume "(a,a') \<in> r \<or> (a,a') \<in> r'"
+ with ** have ?thesis unfolding Osum_def by auto
+ }
+ moreover
+ {assume "a \<in> Field r \<and> a' \<in> Field r'"
+ hence "a \<in> Field(r - Id) \<and> a' \<in> Field (r' - Id)"
+ using assms rel.Total_Id_Field by blast
+ hence ?thesis unfolding Osum_def by auto
+ }
+ ultimately show ?thesis using * unfolding Osum_def by blast
+ qed
+ }
+ thus ?thesis by(auto simp add: Osum_def)
+qed
+
+
+lemma wf_Int_Times:
+assumes "A Int B = {}"
+shows "wf(A \<times> B)"
+proof(unfold wf_def, auto)
+ fix P x
+ assume *: "\<forall>x. (\<forall>y. y \<in> A \<and> x \<in> B \<longrightarrow> P y) \<longrightarrow> P x"
+ moreover have "\<forall>y \<in> A. P y" using assms * by blast
+ ultimately show "P x" using * by (case_tac "x \<in> B", auto)
+qed
+
+lemma Osum_minus_Id1:
+assumes "r \<le> Id"
+shows "(r Osum r') - Id \<le> (r' - Id) \<union> (Field r \<times> Field r')"
+proof-
+ let ?Left = "(r Osum r') - Id"
+ let ?Right = "(r' - Id) \<union> (Field r \<times> Field r')"
+ {fix a::'a and b assume *: "(a,b) \<notin> Id"
+ {assume "(a,b) \<in> r"
+ with * have False using assms by auto
+ }
+ moreover
+ {assume "(a,b) \<in> r'"
+ with * have "(a,b) \<in> r' - Id" by auto
+ }
+ ultimately
+ have "(a,b) \<in> ?Left \<Longrightarrow> (a,b) \<in> ?Right"
+ unfolding Osum_def by auto
+ }
+ thus ?thesis by auto
+qed
+
+lemma Osum_minus_Id2:
+assumes "r' \<le> Id"
+shows "(r Osum r') - Id \<le> (r - Id) \<union> (Field r \<times> Field r')"
+proof-
+ let ?Left = "(r Osum r') - Id"
+ let ?Right = "(r - Id) \<union> (Field r \<times> Field r')"
+ {fix a::'a and b assume *: "(a,b) \<notin> Id"
+ {assume "(a,b) \<in> r'"
+ with * have False using assms by auto
+ }
+ moreover
+ {assume "(a,b) \<in> r"
+ with * have "(a,b) \<in> r - Id" by auto
+ }
+ ultimately
+ have "(a,b) \<in> ?Left \<Longrightarrow> (a,b) \<in> ?Right"
+ unfolding Osum_def by auto
+ }
+ thus ?thesis by auto
+qed
+
+lemma Osum_wf_Id:
+assumes TOT: "Total r" and TOT': "Total r'" and
+ FLD: "Field r Int Field r' = {}" and
+ WF: "wf(r - Id)" and WF': "wf(r' - Id)"
+shows "wf ((r Osum r') - Id)"
+proof(cases "r \<le> Id \<or> r' \<le> Id")
+ assume Case1: "\<not>(r \<le> Id \<or> r' \<le> Id)"
+ have "Field(r - Id) Int Field(r' - Id) = {}"
+ using FLD mono_Field[of "r - Id" r] mono_Field[of "r' - Id" r']
+ Diff_subset[of r Id] Diff_subset[of r' Id] by blast
+ thus ?thesis
+ using Case1 Osum_minus_Id[of r r'] assms Osum_wf[of "r - Id" "r' - Id"]
+ wf_subset[of "(r - Id) \<union>o (r' - Id)" "(r Osum r') - Id"] by auto
+next
+ have 1: "wf(Field r \<times> Field r')"
+ using FLD by (auto simp add: wf_Int_Times)
+ assume Case2: "r \<le> Id \<or> r' \<le> Id"
+ moreover
+ {assume Case21: "r \<le> Id"
+ hence "(r Osum r') - Id \<le> (r' - Id) \<union> (Field r \<times> Field r')"
+ using Osum_minus_Id1[of r r'] by simp
+ moreover
+ {have "Domain(Field r \<times> Field r') Int Range(r' - Id) = {}"
+ using FLD unfolding Field_def by blast
+ hence "wf((r' - Id) \<union> (Field r \<times> Field r'))"
+ using 1 WF' wf_Un[of "Field r \<times> Field r'" "r' - Id"]
+ by (auto simp add: Un_commute)
+ }
+ ultimately have ?thesis by (auto simp add: wf_subset)
+ }
+ moreover
+ {assume Case22: "r' \<le> Id"
+ hence "(r Osum r') - Id \<le> (r - Id) \<union> (Field r \<times> Field r')"
+ using Osum_minus_Id2[of r' r] by simp
+ moreover
+ {have "Range(Field r \<times> Field r') Int Domain(r - Id) = {}"
+ using FLD unfolding Field_def by blast
+ hence "wf((r - Id) \<union> (Field r \<times> Field r'))"
+ using 1 WF wf_Un[of "r - Id" "Field r \<times> Field r'"]
+ by (auto simp add: Un_commute)
+ }
+ ultimately have ?thesis by (auto simp add: wf_subset)
+ }
+ ultimately show ?thesis by blast
+qed
+
+lemma Osum_Well_order:
+assumes FLD: "Field r Int Field r' = {}" and
+ WELL: "Well_order r" and WELL': "Well_order r'"
+shows "Well_order (r Osum r')"
+proof-
+ have "Total r \<and> Total r'" using WELL WELL'
+ by (auto simp add: order_on_defs)
+ thus ?thesis using assms unfolding well_order_on_def
+ using Osum_Linear_order Osum_wf_Id by blast
+qed
+
+lemma Osum_embed:
+assumes FLD: "Field r Int Field r' = {}" and
+ WELL: "Well_order r" and WELL': "Well_order r'"
+shows "embed r (r Osum r') id"
+proof-
+ have 1: "Well_order (r Osum r')"
+ using assms by (auto simp add: Osum_Well_order)
+ moreover
+ have "compat r (r Osum r') id"
+ unfolding compat_def Osum_def by auto
+ moreover
+ have "inj_on id (Field r)" by simp
+ moreover
+ have "ofilter (r Osum r') (Field r)"
+ using 1 proof(auto simp add: wo_rel_def wo_rel.ofilter_def
+ Field_Osum rel.under_def)
+ fix a b assume 2: "a \<in> Field r" and 3: "(b,a) \<in> r Osum r'"
+ moreover
+ {assume "(b,a) \<in> r'"
+ hence "a \<in> Field r'" using Field_def[of r'] by blast
+ hence False using 2 FLD by blast
+ }
+ moreover
+ {assume "a \<in> Field r'"
+ hence False using 2 FLD by blast
+ }
+ ultimately
+ show "b \<in> Field r" by (auto simp add: Osum_def Field_def)
+ qed
+ ultimately show ?thesis
+ using assms by (auto simp add: embed_iff_compat_inj_on_ofilter)
+qed
+
+corollary Osum_ordLeq:
+assumes FLD: "Field r Int Field r' = {}" and
+ WELL: "Well_order r" and WELL': "Well_order r'"
+shows "r \<le>o r Osum r'"
+using assms Osum_embed Osum_Well_order
+unfolding ordLeq_def by blast
+
+lemma Well_order_embed_copy:
+assumes WELL: "well_order_on A r" and
+ INJ: "inj_on f A" and SUB: "f ` A \<le> B"
+shows "\<exists>r'. well_order_on B r' \<and> r \<le>o r'"
+proof-
+ have "bij_betw f A (f ` A)"
+ using INJ inj_on_imp_bij_betw by blast
+ then obtain r'' where "well_order_on (f ` A) r''" and 1: "r =o r''"
+ using WELL Well_order_iso_copy by blast
+ hence 2: "Well_order r'' \<and> Field r'' = (f ` A)"
+ using rel.well_order_on_Well_order by blast
+ (* *)
+ let ?C = "B - (f ` A)"
+ obtain r''' where "well_order_on ?C r'''"
+ using well_order_on by blast
+ hence 3: "Well_order r''' \<and> Field r''' = ?C"
+ using rel.well_order_on_Well_order by blast
+ (* *)
+ let ?r' = "r'' Osum r'''"
+ have "Field r'' Int Field r''' = {}"
+ using 2 3 by auto
+ hence "r'' \<le>o ?r'" using Osum_ordLeq[of r'' r'''] 2 3 by blast
+ hence 4: "r \<le>o ?r'" using 1 ordIso_ordLeq_trans by blast
+ (* *)
+ hence "Well_order ?r'" unfolding ordLeq_def by auto
+ moreover
+ have "Field ?r' = B" using 2 3 SUB by (auto simp add: Field_Osum)
+ ultimately show ?thesis using 4 by blast
+qed
+
+
+subsection {* The maxim among a finite set of ordinals *}
+
+text {* The correct phrasing would be ``a maxim of ...", as @{text "\<le>o"} is only a preorder. *}
+
+definition isOmax :: "'a rel set \<Rightarrow> 'a rel \<Rightarrow> bool"
+where
+"isOmax R r == r \<in> R \<and> (ALL r' : R. r' \<le>o r)"
+
+definition omax :: "'a rel set \<Rightarrow> 'a rel"
+where
+"omax R == SOME r. isOmax R r"
+
+lemma exists_isOmax:
+assumes "finite R" and "R \<noteq> {}" and "\<forall> r \<in> R. Well_order r"
+shows "\<exists> r. isOmax R r"
+proof-
+ have "finite R \<Longrightarrow> R \<noteq> {} \<longrightarrow> (\<forall> r \<in> R. Well_order r) \<longrightarrow> (\<exists> r. isOmax R r)"
+ apply(erule finite_induct) apply(simp add: isOmax_def)
+ proof(clarsimp)
+ fix r :: "('a \<times> 'a) set" and R assume *: "finite R" and **: "r \<notin> R"
+ and ***: "Well_order r" and ****: "\<forall>r\<in>R. Well_order r"
+ and IH: "R \<noteq> {} \<longrightarrow> (\<exists>p. isOmax R p)"
+ let ?R' = "insert r R"
+ show "\<exists>p'. (isOmax ?R' p')"
+ proof(cases "R = {}")
+ assume Case1: "R = {}"
+ thus ?thesis unfolding isOmax_def using ***
+ by (simp add: ordLeq_reflexive)
+ next
+ assume Case2: "R \<noteq> {}"
+ then obtain p where p: "isOmax R p" using IH by auto
+ hence 1: "Well_order p" using **** unfolding isOmax_def by simp
+ {assume Case21: "r \<le>o p"
+ hence "isOmax ?R' p" using p unfolding isOmax_def by simp
+ hence ?thesis by auto
+ }
+ moreover
+ {assume Case22: "p \<le>o r"
+ {fix r' assume "r' \<in> ?R'"
+ moreover
+ {assume "r' \<in> R"
+ hence "r' \<le>o p" using p unfolding isOmax_def by simp
+ hence "r' \<le>o r" using Case22 by(rule ordLeq_transitive)
+ }
+ moreover have "r \<le>o r" using *** by(rule ordLeq_reflexive)
+ ultimately have "r' \<le>o r" by auto
+ }
+ hence "isOmax ?R' r" unfolding isOmax_def by simp
+ hence ?thesis by auto
+ }
+ moreover have "r \<le>o p \<or> p \<le>o r"
+ using 1 *** ordLeq_total by auto
+ ultimately show ?thesis by blast
+ qed
+ qed
+ thus ?thesis using assms by auto
+qed
+
+lemma omax_isOmax:
+assumes "finite R" and "R \<noteq> {}" and "\<forall> r \<in> R. Well_order r"
+shows "isOmax R (omax R)"
+unfolding omax_def using assms
+by(simp add: exists_isOmax someI_ex)
+
+lemma omax_in:
+assumes "finite R" and "R \<noteq> {}" and "\<forall> r \<in> R. Well_order r"
+shows "omax R \<in> R"
+using assms omax_isOmax unfolding isOmax_def by blast
+
+lemma Well_order_omax:
+assumes "finite R" and "R \<noteq> {}" and "\<forall>r\<in>R. Well_order r"
+shows "Well_order (omax R)"
+using assms apply - apply(drule omax_in) by auto
+
+lemma omax_maxim:
+assumes "finite R" and "\<forall> r \<in> R. Well_order r" and "r \<in> R"
+shows "r \<le>o omax R"
+using assms omax_isOmax unfolding isOmax_def by blast
+
+lemma omax_ordLeq:
+assumes "finite R" and "R \<noteq> {}" and *: "\<forall> r \<in> R. r \<le>o p"
+shows "omax R \<le>o p"
+proof-
+ have "\<forall> r \<in> R. Well_order r" using * unfolding ordLeq_def by simp
+ thus ?thesis using assms omax_in by auto
+qed
+
+lemma omax_ordLess:
+assumes "finite R" and "R \<noteq> {}" and *: "\<forall> r \<in> R. r <o p"
+shows "omax R <o p"
+proof-
+ have "\<forall> r \<in> R. Well_order r" using * unfolding ordLess_def by simp
+ thus ?thesis using assms omax_in by auto
+qed
+
+lemma omax_ordLeq_elim:
+assumes "finite R" and "\<forall> r \<in> R. Well_order r"
+and "omax R \<le>o p" and "r \<in> R"
+shows "r \<le>o p"
+using assms omax_maxim[of R r] apply simp
+using ordLeq_transitive by blast
+
+lemma omax_ordLess_elim:
+assumes "finite R" and "\<forall> r \<in> R. Well_order r"
+and "omax R <o p" and "r \<in> R"
+shows "r <o p"
+using assms omax_maxim[of R r] apply simp
+using ordLeq_ordLess_trans by blast
+
+lemma ordLeq_omax:
+assumes "finite R" and "\<forall> r \<in> R. Well_order r"
+and "r \<in> R" and "p \<le>o r"
+shows "p \<le>o omax R"
+using assms omax_maxim[of R r] apply simp
+using ordLeq_transitive by blast
+
+lemma ordLess_omax:
+assumes "finite R" and "\<forall> r \<in> R. Well_order r"
+and "r \<in> R" and "p <o r"
+shows "p <o omax R"
+using assms omax_maxim[of R r] apply simp
+using ordLess_ordLeq_trans by blast
+
+lemma omax_ordLeq_mono:
+assumes P: "finite P" and R: "finite R"
+and NE_P: "P \<noteq> {}" and Well_R: "\<forall> r \<in> R. Well_order r"
+and LEQ: "\<forall> p \<in> P. \<exists> r \<in> R. p \<le>o r"
+shows "omax P \<le>o omax R"
+proof-
+ let ?mp = "omax P" let ?mr = "omax R"
+ {fix p assume "p : P"
+ then obtain r where r: "r : R" and "p \<le>o r"
+ using LEQ by blast
+ moreover have "r <=o ?mr"
+ using r R Well_R omax_maxim by blast
+ ultimately have "p <=o ?mr"
+ using ordLeq_transitive by blast
+ }
+ thus "?mp <=o ?mr"
+ using NE_P P using omax_ordLeq by blast
+qed
+
+lemma omax_ordLess_mono:
+assumes P: "finite P" and R: "finite R"
+and NE_P: "P \<noteq> {}" and Well_R: "\<forall> r \<in> R. Well_order r"
+and LEQ: "\<forall> p \<in> P. \<exists> r \<in> R. p <o r"
+shows "omax P <o omax R"
+proof-
+ let ?mp = "omax P" let ?mr = "omax R"
+ {fix p assume "p : P"
+ then obtain r where r: "r : R" and "p <o r"
+ using LEQ by blast
+ moreover have "r <=o ?mr"
+ using r R Well_R omax_maxim by blast
+ ultimately have "p <o ?mr"
+ using ordLess_ordLeq_trans by blast
+ }
+ thus "?mp <o ?mr"
+ using NE_P P omax_ordLess by blast
+qed
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Ordinals_and_Cardinals/Constructions_on_Wellorders_Base.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,1633 @@
+(* Title: HOL/Ordinals_and_Cardinals/Constructions_on_Wellorders_Base.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Constructions on wellorders (base).
+*)
+
+header {* Constructions on Wellorders (Base) *}
+
+theory Constructions_on_Wellorders_Base
+imports Wellorder_Embedding_Base
+begin
+
+
+text {* In this section, we study basic constructions on well-orders, such as restriction to
+a set/order filter, copy via direct images, ordinal-like sum of disjoint well-orders,
+and bounded square. We also define between well-orders
+the relations @{text "ordLeq"}, of being embedded (abbreviated @{text "\<le>o"}),
+@{text "ordLess"}, of being strictly embedded (abbreviated @{text "<o"}), and
+@{text "ordIso"}, of being isomorphic (abbreviated @{text "=o"}). We study the
+connections between these relations, order filters, and the aforementioned constructions.
+A main result of this section is that @{text "<o"} is well-founded. *}
+
+
+subsection {* Restriction to a set *}
+
+
+abbreviation Restr :: "'a rel \<Rightarrow> 'a set \<Rightarrow> 'a rel"
+where "Restr r A \<equiv> r Int (A \<times> A)"
+
+
+lemma Restr_subset:
+"A \<le> B \<Longrightarrow> Restr (Restr r B) A = Restr r A"
+by blast
+
+
+lemma Restr_Field: "Restr r (Field r) = r"
+unfolding Field_def by auto
+
+
+lemma Refl_Restr: "Refl r \<Longrightarrow> Refl(Restr r A)"
+unfolding refl_on_def Field_def by auto
+
+
+lemma antisym_Restr:
+"antisym r \<Longrightarrow> antisym(Restr r A)"
+unfolding antisym_def Field_def by auto
+
+
+lemma Total_Restr:
+"Total r \<Longrightarrow> Total(Restr r A)"
+unfolding total_on_def Field_def by auto
+
+
+lemma trans_Restr:
+"trans r \<Longrightarrow> trans(Restr r A)"
+unfolding trans_def Field_def by blast
+
+
+lemma Preorder_Restr:
+"Preorder r \<Longrightarrow> Preorder(Restr r A)"
+unfolding preorder_on_def by (simp add: Refl_Restr trans_Restr)
+
+
+lemma Partial_order_Restr:
+"Partial_order r \<Longrightarrow> Partial_order(Restr r A)"
+unfolding partial_order_on_def by (simp add: Preorder_Restr antisym_Restr)
+
+
+lemma Linear_order_Restr:
+"Linear_order r \<Longrightarrow> Linear_order(Restr r A)"
+unfolding linear_order_on_def by (simp add: Partial_order_Restr Total_Restr)
+
+
+lemma Well_order_Restr:
+assumes "Well_order r"
+shows "Well_order(Restr r A)"
+proof-
+ have "Restr r A - Id \<le> r - Id" using Restr_subset by blast
+ hence "wf(Restr r A - Id)" using assms
+ using well_order_on_def wf_subset by blast
+ thus ?thesis using assms unfolding well_order_on_def
+ by (simp add: Linear_order_Restr)
+qed
+
+
+lemma Field_Restr_subset: "Field(Restr r A) \<le> A"
+by (auto simp add: Field_def)
+
+
+lemma Refl_Field_Restr:
+"Refl r \<Longrightarrow> Field(Restr r A) = (Field r) Int A"
+by (auto simp add: refl_on_def Field_def)
+
+
+lemma Refl_Field_Restr2:
+"\<lbrakk>Refl r; A \<le> Field r\<rbrakk> \<Longrightarrow> Field(Restr r A) = A"
+by (auto simp add: Refl_Field_Restr)
+
+
+lemma well_order_on_Restr:
+assumes WELL: "Well_order r" and SUB: "A \<le> Field r"
+shows "well_order_on A (Restr r A)"
+using assms
+using Well_order_Restr[of r A] Refl_Field_Restr2[of r A]
+ order_on_defs[of "Field r" r] by auto
+
+
+subsection {* Order filters versus restrictions and embeddings *}
+
+
+lemma Field_Restr_ofilter:
+"\<lbrakk>Well_order r; wo_rel.ofilter r A\<rbrakk> \<Longrightarrow> Field(Restr r A) = A"
+by (auto simp add: wo_rel_def wo_rel.ofilter_def wo_rel.REFL Refl_Field_Restr2)
+
+
+lemma ofilter_Restr_under:
+assumes WELL: "Well_order r" and OF: "wo_rel.ofilter r A" and IN: "a \<in> A"
+shows "rel.under (Restr r A) a = rel.under r a"
+using assms wo_rel_def
+proof(auto simp add: wo_rel.ofilter_def rel.under_def)
+ fix b assume *: "a \<in> A" and "(b,a) \<in> r"
+ hence "b \<in> rel.under r a \<and> a \<in> Field r"
+ unfolding rel.under_def using Field_def by fastforce
+ thus "b \<in> A" using * assms by (auto simp add: wo_rel_def wo_rel.ofilter_def)
+qed
+
+
+lemma ofilter_embed:
+assumes "Well_order r"
+shows "wo_rel.ofilter r A = (A \<le> Field r \<and> embed (Restr r A) r id)"
+proof
+ assume *: "wo_rel.ofilter r A"
+ show "A \<le> Field r \<and> embed (Restr r A) r id"
+ proof(unfold embed_def, auto)
+ fix a assume "a \<in> A" thus "a \<in> Field r" using assms *
+ by (auto simp add: wo_rel_def wo_rel.ofilter_def)
+ next
+ fix a assume "a \<in> Field (Restr r A)"
+ thus "bij_betw id (rel.under (Restr r A) a) (rel.under r a)" using assms *
+ by (simp add: ofilter_Restr_under Field_Restr_ofilter)
+ qed
+next
+ assume *: "A \<le> Field r \<and> embed (Restr r A) r id"
+ hence "Field(Restr r A) \<le> Field r"
+ using assms embed_Field[of "Restr r A" r id] id_def
+ Well_order_Restr[of r] by auto
+ {fix a assume "a \<in> A"
+ hence "a \<in> Field(Restr r A)" using * assms
+ by (simp add: order_on_defs Refl_Field_Restr2)
+ hence "bij_betw id (rel.under (Restr r A) a) (rel.under r a)"
+ using * unfolding embed_def by auto
+ hence "rel.under r a \<le> rel.under (Restr r A) a"
+ unfolding bij_betw_def by auto
+ also have "\<dots> \<le> Field(Restr r A)" by (simp add: rel.under_Field)
+ also have "\<dots> \<le> A" by (simp add: Field_Restr_subset)
+ finally have "rel.under r a \<le> A" .
+ }
+ thus "wo_rel.ofilter r A" using assms * by (simp add: wo_rel_def wo_rel.ofilter_def)
+qed
+
+
+lemma ofilter_Restr_Int:
+assumes WELL: "Well_order r" and OFA: "wo_rel.ofilter r A"
+shows "wo_rel.ofilter (Restr r B) (A Int B)"
+proof-
+ let ?rB = "Restr r B"
+ have Well: "wo_rel r" unfolding wo_rel_def using WELL .
+ hence Refl: "Refl r" by (simp add: wo_rel.REFL)
+ hence Field: "Field ?rB = Field r Int B"
+ using Refl_Field_Restr by blast
+ have WellB: "wo_rel ?rB \<and> Well_order ?rB" using WELL
+ by (simp add: Well_order_Restr wo_rel_def)
+ (* Main proof *)
+ show ?thesis using WellB assms
+ proof(auto simp add: wo_rel.ofilter_def rel.under_def)
+ fix a assume "a \<in> A" and *: "a \<in> B"
+ hence "a \<in> Field r" using OFA Well by (auto simp add: wo_rel.ofilter_def)
+ with * show "a \<in> Field ?rB" using Field by auto
+ next
+ fix a b assume "a \<in> A" and "(b,a) \<in> r"
+ thus "b \<in> A" using Well OFA by (auto simp add: wo_rel.ofilter_def rel.under_def)
+ qed
+qed
+
+
+lemma ofilter_Restr_subset:
+assumes WELL: "Well_order r" and OFA: "wo_rel.ofilter r A" and SUB: "A \<le> B"
+shows "wo_rel.ofilter (Restr r B) A"
+proof-
+ have "A Int B = A" using SUB by blast
+ thus ?thesis using assms ofilter_Restr_Int[of r A B] by auto
+qed
+
+
+lemma ofilter_subset_embed:
+assumes WELL: "Well_order r" and
+ OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B"
+shows "(A \<le> B) = (embed (Restr r A) (Restr r B) id)"
+proof-
+ let ?rA = "Restr r A" let ?rB = "Restr r B"
+ have Well: "wo_rel r" unfolding wo_rel_def using WELL .
+ hence Refl: "Refl r" by (simp add: wo_rel.REFL)
+ hence FieldA: "Field ?rA = Field r Int A"
+ using Refl_Field_Restr by blast
+ have FieldB: "Field ?rB = Field r Int B"
+ using Refl Refl_Field_Restr by blast
+ have WellA: "wo_rel ?rA \<and> Well_order ?rA" using WELL
+ by (simp add: Well_order_Restr wo_rel_def)
+ have WellB: "wo_rel ?rB \<and> Well_order ?rB" using WELL
+ by (simp add: Well_order_Restr wo_rel_def)
+ (* Main proof *)
+ show ?thesis
+ proof
+ assume *: "A \<le> B"
+ hence "wo_rel.ofilter (Restr r B) A" using assms
+ by (simp add: ofilter_Restr_subset)
+ hence "embed (Restr ?rB A) (Restr r B) id"
+ using WellB ofilter_embed[of "?rB" A] by auto
+ thus "embed (Restr r A) (Restr r B) id"
+ using * by (simp add: Restr_subset)
+ next
+ assume *: "embed (Restr r A) (Restr r B) id"
+ {fix a assume **: "a \<in> A"
+ hence "a \<in> Field r" using Well OFA by (auto simp add: wo_rel.ofilter_def)
+ with ** FieldA have "a \<in> Field ?rA" by auto
+ hence "a \<in> Field ?rB" using * WellA embed_Field[of ?rA ?rB id] by auto
+ hence "a \<in> B" using FieldB by auto
+ }
+ thus "A \<le> B" by blast
+ qed
+qed
+
+
+lemma ofilter_subset_embedS_iso:
+assumes WELL: "Well_order r" and
+ OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B"
+shows "((A < B) = (embedS (Restr r A) (Restr r B) id)) \<and>
+ ((A = B) = (iso (Restr r A) (Restr r B) id))"
+proof-
+ let ?rA = "Restr r A" let ?rB = "Restr r B"
+ have Well: "wo_rel r" unfolding wo_rel_def using WELL .
+ hence Refl: "Refl r" by (simp add: wo_rel.REFL)
+ hence "Field ?rA = Field r Int A"
+ using Refl_Field_Restr by blast
+ hence FieldA: "Field ?rA = A" using OFA Well
+ by (auto simp add: wo_rel.ofilter_def)
+ have "Field ?rB = Field r Int B"
+ using Refl Refl_Field_Restr by blast
+ hence FieldB: "Field ?rB = B" using OFB Well
+ by (auto simp add: wo_rel.ofilter_def)
+ (* Main proof *)
+ show ?thesis unfolding embedS_def iso_def
+ using assms ofilter_subset_embed[of r A B]
+ FieldA FieldB bij_betw_id_iff[of A B] by auto
+qed
+
+
+lemma ofilter_subset_embedS:
+assumes WELL: "Well_order r" and
+ OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B"
+shows "(A < B) = embedS (Restr r A) (Restr r B) id"
+using assms by (simp add: ofilter_subset_embedS_iso)
+
+
+lemma embed_implies_iso_Restr:
+assumes WELL: "Well_order r" and WELL': "Well_order r'" and
+ EMB: "embed r' r f"
+shows "iso r' (Restr r (f ` (Field r'))) f"
+proof-
+ let ?A' = "Field r'"
+ let ?r'' = "Restr r (f ` ?A')"
+ have 0: "Well_order ?r''" using WELL Well_order_Restr by blast
+ have 1: "wo_rel.ofilter r (f ` ?A')" using assms embed_Field_ofilter by blast
+ hence "Field ?r'' = f ` (Field r')" using WELL Field_Restr_ofilter by blast
+ hence "bij_betw f ?A' (Field ?r'')"
+ using EMB embed_inj_on WELL' unfolding bij_betw_def by blast
+ moreover
+ {have "\<forall>a b. (a,b) \<in> r' \<longrightarrow> a \<in> Field r' \<and> b \<in> Field r'"
+ unfolding Field_def by auto
+ hence "compat r' ?r'' f"
+ using assms embed_iff_compat_inj_on_ofilter
+ unfolding compat_def by blast
+ }
+ ultimately show ?thesis using WELL' 0 iso_iff3 by blast
+qed
+
+
+subsection {* The strict inclusion on proper ofilters is well-founded *}
+
+
+definition ofilterIncl :: "'a rel \<Rightarrow> 'a set rel"
+where
+"ofilterIncl r \<equiv> {(A,B). wo_rel.ofilter r A \<and> A \<noteq> Field r \<and>
+ wo_rel.ofilter r B \<and> B \<noteq> Field r \<and> A < B}"
+
+
+lemma wf_ofilterIncl:
+assumes WELL: "Well_order r"
+shows "wf(ofilterIncl r)"
+proof-
+ have Well: "wo_rel r" using WELL by (simp add: wo_rel_def)
+ hence Lo: "Linear_order r" by (simp add: wo_rel.LIN)
+ let ?h = "(\<lambda> A. wo_rel.suc r A)"
+ let ?rS = "r - Id"
+ have "wf ?rS" using WELL by (simp add: order_on_defs)
+ moreover
+ have "compat (ofilterIncl r) ?rS ?h"
+ proof(unfold compat_def ofilterIncl_def,
+ intro allI impI, simp, elim conjE)
+ fix A B
+ assume *: "wo_rel.ofilter r A" "A \<noteq> Field r" and
+ **: "wo_rel.ofilter r B" "B \<noteq> Field r" and ***: "A < B"
+ then obtain a and b where 0: "a \<in> Field r \<and> b \<in> Field r" and
+ 1: "A = rel.underS r a \<and> B = rel.underS r b"
+ using Well by (auto simp add: wo_rel.ofilter_underS_Field)
+ hence "a \<noteq> b" using *** by auto
+ moreover
+ have "(a,b) \<in> r" using 0 1 Lo ***
+ by (auto simp add: rel.underS_incl_iff)
+ moreover
+ have "a = wo_rel.suc r A \<and> b = wo_rel.suc r B"
+ using Well 0 1 by (simp add: wo_rel.suc_underS)
+ ultimately
+ show "(wo_rel.suc r A, wo_rel.suc r B) \<in> r \<and> wo_rel.suc r A \<noteq> wo_rel.suc r B"
+ by simp
+ qed
+ ultimately show "wf (ofilterIncl r)" by (simp add: compat_wf)
+qed
+
+
+
+subsection {* Ordering the well-orders by existence of embeddings *}
+
+
+text {* We define three relations between well-orders:
+\begin{itemize}
+\item @{text "ordLeq"}, of being embedded (abbreviated @{text "\<le>o"});
+\item @{text "ordLess"}, of being strictly embedded (abbreviated @{text "<o"});
+\item @{text "ordIso"}, of being isomorphic (abbreviated @{text "=o"}).
+\end{itemize}
+%
+The prefix "ord" and the index "o" in these names stand for "ordinal-like".
+These relations shall be proved to be inter-connected in a similar fashion as the trio
+@{text "\<le>"}, @{text "<"}, @{text "="} associated to a total order on a set.
+*}
+
+
+definition ordLeq :: "('a rel * 'a' rel) set"
+where
+"ordLeq = {(r,r'). Well_order r \<and> Well_order r' \<and> (\<exists>f. embed r r' f)}"
+
+
+abbreviation ordLeq2 :: "'a rel \<Rightarrow> 'a' rel \<Rightarrow> bool" (infix "<=o" 50)
+where "r <=o r' \<equiv> (r,r') \<in> ordLeq"
+
+
+abbreviation ordLeq3 :: "'a rel \<Rightarrow> 'a' rel \<Rightarrow> bool" (infix "\<le>o" 50)
+where "r \<le>o r' \<equiv> r <=o r'"
+
+
+definition ordLess :: "('a rel * 'a' rel) set"
+where
+"ordLess = {(r,r'). Well_order r \<and> Well_order r' \<and> (\<exists>f. embedS r r' f)}"
+
+abbreviation ordLess2 :: "'a rel \<Rightarrow> 'a' rel \<Rightarrow> bool" (infix "<o" 50)
+where "r <o r' \<equiv> (r,r') \<in> ordLess"
+
+
+definition ordIso :: "('a rel * 'a' rel) set"
+where
+"ordIso = {(r,r'). Well_order r \<and> Well_order r' \<and> (\<exists>f. iso r r' f)}"
+
+abbreviation ordIso2 :: "'a rel \<Rightarrow> 'a' rel \<Rightarrow> bool" (infix "=o" 50)
+where "r =o r' \<equiv> (r,r') \<in> ordIso"
+
+
+lemmas ordRels_def = ordLeq_def ordLess_def ordIso_def
+
+lemma ordLeq_Well_order_simp:
+assumes "r \<le>o r'"
+shows "Well_order r \<and> Well_order r'"
+using assms unfolding ordLeq_def by simp
+
+
+lemma ordLess_Well_order_simp:
+assumes "r <o r'"
+shows "Well_order r \<and> Well_order r'"
+using assms unfolding ordLess_def by simp
+
+
+lemma ordIso_Well_order_simp:
+assumes "r =o r'"
+shows "Well_order r \<and> Well_order r'"
+using assms unfolding ordIso_def by simp
+
+
+text{* Notice that the relations @{text "\<le>o"}, @{text "<o"}, @{text "=o"} connect well-orders
+on potentially {\em distinct} types. However, some of the lemmas below, including the next one,
+restrict implicitly the type of these relations to @{text "(('a rel) * ('a rel)) set"} , i.e.,
+to @{text "'a rel rel"}. *}
+
+
+lemma ordLeq_reflexive:
+"Well_order r \<Longrightarrow> r \<le>o r"
+unfolding ordLeq_def using id_embed[of r] by blast
+
+
+lemma ordLeq_transitive[trans]:
+assumes *: "r \<le>o r'" and **: "r' \<le>o r''"
+shows "r \<le>o r''"
+proof-
+ obtain f and f'
+ where 1: "Well_order r \<and> Well_order r' \<and> Well_order r''" and
+ "embed r r' f" and "embed r' r'' f'"
+ using * ** unfolding ordLeq_def by blast
+ hence "embed r r'' (f' o f)"
+ using comp_embed[of r r' f r'' f'] by auto
+ thus "r \<le>o r''" unfolding ordLeq_def using 1 by auto
+qed
+
+
+lemma ordLeq_total:
+"\<lbrakk>Well_order r; Well_order r'\<rbrakk> \<Longrightarrow> r \<le>o r' \<or> r' \<le>o r"
+unfolding ordLeq_def using wellorders_totally_ordered by blast
+
+
+lemma ordIso_reflexive:
+"Well_order r \<Longrightarrow> r =o r"
+unfolding ordIso_def using id_iso[of r] by blast
+
+
+lemma ordIso_transitive[trans]:
+assumes *: "r =o r'" and **: "r' =o r''"
+shows "r =o r''"
+proof-
+ obtain f and f'
+ where 1: "Well_order r \<and> Well_order r' \<and> Well_order r''" and
+ "iso r r' f" and 3: "iso r' r'' f'"
+ using * ** unfolding ordIso_def by auto
+ hence "iso r r'' (f' o f)"
+ using comp_iso[of r r' f r'' f'] by auto
+ thus "r =o r''" unfolding ordIso_def using 1 by auto
+qed
+
+
+lemma ordIso_symmetric:
+assumes *: "r =o r'"
+shows "r' =o r"
+proof-
+ obtain f where 1: "Well_order r \<and> Well_order r'" and
+ 2: "embed r r' f \<and> bij_betw f (Field r) (Field r')"
+ using * by (auto simp add: ordIso_def iso_def)
+ let ?f' = "inv_into (Field r) f"
+ have "embed r' r ?f' \<and> bij_betw ?f' (Field r') (Field r)"
+ using 1 2 by (simp add: bij_betw_inv_into inv_into_Field_embed_bij_betw)
+ thus "r' =o r" unfolding ordIso_def using 1 by (auto simp add: iso_def)
+qed
+
+
+lemma ordLeq_ordLess_trans[trans]:
+assumes "r \<le>o r'" and " r' <o r''"
+shows "r <o r''"
+proof-
+ have "Well_order r \<and> Well_order r''"
+ using assms unfolding ordLeq_def ordLess_def by auto
+ thus ?thesis using assms unfolding ordLeq_def ordLess_def
+ using embed_comp_embedS by blast
+qed
+
+
+lemma ordLess_ordLeq_trans[trans]:
+assumes "r <o r'" and " r' \<le>o r''"
+shows "r <o r''"
+proof-
+ have "Well_order r \<and> Well_order r''"
+ using assms unfolding ordLeq_def ordLess_def by auto
+ thus ?thesis using assms unfolding ordLeq_def ordLess_def
+ using embedS_comp_embed by blast
+qed
+
+
+lemma ordLeq_ordIso_trans[trans]:
+assumes "r \<le>o r'" and " r' =o r''"
+shows "r \<le>o r''"
+proof-
+ have "Well_order r \<and> Well_order r''"
+ using assms unfolding ordLeq_def ordIso_def by auto
+ thus ?thesis using assms unfolding ordLeq_def ordIso_def
+ using embed_comp_iso by blast
+qed
+
+
+lemma ordIso_ordLeq_trans[trans]:
+assumes "r =o r'" and " r' \<le>o r''"
+shows "r \<le>o r''"
+proof-
+ have "Well_order r \<and> Well_order r''"
+ using assms unfolding ordLeq_def ordIso_def by auto
+ thus ?thesis using assms unfolding ordLeq_def ordIso_def
+ using iso_comp_embed by blast
+qed
+
+
+lemma ordLess_ordIso_trans[trans]:
+assumes "r <o r'" and " r' =o r''"
+shows "r <o r''"
+proof-
+ have "Well_order r \<and> Well_order r''"
+ using assms unfolding ordLess_def ordIso_def by auto
+ thus ?thesis using assms unfolding ordLess_def ordIso_def
+ using embedS_comp_iso by blast
+qed
+
+
+lemma ordIso_ordLess_trans[trans]:
+assumes "r =o r'" and " r' <o r''"
+shows "r <o r''"
+proof-
+ have "Well_order r \<and> Well_order r''"
+ using assms unfolding ordLess_def ordIso_def by auto
+ thus ?thesis using assms unfolding ordLess_def ordIso_def
+ using iso_comp_embedS by blast
+qed
+
+
+lemma ordLess_not_embed:
+assumes "r <o r'"
+shows "\<not>(\<exists>f'. embed r' r f')"
+proof-
+ obtain f where 1: "Well_order r \<and> Well_order r'" and 2: "embed r r' f" and
+ 3: " \<not> bij_betw f (Field r) (Field r')"
+ using assms unfolding ordLess_def by (auto simp add: embedS_def)
+ {fix f' assume *: "embed r' r f'"
+ hence "bij_betw f (Field r) (Field r')" using 1 2
+ by (simp add: embed_bothWays_Field_bij_betw)
+ with 3 have False by contradiction
+ }
+ thus ?thesis by blast
+qed
+
+
+lemma ordLess_Field:
+assumes OL: "r1 <o r2" and EMB: "embed r1 r2 f"
+shows "\<not> (f`(Field r1) = Field r2)"
+proof-
+ let ?A1 = "Field r1" let ?A2 = "Field r2"
+ obtain g where
+ 0: "Well_order r1 \<and> Well_order r2" and
+ 1: "embed r1 r2 g \<and> \<not>(bij_betw g ?A1 ?A2)"
+ using OL unfolding ordLess_def by (auto simp add: embedS_def)
+ hence "\<forall>a \<in> ?A1. f a = g a"
+ using 0 EMB embed_unique[of r1] by auto
+ hence "\<not>(bij_betw f ?A1 ?A2)"
+ using 1 bij_betw_cong[of ?A1] by blast
+ moreover
+ have "inj_on f ?A1" using EMB 0 by (simp add: embed_inj_on)
+ ultimately show ?thesis by (simp add: bij_betw_def)
+qed
+
+
+lemma ordLess_iff:
+"r <o r' = (Well_order r \<and> Well_order r' \<and> \<not>(\<exists>f'. embed r' r f'))"
+proof
+ assume *: "r <o r'"
+ hence "\<not>(\<exists>f'. embed r' r f')" using ordLess_not_embed[of r r'] by simp
+ with * show "Well_order r \<and> Well_order r' \<and> \<not> (\<exists>f'. embed r' r f')"
+ unfolding ordLess_def by auto
+next
+ assume *: "Well_order r \<and> Well_order r' \<and> \<not> (\<exists>f'. embed r' r f')"
+ then obtain f where 1: "embed r r' f"
+ using wellorders_totally_ordered[of r r'] by blast
+ moreover
+ {assume "bij_betw f (Field r) (Field r')"
+ with * 1 have "embed r' r (inv_into (Field r) f) "
+ using inv_into_Field_embed_bij_betw[of r r' f] by auto
+ with * have False by blast
+ }
+ ultimately show "(r,r') \<in> ordLess"
+ unfolding ordLess_def using * by (fastforce simp add: embedS_def)
+qed
+
+
+lemma ordLess_irreflexive: "\<not> r <o r"
+proof
+ assume "r <o r"
+ hence "Well_order r \<and> \<not>(\<exists>f. embed r r f)"
+ unfolding ordLess_iff ..
+ moreover have "embed r r id" using id_embed[of r] .
+ ultimately show False by blast
+qed
+
+
+lemma ordLeq_iff_ordLess_or_ordIso:
+"r \<le>o r' = (r <o r' \<or> r =o r')"
+unfolding ordRels_def embedS_defs iso_defs by blast
+
+
+lemma ordIso_iff_ordLeq:
+"(r =o r') = (r \<le>o r' \<and> r' \<le>o r)"
+proof
+ assume "r =o r'"
+ then obtain f where 1: "Well_order r \<and> Well_order r' \<and>
+ embed r r' f \<and> bij_betw f (Field r) (Field r')"
+ unfolding ordIso_def iso_defs by auto
+ hence "embed r r' f \<and> embed r' r (inv_into (Field r) f)"
+ by (simp add: inv_into_Field_embed_bij_betw)
+ thus "r \<le>o r' \<and> r' \<le>o r"
+ unfolding ordLeq_def using 1 by auto
+next
+ assume "r \<le>o r' \<and> r' \<le>o r"
+ then obtain f and g where 1: "Well_order r \<and> Well_order r' \<and>
+ embed r r' f \<and> embed r' r g"
+ unfolding ordLeq_def by auto
+ hence "iso r r' f" by (auto simp add: embed_bothWays_iso)
+ thus "r =o r'" unfolding ordIso_def using 1 by auto
+qed
+
+
+lemma not_ordLess_ordLeq:
+"r <o r' \<Longrightarrow> \<not> r' \<le>o r"
+using ordLess_ordLeq_trans ordLess_irreflexive by blast
+
+
+lemma ordLess_or_ordLeq:
+assumes WELL: "Well_order r" and WELL': "Well_order r'"
+shows "r <o r' \<or> r' \<le>o r"
+proof-
+ have "r \<le>o r' \<or> r' \<le>o r"
+ using assms by (simp add: ordLeq_total)
+ moreover
+ {assume "\<not> r <o r' \<and> r \<le>o r'"
+ hence "r =o r'" using ordLeq_iff_ordLess_or_ordIso by blast
+ hence "r' \<le>o r" using ordIso_symmetric ordIso_iff_ordLeq by blast
+ }
+ ultimately show ?thesis by blast
+qed
+
+
+lemma not_ordLess_ordIso:
+"r <o r' \<Longrightarrow> \<not> r =o r'"
+using assms ordLess_ordIso_trans ordIso_symmetric ordLess_irreflexive by blast
+
+
+lemma not_ordLeq_iff_ordLess:
+assumes WELL: "Well_order r" and WELL': "Well_order r'"
+shows "(\<not> r' \<le>o r) = (r <o r')"
+using assms not_ordLess_ordLeq ordLess_or_ordLeq by blast
+
+
+lemma not_ordLess_iff_ordLeq:
+assumes WELL: "Well_order r" and WELL': "Well_order r'"
+shows "(\<not> r' <o r) = (r \<le>o r')"
+using assms not_ordLess_ordLeq ordLess_or_ordLeq by blast
+
+
+lemma ordLess_transitive[trans]:
+"\<lbrakk>r <o r'; r' <o r''\<rbrakk> \<Longrightarrow> r <o r''"
+using assms ordLess_ordLeq_trans ordLeq_iff_ordLess_or_ordIso by blast
+
+
+corollary ordLess_trans: "trans ordLess"
+unfolding trans_def using ordLess_transitive by blast
+
+
+lemmas ordIso_equivalence = ordIso_transitive ordIso_reflexive ordIso_symmetric
+
+
+lemma ordIso_imp_ordLeq:
+"r =o r' \<Longrightarrow> r \<le>o r'"
+using ordIso_iff_ordLeq by blast
+
+
+lemma ordLess_imp_ordLeq:
+"r <o r' \<Longrightarrow> r \<le>o r'"
+using ordLeq_iff_ordLess_or_ordIso by blast
+
+
+lemma ofilter_subset_ordLeq:
+assumes WELL: "Well_order r" and
+ OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B"
+shows "(A \<le> B) = (Restr r A \<le>o Restr r B)"
+proof
+ assume "A \<le> B"
+ thus "Restr r A \<le>o Restr r B"
+ unfolding ordLeq_def using assms
+ Well_order_Restr Well_order_Restr ofilter_subset_embed by blast
+next
+ assume *: "Restr r A \<le>o Restr r B"
+ then obtain f where "embed (Restr r A) (Restr r B) f"
+ unfolding ordLeq_def by blast
+ {assume "B < A"
+ hence "Restr r B <o Restr r A"
+ unfolding ordLess_def using assms
+ Well_order_Restr Well_order_Restr ofilter_subset_embedS by blast
+ hence False using * not_ordLess_ordLeq by blast
+ }
+ thus "A \<le> B" using OFA OFB WELL
+ wo_rel_def[of r] wo_rel.ofilter_linord[of r A B] by blast
+qed
+
+
+lemma ofilter_subset_ordLess:
+assumes WELL: "Well_order r" and
+ OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B"
+shows "(A < B) = (Restr r A <o Restr r B)"
+proof-
+ let ?rA = "Restr r A" let ?rB = "Restr r B"
+ have 1: "Well_order ?rA \<and> Well_order ?rB"
+ using WELL Well_order_Restr by blast
+ have "(A < B) = (\<not> B \<le> A)" using assms
+ wo_rel_def wo_rel.ofilter_linord[of r A B] by blast
+ also have "\<dots> = (\<not> Restr r B \<le>o Restr r A)"
+ using assms ofilter_subset_ordLeq by blast
+ also have "\<dots> = (Restr r A <o Restr r B)"
+ using 1 not_ordLeq_iff_ordLess by blast
+ finally show ?thesis .
+qed
+
+
+lemma ofilter_ordLess:
+"\<lbrakk>Well_order r; wo_rel.ofilter r A\<rbrakk> \<Longrightarrow> (A < Field r) = (Restr r A <o r)"
+by (simp add: ofilter_subset_ordLess wo_rel.Field_ofilter
+ wo_rel_def Restr_Field)
+
+
+corollary underS_Restr_ordLess:
+assumes "Well_order r" and "Field r \<noteq> {}"
+shows "Restr r (rel.underS r a) <o r"
+proof-
+ have "rel.underS r a < Field r" using assms
+ by (simp add: rel.underS_Field3)
+ thus ?thesis using assms
+ by (simp add: ofilter_ordLess wo_rel.underS_ofilter wo_rel_def)
+qed
+
+
+lemma embed_ordLess_ofilterIncl:
+assumes
+ OL12: "r1 <o r2" and OL23: "r2 <o r3" and
+ EMB13: "embed r1 r3 f13" and EMB23: "embed r2 r3 f23"
+shows "(f13`(Field r1), f23`(Field r2)) \<in> (ofilterIncl r3)"
+proof-
+ have OL13: "r1 <o r3"
+ using OL12 OL23 using ordLess_transitive by auto
+ let ?A1 = "Field r1" let ?A2 ="Field r2" let ?A3 ="Field r3"
+ obtain f12 g23 where
+ 0: "Well_order r1 \<and> Well_order r2 \<and> Well_order r3" and
+ 1: "embed r1 r2 f12 \<and> \<not>(bij_betw f12 ?A1 ?A2)" and
+ 2: "embed r2 r3 g23 \<and> \<not>(bij_betw g23 ?A2 ?A3)"
+ using OL12 OL23 by (auto simp add: ordLess_def embedS_def)
+ hence "\<forall>a \<in> ?A2. f23 a = g23 a"
+ using EMB23 embed_unique[of r2 r3] by blast
+ hence 3: "\<not>(bij_betw f23 ?A2 ?A3)"
+ using 2 bij_betw_cong[of ?A2 f23 g23] by blast
+ (* *)
+ have 4: "wo_rel.ofilter r2 (f12 ` ?A1) \<and> f12 ` ?A1 \<noteq> ?A2"
+ using 0 1 OL12 by (simp add: embed_Field_ofilter ordLess_Field)
+ have 5: "wo_rel.ofilter r3 (f23 ` ?A2) \<and> f23 ` ?A2 \<noteq> ?A3"
+ using 0 EMB23 OL23 by (simp add: embed_Field_ofilter ordLess_Field)
+ have 6: "wo_rel.ofilter r3 (f13 ` ?A1) \<and> f13 ` ?A1 \<noteq> ?A3"
+ using 0 EMB13 OL13 by (simp add: embed_Field_ofilter ordLess_Field)
+ (* *)
+ have "f12 ` ?A1 < ?A2"
+ using 0 4 by (auto simp add: wo_rel_def wo_rel.ofilter_def)
+ moreover have "inj_on f23 ?A2"
+ using EMB23 0 by (simp add: wo_rel_def embed_inj_on)
+ ultimately
+ have "f23 ` (f12 ` ?A1) < f23 ` ?A2" by (simp add: inj_on_strict_subset)
+ moreover
+ {have "embed r1 r3 (f23 o f12)"
+ using 1 EMB23 0 by (auto simp add: comp_embed)
+ hence "\<forall>a \<in> ?A1. f23(f12 a) = f13 a"
+ using EMB13 0 embed_unique[of r1 r3 "f23 o f12" f13] by auto
+ hence "f23 ` (f12 ` ?A1) = f13 ` ?A1" by force
+ }
+ ultimately
+ have "f13 ` ?A1 < f23 ` ?A2" by simp
+ (* *)
+ with 5 6 show ?thesis
+ unfolding ofilterIncl_def by auto
+qed
+
+
+lemma ordLess_iff_ordIso_Restr:
+assumes WELL: "Well_order r" and WELL': "Well_order r'"
+shows "(r' <o r) = (\<exists>a \<in> Field r. r' =o Restr r (rel.underS r a))"
+proof(auto)
+ fix a assume *: "a \<in> Field r" and **: "r' =o Restr r (rel.underS r a)"
+ hence "Restr r (rel.underS r a) <o r" using WELL underS_Restr_ordLess[of r] by blast
+ thus "r' <o r" using ** ordIso_ordLess_trans by blast
+next
+ assume "r' <o r"
+ then obtain f where 1: "Well_order r \<and> Well_order r'" and
+ 2: "embed r' r f \<and> f ` (Field r') \<noteq> Field r"
+ unfolding ordLess_def embedS_def[abs_def] bij_betw_def using embed_inj_on by blast
+ hence "wo_rel.ofilter r (f ` (Field r'))" using embed_Field_ofilter by blast
+ then obtain a where 3: "a \<in> Field r" and 4: "rel.underS r a = f ` (Field r')"
+ using 1 2 by (auto simp add: wo_rel.ofilter_underS_Field wo_rel_def)
+ have "iso r' (Restr r (f ` (Field r'))) f"
+ using embed_implies_iso_Restr 2 assms by blast
+ moreover have "Well_order (Restr r (f ` (Field r')))"
+ using WELL Well_order_Restr by blast
+ ultimately have "r' =o Restr r (f ` (Field r'))"
+ using WELL' unfolding ordIso_def by auto
+ hence "r' =o Restr r (rel.underS r a)" using 4 by auto
+ thus "\<exists>a \<in> Field r. r' =o Restr r (rel.underS r a)" using 3 by auto
+qed
+
+
+lemma internalize_ordLess:
+"(r' <o r) = (\<exists>p. Field p < Field r \<and> r' =o p \<and> p <o r)"
+proof
+ assume *: "r' <o r"
+ hence 0: "Well_order r \<and> Well_order r'" unfolding ordLess_def by auto
+ with * obtain a where 1: "a \<in> Field r" and 2: "r' =o Restr r (rel.underS r a)"
+ using ordLess_iff_ordIso_Restr by blast
+ let ?p = "Restr r (rel.underS r a)"
+ have "wo_rel.ofilter r (rel.underS r a)" using 0
+ by (simp add: wo_rel_def wo_rel.underS_ofilter)
+ hence "Field ?p = rel.underS r a" using 0 Field_Restr_ofilter by blast
+ hence "Field ?p < Field r" using rel.underS_Field2 1 by fastforce
+ moreover have "?p <o r" using underS_Restr_ordLess[of r a] 0 1 by blast
+ ultimately
+ show "\<exists>p. Field p < Field r \<and> r' =o p \<and> p <o r" using 2 by blast
+next
+ assume "\<exists>p. Field p < Field r \<and> r' =o p \<and> p <o r"
+ thus "r' <o r" using ordIso_ordLess_trans by blast
+qed
+
+
+lemma internalize_ordLeq:
+"(r' \<le>o r) = (\<exists>p. Field p \<le> Field r \<and> r' =o p \<and> p \<le>o r)"
+proof
+ assume *: "r' \<le>o r"
+ moreover
+ {assume "r' <o r"
+ then obtain p where "Field p < Field r \<and> r' =o p \<and> p <o r"
+ using internalize_ordLess[of r' r] by blast
+ hence "\<exists>p. Field p \<le> Field r \<and> r' =o p \<and> p \<le>o r"
+ using ordLeq_iff_ordLess_or_ordIso by blast
+ }
+ moreover
+ have "r \<le>o r" using * ordLeq_def ordLeq_reflexive by blast
+ ultimately show "\<exists>p. Field p \<le> Field r \<and> r' =o p \<and> p \<le>o r"
+ using ordLeq_iff_ordLess_or_ordIso by blast
+next
+ assume "\<exists>p. Field p \<le> Field r \<and> r' =o p \<and> p \<le>o r"
+ thus "r' \<le>o r" using ordIso_ordLeq_trans by blast
+qed
+
+
+lemma ordLeq_iff_ordLess_Restr:
+assumes WELL: "Well_order r" and WELL': "Well_order r'"
+shows "(r \<le>o r') = (\<forall>a \<in> Field r. Restr r (rel.underS r a) <o r')"
+proof(auto)
+ assume *: "r \<le>o r'"
+ fix a assume "a \<in> Field r"
+ hence "Restr r (rel.underS r a) <o r"
+ using WELL underS_Restr_ordLess[of r] by blast
+ thus "Restr r (rel.underS r a) <o r'"
+ using * ordLess_ordLeq_trans by blast
+next
+ assume *: "\<forall>a \<in> Field r. Restr r (rel.underS r a) <o r'"
+ {assume "r' <o r"
+ then obtain a where "a \<in> Field r \<and> r' =o Restr r (rel.underS r a)"
+ using assms ordLess_iff_ordIso_Restr by blast
+ hence False using * not_ordLess_ordIso ordIso_symmetric by blast
+ }
+ thus "r \<le>o r'" using ordLess_or_ordLeq assms by blast
+qed
+
+
+lemma finite_ordLess_infinite:
+assumes WELL: "Well_order r" and WELL': "Well_order r'" and
+ FIN: "finite(Field r)" and INF: "infinite(Field r')"
+shows "r <o r'"
+proof-
+ {assume "r' \<le>o r"
+ then obtain h where "inj_on h (Field r') \<and> h ` (Field r') \<le> Field r"
+ unfolding ordLeq_def using assms embed_inj_on embed_Field by blast
+ hence False using finite_imageD finite_subset FIN INF by blast
+ }
+ thus ?thesis using WELL WELL' ordLess_or_ordLeq by blast
+qed
+
+
+lemma finite_well_order_on_ordIso:
+assumes FIN: "finite A" and
+ WELL: "well_order_on A r" and WELL': "well_order_on A r'"
+shows "r =o r'"
+proof-
+ have 0: "Well_order r \<and> Well_order r' \<and> Field r = A \<and> Field r' = A"
+ using assms rel.well_order_on_Well_order by blast
+ moreover
+ have "\<forall>r r'. well_order_on A r \<and> well_order_on A r' \<and> r \<le>o r'
+ \<longrightarrow> r =o r'"
+ proof(clarify)
+ fix r r' assume *: "well_order_on A r" and **: "well_order_on A r'"
+ have 2: "Well_order r \<and> Well_order r' \<and> Field r = A \<and> Field r' = A"
+ using * ** rel.well_order_on_Well_order by blast
+ assume "r \<le>o r'"
+ then obtain f where 1: "embed r r' f" and
+ "inj_on f A \<and> f ` A \<le> A"
+ unfolding ordLeq_def using 2 embed_inj_on embed_Field by blast
+ hence "bij_betw f A A" unfolding bij_betw_def using FIN endo_inj_surj by blast
+ thus "r =o r'" unfolding ordIso_def iso_def[abs_def] using 1 2 by auto
+ qed
+ ultimately show ?thesis using assms ordLeq_total ordIso_symmetric by blast
+qed
+
+
+subsection{* @{text "<o"} is well-founded *}
+
+
+text {* Of course, it only makes sense to state that the @{text "<o"} is well-founded
+on the restricted type @{text "'a rel rel"}. We prove this by first showing that, for any set
+of well-orders all embedded in a fixed well-order, the function mapping each well-order
+in the set to an order filter of the fixed well-order is compatible w.r.t. to @{text "<o"} versus
+{\em strict inclusion}; and we already know that strict inclusion of order filters is well-founded. *}
+
+
+definition ord_to_filter :: "'a rel \<Rightarrow> 'a rel \<Rightarrow> 'a set"
+where "ord_to_filter r0 r \<equiv> (SOME f. embed r r0 f) ` (Field r)"
+
+
+lemma ord_to_filter_compat:
+"compat (ordLess Int (ordLess^-1``{r0} \<times> ordLess^-1``{r0}))
+ (ofilterIncl r0)
+ (ord_to_filter r0)"
+proof(unfold compat_def ord_to_filter_def, clarify)
+ fix r1::"'a rel" and r2::"'a rel"
+ let ?A1 = "Field r1" let ?A2 ="Field r2" let ?A0 ="Field r0"
+ let ?phi10 = "\<lambda> f10. embed r1 r0 f10" let ?f10 = "SOME f. ?phi10 f"
+ let ?phi20 = "\<lambda> f20. embed r2 r0 f20" let ?f20 = "SOME f. ?phi20 f"
+ assume *: "r1 <o r0" "r2 <o r0" and **: "r1 <o r2"
+ hence "(\<exists>f. ?phi10 f) \<and> (\<exists>f. ?phi20 f)"
+ by (auto simp add: ordLess_def embedS_def)
+ hence "?phi10 ?f10 \<and> ?phi20 ?f20" by (auto simp add: someI_ex)
+ thus "(?f10 ` ?A1, ?f20 ` ?A2) \<in> ofilterIncl r0"
+ using * ** by (simp add: embed_ordLess_ofilterIncl)
+qed
+
+
+theorem wf_ordLess: "wf ordLess"
+proof-
+ {fix r0 :: "('a \<times> 'a) set"
+ (* need to annotate here!*)
+ let ?ordLess = "ordLess::('d rel * 'd rel) set"
+ let ?R = "?ordLess Int (?ordLess^-1``{r0} \<times> ?ordLess^-1``{r0})"
+ {assume Case1: "Well_order r0"
+ hence "wf ?R"
+ using wf_ofilterIncl[of r0]
+ compat_wf[of ?R "ofilterIncl r0" "ord_to_filter r0"]
+ ord_to_filter_compat[of r0] by auto
+ }
+ moreover
+ {assume Case2: "\<not> Well_order r0"
+ hence "?R = {}" unfolding ordLess_def by auto
+ hence "wf ?R" using wf_empty by simp
+ }
+ ultimately have "wf ?R" by blast
+ }
+ thus ?thesis by (simp add: trans_wf_iff ordLess_trans)
+qed
+
+corollary exists_minim_Well_order:
+assumes NE: "R \<noteq> {}" and WELL: "\<forall>r \<in> R. Well_order r"
+shows "\<exists>r \<in> R. \<forall>r' \<in> R. r \<le>o r'"
+proof-
+ obtain r where "r \<in> R \<and> (\<forall>r' \<in> R. \<not> r' <o r)"
+ using NE spec[OF spec[OF subst[OF wf_eq_minimal, of "%x. x", OF wf_ordLess]], of _ R]
+ equals0I[of R] by blast
+ with not_ordLeq_iff_ordLess WELL show ?thesis by blast
+qed
+
+
+
+subsection {* Copy via direct images *}
+
+
+text{* The direct image operator is the dual of the inverse image operator @{text "inv_image"}
+from @{text "Relation.thy"}. It is useful for transporting a well-order between
+different types. *}
+
+
+definition dir_image :: "'a rel \<Rightarrow> ('a \<Rightarrow> 'a') \<Rightarrow> 'a' rel"
+where
+"dir_image r f = {(f a, f b)| a b. (a,b) \<in> r}"
+
+
+lemma dir_image_Field:
+"Field(dir_image r f) \<le> f ` (Field r)"
+unfolding dir_image_def Field_def by auto
+
+
+lemma dir_image_minus_Id:
+"inj_on f (Field r) \<Longrightarrow> (dir_image r f) - Id = dir_image (r - Id) f"
+unfolding inj_on_def Field_def dir_image_def by auto
+
+
+lemma Refl_dir_image:
+assumes "Refl r"
+shows "Refl(dir_image r f)"
+proof-
+ {fix a' b'
+ assume "(a',b') \<in> dir_image r f"
+ then obtain a b where 1: "a' = f a \<and> b' = f b \<and> (a,b) \<in> r"
+ unfolding dir_image_def by blast
+ hence "a \<in> Field r \<and> b \<in> Field r" using Field_def by fastforce
+ hence "(a,a) \<in> r \<and> (b,b) \<in> r" using assms by (simp add: refl_on_def)
+ with 1 have "(a',a') \<in> dir_image r f \<and> (b',b') \<in> dir_image r f"
+ unfolding dir_image_def by auto
+ }
+ thus ?thesis
+ by(unfold refl_on_def Field_def Domain_def Range_def, auto)
+qed
+
+
+lemma trans_dir_image:
+assumes TRANS: "trans r" and INJ: "inj_on f (Field r)"
+shows "trans(dir_image r f)"
+proof(unfold trans_def, auto)
+ fix a' b' c'
+ assume "(a',b') \<in> dir_image r f" "(b',c') \<in> dir_image r f"
+ then obtain a b1 b2 c where 1: "a' = f a \<and> b' = f b1 \<and> b' = f b2 \<and> c' = f c" and
+ 2: "(a,b1) \<in> r \<and> (b2,c) \<in> r"
+ unfolding dir_image_def by blast
+ hence "b1 \<in> Field r \<and> b2 \<in> Field r"
+ unfolding Field_def by auto
+ hence "b1 = b2" using 1 INJ unfolding inj_on_def by auto
+ hence "(a,c): r" using 2 TRANS unfolding trans_def by blast
+ thus "(a',c') \<in> dir_image r f"
+ unfolding dir_image_def using 1 by auto
+qed
+
+
+lemma Preorder_dir_image:
+"\<lbrakk>Preorder r; inj_on f (Field r)\<rbrakk> \<Longrightarrow> Preorder (dir_image r f)"
+by (simp add: preorder_on_def Refl_dir_image trans_dir_image)
+
+
+lemma antisym_dir_image:
+assumes AN: "antisym r" and INJ: "inj_on f (Field r)"
+shows "antisym(dir_image r f)"
+proof(unfold antisym_def, auto)
+ fix a' b'
+ assume "(a',b') \<in> dir_image r f" "(b',a') \<in> dir_image r f"
+ then obtain a1 b1 a2 b2 where 1: "a' = f a1 \<and> a' = f a2 \<and> b' = f b1 \<and> b' = f b2" and
+ 2: "(a1,b1) \<in> r \<and> (b2,a2) \<in> r " and
+ 3: "{a1,a2,b1,b2} \<le> Field r"
+ unfolding dir_image_def Field_def by blast
+ hence "a1 = a2 \<and> b1 = b2" using INJ unfolding inj_on_def by auto
+ hence "a1 = b2" using 2 AN unfolding antisym_def by auto
+ thus "a' = b'" using 1 by auto
+qed
+
+
+lemma Partial_order_dir_image:
+"\<lbrakk>Partial_order r; inj_on f (Field r)\<rbrakk> \<Longrightarrow> Partial_order (dir_image r f)"
+by (simp add: partial_order_on_def Preorder_dir_image antisym_dir_image)
+
+
+lemma Total_dir_image:
+assumes TOT: "Total r" and INJ: "inj_on f (Field r)"
+shows "Total(dir_image r f)"
+proof(unfold total_on_def, intro ballI impI)
+ fix a' b'
+ assume "a' \<in> Field (dir_image r f)" "b' \<in> Field (dir_image r f)"
+ then obtain a and b where 1: "a \<in> Field r \<and> b \<in> Field r \<and> f a = a' \<and> f b = b'"
+ using dir_image_Field[of r f] by blast
+ moreover assume "a' \<noteq> b'"
+ ultimately have "a \<noteq> b" using INJ unfolding inj_on_def by auto
+ hence "(a,b) \<in> r \<or> (b,a) \<in> r" using 1 TOT unfolding total_on_def by auto
+ thus "(a',b') \<in> dir_image r f \<or> (b',a') \<in> dir_image r f"
+ using 1 unfolding dir_image_def by auto
+qed
+
+
+lemma Linear_order_dir_image:
+"\<lbrakk>Linear_order r; inj_on f (Field r)\<rbrakk> \<Longrightarrow> Linear_order (dir_image r f)"
+by (simp add: linear_order_on_def Partial_order_dir_image Total_dir_image)
+
+
+lemma wf_dir_image:
+assumes WF: "wf r" and INJ: "inj_on f (Field r)"
+shows "wf(dir_image r f)"
+proof(unfold wf_eq_minimal2, intro allI impI, elim conjE)
+ fix A'::"'b set"
+ assume SUB: "A' \<le> Field(dir_image r f)" and NE: "A' \<noteq> {}"
+ obtain A where A_def: "A = {a \<in> Field r. f a \<in> A'}" by blast
+ have "A \<noteq> {} \<and> A \<le> Field r"
+ using A_def dir_image_Field[of r f] SUB NE by blast
+ then obtain a where 1: "a \<in> A \<and> (\<forall>b \<in> A. (b,a) \<notin> r)"
+ using WF unfolding wf_eq_minimal2 by blast
+ have "\<forall>b' \<in> A'. (b',f a) \<notin> dir_image r f"
+ proof(clarify)
+ fix b' assume *: "b' \<in> A'" and **: "(b',f a) \<in> dir_image r f"
+ obtain b1 a1 where 2: "b' = f b1 \<and> f a = f a1" and
+ 3: "(b1,a1) \<in> r \<and> {a1,b1} \<le> Field r"
+ using ** unfolding dir_image_def Field_def by blast
+ hence "a = a1" using 1 A_def INJ unfolding inj_on_def by auto
+ hence "b1 \<in> A \<and> (b1,a) \<in> r" using 2 3 A_def * by auto
+ with 1 show False by auto
+ qed
+ thus "\<exists>a'\<in>A'. \<forall>b'\<in>A'. (b', a') \<notin> dir_image r f"
+ using A_def 1 by blast
+qed
+
+
+lemma Well_order_dir_image:
+"\<lbrakk>Well_order r; inj_on f (Field r)\<rbrakk> \<Longrightarrow> Well_order (dir_image r f)"
+using assms unfolding well_order_on_def
+using Linear_order_dir_image[of r f] wf_dir_image[of "r - Id" f]
+ dir_image_minus_Id[of f r]
+ subset_inj_on[of f "Field r" "Field(r - Id)"]
+ mono_Field[of "r - Id" r] by auto
+
+
+lemma dir_image_Field2:
+"Refl r \<Longrightarrow> Field(dir_image r f) = f ` (Field r)"
+unfolding Field_def dir_image_def refl_on_def Domain_def Range_def by blast
+
+
+lemma dir_image_bij_betw:
+"\<lbrakk>Well_order r; inj_on f (Field r)\<rbrakk> \<Longrightarrow> bij_betw f (Field r) (Field (dir_image r f))"
+unfolding bij_betw_def
+by (simp add: dir_image_Field2 order_on_defs)
+
+
+lemma dir_image_compat:
+"compat r (dir_image r f) f"
+unfolding compat_def dir_image_def by auto
+
+
+lemma dir_image_iso:
+"\<lbrakk>Well_order r; inj_on f (Field r)\<rbrakk> \<Longrightarrow> iso r (dir_image r f) f"
+using iso_iff3 dir_image_compat dir_image_bij_betw Well_order_dir_image by blast
+
+
+lemma dir_image_ordIso:
+"\<lbrakk>Well_order r; inj_on f (Field r)\<rbrakk> \<Longrightarrow> r =o dir_image r f"
+unfolding ordIso_def using dir_image_iso Well_order_dir_image by blast
+
+
+lemma Well_order_iso_copy:
+assumes WELL: "well_order_on A r" and BIJ: "bij_betw f A A'"
+shows "\<exists>r'. well_order_on A' r' \<and> r =o r'"
+proof-
+ let ?r' = "dir_image r f"
+ have 1: "A = Field r \<and> Well_order r"
+ using WELL rel.well_order_on_Well_order by blast
+ hence 2: "iso r ?r' f"
+ using dir_image_iso using BIJ unfolding bij_betw_def by auto
+ hence "f ` (Field r) = Field ?r'" using 1 iso_iff[of r ?r'] by blast
+ hence "Field ?r' = A'"
+ using 1 BIJ unfolding bij_betw_def by auto
+ moreover have "Well_order ?r'"
+ using 1 Well_order_dir_image BIJ unfolding bij_betw_def by blast
+ ultimately show ?thesis unfolding ordIso_def using 1 2 by blast
+qed
+
+
+
+subsection {* Bounded square *}
+
+
+text{* This construction essentially defines, for an order relation @{text "r"}, a lexicographic
+order @{text "bsqr r"} on @{text "(Field r) \<times> (Field r)"}, applying the
+following criteria (in this order):
+\begin{itemize}
+\item compare the maximums;
+\item compare the first components;
+\item compare the second components.
+\end{itemize}
+%
+The only application of this construction that we are aware of is
+at proving that the square of an infinite set has the same cardinal
+as that set. The essential property required there (and which is ensured by this
+construction) is that any proper order filter of the product order is included in a rectangle, i.e.,
+in a product of proper filters on the original relation (assumed to be a well-order). *}
+
+
+definition bsqr :: "'a rel => ('a * 'a)rel"
+where
+"bsqr r = {((a1,a2),(b1,b2)).
+ {a1,a2,b1,b2} \<le> Field r \<and>
+ (a1 = b1 \<and> a2 = b2 \<or>
+ (wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r - Id \<or>
+ wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> (a1,b1) \<in> r - Id \<or>
+ wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> a1 = b1 \<and> (a2,b2) \<in> r - Id
+ )}"
+
+
+lemma Field_bsqr:
+"Field (bsqr r) = Field r \<times> Field r"
+proof
+ show "Field (bsqr r) \<le> Field r \<times> Field r"
+ proof-
+ {fix a1 a2 assume "(a1,a2) \<in> Field (bsqr r)"
+ moreover
+ have "\<And> b1 b2. ((a1,a2),(b1,b2)) \<in> bsqr r \<or> ((b1,b2),(a1,a2)) \<in> bsqr r \<Longrightarrow>
+ a1 \<in> Field r \<and> a2 \<in> Field r" unfolding bsqr_def by auto
+ ultimately have "a1 \<in> Field r \<and> a2 \<in> Field r" unfolding Field_def by auto
+ }
+ thus ?thesis unfolding Field_def by force
+ qed
+next
+ show "Field r \<times> Field r \<le> Field (bsqr r)"
+ proof(auto)
+ fix a1 a2 assume "a1 \<in> Field r" and "a2 \<in> Field r"
+ hence "((a1,a2),(a1,a2)) \<in> bsqr r" unfolding bsqr_def by blast
+ thus "(a1,a2) \<in> Field (bsqr r)" unfolding Field_def by auto
+ qed
+qed
+
+
+lemma bsqr_Refl: "Refl(bsqr r)"
+by(unfold refl_on_def Field_bsqr, auto simp add: bsqr_def)
+
+
+lemma bsqr_Trans:
+assumes "Well_order r"
+shows "trans (bsqr r)"
+proof(unfold trans_def, auto)
+ (* Preliminary facts *)
+ have Well: "wo_rel r" using assms wo_rel_def by auto
+ hence Trans: "trans r" using wo_rel.TRANS by auto
+ have Anti: "antisym r" using wo_rel.ANTISYM Well by auto
+ hence TransS: "trans(r - Id)" using Trans by (simp add: trans_diff_Id)
+ (* Main proof *)
+ fix a1 a2 b1 b2 c1 c2
+ assume *: "((a1,a2),(b1,b2)) \<in> bsqr r" and **: "((b1,b2),(c1,c2)) \<in> bsqr r"
+ hence 0: "{a1,a2,b1,b2,c1,c2} \<le> Field r" unfolding bsqr_def by auto
+ have 1: "a1 = b1 \<and> a2 = b2 \<or> (wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r - Id \<or>
+ wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> (a1,b1) \<in> r - Id \<or>
+ wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> a1 = b1 \<and> (a2,b2) \<in> r - Id"
+ using * unfolding bsqr_def by auto
+ have 2: "b1 = c1 \<and> b2 = c2 \<or> (wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \<in> r - Id \<or>
+ wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> (b1,c1) \<in> r - Id \<or>
+ wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> b1 = c1 \<and> (b2,c2) \<in> r - Id"
+ using ** unfolding bsqr_def by auto
+ show "((a1,a2),(c1,c2)) \<in> bsqr r"
+ proof-
+ {assume Case1: "a1 = b1 \<and> a2 = b2"
+ hence ?thesis using ** by simp
+ }
+ moreover
+ {assume Case2: "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r - Id"
+ {assume Case21: "b1 = c1 \<and> b2 = c2"
+ hence ?thesis using * by simp
+ }
+ moreover
+ {assume Case22: "(wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \<in> r - Id"
+ hence "(wo_rel.max2 r a1 a2, wo_rel.max2 r c1 c2) \<in> r - Id"
+ using Case2 TransS trans_def[of "r - Id"] by blast
+ hence ?thesis using 0 unfolding bsqr_def by auto
+ }
+ moreover
+ {assume Case23_4: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2"
+ hence ?thesis using Case2 0 unfolding bsqr_def by auto
+ }
+ ultimately have ?thesis using 0 2 by auto
+ }
+ moreover
+ {assume Case3: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> (a1,b1) \<in> r - Id"
+ {assume Case31: "b1 = c1 \<and> b2 = c2"
+ hence ?thesis using * by simp
+ }
+ moreover
+ {assume Case32: "(wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \<in> r - Id"
+ hence ?thesis using Case3 0 unfolding bsqr_def by auto
+ }
+ moreover
+ {assume Case33: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> (b1,c1) \<in> r - Id"
+ hence "(a1,c1) \<in> r - Id"
+ using Case3 TransS trans_def[of "r - Id"] by blast
+ hence ?thesis using Case3 Case33 0 unfolding bsqr_def by auto
+ }
+ moreover
+ {assume Case33: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> b1 = c1"
+ hence ?thesis using Case3 0 unfolding bsqr_def by auto
+ }
+ ultimately have ?thesis using 0 2 by auto
+ }
+ moreover
+ {assume Case4: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> a1 = b1 \<and> (a2,b2) \<in> r - Id"
+ {assume Case41: "b1 = c1 \<and> b2 = c2"
+ hence ?thesis using * by simp
+ }
+ moreover
+ {assume Case42: "(wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \<in> r - Id"
+ hence ?thesis using Case4 0 unfolding bsqr_def by auto
+ }
+ moreover
+ {assume Case43: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> (b1,c1) \<in> r - Id"
+ hence ?thesis using Case4 0 unfolding bsqr_def by auto
+ }
+ moreover
+ {assume Case44: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> b1 = c1 \<and> (b2,c2) \<in> r - Id"
+ hence "(a2,c2) \<in> r - Id"
+ using Case4 TransS trans_def[of "r - Id"] by blast
+ hence ?thesis using Case4 Case44 0 unfolding bsqr_def by auto
+ }
+ ultimately have ?thesis using 0 2 by auto
+ }
+ ultimately show ?thesis using 0 1 by auto
+ qed
+qed
+
+
+lemma bsqr_antisym:
+assumes "Well_order r"
+shows "antisym (bsqr r)"
+proof(unfold antisym_def, clarify)
+ (* Preliminary facts *)
+ have Well: "wo_rel r" using assms wo_rel_def by auto
+ hence Trans: "trans r" using wo_rel.TRANS by auto
+ have Anti: "antisym r" using wo_rel.ANTISYM Well by auto
+ hence TransS: "trans(r - Id)" using Trans by (simp add: trans_diff_Id)
+ hence IrrS: "\<forall>a b. \<not>((a,b) \<in> r - Id \<and> (b,a) \<in> r - Id)"
+ using Anti trans_def[of "r - Id"] antisym_def[of "r - Id"] by blast
+ (* Main proof *)
+ fix a1 a2 b1 b2
+ assume *: "((a1,a2),(b1,b2)) \<in> bsqr r" and **: "((b1,b2),(a1,a2)) \<in> bsqr r"
+ hence 0: "{a1,a2,b1,b2} \<le> Field r" unfolding bsqr_def by auto
+ have 1: "a1 = b1 \<and> a2 = b2 \<or> (wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r - Id \<or>
+ wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> (a1,b1) \<in> r - Id \<or>
+ wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> a1 = b1 \<and> (a2,b2) \<in> r - Id"
+ using * unfolding bsqr_def by auto
+ have 2: "b1 = a1 \<and> b2 = a2 \<or> (wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \<in> r - Id \<or>
+ wo_rel.max2 r b1 b2 = wo_rel.max2 r a1 a2 \<and> (b1,a1) \<in> r - Id \<or>
+ wo_rel.max2 r b1 b2 = wo_rel.max2 r a1 a2 \<and> b1 = a1 \<and> (b2,a2) \<in> r - Id"
+ using ** unfolding bsqr_def by auto
+ show "a1 = b1 \<and> a2 = b2"
+ proof-
+ {assume Case1: "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r - Id"
+ {assume Case11: "(wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \<in> r - Id"
+ hence False using Case1 IrrS by blast
+ }
+ moreover
+ {assume Case12_3: "wo_rel.max2 r b1 b2 = wo_rel.max2 r a1 a2"
+ hence False using Case1 by auto
+ }
+ ultimately have ?thesis using 0 2 by auto
+ }
+ moreover
+ {assume Case2: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> (a1,b1) \<in> r - Id"
+ {assume Case21: "(wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \<in> r - Id"
+ hence False using Case2 by auto
+ }
+ moreover
+ {assume Case22: "(b1,a1) \<in> r - Id"
+ hence False using Case2 IrrS by blast
+ }
+ moreover
+ {assume Case23: "b1 = a1"
+ hence False using Case2 by auto
+ }
+ ultimately have ?thesis using 0 2 by auto
+ }
+ moreover
+ {assume Case3: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> a1 = b1 \<and> (a2,b2) \<in> r - Id"
+ moreover
+ {assume Case31: "(wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \<in> r - Id"
+ hence False using Case3 by auto
+ }
+ moreover
+ {assume Case32: "(b1,a1) \<in> r - Id"
+ hence False using Case3 by auto
+ }
+ moreover
+ {assume Case33: "(b2,a2) \<in> r - Id"
+ hence False using Case3 IrrS by blast
+ }
+ ultimately have ?thesis using 0 2 by auto
+ }
+ ultimately show ?thesis using 0 1 by blast
+ qed
+qed
+
+
+lemma bsqr_Total:
+assumes "Well_order r"
+shows "Total(bsqr r)"
+proof-
+ (* Preliminary facts *)
+ have Well: "wo_rel r" using assms wo_rel_def by auto
+ hence Total: "\<forall>a \<in> Field r. \<forall>b \<in> Field r. (a,b) \<in> r \<or> (b,a) \<in> r"
+ using wo_rel.TOTALS by auto
+ (* Main proof *)
+ {fix a1 a2 b1 b2 assume "{(a1,a2), (b1,b2)} \<le> Field(bsqr r)"
+ hence 0: "a1 \<in> Field r \<and> a2 \<in> Field r \<and> b1 \<in> Field r \<and> b2 \<in> Field r"
+ using Field_bsqr by blast
+ have "((a1,a2) = (b1,b2) \<or> ((a1,a2),(b1,b2)) \<in> bsqr r \<or> ((b1,b2),(a1,a2)) \<in> bsqr r)"
+ proof(rule wo_rel.cases_Total[of r a1 a2], clarsimp simp add: Well, simp add: 0)
+ (* Why didn't clarsimp simp add: Well 0 do the same job? *)
+ assume Case1: "(a1,a2) \<in> r"
+ hence 1: "wo_rel.max2 r a1 a2 = a2"
+ using Well 0 by (simp add: wo_rel.max2_equals2)
+ show ?thesis
+ proof(rule wo_rel.cases_Total[of r b1 b2], clarsimp simp add: Well, simp add: 0)
+ assume Case11: "(b1,b2) \<in> r"
+ hence 2: "wo_rel.max2 r b1 b2 = b2"
+ using Well 0 by (simp add: wo_rel.max2_equals2)
+ show ?thesis
+ proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0)
+ assume Case111: "(a2,b2) \<in> r - Id \<or> (b2,a2) \<in> r - Id"
+ thus ?thesis using 0 1 2 unfolding bsqr_def by auto
+ next
+ assume Case112: "a2 = b2"
+ show ?thesis
+ proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0)
+ assume Case1121: "(a1,b1) \<in> r - Id \<or> (b1,a1) \<in> r - Id"
+ thus ?thesis using 0 1 2 Case112 unfolding bsqr_def by auto
+ next
+ assume Case1122: "a1 = b1"
+ thus ?thesis using Case112 by auto
+ qed
+ qed
+ next
+ assume Case12: "(b2,b1) \<in> r"
+ hence 3: "wo_rel.max2 r b1 b2 = b1" using Well 0 by (simp add: wo_rel.max2_equals1)
+ show ?thesis
+ proof(rule wo_rel.cases_Total3[of r a2 b1], clarsimp simp add: Well, simp add: 0)
+ assume Case121: "(a2,b1) \<in> r - Id \<or> (b1,a2) \<in> r - Id"
+ thus ?thesis using 0 1 3 unfolding bsqr_def by auto
+ next
+ assume Case122: "a2 = b1"
+ show ?thesis
+ proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0)
+ assume Case1221: "(a1,b1) \<in> r - Id \<or> (b1,a1) \<in> r - Id"
+ thus ?thesis using 0 1 3 Case122 unfolding bsqr_def by auto
+ next
+ assume Case1222: "a1 = b1"
+ show ?thesis
+ proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0)
+ assume Case12221: "(a2,b2) \<in> r - Id \<or> (b2,a2) \<in> r - Id"
+ thus ?thesis using 0 1 3 Case122 Case1222 unfolding bsqr_def by auto
+ next
+ assume Case12222: "a2 = b2"
+ thus ?thesis using Case122 Case1222 by auto
+ qed
+ qed
+ qed
+ qed
+ next
+ assume Case2: "(a2,a1) \<in> r"
+ hence 1: "wo_rel.max2 r a1 a2 = a1" using Well 0 by (simp add: wo_rel.max2_equals1)
+ show ?thesis
+ proof(rule wo_rel.cases_Total[of r b1 b2], clarsimp simp add: Well, simp add: 0)
+ assume Case21: "(b1,b2) \<in> r"
+ hence 2: "wo_rel.max2 r b1 b2 = b2" using Well 0 by (simp add: wo_rel.max2_equals2)
+ show ?thesis
+ proof(rule wo_rel.cases_Total3[of r a1 b2], clarsimp simp add: Well, simp add: 0)
+ assume Case211: "(a1,b2) \<in> r - Id \<or> (b2,a1) \<in> r - Id"
+ thus ?thesis using 0 1 2 unfolding bsqr_def by auto
+ next
+ assume Case212: "a1 = b2"
+ show ?thesis
+ proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0)
+ assume Case2121: "(a1,b1) \<in> r - Id \<or> (b1,a1) \<in> r - Id"
+ thus ?thesis using 0 1 2 Case212 unfolding bsqr_def by auto
+ next
+ assume Case2122: "a1 = b1"
+ show ?thesis
+ proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0)
+ assume Case21221: "(a2,b2) \<in> r - Id \<or> (b2,a2) \<in> r - Id"
+ thus ?thesis using 0 1 2 Case212 Case2122 unfolding bsqr_def by auto
+ next
+ assume Case21222: "a2 = b2"
+ thus ?thesis using Case2122 Case212 by auto
+ qed
+ qed
+ qed
+ next
+ assume Case22: "(b2,b1) \<in> r"
+ hence 3: "wo_rel.max2 r b1 b2 = b1" using Well 0 by (simp add: wo_rel.max2_equals1)
+ show ?thesis
+ proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0)
+ assume Case221: "(a1,b1) \<in> r - Id \<or> (b1,a1) \<in> r - Id"
+ thus ?thesis using 0 1 3 unfolding bsqr_def by auto
+ next
+ assume Case222: "a1 = b1"
+ show ?thesis
+ proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0)
+ assume Case2221: "(a2,b2) \<in> r - Id \<or> (b2,a2) \<in> r - Id"
+ thus ?thesis using 0 1 3 Case222 unfolding bsqr_def by auto
+ next
+ assume Case2222: "a2 = b2"
+ thus ?thesis using Case222 by auto
+ qed
+ qed
+ qed
+ qed
+ }
+ thus ?thesis unfolding total_on_def by fast
+qed
+
+
+lemma bsqr_Linear_order:
+assumes "Well_order r"
+shows "Linear_order(bsqr r)"
+unfolding order_on_defs
+using assms bsqr_Refl bsqr_Trans bsqr_antisym bsqr_Total by blast
+
+
+lemma bsqr_Well_order:
+assumes "Well_order r"
+shows "Well_order(bsqr r)"
+using assms
+proof(simp add: bsqr_Linear_order Linear_order_Well_order_iff, intro allI impI)
+ have 0: "\<forall>A \<le> Field r. A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. (a,a') \<in> r)"
+ using assms well_order_on_def Linear_order_Well_order_iff by blast
+ fix D assume *: "D \<le> Field (bsqr r)" and **: "D \<noteq> {}"
+ hence 1: "D \<le> Field r \<times> Field r" unfolding Field_bsqr by simp
+ (* *)
+ obtain M where M_def: "M = {wo_rel.max2 r a1 a2| a1 a2. (a1,a2) \<in> D}" by blast
+ have "M \<noteq> {}" using 1 M_def ** by auto
+ moreover
+ have "M \<le> Field r" unfolding M_def
+ using 1 assms wo_rel_def[of r] wo_rel.max2_among[of r] by fastforce
+ ultimately obtain m where m_min: "m \<in> M \<and> (\<forall>a \<in> M. (m,a) \<in> r)"
+ using 0 by blast
+ (* *)
+ obtain A1 where A1_def: "A1 = {a1. \<exists>a2. (a1,a2) \<in> D \<and> wo_rel.max2 r a1 a2 = m}" by blast
+ have "A1 \<le> Field r" unfolding A1_def using 1 by auto
+ moreover have "A1 \<noteq> {}" unfolding A1_def using m_min unfolding M_def by blast
+ ultimately obtain a1 where a1_min: "a1 \<in> A1 \<and> (\<forall>a \<in> A1. (a1,a) \<in> r)"
+ using 0 by blast
+ (* *)
+ obtain A2 where A2_def: "A2 = {a2. (a1,a2) \<in> D \<and> wo_rel.max2 r a1 a2 = m}" by blast
+ have "A2 \<le> Field r" unfolding A2_def using 1 by auto
+ moreover have "A2 \<noteq> {}" unfolding A2_def
+ using m_min a1_min unfolding A1_def M_def by blast
+ ultimately obtain a2 where a2_min: "a2 \<in> A2 \<and> (\<forall>a \<in> A2. (a2,a) \<in> r)"
+ using 0 by blast
+ (* *)
+ have 2: "wo_rel.max2 r a1 a2 = m"
+ using a1_min a2_min unfolding A1_def A2_def by auto
+ have 3: "(a1,a2) \<in> D" using a2_min unfolding A2_def by auto
+ (* *)
+ moreover
+ {fix b1 b2 assume ***: "(b1,b2) \<in> D"
+ hence 4: "{a1,a2,b1,b2} \<le> Field r" using 1 3 by blast
+ have 5: "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r"
+ using *** a1_min a2_min m_min unfolding A1_def A2_def M_def by auto
+ have "((a1,a2),(b1,b2)) \<in> bsqr r"
+ proof(cases "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2")
+ assume Case1: "wo_rel.max2 r a1 a2 \<noteq> wo_rel.max2 r b1 b2"
+ thus ?thesis unfolding bsqr_def using 4 5 by auto
+ next
+ assume Case2: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2"
+ hence "b1 \<in> A1" unfolding A1_def using 2 *** by auto
+ hence 6: "(a1,b1) \<in> r" using a1_min by auto
+ show ?thesis
+ proof(cases "a1 = b1")
+ assume Case21: "a1 \<noteq> b1"
+ thus ?thesis unfolding bsqr_def using 4 Case2 6 by auto
+ next
+ assume Case22: "a1 = b1"
+ hence "b2 \<in> A2" unfolding A2_def using 2 *** Case2 by auto
+ hence 7: "(a2,b2) \<in> r" using a2_min by auto
+ thus ?thesis unfolding bsqr_def using 4 7 Case2 Case22 by auto
+ qed
+ qed
+ }
+ (* *)
+ ultimately show "\<exists>d \<in> D. \<forall>d' \<in> D. (d,d') \<in> bsqr r" by fastforce
+qed
+
+
+lemma bsqr_max2:
+assumes WELL: "Well_order r" and LEQ: "((a1,a2),(b1,b2)) \<in> bsqr r"
+shows "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r"
+proof-
+ have "{(a1,a2),(b1,b2)} \<le> Field(bsqr r)"
+ using LEQ unfolding Field_def by auto
+ hence "{a1,a2,b1,b2} \<le> Field r" unfolding Field_bsqr by auto
+ hence "{wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2} \<le> Field r"
+ using WELL wo_rel_def[of r] wo_rel.max2_among[of r] by fastforce
+ moreover have "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r \<or> wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2"
+ using LEQ unfolding bsqr_def by auto
+ ultimately show ?thesis using WELL unfolding order_on_defs refl_on_def by auto
+qed
+
+
+lemma bsqr_ofilter:
+assumes WELL: "Well_order r" and
+ OF: "wo_rel.ofilter (bsqr r) D" and SUB: "D < Field r \<times> Field r" and
+ NE: "\<not> (\<exists>a. Field r = rel.under r a)"
+shows "\<exists>A. wo_rel.ofilter r A \<and> A < Field r \<and> D \<le> A \<times> A"
+proof-
+ let ?r' = "bsqr r"
+ have Well: "wo_rel r" using WELL wo_rel_def by blast
+ hence Trans: "trans r" using wo_rel.TRANS by blast
+ have Well': "Well_order ?r' \<and> wo_rel ?r'"
+ using WELL bsqr_Well_order wo_rel_def by blast
+ (* *)
+ have "D < Field ?r'" unfolding Field_bsqr using SUB .
+ with OF obtain a1 and a2 where
+ "(a1,a2) \<in> Field ?r'" and 1: "D = rel.underS ?r' (a1,a2)"
+ using Well' wo_rel.ofilter_underS_Field[of ?r' D] by auto
+ hence 2: "{a1,a2} \<le> Field r" unfolding Field_bsqr by auto
+ let ?m = "wo_rel.max2 r a1 a2"
+ have "D \<le> (rel.under r ?m) \<times> (rel.under r ?m)"
+ proof(unfold 1)
+ {fix b1 b2
+ let ?n = "wo_rel.max2 r b1 b2"
+ assume "(b1,b2) \<in> rel.underS ?r' (a1,a2)"
+ hence 3: "((b1,b2),(a1,a2)) \<in> ?r'"
+ unfolding rel.underS_def by blast
+ hence "(?n,?m) \<in> r" using WELL by (simp add: bsqr_max2)
+ moreover
+ {have "(b1,b2) \<in> Field ?r'" using 3 unfolding Field_def by auto
+ hence "{b1,b2} \<le> Field r" unfolding Field_bsqr by auto
+ hence "(b1,?n) \<in> r \<and> (b2,?n) \<in> r"
+ using Well by (simp add: wo_rel.max2_greater)
+ }
+ ultimately have "(b1,?m) \<in> r \<and> (b2,?m) \<in> r"
+ using Trans trans_def[of r] by blast
+ hence "(b1,b2) \<in> (rel.under r ?m) \<times> (rel.under r ?m)" unfolding rel.under_def by simp}
+ thus "rel.underS ?r' (a1,a2) \<le> (rel.under r ?m) \<times> (rel.under r ?m)" by auto
+ qed
+ moreover have "wo_rel.ofilter r (rel.under r ?m)"
+ using Well by (simp add: wo_rel.under_ofilter)
+ moreover have "rel.under r ?m < Field r"
+ using NE rel.under_Field[of r ?m] by blast
+ ultimately show ?thesis by blast
+qed
+
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Ordinals_and_Cardinals/Fun_More.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,183 @@
+(* Title: HOL/Ordinals_and_Cardinals/Fun_More.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+More on injections, bijections and inverses.
+*)
+
+header {* More on Injections, Bijections and Inverses *}
+
+theory Fun_More
+imports Fun_More_Base
+begin
+
+
+subsection {* Purely functional properties *}
+
+(* unused *)
+(*1*)lemma notIn_Un_bij_betw2:
+assumes NIN: "b \<notin> A" and NIN': "b' \<notin> A'" and
+ BIJ: "bij_betw f A A'"
+shows "bij_betw f (A \<union> {b}) (A' \<union> {b'}) = (f b = b')"
+proof
+ assume "f b = b'"
+ thus "bij_betw f (A \<union> {b}) (A' \<union> {b'})"
+ using assms notIn_Un_bij_betw[of b A f A'] by auto
+next
+ assume *: "bij_betw f (A \<union> {b}) (A' \<union> {b'})"
+ hence "f b \<in> A' \<union> {b'}"
+ unfolding bij_betw_def by auto
+ moreover
+ {assume "f b \<in> A'"
+ then obtain b1 where 1: "b1 \<in> A" and 2: "f b1 = f b" using BIJ
+ by (auto simp add: bij_betw_def)
+ hence "b = b1" using *
+ by (auto simp add: bij_betw_def inj_on_def)
+ with 1 NIN have False by auto
+ }
+ ultimately show "f b = b'" by blast
+qed
+
+(* unused *)
+(*1*)lemma bij_betw_ball:
+assumes BIJ: "bij_betw f A B"
+shows "(\<forall>b \<in> B. phi b) = (\<forall>a \<in> A. phi(f a))"
+using assms unfolding bij_betw_def inj_on_def by blast
+
+(* unused *)
+(*1*)lemma bij_betw_diff_singl:
+assumes BIJ: "bij_betw f A A'" and IN: "a \<in> A"
+shows "bij_betw f (A - {a}) (A' - {f a})"
+proof-
+ let ?B = "A - {a}" let ?B' = "A' - {f a}"
+ have "f a \<in> A'" using IN BIJ unfolding bij_betw_def by blast
+ hence "a \<notin> ?B \<and> f a \<notin> ?B' \<and> A = ?B \<union> {a} \<and> A' = ?B' \<union> {f a}"
+ using IN by blast
+ thus ?thesis using notIn_Un_bij_betw3[of a ?B f ?B'] BIJ by simp
+qed
+
+
+subsection {* Properties involving finite and infinite sets *}
+
+(*3*)lemma inj_on_image_Pow:
+assumes "inj_on f A"
+shows "inj_on (image f) (Pow A)"
+unfolding Pow_def inj_on_def proof(clarsimp)
+ fix X Y assume *: "X \<le> A" and **: "Y \<le> A" and
+ ***: "f ` X = f ` Y"
+ show "X = Y"
+ proof(auto)
+ fix x assume ****: "x \<in> X"
+ with *** obtain y where "y \<in> Y \<and> f x = f y" by blast
+ with **** * ** assms show "x \<in> Y"
+ unfolding inj_on_def by auto
+ next
+ fix y assume ****: "y \<in> Y"
+ with *** obtain x where "x \<in> X \<and> f x = f y" by force
+ with **** * ** assms show "y \<in> X"
+ unfolding inj_on_def by auto
+ qed
+qed
+
+(*2*)lemma bij_betw_image_Pow:
+assumes "bij_betw f A B"
+shows "bij_betw (image f) (Pow A) (Pow B)"
+using assms unfolding bij_betw_def
+by (auto simp add: inj_on_image_Pow image_Pow_surj)
+
+(* unused *)
+(*1*)lemma bij_betw_inv_into_RIGHT:
+assumes BIJ: "bij_betw f A A'" and SUB: "B' \<le> A'"
+shows "f `((inv_into A f)`B') = B'"
+using assms
+proof(auto simp add: bij_betw_inv_into_right)
+ let ?f' = "(inv_into A f)"
+ fix a' assume *: "a' \<in> B'"
+ hence "a' \<in> A'" using SUB by auto
+ hence "a' = f (?f' a')"
+ using BIJ by (auto simp add: bij_betw_inv_into_right)
+ thus "a' \<in> f ` (?f' ` B')" using * by blast
+qed
+
+(* unused *)
+(*1*)lemma bij_betw_inv_into_RIGHT_LEFT:
+assumes BIJ: "bij_betw f A A'" and SUB: "B' \<le> A'" and
+ IM: "(inv_into A f) ` B' = B"
+shows "f ` B = B'"
+proof-
+ have "f`((inv_into A f)` B') = B'"
+ using assms bij_betw_inv_into_RIGHT[of f A A' B'] by auto
+ thus ?thesis using IM by auto
+qed
+
+(* unused *)
+(*2*)lemma bij_betw_inv_into_twice:
+assumes "bij_betw f A A'"
+shows "\<forall>a \<in> A. inv_into A' (inv_into A f) a = f a"
+proof
+ let ?f' = "inv_into A f" let ?f'' = "inv_into A' ?f'"
+ have 1: "bij_betw ?f' A' A" using assms
+ by (auto simp add: bij_betw_inv_into)
+ fix a assume *: "a \<in> A"
+ then obtain a' where 2: "a' \<in> A'" and 3: "?f' a' = a"
+ using 1 unfolding bij_betw_def by force
+ hence "?f'' a = a'"
+ using * 1 3 by (auto simp add: bij_betw_inv_into_left)
+ moreover have "f a = a'" using assms 2 3
+ by (auto simp add: bij_betw_inv_into_right)
+ ultimately show "?f'' a = f a" by simp
+qed
+
+
+subsection {* Properties involving Hilbert choice *}
+
+
+subsection {* Other facts *}
+
+(*3*)lemma atLeastLessThan_injective:
+assumes "{0 ..< m::nat} = {0 ..< n}"
+shows "m = n"
+proof-
+ {assume "m < n"
+ hence "m \<in> {0 ..< n}" by auto
+ hence "{0 ..< m} < {0 ..< n}" by auto
+ hence False using assms by blast
+ }
+ moreover
+ {assume "n < m"
+ hence "n \<in> {0 ..< m}" by auto
+ hence "{0 ..< n} < {0 ..< m}" by auto
+ hence False using assms by blast
+ }
+ ultimately show ?thesis by force
+qed
+
+(*2*)lemma atLeastLessThan_injective2:
+"bij_betw f {0 ..< m::nat} {0 ..< n} \<Longrightarrow> m = n"
+using finite_atLeastLessThan[of m] finite_atLeastLessThan[of n]
+ card_atLeastLessThan[of m] card_atLeastLessThan[of n]
+ bij_betw_iff_card[of "{0 ..< m}" "{0 ..< n}"] by auto
+
+(* unused *)
+(*2*)lemma atLeastLessThan_less_eq3:
+"(\<exists>f. inj_on f {0..<(m::nat)} \<and> f ` {0..<m} \<le> {0..<n}) = (m \<le> n)"
+using atLeastLessThan_less_eq2
+proof(auto)
+ assume "m \<le> n"
+ hence "inj_on id {0..<m} \<and> id ` {0..<m} \<subseteq> {0..<n}" unfolding inj_on_def by force
+ thus "\<exists>f. inj_on f {0..<m} \<and> f ` {0..<m} \<subseteq> {0..<n}" by blast
+qed
+
+(* unused *)
+(*3*)lemma atLeastLessThan_less:
+"({0..<m} < {0..<n}) = ((m::nat) < n)"
+proof-
+ have "({0..<m} < {0..<n}) = ({0..<m} \<le> {0..<n} \<and> {0..<m} ~= {0..<n})"
+ using subset_iff_psubset_eq by blast
+ also have "\<dots> = (m \<le> n \<and> m ~= n)"
+ using atLeastLessThan_less_eq atLeastLessThan_injective by blast
+ also have "\<dots> = (m < n)" by auto
+ finally show ?thesis .
+qed
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Ordinals_and_Cardinals/Fun_More_Base.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,252 @@
+(* Title: HOL/Ordinals_and_Cardinals/Fun_More_Base.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+More on injections, bijections and inverses (base).
+*)
+
+header {* More on Injections, Bijections and Inverses (Base) *}
+
+theory Fun_More_Base
+imports "~~/src/HOL/Library/Infinite_Set"
+begin
+
+
+text {* This section proves more facts (additional to those in @{text "Fun.thy"},
+@{text "Hilbert_Choice.thy"}, @{text "Finite_Set.thy"} and @{text "Infinite_Set.thy"}),
+mainly concerning injections, bijections, inverses and (numeric) cardinals of
+finite sets. *}
+
+
+subsection {* Purely functional properties *}
+
+
+(*2*)lemma bij_betw_id_iff:
+"(A = B) = (bij_betw id A B)"
+by(simp add: bij_betw_def)
+
+
+(*2*)lemma bij_betw_byWitness:
+assumes LEFT: "\<forall>a \<in> A. f'(f a) = a" and
+ RIGHT: "\<forall>a' \<in> A'. f(f' a') = a'" and
+ IM1: "f ` A \<le> A'" and IM2: "f' ` A' \<le> A"
+shows "bij_betw f A A'"
+using assms
+proof(unfold bij_betw_def inj_on_def, auto)
+ fix a b assume *: "a \<in> A" "b \<in> A" and **: "f a = f b"
+ have "a = f'(f a) \<and> b = f'(f b)" using * LEFT by simp
+ with ** show "a = b" by simp
+next
+ fix a' assume *: "a' \<in> A'"
+ hence "f' a' \<in> A" using IM2 by blast
+ moreover
+ have "a' = f(f' a')" using * RIGHT by simp
+ ultimately show "a' \<in> f ` A" by blast
+qed
+
+
+(*3*)corollary notIn_Un_bij_betw:
+assumes NIN: "b \<notin> A" and NIN': "f b \<notin> A'" and
+ BIJ: "bij_betw f A A'"
+shows "bij_betw f (A \<union> {b}) (A' \<union> {f b})"
+proof-
+ have "bij_betw f {b} {f b}"
+ unfolding bij_betw_def inj_on_def by simp
+ with assms show ?thesis
+ using bij_betw_combine[of f A A' "{b}" "{f b}"] by blast
+qed
+
+
+(*1*)lemma notIn_Un_bij_betw3:
+assumes NIN: "b \<notin> A" and NIN': "f b \<notin> A'"
+shows "bij_betw f A A' = bij_betw f (A \<union> {b}) (A' \<union> {f b})"
+proof
+ assume "bij_betw f A A'"
+ thus "bij_betw f (A \<union> {b}) (A' \<union> {f b})"
+ using assms notIn_Un_bij_betw[of b A f A'] by blast
+next
+ assume *: "bij_betw f (A \<union> {b}) (A' \<union> {f b})"
+ have "f ` A = A'"
+ proof(auto)
+ fix a assume **: "a \<in> A"
+ hence "f a \<in> A' \<union> {f b}" using * unfolding bij_betw_def by blast
+ moreover
+ {assume "f a = f b"
+ hence "a = b" using * ** unfolding bij_betw_def inj_on_def by blast
+ with NIN ** have False by blast
+ }
+ ultimately show "f a \<in> A'" by blast
+ next
+ fix a' assume **: "a' \<in> A'"
+ hence "a' \<in> f`(A \<union> {b})"
+ using * by (auto simp add: bij_betw_def)
+ then obtain a where 1: "a \<in> A \<union> {b} \<and> f a = a'" by blast
+ moreover
+ {assume "a = b" with 1 ** NIN' have False by blast
+ }
+ ultimately have "a \<in> A" by blast
+ with 1 show "a' \<in> f ` A" by blast
+ qed
+ thus "bij_betw f A A'" using * bij_betw_subset[of f "A \<union> {b}" _ A] by blast
+qed
+
+
+subsection {* Properties involving finite and infinite sets *}
+
+
+(*3*)lemma inj_on_finite:
+assumes "inj_on f A" "f ` A \<le> B" "finite B"
+shows "finite A"
+using assms infinite_super by (fast dest: finite_imageD)
+
+
+(*3*)lemma infinite_imp_bij_betw:
+assumes INF: "infinite A"
+shows "\<exists>h. bij_betw h A (A - {a})"
+proof(cases "a \<in> A")
+ assume Case1: "a \<notin> A" hence "A - {a} = A" by blast
+ thus ?thesis using bij_betw_id[of A] by auto
+next
+ assume Case2: "a \<in> A"
+ have "infinite (A - {a})" using INF infinite_remove by auto
+ with infinite_iff_countable_subset[of "A - {a}"] obtain f::"nat \<Rightarrow> 'a"
+ where 1: "inj f" and 2: "f ` UNIV \<le> A - {a}" by blast
+ obtain g where g_def: "g = (\<lambda> n. if n = 0 then a else f (Suc n))" by blast
+ obtain A' where A'_def: "A' = g ` UNIV" by blast
+ have temp: "\<forall>y. f y \<noteq> a" using 2 by blast
+ have 3: "inj_on g UNIV \<and> g ` UNIV \<le> A \<and> a \<in> g ` UNIV"
+ proof(auto simp add: Case2 g_def, unfold inj_on_def, intro ballI impI,
+ case_tac "x = 0", auto simp add: 2)
+ fix y assume "a = (if y = 0 then a else f (Suc y))"
+ thus "y = 0" using temp by (case_tac "y = 0", auto)
+ next
+ fix x y
+ assume "f (Suc x) = (if y = 0 then a else f (Suc y))"
+ thus "x = y" using 1 temp unfolding inj_on_def by (case_tac "y = 0", auto)
+ next
+ fix n show "f (Suc n) \<in> A" using 2 by blast
+ qed
+ hence 4: "bij_betw g UNIV A' \<and> a \<in> A' \<and> A' \<le> A"
+ using inj_on_imp_bij_betw[of g] unfolding A'_def by auto
+ hence 5: "bij_betw (inv g) A' UNIV"
+ by (auto simp add: bij_betw_inv_into)
+ (* *)
+ obtain n where "g n = a" using 3 by auto
+ hence 6: "bij_betw g (UNIV - {n}) (A' - {a})"
+ using 3 4 unfolding A'_def
+ by clarify (rule bij_betw_subset, auto simp: image_set_diff)
+ (* *)
+ obtain v where v_def: "v = (\<lambda> m. if m < n then m else Suc m)" by blast
+ have 7: "bij_betw v UNIV (UNIV - {n})"
+ proof(unfold bij_betw_def inj_on_def, intro conjI, clarify)
+ fix m1 m2 assume "v m1 = v m2"
+ thus "m1 = m2"
+ by(case_tac "m1 < n", case_tac "m2 < n",
+ auto simp add: inj_on_def v_def, case_tac "m2 < n", auto)
+ next
+ show "v ` UNIV = UNIV - {n}"
+ proof(auto simp add: v_def)
+ fix m assume *: "m \<noteq> n" and **: "m \<notin> Suc ` {m'. \<not> m' < n}"
+ {assume "n \<le> m" with * have 71: "Suc n \<le> m" by auto
+ then obtain m' where 72: "m = Suc m'" using Suc_le_D by auto
+ with 71 have "n \<le> m'" by auto
+ with 72 ** have False by auto
+ }
+ thus "m < n" by force
+ qed
+ qed
+ (* *)
+ obtain h' where h'_def: "h' = g o v o (inv g)" by blast
+ hence 8: "bij_betw h' A' (A' - {a})" using 5 7 6
+ by (auto simp add: bij_betw_trans)
+ (* *)
+ obtain h where h_def: "h = (\<lambda> b. if b \<in> A' then h' b else b)" by blast
+ have "\<forall>b \<in> A'. h b = h' b" unfolding h_def by auto
+ hence "bij_betw h A' (A' - {a})" using 8 bij_betw_cong[of A' h] by auto
+ moreover
+ {have "\<forall>b \<in> A - A'. h b = b" unfolding h_def by auto
+ hence "bij_betw h (A - A') (A - A')"
+ using bij_betw_cong[of "A - A'" h id] bij_betw_id[of "A - A'"] by auto
+ }
+ moreover
+ have "(A' Int (A - A') = {} \<and> A' \<union> (A - A') = A) \<and>
+ ((A' - {a}) Int (A - A') = {} \<and> (A' - {a}) \<union> (A - A') = A - {a})"
+ using 4 by blast
+ ultimately have "bij_betw h A (A - {a})"
+ using bij_betw_combine[of h A' "A' - {a}" "A - A'" "A - A'"] by simp
+ thus ?thesis by blast
+qed
+
+
+(*3*)lemma infinite_imp_bij_betw2:
+assumes INF: "infinite A"
+shows "\<exists>h. bij_betw h A (A \<union> {a})"
+proof(cases "a \<in> A")
+ assume Case1: "a \<in> A" hence "A \<union> {a} = A" by blast
+ thus ?thesis using bij_betw_id[of A] by auto
+next
+ let ?A' = "A \<union> {a}"
+ assume Case2: "a \<notin> A" hence "A = ?A' - {a}" by blast
+ moreover have "infinite ?A'" using INF by auto
+ ultimately obtain f where "bij_betw f ?A' A"
+ using infinite_imp_bij_betw[of ?A' a] by auto
+ hence "bij_betw(inv_into ?A' f) A ?A'" using bij_betw_inv_into by blast
+ thus ?thesis by auto
+qed
+
+
+subsection {* Properties involving Hilbert choice *}
+
+
+(*2*)lemma bij_betw_inv_into_left:
+assumes BIJ: "bij_betw f A A'" and IN: "a \<in> A"
+shows "(inv_into A f) (f a) = a"
+using assms unfolding bij_betw_def
+by clarify (rule inv_into_f_f)
+
+(*2*)lemma bij_betw_inv_into_right:
+assumes "bij_betw f A A'" "a' \<in> A'"
+shows "f(inv_into A f a') = a'"
+using assms unfolding bij_betw_def using f_inv_into_f by force
+
+
+(*1*)lemma bij_betw_inv_into_LEFT:
+assumes BIJ: "bij_betw f A A'" and SUB: "B \<le> A"
+shows "(inv_into A f)`(f ` B) = B"
+using assms unfolding bij_betw_def using inv_into_image_cancel by force
+
+
+(*1*)lemma bij_betw_inv_into_LEFT_RIGHT:
+assumes BIJ: "bij_betw f A A'" and SUB: "B \<le> A" and
+ IM: "f ` B = B'"
+shows "(inv_into A f) ` B' = B"
+using assms bij_betw_inv_into_LEFT[of f A A' B] by fast
+
+
+(*1*)lemma bij_betw_inv_into_subset:
+assumes BIJ: "bij_betw f A A'" and
+ SUB: "B \<le> A" and IM: "f ` B = B'"
+shows "bij_betw (inv_into A f) B' B"
+using assms unfolding bij_betw_def
+by (auto intro: inj_on_inv_into)
+
+
+subsection {* Other facts *}
+
+
+(*2*)lemma atLeastLessThan_less_eq:
+"({0..<m} \<le> {0..<n}) = ((m::nat) \<le> n)"
+unfolding ivl_subset by arith
+
+
+(*2*)lemma atLeastLessThan_less_eq2:
+assumes "inj_on f {0..<(m::nat)} \<and> f ` {0..<m} \<le> {0..<n}"
+shows "m \<le> n"
+using assms
+using finite_atLeastLessThan[of m] finite_atLeastLessThan[of n]
+ card_atLeastLessThan[of m] card_atLeastLessThan[of n]
+ card_inj_on_le[of f "{0 ..< m}" "{0 ..< n}"] by auto
+
+
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Ordinals_and_Cardinals/Order_Relation_More.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,582 @@
+(* Title: HOL/Ordinals_and_Cardinals/Order_Relation_More.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Basics on order-like relations.
+*)
+
+header {* Basics on Order-Like Relations *}
+
+theory Order_Relation_More
+imports Order_Relation_More_Base
+begin
+
+
+subsection {* The upper and lower bounds operators *}
+
+lemma (in rel) aboveS_subset_above: "aboveS a \<le> above a"
+by(auto simp add: aboveS_def above_def)
+
+lemma (in rel) AboveS_subset_Above: "AboveS A \<le> Above A"
+by(auto simp add: AboveS_def Above_def)
+
+lemma (in rel) UnderS_disjoint: "A Int (UnderS A) = {}"
+by(auto simp add: UnderS_def)
+
+lemma (in rel) aboveS_notIn: "a \<notin> aboveS a"
+by(auto simp add: aboveS_def)
+
+lemma (in rel) Refl_above_in: "\<lbrakk>Refl r; a \<in> Field r\<rbrakk> \<Longrightarrow> a \<in> above a"
+by(auto simp add: refl_on_def above_def)
+
+lemma (in rel) in_Above_under: "a \<in> Field r \<Longrightarrow> a \<in> Above (under a)"
+by(auto simp add: Above_def under_def)
+
+lemma (in rel) in_Under_above: "a \<in> Field r \<Longrightarrow> a \<in> Under (above a)"
+by(auto simp add: Under_def above_def)
+
+lemma (in rel) in_UnderS_aboveS: "a \<in> Field r \<Longrightarrow> a \<in> UnderS (aboveS a)"
+by(auto simp add: UnderS_def aboveS_def)
+
+lemma (in rel) subset_Above_Under: "B \<le> Field r \<Longrightarrow> B \<le> Above (Under B)"
+by(auto simp add: Above_def Under_def)
+
+lemma (in rel) subset_Under_Above: "B \<le> Field r \<Longrightarrow> B \<le> Under (Above B)"
+by(auto simp add: Under_def Above_def)
+
+lemma (in rel) subset_AboveS_UnderS: "B \<le> Field r \<Longrightarrow> B \<le> AboveS (UnderS B)"
+by(auto simp add: AboveS_def UnderS_def)
+
+lemma (in rel) subset_UnderS_AboveS: "B \<le> Field r \<Longrightarrow> B \<le> UnderS (AboveS B)"
+by(auto simp add: UnderS_def AboveS_def)
+
+lemma (in rel) Under_Above_Galois:
+"\<lbrakk>B \<le> Field r; C \<le> Field r\<rbrakk> \<Longrightarrow> (B \<le> Above C) = (C \<le> Under B)"
+by(unfold Above_def Under_def, blast)
+
+lemma (in rel) UnderS_AboveS_Galois:
+"\<lbrakk>B \<le> Field r; C \<le> Field r\<rbrakk> \<Longrightarrow> (B \<le> AboveS C) = (C \<le> UnderS B)"
+by(unfold AboveS_def UnderS_def, blast)
+
+lemma (in rel) Refl_above_aboveS:
+assumes REFL: "Refl r" and IN: "a \<in> Field r"
+shows "above a = aboveS a \<union> {a}"
+proof(unfold above_def aboveS_def, auto)
+ show "(a,a) \<in> r" using REFL IN refl_on_def[of _ r] by blast
+qed
+
+lemma (in rel) Linear_order_under_aboveS_Field:
+assumes LIN: "Linear_order r" and IN: "a \<in> Field r"
+shows "Field r = under a \<union> aboveS a"
+proof(unfold under_def aboveS_def, auto)
+ assume "a \<in> Field r" "(a, a) \<notin> r"
+ with LIN IN order_on_defs[of _ r] refl_on_def[of _ r]
+ show False by auto
+next
+ fix b assume "b \<in> Field r" "(b, a) \<notin> r"
+ with LIN IN order_on_defs[of "Field r" r] total_on_def[of "Field r" r]
+ have "(a,b) \<in> r \<or> a = b" by blast
+ thus "(a,b) \<in> r"
+ using LIN IN order_on_defs[of _ r] refl_on_def[of _ r] by auto
+next
+ fix b assume "(b, a) \<in> r"
+ thus "b \<in> Field r"
+ using LIN order_on_defs[of _ r] refl_on_def[of _ r] by blast
+next
+ fix b assume "b \<noteq> a" "(a, b) \<in> r"
+ thus "b \<in> Field r"
+ using LIN order_on_defs[of "Field r" r] refl_on_def[of "Field r" r] by blast
+qed
+
+lemma (in rel) Linear_order_underS_above_Field:
+assumes LIN: "Linear_order r" and IN: "a \<in> Field r"
+shows "Field r = underS a \<union> above a"
+proof(unfold underS_def above_def, auto)
+ assume "a \<in> Field r" "(a, a) \<notin> r"
+ with LIN IN order_on_defs[of _ r] refl_on_def[of _ r]
+ show False by auto
+next
+ fix b assume "b \<in> Field r" "(a, b) \<notin> r"
+ with LIN IN order_on_defs[of "Field r" r] total_on_def[of "Field r" r]
+ have "(b,a) \<in> r \<or> b = a" by blast
+ thus "(b,a) \<in> r"
+ using LIN IN order_on_defs[of _ r] refl_on_def[of _ r] by auto
+next
+ fix b assume "b \<noteq> a" "(b, a) \<in> r"
+ thus "b \<in> Field r"
+ using LIN order_on_defs[of _ r] refl_on_def[of _ r] by blast
+next
+ fix b assume "(a, b) \<in> r"
+ thus "b \<in> Field r"
+ using LIN order_on_defs[of "Field r" r] refl_on_def[of "Field r" r] by blast
+qed
+
+lemma (in rel) under_empty: "a \<notin> Field r \<Longrightarrow> under a = {}"
+unfolding Field_def under_def by auto
+
+lemma (in rel) above_Field: "above a \<le> Field r"
+by(unfold above_def Field_def, auto)
+
+lemma (in rel) aboveS_Field: "aboveS a \<le> Field r"
+by(unfold aboveS_def Field_def, auto)
+
+lemma (in rel) Above_Field: "Above A \<le> Field r"
+by(unfold Above_def Field_def, auto)
+
+lemma (in rel) Refl_under_Under:
+assumes REFL: "Refl r" and NE: "A \<noteq> {}"
+shows "Under A = (\<Inter> a \<in> A. under a)"
+proof
+ show "Under A \<subseteq> (\<Inter> a \<in> A. under a)"
+ by(unfold Under_def under_def, auto)
+next
+ show "(\<Inter> a \<in> A. under a) \<subseteq> Under A"
+ proof(auto)
+ fix x
+ assume *: "\<forall>xa \<in> A. x \<in> under xa"
+ hence "\<forall>xa \<in> A. (x,xa) \<in> r"
+ by (simp add: under_def)
+ moreover
+ {from NE obtain a where "a \<in> A" by blast
+ with * have "x \<in> under a" by simp
+ hence "x \<in> Field r"
+ using under_Field[of a] by auto
+ }
+ ultimately show "x \<in> Under A"
+ unfolding Under_def by auto
+ qed
+qed
+
+lemma (in rel) Refl_underS_UnderS:
+assumes REFL: "Refl r" and NE: "A \<noteq> {}"
+shows "UnderS A = (\<Inter> a \<in> A. underS a)"
+proof
+ show "UnderS A \<subseteq> (\<Inter> a \<in> A. underS a)"
+ by(unfold UnderS_def underS_def, auto)
+next
+ show "(\<Inter> a \<in> A. underS a) \<subseteq> UnderS A"
+ proof(auto)
+ fix x
+ assume *: "\<forall>xa \<in> A. x \<in> underS xa"
+ hence "\<forall>xa \<in> A. x \<noteq> xa \<and> (x,xa) \<in> r"
+ by (auto simp add: underS_def)
+ moreover
+ {from NE obtain a where "a \<in> A" by blast
+ with * have "x \<in> underS a" by simp
+ hence "x \<in> Field r"
+ using underS_Field[of a] by auto
+ }
+ ultimately show "x \<in> UnderS A"
+ unfolding UnderS_def by auto
+ qed
+qed
+
+lemma (in rel) Refl_above_Above:
+assumes REFL: "Refl r" and NE: "A \<noteq> {}"
+shows "Above A = (\<Inter> a \<in> A. above a)"
+proof
+ show "Above A \<subseteq> (\<Inter> a \<in> A. above a)"
+ by(unfold Above_def above_def, auto)
+next
+ show "(\<Inter> a \<in> A. above a) \<subseteq> Above A"
+ proof(auto)
+ fix x
+ assume *: "\<forall>xa \<in> A. x \<in> above xa"
+ hence "\<forall>xa \<in> A. (xa,x) \<in> r"
+ by (simp add: above_def)
+ moreover
+ {from NE obtain a where "a \<in> A" by blast
+ with * have "x \<in> above a" by simp
+ hence "x \<in> Field r"
+ using above_Field[of a] by auto
+ }
+ ultimately show "x \<in> Above A"
+ unfolding Above_def by auto
+ qed
+qed
+
+lemma (in rel) Refl_aboveS_AboveS:
+assumes REFL: "Refl r" and NE: "A \<noteq> {}"
+shows "AboveS A = (\<Inter> a \<in> A. aboveS a)"
+proof
+ show "AboveS A \<subseteq> (\<Inter> a \<in> A. aboveS a)"
+ by(unfold AboveS_def aboveS_def, auto)
+next
+ show "(\<Inter> a \<in> A. aboveS a) \<subseteq> AboveS A"
+ proof(auto)
+ fix x
+ assume *: "\<forall>xa \<in> A. x \<in> aboveS xa"
+ hence "\<forall>xa \<in> A. xa \<noteq> x \<and> (xa,x) \<in> r"
+ by (auto simp add: aboveS_def)
+ moreover
+ {from NE obtain a where "a \<in> A" by blast
+ with * have "x \<in> aboveS a" by simp
+ hence "x \<in> Field r"
+ using aboveS_Field[of a] by auto
+ }
+ ultimately show "x \<in> AboveS A"
+ unfolding AboveS_def by auto
+ qed
+qed
+
+lemma (in rel) under_Under_singl: "under a = Under {a}"
+by(unfold Under_def under_def, auto simp add: Field_def)
+
+lemma (in rel) underS_UnderS_singl: "underS a = UnderS {a}"
+by(unfold UnderS_def underS_def, auto simp add: Field_def)
+
+lemma (in rel) above_Above_singl: "above a = Above {a}"
+by(unfold Above_def above_def, auto simp add: Field_def)
+
+lemma (in rel) aboveS_AboveS_singl: "aboveS a = AboveS {a}"
+by(unfold AboveS_def aboveS_def, auto simp add: Field_def)
+
+lemma (in rel) Under_decr: "A \<le> B \<Longrightarrow> Under B \<le> Under A"
+by(unfold Under_def, auto)
+
+lemma (in rel) UnderS_decr: "A \<le> B \<Longrightarrow> UnderS B \<le> UnderS A"
+by(unfold UnderS_def, auto)
+
+lemma (in rel) Above_decr: "A \<le> B \<Longrightarrow> Above B \<le> Above A"
+by(unfold Above_def, auto)
+
+lemma (in rel) AboveS_decr: "A \<le> B \<Longrightarrow> AboveS B \<le> AboveS A"
+by(unfold AboveS_def, auto)
+
+lemma (in rel) under_incl_iff:
+assumes TRANS: "trans r" and REFL: "Refl r" and IN: "a \<in> Field r"
+shows "(under a \<le> under b) = ((a,b) \<in> r)"
+proof
+ assume "(a,b) \<in> r"
+ thus "under a \<le> under b" using TRANS
+ by (auto simp add: under_incr)
+next
+ assume "under a \<le> under b"
+ moreover
+ have "a \<in> under a" using REFL IN
+ by (auto simp add: Refl_under_in)
+ ultimately show "(a,b) \<in> r"
+ by (auto simp add: under_def)
+qed
+
+lemma (in rel) above_decr:
+assumes TRANS: "trans r" and REL: "(a,b) \<in> r"
+shows "above b \<le> above a"
+proof(unfold above_def, auto)
+ fix x assume "(b,x) \<in> r"
+ with REL TRANS trans_def[of r]
+ show "(a,x) \<in> r" by blast
+qed
+
+lemma (in rel) aboveS_decr:
+assumes TRANS: "trans r" and ANTISYM: "antisym r" and
+ REL: "(a,b) \<in> r"
+shows "aboveS b \<le> aboveS a"
+proof(unfold aboveS_def, auto)
+ assume *: "a \<noteq> b" and **: "(b,a) \<in> r"
+ with ANTISYM antisym_def[of r] REL
+ show False by auto
+next
+ fix x assume "x \<noteq> b" "(b,x) \<in> r"
+ with REL TRANS trans_def[of r]
+ show "(a,x) \<in> r" by blast
+qed
+
+lemma (in rel) under_trans:
+assumes TRANS: "trans r" and
+ IN1: "a \<in> under b" and IN2: "b \<in> under c"
+shows "a \<in> under c"
+proof-
+ have "(a,b) \<in> r \<and> (b,c) \<in> r"
+ using IN1 IN2 under_def by auto
+ hence "(a,c) \<in> r"
+ using TRANS trans_def[of r] by blast
+ thus ?thesis unfolding under_def by simp
+qed
+
+lemma (in rel) under_underS_trans:
+assumes TRANS: "trans r" and ANTISYM: "antisym r" and
+ IN1: "a \<in> under b" and IN2: "b \<in> underS c"
+shows "a \<in> underS c"
+proof-
+ have 0: "(a,b) \<in> r \<and> (b,c) \<in> r"
+ using IN1 IN2 under_def underS_def by auto
+ hence 1: "(a,c) \<in> r"
+ using TRANS trans_def[of r] by blast
+ have 2: "b \<noteq> c" using IN2 underS_def by auto
+ have 3: "a \<noteq> c"
+ proof
+ assume "a = c" with 0 2 ANTISYM antisym_def[of r]
+ show False by auto
+ qed
+ from 1 3 show ?thesis unfolding underS_def by simp
+qed
+
+lemma (in rel) underS_under_trans:
+assumes TRANS: "trans r" and ANTISYM: "antisym r" and
+ IN1: "a \<in> underS b" and IN2: "b \<in> under c"
+shows "a \<in> underS c"
+proof-
+ have 0: "(a,b) \<in> r \<and> (b,c) \<in> r"
+ using IN1 IN2 under_def underS_def by auto
+ hence 1: "(a,c) \<in> r"
+ using TRANS trans_def[of r] by blast
+ have 2: "a \<noteq> b" using IN1 underS_def by auto
+ have 3: "a \<noteq> c"
+ proof
+ assume "a = c" with 0 2 ANTISYM antisym_def[of r]
+ show False by auto
+ qed
+ from 1 3 show ?thesis unfolding underS_def by simp
+qed
+
+lemma (in rel) underS_underS_trans:
+assumes TRANS: "trans r" and ANTISYM: "antisym r" and
+ IN1: "a \<in> underS b" and IN2: "b \<in> underS c"
+shows "a \<in> underS c"
+proof-
+ have "a \<in> under b"
+ using IN1 underS_subset_under by auto
+ with assms under_underS_trans show ?thesis by auto
+qed
+
+lemma (in rel) above_trans:
+assumes TRANS: "trans r" and
+ IN1: "b \<in> above a" and IN2: "c \<in> above b"
+shows "c \<in> above a"
+proof-
+ have "(a,b) \<in> r \<and> (b,c) \<in> r"
+ using IN1 IN2 above_def by auto
+ hence "(a,c) \<in> r"
+ using TRANS trans_def[of r] by blast
+ thus ?thesis unfolding above_def by simp
+qed
+
+lemma (in rel) above_aboveS_trans:
+assumes TRANS: "trans r" and ANTISYM: "antisym r" and
+ IN1: "b \<in> above a" and IN2: "c \<in> aboveS b"
+shows "c \<in> aboveS a"
+proof-
+ have 0: "(a,b) \<in> r \<and> (b,c) \<in> r"
+ using IN1 IN2 above_def aboveS_def by auto
+ hence 1: "(a,c) \<in> r"
+ using TRANS trans_def[of r] by blast
+ have 2: "b \<noteq> c" using IN2 aboveS_def by auto
+ have 3: "a \<noteq> c"
+ proof
+ assume "a = c" with 0 2 ANTISYM antisym_def[of r]
+ show False by auto
+ qed
+ from 1 3 show ?thesis unfolding aboveS_def by simp
+qed
+
+lemma (in rel) aboveS_above_trans:
+assumes TRANS: "trans r" and ANTISYM: "antisym r" and
+ IN1: "b \<in> aboveS a" and IN2: "c \<in> above b"
+shows "c \<in> aboveS a"
+proof-
+ have 0: "(a,b) \<in> r \<and> (b,c) \<in> r"
+ using IN1 IN2 above_def aboveS_def by auto
+ hence 1: "(a,c) \<in> r"
+ using TRANS trans_def[of r] by blast
+ have 2: "a \<noteq> b" using IN1 aboveS_def by auto
+ have 3: "a \<noteq> c"
+ proof
+ assume "a = c" with 0 2 ANTISYM antisym_def[of r]
+ show False by auto
+ qed
+ from 1 3 show ?thesis unfolding aboveS_def by simp
+qed
+
+lemma (in rel) aboveS_aboveS_trans:
+assumes TRANS: "trans r" and ANTISYM: "antisym r" and
+ IN1: "b \<in> aboveS a" and IN2: "c \<in> aboveS b"
+shows "c \<in> aboveS a"
+proof-
+ have "b \<in> above a"
+ using IN1 aboveS_subset_above by auto
+ with assms above_aboveS_trans show ?thesis by auto
+qed
+
+lemma (in rel) underS_Under_trans:
+assumes TRANS: "trans r" and ANTISYM: "antisym r" and
+ IN1: "a \<in> underS b" and IN2: "b \<in> Under C"
+shows "a \<in> UnderS C"
+proof-
+ from IN1 have "a \<in> under b"
+ using underS_subset_under[of b] by blast
+ with assms under_Under_trans
+ have "a \<in> Under C" by auto
+ (* *)
+ moreover
+ have "a \<notin> C"
+ proof
+ assume *: "a \<in> C"
+ have 1: "b \<noteq> a \<and> (a,b) \<in> r"
+ using IN1 underS_def[of b] by auto
+ have "\<forall>c \<in> C. (b,c) \<in> r"
+ using IN2 Under_def[of C] by auto
+ with * have "(b,a) \<in> r" by simp
+ with 1 ANTISYM antisym_def[of r]
+ show False by blast
+ qed
+ (* *)
+ ultimately
+ show ?thesis unfolding UnderS_def
+ using Under_def by auto
+qed
+
+lemma (in rel) underS_UnderS_trans:
+assumes TRANS: "trans r" and ANTISYM: "antisym r" and
+ IN1: "a \<in> underS b" and IN2: "b \<in> UnderS C"
+shows "a \<in> UnderS C"
+proof-
+ from IN2 have "b \<in> Under C"
+ using UnderS_subset_Under[of C] by blast
+ with underS_Under_trans assms
+ show ?thesis by auto
+qed
+
+lemma (in rel) above_Above_trans:
+assumes TRANS: "trans r" and
+ IN1: "a \<in> above b" and IN2: "b \<in> Above C"
+shows "a \<in> Above C"
+proof-
+ have "(b,a) \<in> r \<and> (\<forall>c \<in> C. (c,b) \<in> r)"
+ using IN1 IN2 above_def Above_def by auto
+ hence "\<forall>c \<in> C. (c,a) \<in> r"
+ using TRANS trans_def[of r] by blast
+ moreover
+ have "a \<in> Field r" using IN1 Field_def above_def by force
+ ultimately
+ show ?thesis unfolding Above_def by auto
+qed
+
+lemma (in rel) aboveS_Above_trans:
+assumes TRANS: "trans r" and ANTISYM: "antisym r" and
+ IN1: "a \<in> aboveS b" and IN2: "b \<in> Above C"
+shows "a \<in> AboveS C"
+proof-
+ from IN1 have "a \<in> above b"
+ using aboveS_subset_above[of b] by blast
+ with assms above_Above_trans
+ have "a \<in> Above C" by auto
+ (* *)
+ moreover
+ have "a \<notin> C"
+ proof
+ assume *: "a \<in> C"
+ have 1: "b \<noteq> a \<and> (b,a) \<in> r"
+ using IN1 aboveS_def[of b] by auto
+ have "\<forall>c \<in> C. (c,b) \<in> r"
+ using IN2 Above_def[of C] by auto
+ with * have "(a,b) \<in> r" by simp
+ with 1 ANTISYM antisym_def[of r]
+ show False by blast
+ qed
+ (* *)
+ ultimately
+ show ?thesis unfolding AboveS_def
+ using Above_def by auto
+qed
+
+lemma (in rel) above_AboveS_trans:
+assumes TRANS: "trans r" and ANTISYM: "antisym r" and
+ IN1: "a \<in> above b" and IN2: "b \<in> AboveS C"
+shows "a \<in> AboveS C"
+proof-
+ from IN2 have "b \<in> Above C"
+ using AboveS_subset_Above[of C] by blast
+ with assms above_Above_trans
+ have "a \<in> Above C" by auto
+ (* *)
+ moreover
+ have "a \<notin> C"
+ proof
+ assume *: "a \<in> C"
+ have 1: "(b,a) \<in> r"
+ using IN1 above_def[of b] by auto
+ have "\<forall>c \<in> C. b \<noteq> c \<and> (c,b) \<in> r"
+ using IN2 AboveS_def[of C] by auto
+ with * have "b \<noteq> a \<and> (a,b) \<in> r" by simp
+ with 1 ANTISYM antisym_def[of r]
+ show False by blast
+ qed
+ (* *)
+ ultimately
+ show ?thesis unfolding AboveS_def
+ using Above_def by auto
+qed
+
+lemma (in rel) aboveS_AboveS_trans:
+assumes TRANS: "trans r" and ANTISYM: "antisym r" and
+ IN1: "a \<in> aboveS b" and IN2: "b \<in> AboveS C"
+shows "a \<in> AboveS C"
+proof-
+ from IN2 have "b \<in> Above C"
+ using AboveS_subset_Above[of C] by blast
+ with aboveS_Above_trans assms
+ show ?thesis by auto
+qed
+
+
+subsection {* Properties depending on more than one relation *}
+
+abbreviation "under \<equiv> rel.under"
+abbreviation "underS \<equiv> rel.underS"
+abbreviation "Under \<equiv> rel.Under"
+abbreviation "UnderS \<equiv> rel.UnderS"
+abbreviation "above \<equiv> rel.above"
+abbreviation "aboveS \<equiv> rel.aboveS"
+abbreviation "Above \<equiv> rel.Above"
+abbreviation "AboveS \<equiv> rel.AboveS"
+
+lemma under_incr2:
+"r \<le> r' \<Longrightarrow> under r a \<le> under r' a"
+unfolding rel.under_def by blast
+
+lemma underS_incr2:
+"r \<le> r' \<Longrightarrow> underS r a \<le> underS r' a"
+unfolding rel.underS_def by blast
+
+lemma Under_incr:
+"r \<le> r' \<Longrightarrow> Under r A \<le> Under r A"
+unfolding rel.Under_def by blast
+
+lemma UnderS_incr:
+"r \<le> r' \<Longrightarrow> UnderS r A \<le> UnderS r A"
+unfolding rel.UnderS_def by blast
+
+lemma Under_incr_decr:
+"\<lbrakk>r \<le> r'; A' \<le> A\<rbrakk> \<Longrightarrow> Under r A \<le> Under r A'"
+unfolding rel.Under_def by blast
+
+lemma UnderS_incr_decr:
+"\<lbrakk>r \<le> r'; A' \<le> A\<rbrakk> \<Longrightarrow> UnderS r A \<le> UnderS r A'"
+unfolding rel.UnderS_def by blast
+
+lemma above_incr2:
+"r \<le> r' \<Longrightarrow> above r a \<le> above r' a"
+unfolding rel.above_def by blast
+
+lemma aboveS_incr2:
+"r \<le> r' \<Longrightarrow> aboveS r a \<le> aboveS r' a"
+unfolding rel.aboveS_def by blast
+
+lemma Above_incr:
+"r \<le> r' \<Longrightarrow> Above r A \<le> Above r A"
+unfolding rel.Above_def by blast
+
+lemma AboveS_incr:
+"r \<le> r' \<Longrightarrow> AboveS r A \<le> AboveS r A"
+unfolding rel.AboveS_def by blast
+
+lemma Above_incr_decr:
+"\<lbrakk>r \<le> r'; A' \<le> A\<rbrakk> \<Longrightarrow> Above r A \<le> Above r A'"
+unfolding rel.Above_def by blast
+
+lemma AboveS_incr_decr:
+"\<lbrakk>r \<le> r'; A' \<le> A\<rbrakk> \<Longrightarrow> AboveS r A \<le> AboveS r A'"
+unfolding rel.AboveS_def by blast
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Ordinals_and_Cardinals/Order_Relation_More_Base.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,304 @@
+(* Title: HOL/Ordinals_and_Cardinals/Order_Relation_More_Base.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Basics on order-like relations (base).
+*)
+
+header {* Basics on Order-Like Relations (Base) *}
+
+theory Order_Relation_More_Base
+imports "~~/src/HOL/Library/Order_Relation"
+begin
+
+
+text{* In this section, we develop basic concepts and results pertaining
+to order-like relations, i.e., to reflexive and/or transitive and/or symmetric and/or
+total relations. The development is placed on top of the definitions
+from the theory @{text "Order_Relation"}. We also
+further define upper and lower bounds operators. *}
+
+
+locale rel = fixes r :: "'a rel"
+
+text{* The following context encompasses all this section, except
+for its last subsection. In other words, for the rest of this section except its last
+subsection, we consider a fixed relation @{text "r"}. *}
+
+context rel
+begin
+
+
+subsection {* Auxiliaries *}
+
+
+lemma refl_on_domain:
+"\<lbrakk>refl_on A r; (a,b) : r\<rbrakk> \<Longrightarrow> a \<in> A \<and> b \<in> A"
+by(auto simp add: refl_on_def)
+
+
+corollary well_order_on_domain:
+"\<lbrakk>well_order_on A r; (a,b) \<in> r\<rbrakk> \<Longrightarrow> a \<in> A \<and> b \<in> A"
+by(simp add: refl_on_domain order_on_defs)
+
+
+lemma well_order_on_Field:
+"well_order_on A r \<Longrightarrow> A = Field r"
+by(auto simp add: refl_on_def Field_def order_on_defs)
+
+
+lemma well_order_on_Well_order:
+"well_order_on A r \<Longrightarrow> A = Field r \<and> Well_order r"
+using well_order_on_Field by simp
+
+
+lemma Total_Id_Field:
+assumes TOT: "Total r" and NID: "\<not> (r <= Id)"
+shows "Field r = Field(r - Id)"
+using mono_Field[of "r - Id" r] Diff_subset[of r Id]
+proof(auto)
+ have "r \<noteq> {}" using NID by fast
+ then obtain b and c where "b \<noteq> c \<and> (b,c) \<in> r" using NID by fast
+ hence 1: "b \<noteq> c \<and> {b,c} \<le> Field r" by (auto simp: Field_def)
+ (* *)
+ fix a assume *: "a \<in> Field r"
+ obtain d where 2: "d \<in> Field r" and 3: "d \<noteq> a"
+ using * 1 by blast
+ hence "(a,d) \<in> r \<or> (d,a) \<in> r" using * TOT
+ by (simp add: total_on_def)
+ thus "a \<in> Field(r - Id)" using 3 unfolding Field_def by blast
+qed
+
+
+lemma Total_subset_Id:
+assumes TOT: "Total r" and SUB: "r \<le> Id"
+shows "r = {} \<or> (\<exists>a. r = {(a,a)})"
+proof-
+ {assume "r \<noteq> {}"
+ then obtain a b where 1: "(a,b) \<in> r" by fast
+ hence "a = b" using SUB by blast
+ hence 2: "(a,a) \<in> r" using 1 by simp
+ {fix c d assume "(c,d) \<in> r"
+ hence "{a,c,d} \<le> Field r" using 1 unfolding Field_def by blast
+ hence "((a,c) \<in> r \<or> (c,a) \<in> r \<or> a = c) \<and>
+ ((a,d) \<in> r \<or> (d,a) \<in> r \<or> a = d)"
+ using TOT unfolding total_on_def by blast
+ hence "a = c \<and> a = d" using SUB by blast
+ }
+ hence "r \<le> {(a,a)}" by auto
+ with 2 have "\<exists>a. r = {(a,a)}" by blast
+ }
+ thus ?thesis by blast
+qed
+
+
+lemma Linear_order_in_diff_Id:
+assumes LI: "Linear_order r" and
+ IN1: "a \<in> Field r" and IN2: "b \<in> Field r"
+shows "((a,b) \<in> r) = ((b,a) \<notin> r - Id)"
+using assms unfolding order_on_defs total_on_def antisym_def Id_def refl_on_def by force
+
+
+subsection {* The upper and lower bounds operators *}
+
+
+text{* Here we define upper (``above") and lower (``below") bounds operators.
+We think of @{text "r"} as a {\em non-strict} relation. The suffix ``S"
+at the names of some operators indicates that the bounds are strict -- e.g.,
+@{text "underS a"} is the set of all strict lower bounds of @{text "a"} (w.r.t. @{text "r"}).
+Capitalization of the first letter in the name reminds that the operator acts on sets, rather
+than on individual elements. *}
+
+definition under::"'a \<Rightarrow> 'a set"
+where "under a \<equiv> {b. (b,a) \<in> r}"
+
+definition underS::"'a \<Rightarrow> 'a set"
+where "underS a \<equiv> {b. b \<noteq> a \<and> (b,a) \<in> r}"
+
+definition Under::"'a set \<Rightarrow> 'a set"
+where "Under A \<equiv> {b \<in> Field r. \<forall>a \<in> A. (b,a) \<in> r}"
+
+definition UnderS::"'a set \<Rightarrow> 'a set"
+where "UnderS A \<equiv> {b \<in> Field r. \<forall>a \<in> A. b \<noteq> a \<and> (b,a) \<in> r}"
+
+definition above::"'a \<Rightarrow> 'a set"
+where "above a \<equiv> {b. (a,b) \<in> r}"
+
+definition aboveS::"'a \<Rightarrow> 'a set"
+where "aboveS a \<equiv> {b. b \<noteq> a \<and> (a,b) \<in> r}"
+
+definition Above::"'a set \<Rightarrow> 'a set"
+where "Above A \<equiv> {b \<in> Field r. \<forall>a \<in> A. (a,b) \<in> r}"
+
+definition AboveS::"'a set \<Rightarrow> 'a set"
+where "AboveS A \<equiv> {b \<in> Field r. \<forall>a \<in> A. b \<noteq> a \<and> (a,b) \<in> r}"
+(* *)
+
+text{* Note: In the definitions of @{text "Above[S]"} and @{text "Under[S]"},
+ we bounded comprehension by @{text "Field r"} in order to properly cover
+ the case of @{text "A"} being empty. *}
+
+
+lemma UnderS_subset_Under: "UnderS A \<le> Under A"
+by(auto simp add: UnderS_def Under_def)
+
+
+lemma underS_subset_under: "underS a \<le> under a"
+by(auto simp add: underS_def under_def)
+
+
+lemma underS_notIn: "a \<notin> underS a"
+by(simp add: underS_def)
+
+
+lemma Refl_under_in: "\<lbrakk>Refl r; a \<in> Field r\<rbrakk> \<Longrightarrow> a \<in> under a"
+by(simp add: refl_on_def under_def)
+
+
+lemma AboveS_disjoint: "A Int (AboveS A) = {}"
+by(auto simp add: AboveS_def)
+
+
+lemma in_AboveS_underS: "a \<in> Field r \<Longrightarrow> a \<in> AboveS (underS a)"
+by(auto simp add: AboveS_def underS_def)
+
+
+lemma Refl_under_underS:
+assumes "Refl r" "a \<in> Field r"
+shows "under a = underS a \<union> {a}"
+unfolding under_def underS_def
+using assms refl_on_def[of _ r] by fastforce
+
+
+lemma underS_empty: "a \<notin> Field r \<Longrightarrow> underS a = {}"
+by (auto simp: Field_def underS_def)
+
+
+lemma under_Field: "under a \<le> Field r"
+by(unfold under_def Field_def, auto)
+
+
+lemma underS_Field: "underS a \<le> Field r"
+by(unfold underS_def Field_def, auto)
+
+
+lemma underS_Field2:
+"a \<in> Field r \<Longrightarrow> underS a < Field r"
+using assms underS_notIn underS_Field by blast
+
+
+lemma underS_Field3:
+"Field r \<noteq> {} \<Longrightarrow> underS a < Field r"
+by(cases "a \<in> Field r", simp add: underS_Field2, auto simp add: underS_empty)
+
+
+lemma Under_Field: "Under A \<le> Field r"
+by(unfold Under_def Field_def, auto)
+
+
+lemma UnderS_Field: "UnderS A \<le> Field r"
+by(unfold UnderS_def Field_def, auto)
+
+
+lemma AboveS_Field: "AboveS A \<le> Field r"
+by(unfold AboveS_def Field_def, auto)
+
+
+lemma under_incr:
+assumes TRANS: "trans r" and REL: "(a,b) \<in> r"
+shows "under a \<le> under b"
+proof(unfold under_def, auto)
+ fix x assume "(x,a) \<in> r"
+ with REL TRANS trans_def[of r]
+ show "(x,b) \<in> r" by blast
+qed
+
+
+lemma underS_incr:
+assumes TRANS: "trans r" and ANTISYM: "antisym r" and
+ REL: "(a,b) \<in> r"
+shows "underS a \<le> underS b"
+proof(unfold underS_def, auto)
+ assume *: "b \<noteq> a" and **: "(b,a) \<in> r"
+ with ANTISYM antisym_def[of r] REL
+ show False by blast
+next
+ fix x assume "x \<noteq> a" "(x,a) \<in> r"
+ with REL TRANS trans_def[of r]
+ show "(x,b) \<in> r" by blast
+qed
+
+
+lemma underS_incl_iff:
+assumes LO: "Linear_order r" and
+ INa: "a \<in> Field r" and INb: "b \<in> Field r"
+shows "(underS a \<le> underS b) = ((a,b) \<in> r)"
+proof
+ assume "(a,b) \<in> r"
+ thus "underS a \<le> underS b" using LO
+ by (simp add: order_on_defs underS_incr)
+next
+ assume *: "underS a \<le> underS b"
+ {assume "a = b"
+ hence "(a,b) \<in> r" using assms
+ by (simp add: order_on_defs refl_on_def)
+ }
+ moreover
+ {assume "a \<noteq> b \<and> (b,a) \<in> r"
+ hence "b \<in> underS a" unfolding underS_def by blast
+ hence "b \<in> underS b" using * by blast
+ hence False by (simp add: underS_notIn)
+ }
+ ultimately
+ show "(a,b) \<in> r" using assms
+ order_on_defs[of "Field r" r] total_on_def[of "Field r" r] by blast
+qed
+
+
+lemma under_Under_trans:
+assumes TRANS: "trans r" and
+ IN1: "a \<in> under b" and IN2: "b \<in> Under C"
+shows "a \<in> Under C"
+proof-
+ have "(a,b) \<in> r \<and> (\<forall>c \<in> C. (b,c) \<in> r)"
+ using IN1 IN2 under_def Under_def by blast
+ hence "\<forall>c \<in> C. (a,c) \<in> r"
+ using TRANS trans_def[of r] by blast
+ moreover
+ have "a \<in> Field r" using IN1 unfolding Field_def under_def by blast
+ ultimately
+ show ?thesis unfolding Under_def by blast
+qed
+
+
+lemma under_UnderS_trans:
+assumes TRANS: "trans r" and ANTISYM: "antisym r" and
+ IN1: "a \<in> under b" and IN2: "b \<in> UnderS C"
+shows "a \<in> UnderS C"
+proof-
+ from IN2 have "b \<in> Under C"
+ using UnderS_subset_Under[of C] by blast
+ with assms under_Under_trans
+ have "a \<in> Under C" by blast
+ (* *)
+ moreover
+ have "a \<notin> C"
+ proof
+ assume *: "a \<in> C"
+ have 1: "(a,b) \<in> r"
+ using IN1 under_def[of b] by auto
+ have "\<forall>c \<in> C. b \<noteq> c \<and> (b,c) \<in> r"
+ using IN2 UnderS_def[of C] by blast
+ with * have "b \<noteq> a \<and> (b,a) \<in> r" by blast
+ with 1 ANTISYM antisym_def[of r]
+ show False by blast
+ qed
+ (* *)
+ ultimately
+ show ?thesis unfolding UnderS_def Under_def by fast
+qed
+
+
+end (* context rel *)
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Ordinals_and_Cardinals/README.txt Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,218 @@
+Authors: Andrei Popescu & Dmitriy Traytel
+
+
+PDF Documentation:
+-----------------
+
+See "document/root.pdf".
+
+
+
+Short description of the theories' main content:
+-----------------------------------------------
+
+The "minor" theories Fun_More, Wellfounded_More and Order_Relation_More are
+extensions of the existing theories Fun, Wellfounded and
+Order_Relation:
+-- Fun_More states more facts (mainly) concerning injections, bijections,
+inverses, and (numeric) cardinals of finite sets.
+-- Wellfounded_More states variations of well-founded
+recursion and well-founded recursion.
+-- Order_Relation_More fixes a relation, defines upper and lower bounds
+operators and proves many basic properties for these
+(depending on assumptions such as reflexivity or transitivity).
+
+The "major" theories are:
+-- Wellorder_Relation: Here one fixes a well-order relation, and then:
+----- 1) Defines the concepts of maximum (of two elements), minimum (of a set), supremum,
+ successor (of a set), and order filter (i.e., downwards closed set, a.k.a.
+ initial segment).
+-- Wellorder_Embedding:
+----- 2) For two well-order relations,
+ defines *well-order embeddings* as injective functions copying
+ the source into an order filter of the target and *compatible functions*
+ as those preserving the order. Also, *isomorphisms*
+ and *strict embeddings* are defined to be embeddings that are, and respectively
+ are not, bijections.
+
+-- Constructions_on_Wellorders:
+----- 1) Defines direct images, restrictions, disjoint unions and
+ bounded squares of well-orders.
+----- 2) Defines the relations "ordLeq", "ordLess" and "ordIso"
+ between well-order relations (concrete syntax: r <=o r', r <o r' and
+ r =o r', respectively), defined by the existence of an embedding,
+ strict embedding and isomorphism, respectively between the two members.
+ Among the properties proved for these relations:
+--------- ordLeq is total;
+--------- ordLess (on a fixed type) is well-founded.
+
+-- Cardinal_Order_Relation:
+---- 1) Defines a *cardinal order* to be a well-order minim w.r.t. "ordLeq"
+ (or, equivalently, minimal w.r.t. "ordLess").
+ ordLess being well-founded together with the well-ordering theorem (from theory Zorn.thy)
+ ensures the existence of a cardinal relation on any given set. In addition,
+ a cardinal relation on a set is unique up to order isomorphism.
+---- 2) Defines the cardinal of a set A, |A|, to be SOME cardinal
+ order on it (unique up to =o, according to the above).
+---- 3) Proves properties of cardinals, including their
+ interactions with sums, products, unions, lists,
+ powersets, sets of finite sets. Among them, nontrivial
+ facts concerning the invariance of infinite cardinals
+ under some of these constructs -- e.g., if A is infinite,
+ than the cardinal of lists/words over A is the same (up to
+ the "cardinal equality" =o) as that of A.
+---- 5) Studies the connection between the introduced (order-based) notion
+ of cardinal and the numeric one previously defined for
+ finite sets (operator "card"). On the way, one introduces the cardinal omega
+ (natLeq) and the finite cardinals (natLeq_on n).
+---- 6) Defines and proves the existence of successor cardinals.
+
+-- Cardinal_Arithmetic
+
+
+Here is a list of names of proved facts concerning cardinalities that are
+expressed independently of notions of order, and of potential interest
+for "working mathematicians":
+--- one_set_greater, one_type_greater (their proofs use the
+ fact that ordinals are totally ordered)
+--- Plus_into_Times, Plus_into_Times_types,
+ Plus_infinite_bij_betw, Plus_infinite_bij_betw_types,
+ Times_same_infinite_bij_betw, Times_same_infinite_bij_betw_types,
+ Times_infinite_bij_betw, Times_infinite_bij_betw_types
+ inj_on_UNION_infinite, List_infinite_bij_betw, List_infinite_bij_betw_types
+ Fpow_infinite_bij_betw
+ (their proofs employ cardinals)
+
+
+
+
+Minor technicalities and naming issues:
+---------------------------------------
+
+1. Most of the definitions and theorems are proved in files suffixed with
+"_Base". Bootstrapping considerations (for the (co)datatype package) made this
+division desirable.
+
+
+2. Even though we would have preferred to use "initial segment" instead of
+"order filter", we chose the latter to avoid terminological clash with
+the operator "init_seg_of" from Zorn.thy. The latter expresses a related, but different
+concept -- it considers a relation, rather than a set, as initial segment of a relation.
+
+
+3. We prefer to define the upper-bound operations under, underS,
+etc., as opposed to working with combinations of relation image,
+converse and diagonal, because the former seem more intuitive
+notations when we think of orderings (but of course we cannot
+define them as abbreviations, as this would have a global
+effect, also affecting cases where one does not think of relations
+as orders). Moreover, in my locales the relation parameter r for
+under, underS etc. is fixed, hence these operations can keep r
+implicit. To get a concrete glimpse at my aesthetic reason for
+introducing these operations: otherwise, instead of "underS a",
+we would have to write "(r - Id)^-1 `` {a}" or "r^-1 `` {a} - Id".
+
+
+4. Even though the main focus of this development are
+the well-order relations, we prove the basic results on order relations
+and bounds as generally as possible.
+To the contrary, the results concerning minima, suprema and successors
+are stated for well-order relations, not maximally generally.
+
+
+5. "Refl_on A r" requires in particular that "r <= A <*> A",
+and therefore whenever "Refl_on A r", we have that necessarily
+"A = Field r". This means that in a theory of orders the domain
+A would be redundant -- we decided not to include it explicitly
+for most of the tehory.
+
+
+6. An infinite ordinal/cardinal is one for which the field is infinite.
+We always prefer the slightly more verbose "finite (Field r)" to the more
+compact but less standard equivalent condition "finite r".
+
+
+7. After we proved lots of facts about injections and
+bijections, we discovered that a couple of
+fancier (set-restricted) version of some of them are proved in
+the theory FuncSet. However, we did not need here restricted
+abstraction and such, and felt we should not import the whole
+theory for just a couple of minor facts.
+
+
+
+
+
+
+
+Notes for anyone who would like to enrich these theories in the future
+--------------------------------------------------------------------------------------
+
+Theory Fun_More (and Fun_More_Base):
+- Careful: "inj" is an abbreviation for "inj_on UNIV", while
+ "bij" is not an abreviation for "bij_betw UNIV UNIV", but
+ a defined constant; there is no "surj_betw", but only "surj".
+ "inv" is an abbreviation for "inv_into UNIV"
+- In subsection "Purely functional properties":
+-- Recall lemma "comp_inj_on".
+-- A lemma for inj_on corresponding to "bij_betw_same_card" already exists, and is called "card_inj_on_le".
+- In subsection "Properties involving Hilbert choice":
+-- One should refrain from trying to prove "intuitive" properties of f conditioned
+ by properties of (inv_into f A), such as "bij_betw A' A (inv_into f A) ==> bij_betw A A' f".
+ They usually do not hold, since one cannot usually infer the well-definedness of "inv_into f A".
+- A lemma "bij_betw_inv_into_LEFT" -- why didn't
+"proof(auto simp add: bij_betw_inv_into_left)" finish the proof?
+-- Recall lemma "bij_betw_inv_into".
+- In subsection "Other facts":
+-- Does the lemma "atLeastLessThan_injective" already exist anywhere?
+
+Theory Order_Relation_More (and Order_Relation_More_Base):
+- In subsection "Auxiliaries":
+-- Recall the lemmas "order_on_defs", "Field_def", "Domain_def", "Range_def", "converse_def".
+-- Recall that "refl_on r A" forces r to not be defined outside A.
+ This is why "partial_order_def"
+ can afford to use non-parameterized versions of antisym and trans.
+-- Recall the ASCII notation for "converse r": "r ^-1".
+-- Recall the abbreviations:
+ abbreviation "Refl r ≡ refl_on (Field r) r"
+ abbreviation "Preorder r ≡ preorder_on (Field r) r"
+ abbreviation "Partial_order r ≡ partial_order_on (Field r) r"
+ abbreviation "Total r ≡ total_on (Field r) r"
+ abbreviation "Linear_order r ≡ linear_order_on (Field r) r"
+ abbreviation "Well_order r ≡ well_order_on (Field r) r"
+
+Theory Wellorder_Relation (and Wellorder_Relation_Base):
+- In subsection "Auxiliaries": recall lemmas "order_on_defs"
+- In subsection "The notions of maximum, minimum, supremum, successor and order filter":
+ Should we define all constants from "wo_rel" in "rel" instead,
+ so that their outside definition not be conditional in "wo_rel r"?
+
+Theory Wellfounded_More (and Wellfounded_More_Base):
+ Recall the lemmas "wfrec" and "wf_induct".
+
+Theory Wellorder_Embedding (and Wellorder_Embedding_Base):
+- Recall "inj_on_def" and "bij_betw_def".
+- Line 5 in the proof of lemma embed_in_Field: we have to figure out for this and many other
+ situations: Why did it work without annotations to Refl_under_in?
+- At the proof of theorem "wellorders_totally_ordered" (and, similarly, elsewhere):
+ Had we used metavariables instead of local definitions for H, f, g and test, the
+ goals (in the goal window) would have become unreadable,
+ making impossible to debug theorem instantiations.
+- At lemma "embed_unique": If we add the attribute "rule format" at lemma, we get an error at qed.
+
+Theory Constructions_on_Wellorders (and Constructions_on_Wellorders_Base):
+- Some of the lemmas in this section are about more general kinds of relations than
+ well-orders, but it is not clear whether they are useful in such more general contexts.
+- Recall that "equiv" does not have the "equiv_on" and "Equiv" versions,
+ like the order relation. "equiv" corresponds, for instance, to "well_order_on".
+- The lemmas "ord_trans" are not clearly useful together, as their employment within blast or auto
+tends to diverge.
+
+Theory Cardinal_Order_Relation (and Cardinal_Order_Relation_Base):
+- Careful: if "|..|" meets an outer parehthesis, an extra space needs to be inserted, as in
+ "( |A| )".
+- At lemma like ordLeq_Sigma_mono1: Not worth stating something like ordLeq_Sigma_mono2 --
+ would be a mere instance of card_of_Sigma_mono2.
+- At lemma ordLeq_Sigma_cong2: Again, no reason for stating something like ordLeq_Sigma_cong2.
+- At lemma Fpow_Pow_finite: why wouldn't a version of this lemma with "... Int finite"
+also be proved by blast?
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Ordinals_and_Cardinals/TODO.txt Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,59 @@
+-1. At the new Isabelle release (after 2009-2), replace the operator List
+from Cardinal_Order_Relation with "lists" from List.thy in standard library.
+A lemma for "lists" is the actual definition of "List".
+
+0. Add:
+
+lemma finite_iff_cardOf_nat:
+"finite A = ( |A| <o |UNIV :: nat set| )"
+using infinite_iff_card_of_nat[of A]
+by (auto simp add: card_of_Well_order ordLess_iff_not_ordLeq)
+
+
+
+
+1. Remember that the version from my computer has a few more theorems than the one
+from the Isabelle site.
+
+2. Many basic cardinal arithmetic facts can be listed as simps!
+The the proofs can be simplified.
+
+3. Add:
+
+lemma card_of_Plus_ordLeq_infinite:
+assumes "infinite C" and "|A| );£o |C|" and "|B| £o |C|"-A
+shows "|A <+> B| );£o |C|"-A
+proof-
+ let ?phi = "(EX a1 a2. a1 ~= a2 );Ù {a1,a2} <= A) Ù -A
+ (EX b1 b2. b1 ~= b2 );Ù {b1,b2} <= B)"-A
+ show ?thesis
+ proof(cases ?phi, auto)
+ fix a1 b1 a2 b2
+ assume "a1 );¹ a2" "b1 ¹ b2" "a1 Î A" "a2 Î A" "b1 Î B" "b2 Î B"-A
+ hence "|A <+> B| <=o |A <*> B|"
+ apply - apply(rule card_of_Plus_Times) by auto
+ moreover have "|A <*> B| );£o |C|"-A
+ using assms by (auto simp add: card_of_Times_ordLeq_infinite)
+ ultimately show ?thesis using ordLeq_transitive by blast
+ next
+ assume "-=Õa1. a1 );Î A (<_(B (-=Õa2. a1 = a2 );Ú a2 Ï A)"-A
+ then obtain a where "A <= {a}" by blast
+ {assume "A = {}"
+ hence "|A <+> B| =o |B|"
+ using ordIso_symmetric card_of_Plus_empty2 by blast
+ hence ?thesis using assms ordIso_ordLeq_trans by blast
+ }
+ moreover
+ {assume "A = {a}"
+
+ }
+ qed
+qed
+
+
+lemma card_of_Plus_ordLess_infinite:
+assumes "infinite C" and "|A| <o |C|" and "|B| <o |C|"
+shows "|A <+> B| <o |C|"
+
+
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Ordinals_and_Cardinals/Wellfounded_More.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,50 @@
+(* Title: HOL/Ordinals_and_Cardinals/Wellfounded_More.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+More on well-founded relations.
+*)
+
+header {* More on Well-Founded Relations *}
+
+theory Wellfounded_More
+imports Wellfounded_More_Base Order_Relation_More
+begin
+
+
+subsection {* Well-founded recursion via genuine fixpoints *}
+
+(*2*)lemma adm_wf_unique_fixpoint:
+fixes r :: "('a * 'a) set" and
+ H :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" and
+ f :: "'a \<Rightarrow> 'b" and g :: "'a \<Rightarrow> 'b"
+assumes WF: "wf r" and ADM: "adm_wf r H" and fFP: "f = H f" and gFP: "g = H g"
+shows "f = g"
+proof-
+ {fix x
+ have "f x = g x"
+ proof(rule wf_induct[of r "(\<lambda>x. f x = g x)"],
+ auto simp add: WF)
+ fix x assume "\<forall>y. (y, x) \<in> r \<longrightarrow> f y = g y"
+ hence "H f x = H g x" using ADM adm_wf_def[of r H] by auto
+ thus "f x = g x" using fFP and gFP by simp
+ qed
+ }
+ thus ?thesis by (simp add: ext)
+qed
+
+(*2*)lemma wfrec_unique_fixpoint:
+fixes r :: "('a * 'a) set" and
+ H :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" and
+ f :: "'a \<Rightarrow> 'b"
+assumes WF: "wf r" and ADM: "adm_wf r H" and
+ fp: "f = H f"
+shows "f = wfrec r H"
+proof-
+ have "H (wfrec r H) = wfrec r H"
+ using assms wfrec_fixpoint[of r H] by simp
+ thus ?thesis
+ using assms adm_wf_unique_fixpoint[of r H "wfrec r H"] by simp
+qed
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Ordinals_and_Cardinals/Wellfounded_More_Base.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,194 @@
+(* Title: HOL/Ordinals_and_Cardinals/Wellfounded_More_Base.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+More on well-founded relations (base).
+*)
+
+header {* More on Well-Founded Relations (Base) *}
+
+theory Wellfounded_More_Base
+imports Wellfounded Order_Relation_More_Base "~~/src/HOL/Library/Wfrec"
+begin
+
+
+text {* This section contains some variations of results in the theory
+@{text "Wellfounded.thy"}:
+\begin{itemize}
+\item means for slightly more direct definitions by well-founded recursion;
+\item variations of well-founded induction;
+\item means for proving a linear order to be a well-order.
+\end{itemize} *}
+
+
+subsection {* Well-founded recursion via genuine fixpoints *}
+
+
+(*2*)lemma wfrec_fixpoint:
+fixes r :: "('a * 'a) set" and
+ H :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
+assumes WF: "wf r" and ADM: "adm_wf r H"
+shows "wfrec r H = H (wfrec r H)"
+proof(rule ext)
+ fix x
+ have "wfrec r H x = H (cut (wfrec r H) r x) x"
+ using wfrec[of r H] WF by simp
+ also
+ {have "\<And> y. (y,x) : r \<Longrightarrow> (cut (wfrec r H) r x) y = (wfrec r H) y"
+ by (auto simp add: cut_apply)
+ hence "H (cut (wfrec r H) r x) x = H (wfrec r H) x"
+ using ADM adm_wf_def[of r H] by auto
+ }
+ finally show "wfrec r H x = H (wfrec r H) x" .
+qed
+
+
+
+subsection {* Characterizations of well-founded-ness *}
+
+
+text {* A transitive relation is well-founded iff it is ``locally" well-founded,
+i.e., iff its restriction to the lower bounds of of any element is well-founded. *}
+
+(*3*)lemma trans_wf_iff:
+assumes "trans r"
+shows "wf r = (\<forall>a. wf(r Int (r^-1``{a} \<times> r^-1``{a})))"
+proof-
+ obtain R where R_def: "R = (\<lambda> a. r Int (r^-1``{a} \<times> r^-1``{a}))" by blast
+ {assume *: "wf r"
+ {fix a
+ have "wf(R a)"
+ using * R_def wf_subset[of r "R a"] by auto
+ }
+ }
+ (* *)
+ moreover
+ {assume *: "\<forall>a. wf(R a)"
+ have "wf r"
+ proof(unfold wf_def, clarify)
+ fix phi a
+ assume **: "\<forall>a. (\<forall>b. (b,a) \<in> r \<longrightarrow> phi b) \<longrightarrow> phi a"
+ obtain chi where chi_def: "chi = (\<lambda>b. (b,a) \<in> r \<longrightarrow> phi b)" by blast
+ with * have "wf (R a)" by auto
+ hence "(\<forall>b. (\<forall>c. (c,b) \<in> R a \<longrightarrow> chi c) \<longrightarrow> chi b) \<longrightarrow> (\<forall>b. chi b)"
+ unfolding wf_def by blast
+ moreover
+ have "\<forall>b. (\<forall>c. (c,b) \<in> R a \<longrightarrow> chi c) \<longrightarrow> chi b"
+ proof(auto simp add: chi_def R_def)
+ fix b
+ assume 1: "(b,a) \<in> r" and 2: "\<forall>c. (c, b) \<in> r \<and> (c, a) \<in> r \<longrightarrow> phi c"
+ hence "\<forall>c. (c, b) \<in> r \<longrightarrow> phi c"
+ using assms trans_def[of r] by blast
+ thus "phi b" using ** by blast
+ qed
+ ultimately have "\<forall>b. chi b" by (rule mp)
+ with ** chi_def show "phi a" by blast
+ qed
+ }
+ ultimately show ?thesis using R_def by blast
+qed
+
+
+text {* The next lemma is a variation of @{text "wf_eq_minimal"} from Wellfounded,
+allowing one to assume the set included in the field. *}
+
+(*2*)lemma wf_eq_minimal2:
+"wf r = (\<forall>A. A <= Field r \<and> A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. \<not> (a',a) \<in> r))"
+proof-
+ let ?phi = "\<lambda> A. A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. \<not> (a',a) \<in> r)"
+ have "wf r = (\<forall>A. ?phi A)"
+ by (auto simp: ex_in_conv [THEN sym], erule wfE_min, assumption, blast)
+ (rule wfI_min, metis)
+ (* *)
+ also have "(\<forall>A. ?phi A) = (\<forall>B \<le> Field r. ?phi B)"
+ proof
+ assume "\<forall>A. ?phi A"
+ thus "\<forall>B \<le> Field r. ?phi B" by simp
+ next
+ assume *: "\<forall>B \<le> Field r. ?phi B"
+ show "\<forall>A. ?phi A"
+ proof(clarify)
+ fix A::"'a set" assume **: "A \<noteq> {}"
+ obtain B where B_def: "B = A Int (Field r)" by blast
+ show "\<exists>a \<in> A. \<forall>a' \<in> A. (a',a) \<notin> r"
+ proof(cases "B = {}")
+ assume Case1: "B = {}"
+ obtain a where 1: "a \<in> A \<and> a \<notin> Field r"
+ using ** Case1 unfolding B_def by blast
+ hence "\<forall>a' \<in> A. (a',a) \<notin> r" using 1 unfolding Field_def by blast
+ thus ?thesis using 1 by blast
+ next
+ assume Case2: "B \<noteq> {}" have 1: "B \<le> Field r" unfolding B_def by blast
+ obtain a where 2: "a \<in> B \<and> (\<forall>a' \<in> B. (a',a) \<notin> r)"
+ using Case2 1 * by blast
+ have "\<forall>a' \<in> A. (a',a) \<notin> r"
+ proof(clarify)
+ fix a' assume "a' \<in> A" and **: "(a',a) \<in> r"
+ hence "a' \<in> B" unfolding B_def Field_def by blast
+ thus False using 2 ** by blast
+ qed
+ thus ?thesis using 2 unfolding B_def by blast
+ qed
+ qed
+ qed
+ finally show ?thesis by blast
+qed
+
+subsection {* Characterizations of well-founded-ness *}
+
+text {* The next lemma and its corollary enable one to prove that
+a linear order is a well-order in a way which is more standard than
+via well-founded-ness of the strict version of the relation. *}
+
+(*3*)
+lemma Linear_order_wf_diff_Id:
+assumes LI: "Linear_order r"
+shows "wf(r - Id) = (\<forall>A \<le> Field r. A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. (a,a') \<in> r))"
+proof(cases "r \<le> Id")
+ assume Case1: "r \<le> Id"
+ hence temp: "r - Id = {}" by blast
+ hence "wf(r - Id)" by (simp add: temp)
+ moreover
+ {fix A assume *: "A \<le> Field r" and **: "A \<noteq> {}"
+ obtain a where 1: "r = {} \<or> r = {(a,a)}" using LI
+ unfolding order_on_defs using Case1 rel.Total_subset_Id by metis
+ hence "A = {a} \<and> r = {(a,a)}" using * ** unfolding Field_def by blast
+ hence "\<exists>a \<in> A. \<forall>a' \<in> A. (a,a') \<in> r" using 1 by blast
+ }
+ ultimately show ?thesis by blast
+next
+ assume Case2: "\<not> r \<le> Id"
+ hence 1: "Field r = Field(r - Id)" using rel.Total_Id_Field LI
+ unfolding order_on_defs by blast
+ show ?thesis
+ proof
+ assume *: "wf(r - Id)"
+ show "\<forall>A \<le> Field r. A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. (a,a') \<in> r)"
+ proof(clarify)
+ fix A assume **: "A \<le> Field r" and ***: "A \<noteq> {}"
+ hence "\<exists>a \<in> A. \<forall>a' \<in> A. (a',a) \<notin> r - Id"
+ using 1 * unfolding wf_eq_minimal2 by simp
+ moreover have "\<forall>a \<in> A. \<forall>a' \<in> A. ((a,a') \<in> r) = ((a',a) \<notin> r - Id)"
+ using rel.Linear_order_in_diff_Id[of r] ** LI by blast
+ ultimately show "\<exists>a \<in> A. \<forall>a' \<in> A. (a,a') \<in> r" by blast
+ qed
+ next
+ assume *: "\<forall>A \<le> Field r. A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. (a,a') \<in> r)"
+ show "wf(r - Id)"
+ proof(unfold wf_eq_minimal2, clarify)
+ fix A assume **: "A \<le> Field(r - Id)" and ***: "A \<noteq> {}"
+ hence "\<exists>a \<in> A. \<forall>a' \<in> A. (a,a') \<in> r"
+ using 1 * by simp
+ moreover have "\<forall>a \<in> A. \<forall>a' \<in> A. ((a,a') \<in> r) = ((a',a) \<notin> r - Id)"
+ using rel.Linear_order_in_diff_Id[of r] ** LI mono_Field[of "r - Id" r] by blast
+ ultimately show "\<exists>a \<in> A. \<forall>a' \<in> A. (a',a) \<notin> r - Id" by blast
+ qed
+ qed
+qed
+
+(*3*)corollary Linear_order_Well_order_iff:
+assumes "Linear_order r"
+shows "Well_order r = (\<forall>A \<le> Field r. A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. (a,a') \<in> r))"
+using assms unfolding well_order_on_def using Linear_order_wf_diff_Id[of r] by blast
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Ordinals_and_Cardinals/Wellorder_Embedding.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,175 @@
+(* Title: HOL/Ordinals_and_Cardinals/Wellorder_Embedding.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Well-order embeddings.
+*)
+
+header {* Well-Order Embeddings *}
+
+theory Wellorder_Embedding
+imports "../Ordinals_and_Cardinals_Base/Wellorder_Embedding_Base" Fun_More Wellorder_Relation
+begin
+
+
+subsection {* Auxiliaries *}
+
+lemma UNION_bij_betw_ofilter:
+assumes WELL: "Well_order r" and
+ OF: "\<And> i. i \<in> I \<Longrightarrow> ofilter r (A i)" and
+ BIJ: "\<And> i. i \<in> I \<Longrightarrow> bij_betw f (A i) (A' i)"
+shows "bij_betw f (\<Union> i \<in> I. A i) (\<Union> i \<in> I. A' i)"
+proof-
+ have "wo_rel r" using WELL by (simp add: wo_rel_def)
+ hence "\<And> i j. \<lbrakk>i \<in> I; j \<in> I\<rbrakk> \<Longrightarrow> A i \<le> A j \<or> A j \<le> A i"
+ using wo_rel.ofilter_linord[of r] OF by blast
+ with WELL BIJ show ?thesis
+ by (auto simp add: bij_betw_UNION_chain)
+qed
+
+
+subsection {* (Well-order) embeddings, strict embeddings, isomorphisms and order-compatible
+functions *}
+
+lemma embed_halfcong:
+assumes EQ: "\<And> a. a \<in> Field r \<Longrightarrow> f a = g a" and
+ EMB: "embed r r' f"
+shows "embed r r' g"
+proof(unfold embed_def, auto)
+ fix a assume *: "a \<in> Field r"
+ hence "bij_betw f (under r a) (under r' (f a))"
+ using EMB unfolding embed_def by simp
+ moreover
+ {have "under r a \<le> Field r"
+ by (auto simp add: rel.under_Field)
+ hence "\<And> b. b \<in> under r a \<Longrightarrow> f b = g b"
+ using EQ by blast
+ }
+ moreover have "f a = g a" using * EQ by auto
+ ultimately show "bij_betw g (under r a) (under r' (g a))"
+ using bij_betw_cong[of "under r a" f g "under r' (f a)"] by auto
+qed
+
+lemma embed_cong[fundef_cong]:
+assumes "\<And> a. a \<in> Field r \<Longrightarrow> f a = g a"
+shows "embed r r' f = embed r r' g"
+using assms embed_halfcong[of r f g r']
+ embed_halfcong[of r g f r'] by auto
+
+lemma embedS_cong[fundef_cong]:
+assumes "\<And> a. a \<in> Field r \<Longrightarrow> f a = g a"
+shows "embedS r r' f = embedS r r' g"
+unfolding embedS_def using assms
+embed_cong[of r f g r'] bij_betw_cong[of "Field r" f g "Field r'"] by blast
+
+lemma iso_cong[fundef_cong]:
+assumes "\<And> a. a \<in> Field r \<Longrightarrow> f a = g a"
+shows "iso r r' f = iso r r' g"
+unfolding iso_def using assms
+embed_cong[of r f g r'] bij_betw_cong[of "Field r" f g "Field r'"] by blast
+
+lemma id_compat: "compat r r id"
+by(auto simp add: id_def compat_def)
+
+lemma comp_compat:
+"\<lbrakk>compat r r' f; compat r' r'' f'\<rbrakk> \<Longrightarrow> compat r r'' (f' o f)"
+by(auto simp add: comp_def compat_def)
+
+corollary one_set_greater:
+"(\<exists>f::'a \<Rightarrow> 'a'. f ` A \<le> A' \<and> inj_on f A) \<or> (\<exists>g::'a' \<Rightarrow> 'a. g ` A' \<le> A \<and> inj_on g A')"
+proof-
+ obtain r where "well_order_on A r" by (fastforce simp add: well_order_on)
+ hence 1: "A = Field r \<and> Well_order r"
+ using rel.well_order_on_Well_order by auto
+ obtain r' where 2: "well_order_on A' r'" by (fastforce simp add: well_order_on)
+ hence 2: "A' = Field r' \<and> Well_order r'"
+ using rel.well_order_on_Well_order by auto
+ hence "(\<exists>f. embed r r' f) \<or> (\<exists>g. embed r' r g)"
+ using 1 2 by (auto simp add: wellorders_totally_ordered)
+ moreover
+ {fix f assume "embed r r' f"
+ hence "f`A \<le> A' \<and> inj_on f A"
+ using 1 2 by (auto simp add: embed_Field embed_inj_on)
+ }
+ moreover
+ {fix g assume "embed r' r g"
+ hence "g`A' \<le> A \<and> inj_on g A'"
+ using 1 2 by (auto simp add: embed_Field embed_inj_on)
+ }
+ ultimately show ?thesis by blast
+qed
+
+corollary one_type_greater:
+"(\<exists>f::'a \<Rightarrow> 'a'. inj f) \<or> (\<exists>g::'a' \<Rightarrow> 'a. inj g)"
+using one_set_greater[of UNIV UNIV] by auto
+
+
+subsection {* Uniqueness of embeddings *}
+
+lemma comp_embedS:
+assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''"
+ and EMB: "embedS r r' f" and EMB': "embedS r' r'' f'"
+shows "embedS r r'' (f' o f)"
+proof-
+ have "embed r' r'' f'" using EMB' unfolding embedS_def by simp
+ thus ?thesis using assms by (auto simp add: embedS_comp_embed)
+qed
+
+lemma iso_iff4:
+assumes WELL: "Well_order r" and WELL': "Well_order r'"
+shows "iso r r' f = (embed r r' f \<and> embed r' r (inv_into (Field r) f))"
+using assms embed_bothWays_iso
+by(unfold iso_def, auto simp add: inv_into_Field_embed_bij_betw)
+
+lemma embed_embedS_iso:
+"embed r r' f = (embedS r r' f \<or> iso r r' f)"
+unfolding embedS_def iso_def by blast
+
+lemma not_embedS_iso:
+"\<not> (embedS r r' f \<and> iso r r' f)"
+unfolding embedS_def iso_def by blast
+
+lemma embed_embedS_iff_not_iso:
+assumes "embed r r' f"
+shows "embedS r r' f = (\<not> iso r r' f)"
+using assms unfolding embedS_def iso_def by blast
+
+lemma iso_inv_into:
+assumes WELL: "Well_order r" and ISO: "iso r r' f"
+shows "iso r' r (inv_into (Field r) f)"
+using assms unfolding iso_def
+using bij_betw_inv_into inv_into_Field_embed_bij_betw by blast
+
+lemma embedS_or_iso:
+assumes WELL: "Well_order r" and WELL': "Well_order r'"
+shows "(\<exists>g. embedS r r' g) \<or> (\<exists>h. embedS r' r h) \<or> (\<exists>f. iso r r' f)"
+proof-
+ {fix f assume *: "embed r r' f"
+ {assume "bij_betw f (Field r) (Field r')"
+ hence ?thesis using * by (auto simp add: iso_def)
+ }
+ moreover
+ {assume "\<not> bij_betw f (Field r) (Field r')"
+ hence ?thesis using * by (auto simp add: embedS_def)
+ }
+ ultimately have ?thesis by auto
+ }
+ moreover
+ {fix f assume *: "embed r' r f"
+ {assume "bij_betw f (Field r') (Field r)"
+ hence "iso r' r f" using * by (auto simp add: iso_def)
+ hence "iso r r' (inv_into (Field r') f)"
+ using WELL' by (auto simp add: iso_inv_into)
+ hence ?thesis by blast
+ }
+ moreover
+ {assume "\<not> bij_betw f (Field r') (Field r)"
+ hence ?thesis using * by (auto simp add: embedS_def)
+ }
+ ultimately have ?thesis by auto
+ }
+ ultimately show ?thesis using WELL WELL'
+ wellorders_totally_ordered[of r r'] by blast
+qed
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Ordinals_and_Cardinals/Wellorder_Embedding_Base.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,1146 @@
+(* Title: HOL/Ordinals_and_Cardinals/Wellorder_Embedding_Base.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Well-order embeddings (base).
+*)
+
+header {* Well-Order Embeddings (Base) *}
+
+theory Wellorder_Embedding_Base
+imports "~~/src/HOL/Library/Zorn" Fun_More_Base Wellorder_Relation_Base
+begin
+
+
+text{* In this section, we introduce well-order {\em embeddings} and {\em isomorphisms} and
+prove their basic properties. The notion of embedding is considered from the point
+of view of the theory of ordinals, and therefore requires the source to be injected
+as an {\em initial segment} (i.e., {\em order filter}) of the target. A main result
+of this section is the existence of embeddings (in one direction or another) between
+any two well-orders, having as a consequence the fact that, given any two sets on
+any two types, one is smaller than (i.e., can be injected into) the other. *}
+
+
+subsection {* Auxiliaries *}
+
+lemma UNION_inj_on_ofilter:
+assumes WELL: "Well_order r" and
+ OF: "\<And> i. i \<in> I \<Longrightarrow> wo_rel.ofilter r (A i)" and
+ INJ: "\<And> i. i \<in> I \<Longrightarrow> inj_on f (A i)"
+shows "inj_on f (\<Union> i \<in> I. A i)"
+proof-
+ have "wo_rel r" using WELL by (simp add: wo_rel_def)
+ hence "\<And> i j. \<lbrakk>i \<in> I; j \<in> I\<rbrakk> \<Longrightarrow> A i <= A j \<or> A j <= A i"
+ using wo_rel.ofilter_linord[of r] OF by blast
+ with WELL INJ show ?thesis
+ by (auto simp add: inj_on_UNION_chain)
+qed
+
+
+lemma under_underS_bij_betw:
+assumes WELL: "Well_order r" and WELL': "Well_order r'" and
+ IN: "a \<in> Field r" and IN': "f a \<in> Field r'" and
+ BIJ: "bij_betw f (rel.underS r a) (rel.underS r' (f a))"
+shows "bij_betw f (rel.under r a) (rel.under r' (f a))"
+proof-
+ have "a \<notin> rel.underS r a \<and> f a \<notin> rel.underS r' (f a)"
+ unfolding rel.underS_def by auto
+ moreover
+ {have "Refl r \<and> Refl r'" using WELL WELL'
+ by (auto simp add: order_on_defs)
+ hence "rel.under r a = rel.underS r a \<union> {a} \<and>
+ rel.under r' (f a) = rel.underS r' (f a) \<union> {f a}"
+ using IN IN' by(auto simp add: rel.Refl_under_underS)
+ }
+ ultimately show ?thesis
+ using BIJ notIn_Un_bij_betw[of a "rel.underS r a" f "rel.underS r' (f a)"] by auto
+qed
+
+
+
+subsection {* (Well-order) embeddings, strict embeddings, isomorphisms and order-compatible
+functions *}
+
+
+text{* Standardly, a function is an embedding of a well-order in another if it injectively and
+order-compatibly maps the former into an order filter of the latter.
+Here we opt for a more succinct definition (operator @{text "embed"}),
+asking that, for any element in the source, the function should be a bijection
+between the set of strict lower bounds of that element
+and the set of strict lower bounds of its image. (Later we prove equivalence with
+the standard definition -- lemma @{text "embed_iff_compat_inj_on_ofilter"}.)
+A {\em strict embedding} (operator @{text "embedS"}) is a non-bijective embedding
+and an isomorphism (operator @{text "iso"}) is a bijective embedding. *}
+
+
+definition embed :: "'a rel \<Rightarrow> 'a' rel \<Rightarrow> ('a \<Rightarrow> 'a') \<Rightarrow> bool"
+where
+"embed r r' f \<equiv> \<forall>a \<in> Field r. bij_betw f (rel.under r a) (rel.under r' (f a))"
+
+
+lemmas embed_defs = embed_def embed_def[abs_def]
+
+
+text {* Strict embeddings: *}
+
+definition embedS :: "'a rel \<Rightarrow> 'a' rel \<Rightarrow> ('a \<Rightarrow> 'a') \<Rightarrow> bool"
+where
+"embedS r r' f \<equiv> embed r r' f \<and> \<not> bij_betw f (Field r) (Field r')"
+
+
+lemmas embedS_defs = embedS_def embedS_def[abs_def]
+
+
+definition iso :: "'a rel \<Rightarrow> 'a' rel \<Rightarrow> ('a \<Rightarrow> 'a') \<Rightarrow> bool"
+where
+"iso r r' f \<equiv> embed r r' f \<and> bij_betw f (Field r) (Field r')"
+
+
+lemmas iso_defs = iso_def iso_def[abs_def]
+
+
+definition compat :: "'a rel \<Rightarrow> 'a' rel \<Rightarrow> ('a \<Rightarrow> 'a') \<Rightarrow> bool"
+where
+"compat r r' f \<equiv> \<forall>a b. (a,b) \<in> r \<longrightarrow> (f a, f b) \<in> r'"
+
+
+lemma compat_wf:
+assumes CMP: "compat r r' f" and WF: "wf r'"
+shows "wf r"
+proof-
+ have "r \<le> inv_image r' f"
+ unfolding inv_image_def using CMP
+ by (auto simp add: compat_def)
+ with WF show ?thesis
+ using wf_inv_image[of r' f] wf_subset[of "inv_image r' f"] by auto
+qed
+
+
+lemma id_embed: "embed r r id"
+by(auto simp add: id_def embed_def bij_betw_def)
+
+
+lemma id_iso: "iso r r id"
+by(auto simp add: id_def embed_def iso_def bij_betw_def)
+
+
+lemma embed_in_Field:
+assumes WELL: "Well_order r" and
+ EMB: "embed r r' f" and IN: "a \<in> Field r"
+shows "f a \<in> Field r'"
+proof-
+ have Well: "wo_rel r"
+ using WELL by (auto simp add: wo_rel_def)
+ hence 1: "Refl r"
+ by (auto simp add: wo_rel.REFL)
+ hence "a \<in> rel.under r a" using IN rel.Refl_under_in by fastforce
+ hence "f a \<in> rel.under r' (f a)"
+ using EMB IN by (auto simp add: embed_def bij_betw_def)
+ thus ?thesis unfolding Field_def
+ by (auto simp: rel.under_def)
+qed
+
+
+lemma comp_embed:
+assumes WELL: "Well_order r" and
+ EMB: "embed r r' f" and EMB': "embed r' r'' f'"
+shows "embed r r'' (f' o f)"
+proof(unfold embed_def, auto)
+ fix a assume *: "a \<in> Field r"
+ hence "bij_betw f (rel.under r a) (rel.under r' (f a))"
+ using embed_def[of r] EMB by auto
+ moreover
+ {have "f a \<in> Field r'"
+ using EMB WELL * by (auto simp add: embed_in_Field)
+ hence "bij_betw f' (rel.under r' (f a)) (rel.under r'' (f' (f a)))"
+ using embed_def[of r'] EMB' by auto
+ }
+ ultimately
+ show "bij_betw (f' \<circ> f) (rel.under r a) (rel.under r'' (f'(f a)))"
+ by(auto simp add: bij_betw_trans)
+qed
+
+
+lemma comp_iso:
+assumes WELL: "Well_order r" and
+ EMB: "iso r r' f" and EMB': "iso r' r'' f'"
+shows "iso r r'' (f' o f)"
+using assms unfolding iso_def
+by (auto simp add: comp_embed bij_betw_trans)
+
+
+text{* That @{text "embedS"} is also preserved by function composition shall be proved only later. *}
+
+
+lemma embed_Field:
+"\<lbrakk>Well_order r; embed r r' f\<rbrakk> \<Longrightarrow> f`(Field r) \<le> Field r'"
+by (auto simp add: embed_in_Field)
+
+
+lemma embed_preserves_ofilter:
+assumes WELL: "Well_order r" and WELL': "Well_order r'" and
+ EMB: "embed r r' f" and OF: "wo_rel.ofilter r A"
+shows "wo_rel.ofilter r' (f`A)"
+proof-
+ (* Preliminary facts *)
+ from WELL have Well: "wo_rel r" unfolding wo_rel_def .
+ from WELL' have Well': "wo_rel r'" unfolding wo_rel_def .
+ from OF have 0: "A \<le> Field r" by(auto simp add: Well wo_rel.ofilter_def)
+ (* Main proof *)
+ show ?thesis using Well' WELL EMB 0 embed_Field[of r r' f]
+ proof(unfold wo_rel.ofilter_def, auto simp add: image_def)
+ fix a b'
+ assume *: "a \<in> A" and **: "b' \<in> rel.under r' (f a)"
+ hence "a \<in> Field r" using 0 by auto
+ hence "bij_betw f (rel.under r a) (rel.under r' (f a))"
+ using * EMB by (auto simp add: embed_def)
+ hence "f`(rel.under r a) = rel.under r' (f a)"
+ by (simp add: bij_betw_def)
+ with ** image_def[of f "rel.under r a"] obtain b where
+ 1: "b \<in> rel.under r a \<and> b' = f b" by blast
+ hence "b \<in> A" using Well * OF
+ by (auto simp add: wo_rel.ofilter_def)
+ with 1 show "\<exists>b \<in> A. b' = f b" by blast
+ qed
+qed
+
+
+lemma embed_Field_ofilter:
+assumes WELL: "Well_order r" and WELL': "Well_order r'" and
+ EMB: "embed r r' f"
+shows "wo_rel.ofilter r' (f`(Field r))"
+proof-
+ have "wo_rel.ofilter r (Field r)"
+ using WELL by (auto simp add: wo_rel_def wo_rel.Field_ofilter)
+ with WELL WELL' EMB
+ show ?thesis by (auto simp add: embed_preserves_ofilter)
+qed
+
+
+lemma embed_compat:
+assumes EMB: "embed r r' f"
+shows "compat r r' f"
+proof(unfold compat_def, clarify)
+ fix a b
+ assume *: "(a,b) \<in> r"
+ hence 1: "b \<in> Field r" using Field_def[of r] by blast
+ have "a \<in> rel.under r b"
+ using * rel.under_def[of r] by simp
+ hence "f a \<in> rel.under r' (f b)"
+ using EMB embed_def[of r r' f]
+ bij_betw_def[of f "rel.under r b" "rel.under r' (f b)"]
+ image_def[of f "rel.under r b"] 1 by auto
+ thus "(f a, f b) \<in> r'"
+ by (auto simp add: rel.under_def)
+qed
+
+
+lemma embed_inj_on:
+assumes WELL: "Well_order r" and EMB: "embed r r' f"
+shows "inj_on f (Field r)"
+proof(unfold inj_on_def, clarify)
+ (* Preliminary facts *)
+ from WELL have Well: "wo_rel r" unfolding wo_rel_def .
+ with wo_rel.TOTAL[of r]
+ have Total: "Total r" by simp
+ from Well wo_rel.REFL[of r]
+ have Refl: "Refl r" by simp
+ (* Main proof *)
+ fix a b
+ assume *: "a \<in> Field r" and **: "b \<in> Field r" and
+ ***: "f a = f b"
+ hence 1: "a \<in> Field r \<and> b \<in> Field r"
+ unfolding Field_def by auto
+ {assume "(a,b) \<in> r"
+ hence "a \<in> rel.under r b \<and> b \<in> rel.under r b"
+ using Refl by(auto simp add: rel.under_def refl_on_def)
+ hence "a = b"
+ using EMB 1 ***
+ by (auto simp add: embed_def bij_betw_def inj_on_def)
+ }
+ moreover
+ {assume "(b,a) \<in> r"
+ hence "a \<in> rel.under r a \<and> b \<in> rel.under r a"
+ using Refl by(auto simp add: rel.under_def refl_on_def)
+ hence "a = b"
+ using EMB 1 ***
+ by (auto simp add: embed_def bij_betw_def inj_on_def)
+ }
+ ultimately
+ show "a = b" using Total 1
+ by (auto simp add: total_on_def)
+qed
+
+
+lemma embed_underS:
+assumes WELL: "Well_order r" and WELL: "Well_order r'" and
+ EMB: "embed r r' f" and IN: "a \<in> Field r"
+shows "bij_betw f (rel.underS r a) (rel.underS r' (f a))"
+proof-
+ have "bij_betw f (rel.under r a) (rel.under r' (f a))"
+ using assms by (auto simp add: embed_def)
+ moreover
+ {have "f a \<in> Field r'" using assms embed_Field[of r r' f] by auto
+ hence "rel.under r a = rel.underS r a \<union> {a} \<and>
+ rel.under r' (f a) = rel.underS r' (f a) \<union> {f a}"
+ using assms by (auto simp add: order_on_defs rel.Refl_under_underS)
+ }
+ moreover
+ {have "a \<notin> rel.underS r a \<and> f a \<notin> rel.underS r' (f a)"
+ unfolding rel.underS_def by blast
+ }
+ ultimately show ?thesis
+ by (auto simp add: notIn_Un_bij_betw3)
+qed
+
+
+lemma embed_iff_compat_inj_on_ofilter:
+assumes WELL: "Well_order r" and WELL': "Well_order r'"
+shows "embed r r' f = (compat r r' f \<and> inj_on f (Field r) \<and> wo_rel.ofilter r' (f`(Field r)))"
+using assms
+proof(auto simp add: embed_compat embed_inj_on embed_Field_ofilter,
+ unfold embed_def, auto) (* get rid of one implication *)
+ fix a
+ assume *: "inj_on f (Field r)" and
+ **: "compat r r' f" and
+ ***: "wo_rel.ofilter r' (f`(Field r))" and
+ ****: "a \<in> Field r"
+ (* Preliminary facts *)
+ have Well: "wo_rel r"
+ using WELL wo_rel_def[of r] by simp
+ hence Refl: "Refl r"
+ using wo_rel.REFL[of r] by simp
+ have Total: "Total r"
+ using Well wo_rel.TOTAL[of r] by simp
+ have Well': "wo_rel r'"
+ using WELL' wo_rel_def[of r'] by simp
+ hence Antisym': "antisym r'"
+ using wo_rel.ANTISYM[of r'] by simp
+ have "(a,a) \<in> r"
+ using **** Well wo_rel.REFL[of r]
+ refl_on_def[of _ r] by auto
+ hence "(f a, f a) \<in> r'"
+ using ** by(auto simp add: compat_def)
+ hence 0: "f a \<in> Field r'"
+ unfolding Field_def by auto
+ have "f a \<in> f`(Field r)"
+ using **** by auto
+ hence 2: "rel.under r' (f a) \<le> f`(Field r)"
+ using Well' *** wo_rel.ofilter_def[of r' "f`(Field r)"] by fastforce
+ (* Main proof *)
+ show "bij_betw f (rel.under r a) (rel.under r' (f a))"
+ proof(unfold bij_betw_def, auto)
+ show "inj_on f (rel.under r a)"
+ using *
+ by (auto simp add: rel.under_Field subset_inj_on)
+ next
+ fix b assume "b \<in> rel.under r a"
+ thus "f b \<in> rel.under r' (f a)"
+ unfolding rel.under_def using **
+ by (auto simp add: compat_def)
+ next
+ fix b' assume *****: "b' \<in> rel.under r' (f a)"
+ hence "b' \<in> f`(Field r)"
+ using 2 by auto
+ with Field_def[of r] obtain b where
+ 3: "b \<in> Field r" and 4: "b' = f b" by auto
+ have "(b,a): r"
+ proof-
+ {assume "(a,b) \<in> r"
+ with ** 4 have "(f a, b'): r'"
+ by (auto simp add: compat_def)
+ with ***** Antisym' have "f a = b'"
+ by(auto simp add: rel.under_def antisym_def)
+ with 3 **** 4 * have "a = b"
+ by(auto simp add: inj_on_def)
+ }
+ moreover
+ {assume "a = b"
+ hence "(b,a) \<in> r" using Refl **** 3
+ by (auto simp add: refl_on_def)
+ }
+ ultimately
+ show ?thesis using Total **** 3 by (fastforce simp add: total_on_def)
+ qed
+ with 4 show "b' \<in> f`(rel.under r a)"
+ unfolding rel.under_def by auto
+ qed
+qed
+
+
+lemma inv_into_ofilter_embed:
+assumes WELL: "Well_order r" and OF: "wo_rel.ofilter r A" and
+ BIJ: "\<forall>b \<in> A. bij_betw f (rel.under r b) (rel.under r' (f b))" and
+ IMAGE: "f ` A = Field r'"
+shows "embed r' r (inv_into A f)"
+proof-
+ (* Preliminary facts *)
+ have Well: "wo_rel r"
+ using WELL wo_rel_def[of r] by simp
+ have Refl: "Refl r"
+ using Well wo_rel.REFL[of r] by simp
+ have Total: "Total r"
+ using Well wo_rel.TOTAL[of r] by simp
+ (* Main proof *)
+ have 1: "bij_betw f A (Field r')"
+ proof(unfold bij_betw_def inj_on_def, auto simp add: IMAGE)
+ fix b1 b2
+ assume *: "b1 \<in> A" and **: "b2 \<in> A" and
+ ***: "f b1 = f b2"
+ have 11: "b1 \<in> Field r \<and> b2 \<in> Field r"
+ using * ** Well OF by (auto simp add: wo_rel.ofilter_def)
+ moreover
+ {assume "(b1,b2) \<in> r"
+ hence "b1 \<in> rel.under r b2 \<and> b2 \<in> rel.under r b2"
+ unfolding rel.under_def using 11 Refl
+ by (auto simp add: refl_on_def)
+ hence "b1 = b2" using BIJ * ** ***
+ by (auto simp add: bij_betw_def inj_on_def)
+ }
+ moreover
+ {assume "(b2,b1) \<in> r"
+ hence "b1 \<in> rel.under r b1 \<and> b2 \<in> rel.under r b1"
+ unfolding rel.under_def using 11 Refl
+ by (auto simp add: refl_on_def)
+ hence "b1 = b2" using BIJ * ** ***
+ by (auto simp add: bij_betw_def inj_on_def)
+ }
+ ultimately
+ show "b1 = b2"
+ using Total by (auto simp add: total_on_def)
+ qed
+ (* *)
+ let ?f' = "(inv_into A f)"
+ (* *)
+ have 2: "\<forall>b \<in> A. bij_betw ?f' (rel.under r' (f b)) (rel.under r b)"
+ proof(clarify)
+ fix b assume *: "b \<in> A"
+ hence "rel.under r b \<le> A"
+ using Well OF by(auto simp add: wo_rel.ofilter_def)
+ moreover
+ have "f ` (rel.under r b) = rel.under r' (f b)"
+ using * BIJ by (auto simp add: bij_betw_def)
+ ultimately
+ show "bij_betw ?f' (rel.under r' (f b)) (rel.under r b)"
+ using 1 by (auto simp add: bij_betw_inv_into_subset)
+ qed
+ (* *)
+ have 3: "\<forall>b' \<in> Field r'. bij_betw ?f' (rel.under r' b') (rel.under r (?f' b'))"
+ proof(clarify)
+ fix b' assume *: "b' \<in> Field r'"
+ have "b' = f (?f' b')" using * 1
+ by (auto simp add: bij_betw_inv_into_right)
+ moreover
+ {obtain b where 31: "b \<in> A" and "f b = b'" using IMAGE * by force
+ hence "?f' b' = b" using 1 by (auto simp add: bij_betw_inv_into_left)
+ with 31 have "?f' b' \<in> A" by auto
+ }
+ ultimately
+ show "bij_betw ?f' (rel.under r' b') (rel.under r (?f' b'))"
+ using 2 by auto
+ qed
+ (* *)
+ thus ?thesis unfolding embed_def .
+qed
+
+
+lemma inv_into_underS_embed:
+assumes WELL: "Well_order r" and
+ BIJ: "\<forall>b \<in> rel.underS r a. bij_betw f (rel.under r b) (rel.under r' (f b))" and
+ IN: "a \<in> Field r" and
+ IMAGE: "f ` (rel.underS r a) = Field r'"
+shows "embed r' r (inv_into (rel.underS r a) f)"
+using assms
+by(auto simp add: wo_rel_def wo_rel.underS_ofilter inv_into_ofilter_embed)
+
+
+lemma inv_into_Field_embed:
+assumes WELL: "Well_order r" and EMB: "embed r r' f" and
+ IMAGE: "Field r' \<le> f ` (Field r)"
+shows "embed r' r (inv_into (Field r) f)"
+proof-
+ have "(\<forall>b \<in> Field r. bij_betw f (rel.under r b) (rel.under r' (f b)))"
+ using EMB by (auto simp add: embed_def)
+ moreover
+ have "f ` (Field r) \<le> Field r'"
+ using EMB WELL by (auto simp add: embed_Field)
+ ultimately
+ show ?thesis using assms
+ by(auto simp add: wo_rel_def wo_rel.Field_ofilter inv_into_ofilter_embed)
+qed
+
+
+lemma inv_into_Field_embed_bij_betw:
+assumes WELL: "Well_order r" and
+ EMB: "embed r r' f" and BIJ: "bij_betw f (Field r) (Field r')"
+shows "embed r' r (inv_into (Field r) f)"
+proof-
+ have "Field r' \<le> f ` (Field r)"
+ using BIJ by (auto simp add: bij_betw_def)
+ thus ?thesis using assms
+ by(auto simp add: inv_into_Field_embed)
+qed
+
+
+
+
+
+subsection {* Given any two well-orders, one can be embedded in the other *}
+
+
+text{* Here is an overview of the proof of of this fact, stated in theorem
+@{text "wellorders_totally_ordered"}:
+
+ Fix the well-orders @{text "r::'a rel"} and @{text "r'::'a' rel"}.
+ Attempt to define an embedding @{text "f::'a \<Rightarrow> 'a'"} from @{text "r"} to @{text "r'"} in the
+ natural way by well-order recursion ("hoping" that @{text "Field r"} turns out to be smaller
+ than @{text "Field r'"}), but also record, at the recursive step, in a function
+ @{text "g::'a \<Rightarrow> bool"}, the extra information of whether @{text "Field r'"}
+ gets exhausted or not.
+
+ If @{text "Field r'"} does not get exhausted, then @{text "Field r"} is indeed smaller
+ and @{text "f"} is the desired embedding from @{text "r"} to @{text "r'"}
+ (lemma @{text "wellorders_totally_ordered_aux"}).
+
+ Otherwise, it means that @{text "Field r'"} is the smaller one, and the inverse of
+ (the "good" segment of) @{text "f"} is the desired embedding from @{text "r'"} to @{text "r"}
+ (lemma @{text "wellorders_totally_ordered_aux2"}).
+*}
+
+
+lemma wellorders_totally_ordered_aux:
+fixes r ::"'a rel" and r'::"'a' rel" and
+ f :: "'a \<Rightarrow> 'a'" and a::'a
+assumes WELL: "Well_order r" and WELL': "Well_order r'" and IN: "a \<in> Field r" and
+ IH: "\<forall>b \<in> rel.underS r a. bij_betw f (rel.under r b) (rel.under r' (f b))" and
+ NOT: "f ` (rel.underS r a) \<noteq> Field r'" and SUC: "f a = wo_rel.suc r' (f`(rel.underS r a))"
+shows "bij_betw f (rel.under r a) (rel.under r' (f a))"
+proof-
+ (* Preliminary facts *)
+ have Well: "wo_rel r" using WELL unfolding wo_rel_def .
+ hence Refl: "Refl r" using wo_rel.REFL[of r] by auto
+ have Trans: "trans r" using Well wo_rel.TRANS[of r] by auto
+ have Well': "wo_rel r'" using WELL' unfolding wo_rel_def .
+ have OF: "wo_rel.ofilter r (rel.underS r a)"
+ by (auto simp add: Well wo_rel.underS_ofilter)
+ hence UN: "rel.underS r a = (\<Union> b \<in> rel.underS r a. rel.under r b)"
+ using Well wo_rel.ofilter_under_UNION[of r "rel.underS r a"] by blast
+ (* Gather facts about elements of rel.underS r a *)
+ {fix b assume *: "b \<in> rel.underS r a"
+ hence t0: "(b,a) \<in> r \<and> b \<noteq> a" unfolding rel.underS_def by auto
+ have t1: "b \<in> Field r"
+ using * rel.underS_Field[of r a] by auto
+ have t2: "f`(rel.under r b) = rel.under r' (f b)"
+ using IH * by (auto simp add: bij_betw_def)
+ hence t3: "wo_rel.ofilter r' (f`(rel.under r b))"
+ using Well' by (auto simp add: wo_rel.under_ofilter)
+ have "f`(rel.under r b) \<le> Field r'"
+ using t2 by (auto simp add: rel.under_Field)
+ moreover
+ have "b \<in> rel.under r b"
+ using t1 by(auto simp add: Refl rel.Refl_under_in)
+ ultimately
+ have t4: "f b \<in> Field r'" by auto
+ have "f`(rel.under r b) = rel.under r' (f b) \<and>
+ wo_rel.ofilter r' (f`(rel.under r b)) \<and>
+ f b \<in> Field r'"
+ using t2 t3 t4 by auto
+ }
+ hence bFact:
+ "\<forall>b \<in> rel.underS r a. f`(rel.under r b) = rel.under r' (f b) \<and>
+ wo_rel.ofilter r' (f`(rel.under r b)) \<and>
+ f b \<in> Field r'" by blast
+ (* *)
+ have subField: "f`(rel.underS r a) \<le> Field r'"
+ using bFact by blast
+ (* *)
+ have OF': "wo_rel.ofilter r' (f`(rel.underS r a))"
+ proof-
+ have "f`(rel.underS r a) = f`(\<Union> b \<in> rel.underS r a. rel.under r b)"
+ using UN by auto
+ also have "\<dots> = (\<Union> b \<in> rel.underS r a. f`(rel.under r b))" by blast
+ also have "\<dots> = (\<Union> b \<in> rel.underS r a. (rel.under r' (f b)))"
+ using bFact by auto
+ finally
+ have "f`(rel.underS r a) = (\<Union> b \<in> rel.underS r a. (rel.under r' (f b)))" .
+ thus ?thesis
+ using Well' bFact
+ wo_rel.ofilter_UNION[of r' "rel.underS r a" "\<lambda> b. rel.under r' (f b)"] by fastforce
+ qed
+ (* *)
+ have "f`(rel.underS r a) \<union> rel.AboveS r' (f`(rel.underS r a)) = Field r'"
+ using Well' OF' by (auto simp add: wo_rel.ofilter_AboveS_Field)
+ hence NE: "rel.AboveS r' (f`(rel.underS r a)) \<noteq> {}"
+ using subField NOT by blast
+ (* Main proof *)
+ have INCL1: "f`(rel.underS r a) \<le> rel.underS r' (f a) "
+ proof(auto)
+ fix b assume *: "b \<in> rel.underS r a"
+ have "f b \<noteq> f a \<and> (f b, f a) \<in> r'"
+ using subField Well' SUC NE *
+ wo_rel.suc_greater[of r' "f`(rel.underS r a)" "f b"] by auto
+ thus "f b \<in> rel.underS r' (f a)"
+ unfolding rel.underS_def by simp
+ qed
+ (* *)
+ have INCL2: "rel.underS r' (f a) \<le> f`(rel.underS r a)"
+ proof
+ fix b' assume "b' \<in> rel.underS r' (f a)"
+ hence "b' \<noteq> f a \<and> (b', f a) \<in> r'"
+ unfolding rel.underS_def by simp
+ thus "b' \<in> f`(rel.underS r a)"
+ using Well' SUC NE OF'
+ wo_rel.suc_ofilter_in[of r' "f ` rel.underS r a" b'] by auto
+ qed
+ (* *)
+ have INJ: "inj_on f (rel.underS r a)"
+ proof-
+ have "\<forall>b \<in> rel.underS r a. inj_on f (rel.under r b)"
+ using IH by (auto simp add: bij_betw_def)
+ moreover
+ have "\<forall>b. wo_rel.ofilter r (rel.under r b)"
+ using Well by (auto simp add: wo_rel.under_ofilter)
+ ultimately show ?thesis
+ using WELL bFact UN
+ UNION_inj_on_ofilter[of r "rel.underS r a" "\<lambda>b. rel.under r b" f]
+ by auto
+ qed
+ (* *)
+ have BIJ: "bij_betw f (rel.underS r a) (rel.underS r' (f a))"
+ unfolding bij_betw_def
+ using INJ INCL1 INCL2 by auto
+ (* *)
+ have "f a \<in> Field r'"
+ using Well' subField NE SUC
+ by (auto simp add: wo_rel.suc_inField)
+ thus ?thesis
+ using WELL WELL' IN BIJ under_underS_bij_betw[of r r' a f] by auto
+qed
+
+
+lemma wellorders_totally_ordered_aux2:
+fixes r ::"'a rel" and r'::"'a' rel" and
+ f :: "'a \<Rightarrow> 'a'" and g :: "'a \<Rightarrow> bool" and a::'a
+assumes WELL: "Well_order r" and WELL': "Well_order r'" and
+MAIN1:
+ "\<And> a. (False \<notin> g`(rel.underS r a) \<and> f`(rel.underS r a) \<noteq> Field r'
+ \<longrightarrow> f a = wo_rel.suc r' (f`(rel.underS r a)) \<and> g a = True)
+ \<and>
+ (\<not>(False \<notin> (g`(rel.underS r a)) \<and> f`(rel.underS r a) \<noteq> Field r')
+ \<longrightarrow> g a = False)" and
+MAIN2: "\<And> a. a \<in> Field r \<and> False \<notin> g`(rel.under r a) \<longrightarrow>
+ bij_betw f (rel.under r a) (rel.under r' (f a))" and
+Case: "a \<in> Field r \<and> False \<in> g`(rel.under r a)"
+shows "\<exists>f'. embed r' r f'"
+proof-
+ have Well: "wo_rel r" using WELL unfolding wo_rel_def .
+ hence Refl: "Refl r" using wo_rel.REFL[of r] by auto
+ have Trans: "trans r" using Well wo_rel.TRANS[of r] by auto
+ have Antisym: "antisym r" using Well wo_rel.ANTISYM[of r] by auto
+ have Well': "wo_rel r'" using WELL' unfolding wo_rel_def .
+ (* *)
+ have 0: "rel.under r a = rel.underS r a \<union> {a}"
+ using Refl Case by(auto simp add: rel.Refl_under_underS)
+ (* *)
+ have 1: "g a = False"
+ proof-
+ {assume "g a \<noteq> False"
+ with 0 Case have "False \<in> g`(rel.underS r a)" by blast
+ with MAIN1 have "g a = False" by blast}
+ thus ?thesis by blast
+ qed
+ let ?A = "{a \<in> Field r. g a = False}"
+ let ?a = "(wo_rel.minim r ?A)"
+ (* *)
+ have 2: "?A \<noteq> {} \<and> ?A \<le> Field r" using Case 1 by blast
+ (* *)
+ have 3: "False \<notin> g`(rel.underS r ?a)"
+ proof
+ assume "False \<in> g`(rel.underS r ?a)"
+ then obtain b where "b \<in> rel.underS r ?a" and 31: "g b = False" by auto
+ hence 32: "(b,?a) \<in> r \<and> b \<noteq> ?a"
+ by (auto simp add: rel.underS_def)
+ hence "b \<in> Field r" unfolding Field_def by auto
+ with 31 have "b \<in> ?A" by auto
+ hence "(?a,b) \<in> r" using wo_rel.minim_least 2 Well by fastforce
+ (* again: why worked without type annotations? *)
+ with 32 Antisym show False
+ by (auto simp add: antisym_def)
+ qed
+ have temp: "?a \<in> ?A"
+ using Well 2 wo_rel.minim_in[of r ?A] by auto
+ hence 4: "?a \<in> Field r" by auto
+ (* *)
+ have 5: "g ?a = False" using temp by blast
+ (* *)
+ have 6: "f`(rel.underS r ?a) = Field r'"
+ using MAIN1[of ?a] 3 5 by blast
+ (* *)
+ have 7: "\<forall>b \<in> rel.underS r ?a. bij_betw f (rel.under r b) (rel.under r' (f b))"
+ proof
+ fix b assume as: "b \<in> rel.underS r ?a"
+ moreover
+ have "wo_rel.ofilter r (rel.underS r ?a)"
+ using Well by (auto simp add: wo_rel.underS_ofilter)
+ ultimately
+ have "False \<notin> g`(rel.under r b)" using 3 Well by (auto simp add: wo_rel.ofilter_def)
+ moreover have "b \<in> Field r"
+ unfolding Field_def using as by (auto simp add: rel.underS_def)
+ ultimately
+ show "bij_betw f (rel.under r b) (rel.under r' (f b))"
+ using MAIN2 by auto
+ qed
+ (* *)
+ have "embed r' r (inv_into (rel.underS r ?a) f)"
+ using WELL WELL' 7 4 6 inv_into_underS_embed[of r ?a f r'] by auto
+ thus ?thesis
+ unfolding embed_def by blast
+qed
+
+
+theorem wellorders_totally_ordered:
+fixes r ::"'a rel" and r'::"'a' rel"
+assumes WELL: "Well_order r" and WELL': "Well_order r'"
+shows "(\<exists>f. embed r r' f) \<or> (\<exists>f'. embed r' r f')"
+proof-
+ (* Preliminary facts *)
+ have Well: "wo_rel r" using WELL unfolding wo_rel_def .
+ hence Refl: "Refl r" using wo_rel.REFL[of r] by auto
+ have Trans: "trans r" using Well wo_rel.TRANS[of r] by auto
+ have Well': "wo_rel r'" using WELL' unfolding wo_rel_def .
+ (* Main proof *)
+ obtain H where H_def: "H =
+ (\<lambda>h a. if False \<notin> (snd o h)`(rel.underS r a) \<and> (fst o h)`(rel.underS r a) \<noteq> Field r'
+ then (wo_rel.suc r' ((fst o h)`(rel.underS r a)), True)
+ else (undefined, False))" by blast
+ have Adm: "wo_rel.adm_wo r H"
+ using Well
+ proof(unfold wo_rel.adm_wo_def, clarify)
+ fix h1::"'a \<Rightarrow> 'a' * bool" and h2::"'a \<Rightarrow> 'a' * bool" and x
+ assume "\<forall>y\<in>rel.underS r x. h1 y = h2 y"
+ hence "\<forall>y\<in>rel.underS r x. (fst o h1) y = (fst o h2) y \<and>
+ (snd o h1) y = (snd o h2) y" by auto
+ hence "(fst o h1)`(rel.underS r x) = (fst o h2)`(rel.underS r x) \<and>
+ (snd o h1)`(rel.underS r x) = (snd o h2)`(rel.underS r x)"
+ by (auto simp add: image_def)
+ thus "H h1 x = H h2 x" by (simp add: H_def)
+ qed
+ (* More constant definitions: *)
+ obtain h::"'a \<Rightarrow> 'a' * bool" and f::"'a \<Rightarrow> 'a'" and g::"'a \<Rightarrow> bool"
+ where h_def: "h = wo_rel.worec r H" and
+ f_def: "f = fst o h" and g_def: "g = snd o h" by blast
+ obtain test where test_def:
+ "test = (\<lambda> a. False \<notin> (g`(rel.underS r a)) \<and> f`(rel.underS r a) \<noteq> Field r')" by blast
+ (* *)
+ have *: "\<And> a. h a = H h a"
+ using Adm Well wo_rel.worec_fixpoint[of r H] by (simp add: h_def)
+ have Main1:
+ "\<And> a. (test a \<longrightarrow> f a = wo_rel.suc r' (f`(rel.underS r a)) \<and> g a = True) \<and>
+ (\<not>(test a) \<longrightarrow> g a = False)"
+ proof- (* How can I prove this withou fixing a? *)
+ fix a show "(test a \<longrightarrow> f a = wo_rel.suc r' (f`(rel.underS r a)) \<and> g a = True) \<and>
+ (\<not>(test a) \<longrightarrow> g a = False)"
+ using *[of a] test_def f_def g_def H_def by auto
+ qed
+ (* *)
+ let ?phi = "\<lambda> a. a \<in> Field r \<and> False \<notin> g`(rel.under r a) \<longrightarrow>
+ bij_betw f (rel.under r a) (rel.under r' (f a))"
+ have Main2: "\<And> a. ?phi a"
+ proof-
+ fix a show "?phi a"
+ proof(rule wo_rel.well_order_induct[of r ?phi],
+ simp only: Well, clarify)
+ fix a
+ assume IH: "\<forall>b. b \<noteq> a \<and> (b,a) \<in> r \<longrightarrow> ?phi b" and
+ *: "a \<in> Field r" and
+ **: "False \<notin> g`(rel.under r a)"
+ have 1: "\<forall>b \<in> rel.underS r a. bij_betw f (rel.under r b) (rel.under r' (f b))"
+ proof(clarify)
+ fix b assume ***: "b \<in> rel.underS r a"
+ hence 0: "(b,a) \<in> r \<and> b \<noteq> a" unfolding rel.underS_def by auto
+ moreover have "b \<in> Field r"
+ using *** rel.underS_Field[of r a] by auto
+ moreover have "False \<notin> g`(rel.under r b)"
+ using 0 ** Trans rel.under_incr[of r b a] by auto
+ ultimately show "bij_betw f (rel.under r b) (rel.under r' (f b))"
+ using IH by auto
+ qed
+ (* *)
+ have 21: "False \<notin> g`(rel.underS r a)"
+ using ** rel.underS_subset_under[of r a] by auto
+ have 22: "g`(rel.under r a) \<le> {True}" using ** by auto
+ moreover have 23: "a \<in> rel.under r a"
+ using Refl * by (auto simp add: rel.Refl_under_in)
+ ultimately have 24: "g a = True" by blast
+ have 2: "f`(rel.underS r a) \<noteq> Field r'"
+ proof
+ assume "f`(rel.underS r a) = Field r'"
+ hence "g a = False" using Main1 test_def by blast
+ with 24 show False using ** by blast
+ qed
+ (* *)
+ have 3: "f a = wo_rel.suc r' (f`(rel.underS r a))"
+ using 21 2 Main1 test_def by blast
+ (* *)
+ show "bij_betw f (rel.under r a) (rel.under r' (f a))"
+ using WELL WELL' 1 2 3 *
+ wellorders_totally_ordered_aux[of r r' a f] by auto
+ qed
+ qed
+ (* *)
+ let ?chi = "(\<lambda> a. a \<in> Field r \<and> False \<in> g`(rel.under r a))"
+ show ?thesis
+ proof(cases "\<exists>a. ?chi a")
+ assume "\<not> (\<exists>a. ?chi a)"
+ hence "\<forall>a \<in> Field r. bij_betw f (rel.under r a) (rel.under r' (f a))"
+ using Main2 by blast
+ thus ?thesis unfolding embed_def by blast
+ next
+ assume "\<exists>a. ?chi a"
+ then obtain a where "?chi a" by blast
+ hence "\<exists>f'. embed r' r f'"
+ using wellorders_totally_ordered_aux2[of r r' g f a]
+ WELL WELL' Main1 Main2 test_def by blast
+ thus ?thesis by blast
+ qed
+qed
+
+
+subsection {* Uniqueness of embeddings *}
+
+
+text{* Here we show a fact complementary to the one from the previous subsection -- namely,
+that between any two well-orders there is {\em at most} one embedding, and is the one
+definable by the expected well-order recursive equation. As a consequence, any two
+embeddings of opposite directions are mutually inverse. *}
+
+
+lemma embed_determined:
+assumes WELL: "Well_order r" and WELL': "Well_order r'" and
+ EMB: "embed r r' f" and IN: "a \<in> Field r"
+shows "f a = wo_rel.suc r' (f`(rel.underS r a))"
+proof-
+ have "bij_betw f (rel.underS r a) (rel.underS r' (f a))"
+ using assms by (auto simp add: embed_underS)
+ hence "f`(rel.underS r a) = rel.underS r' (f a)"
+ by (auto simp add: bij_betw_def)
+ moreover
+ {have "f a \<in> Field r'" using IN
+ using EMB WELL embed_Field[of r r' f] by auto
+ hence "f a = wo_rel.suc r' (rel.underS r' (f a))"
+ using WELL' by (auto simp add: wo_rel_def wo_rel.suc_underS)
+ }
+ ultimately show ?thesis by simp
+qed
+
+
+lemma embed_unique:
+assumes WELL: "Well_order r" and WELL': "Well_order r'" and
+ EMBf: "embed r r' f" and EMBg: "embed r r' g"
+shows "a \<in> Field r \<longrightarrow> f a = g a"
+proof(rule wo_rel.well_order_induct[of r], auto simp add: WELL wo_rel_def)
+ fix a
+ assume IH: "\<forall>b. b \<noteq> a \<and> (b,a): r \<longrightarrow> b \<in> Field r \<longrightarrow> f b = g b" and
+ *: "a \<in> Field r"
+ hence "\<forall>b \<in> rel.underS r a. f b = g b"
+ unfolding rel.underS_def by (auto simp add: Field_def)
+ hence "f`(rel.underS r a) = g`(rel.underS r a)" by force
+ thus "f a = g a"
+ using assms * embed_determined[of r r' f a] embed_determined[of r r' g a] by auto
+qed
+
+
+lemma embed_bothWays_inverse:
+assumes WELL: "Well_order r" and WELL': "Well_order r'" and
+ EMB: "embed r r' f" and EMB': "embed r' r f'"
+shows "(\<forall>a \<in> Field r. f'(f a) = a) \<and> (\<forall>a' \<in> Field r'. f(f' a') = a')"
+proof-
+ have "embed r r (f' o f)" using assms
+ by(auto simp add: comp_embed)
+ moreover have "embed r r id" using assms
+ by (auto simp add: id_embed)
+ ultimately have "\<forall>a \<in> Field r. f'(f a) = a"
+ using assms embed_unique[of r r "f' o f" id] id_def by auto
+ moreover
+ {have "embed r' r' (f o f')" using assms
+ by(auto simp add: comp_embed)
+ moreover have "embed r' r' id" using assms
+ by (auto simp add: id_embed)
+ ultimately have "\<forall>a' \<in> Field r'. f(f' a') = a'"
+ using assms embed_unique[of r' r' "f o f'" id] id_def by auto
+ }
+ ultimately show ?thesis by blast
+qed
+
+
+lemma embed_bothWays_bij_betw:
+assumes WELL: "Well_order r" and WELL': "Well_order r'" and
+ EMB: "embed r r' f" and EMB': "embed r' r g"
+shows "bij_betw f (Field r) (Field r')"
+proof-
+ let ?A = "Field r" let ?A' = "Field r'"
+ have "embed r r (g o f) \<and> embed r' r' (f o g)"
+ using assms by (auto simp add: comp_embed)
+ hence 1: "(\<forall>a \<in> ?A. g(f a) = a) \<and> (\<forall>a' \<in> ?A'. f(g a') = a')"
+ using WELL id_embed[of r] embed_unique[of r r "g o f" id]
+ WELL' id_embed[of r'] embed_unique[of r' r' "f o g" id]
+ id_def by auto
+ have 2: "(\<forall>a \<in> ?A. f a \<in> ?A') \<and> (\<forall>a' \<in> ?A'. g a' \<in> ?A)"
+ using assms embed_Field[of r r' f] embed_Field[of r' r g] by blast
+ (* *)
+ show ?thesis
+ proof(unfold bij_betw_def inj_on_def, auto simp add: 2)
+ fix a b assume *: "a \<in> ?A" "b \<in> ?A" and **: "f a = f b"
+ have "a = g(f a) \<and> b = g(f b)" using * 1 by auto
+ with ** show "a = b" by auto
+ next
+ fix a' assume *: "a' \<in> ?A'"
+ hence "g a' \<in> ?A \<and> f(g a') = a'" using 1 2 by auto
+ thus "a' \<in> f ` ?A" by force
+ qed
+qed
+
+
+lemma embed_bothWays_iso:
+assumes WELL: "Well_order r" and WELL': "Well_order r'" and
+ EMB: "embed r r' f" and EMB': "embed r' r g"
+shows "iso r r' f"
+unfolding iso_def using assms by (auto simp add: embed_bothWays_bij_betw)
+
+
+subsection {* More properties of embeddings, strict embeddings and isomorphisms *}
+
+
+lemma embed_bothWays_Field_bij_betw:
+assumes WELL: "Well_order r" and WELL': "Well_order r'" and
+ EMB: "embed r r' f" and EMB': "embed r' r f'"
+shows "bij_betw f (Field r) (Field r')"
+proof-
+ have "(\<forall>a \<in> Field r. f'(f a) = a) \<and> (\<forall>a' \<in> Field r'. f(f' a') = a')"
+ using assms by (auto simp add: embed_bothWays_inverse)
+ moreover
+ have "f`(Field r) \<le> Field r' \<and> f' ` (Field r') \<le> Field r"
+ using assms by (auto simp add: embed_Field)
+ ultimately
+ show ?thesis using bij_betw_byWitness[of "Field r" f' f "Field r'"] by auto
+qed
+
+
+lemma embedS_comp_embed:
+assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''"
+ and EMB: "embedS r r' f" and EMB': "embed r' r'' f'"
+shows "embedS r r'' (f' o f)"
+proof-
+ let ?g = "(f' o f)" let ?h = "inv_into (Field r) ?g"
+ have 1: "embed r r' f \<and> \<not> (bij_betw f (Field r) (Field r'))"
+ using EMB by (auto simp add: embedS_def)
+ hence 2: "embed r r'' ?g"
+ using WELL EMB' comp_embed[of r r' f r'' f'] by auto
+ moreover
+ {assume "bij_betw ?g (Field r) (Field r'')"
+ hence "embed r'' r ?h" using 2 WELL
+ by (auto simp add: inv_into_Field_embed_bij_betw)
+ hence "embed r' r (?h o f')" using WELL' EMB'
+ by (auto simp add: comp_embed)
+ hence "bij_betw f (Field r) (Field r')" using WELL WELL' 1
+ by (auto simp add: embed_bothWays_Field_bij_betw)
+ with 1 have False by blast
+ }
+ ultimately show ?thesis unfolding embedS_def by auto
+qed
+
+
+lemma embed_comp_embedS:
+assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''"
+ and EMB: "embed r r' f" and EMB': "embedS r' r'' f'"
+shows "embedS r r'' (f' o f)"
+proof-
+ let ?g = "(f' o f)" let ?h = "inv_into (Field r) ?g"
+ have 1: "embed r' r'' f' \<and> \<not> (bij_betw f' (Field r') (Field r''))"
+ using EMB' by (auto simp add: embedS_def)
+ hence 2: "embed r r'' ?g"
+ using WELL EMB comp_embed[of r r' f r'' f'] by auto
+ moreover
+ {assume "bij_betw ?g (Field r) (Field r'')"
+ hence "embed r'' r ?h" using 2 WELL
+ by (auto simp add: inv_into_Field_embed_bij_betw)
+ hence "embed r'' r' (f o ?h)" using WELL'' EMB
+ by (auto simp add: comp_embed)
+ hence "bij_betw f' (Field r') (Field r'')" using WELL' WELL'' 1
+ by (auto simp add: embed_bothWays_Field_bij_betw)
+ with 1 have False by blast
+ }
+ ultimately show ?thesis unfolding embedS_def by auto
+qed
+
+
+lemma embed_comp_iso:
+assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''"
+ and EMB: "embed r r' f" and EMB': "iso r' r'' f'"
+shows "embed r r'' (f' o f)"
+using assms unfolding iso_def
+by (auto simp add: comp_embed)
+
+
+lemma iso_comp_embed:
+assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''"
+ and EMB: "iso r r' f" and EMB': "embed r' r'' f'"
+shows "embed r r'' (f' o f)"
+using assms unfolding iso_def
+by (auto simp add: comp_embed)
+
+
+lemma embedS_comp_iso:
+assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''"
+ and EMB: "embedS r r' f" and EMB': "iso r' r'' f'"
+shows "embedS r r'' (f' o f)"
+using assms unfolding iso_def
+by (auto simp add: embedS_comp_embed)
+
+
+lemma iso_comp_embedS:
+assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''"
+ and EMB: "iso r r' f" and EMB': "embedS r' r'' f'"
+shows "embedS r r'' (f' o f)"
+using assms unfolding iso_def using embed_comp_embedS
+by (auto simp add: embed_comp_embedS)
+
+
+lemma embedS_Field:
+assumes WELL: "Well_order r" and EMB: "embedS r r' f"
+shows "f ` (Field r) < Field r'"
+proof-
+ have "f`(Field r) \<le> Field r'" using assms
+ by (auto simp add: embed_Field embedS_def)
+ moreover
+ {have "inj_on f (Field r)" using assms
+ by (auto simp add: embedS_def embed_inj_on)
+ hence "f`(Field r) \<noteq> Field r'" using EMB
+ by (auto simp add: embedS_def bij_betw_def)
+ }
+ ultimately show ?thesis by blast
+qed
+
+
+lemma embedS_iff:
+assumes WELL: "Well_order r" and ISO: "embed r r' f"
+shows "embedS r r' f = (f ` (Field r) < Field r')"
+proof
+ assume "embedS r r' f"
+ thus "f ` Field r \<subset> Field r'"
+ using WELL by (auto simp add: embedS_Field)
+next
+ assume "f ` Field r \<subset> Field r'"
+ hence "\<not> bij_betw f (Field r) (Field r')"
+ unfolding bij_betw_def by blast
+ thus "embedS r r' f" unfolding embedS_def
+ using ISO by auto
+qed
+
+
+lemma iso_Field:
+"iso r r' f \<Longrightarrow> f ` (Field r) = Field r'"
+using assms by (auto simp add: iso_def bij_betw_def)
+
+
+lemma iso_iff:
+assumes "Well_order r"
+shows "iso r r' f = (embed r r' f \<and> f ` (Field r) = Field r')"
+proof
+ assume "iso r r' f"
+ thus "embed r r' f \<and> f ` (Field r) = Field r'"
+ by (auto simp add: iso_Field iso_def)
+next
+ assume *: "embed r r' f \<and> f ` Field r = Field r'"
+ hence "inj_on f (Field r)" using assms by (auto simp add: embed_inj_on)
+ with * have "bij_betw f (Field r) (Field r')"
+ unfolding bij_betw_def by simp
+ with * show "iso r r' f" unfolding iso_def by auto
+qed
+
+
+lemma iso_iff2:
+assumes "Well_order r"
+shows "iso r r' f = (bij_betw f (Field r) (Field r') \<and>
+ (\<forall>a \<in> Field r. \<forall>b \<in> Field r.
+ (((a,b) \<in> r) = ((f a, f b) \<in> r'))))"
+using assms
+proof(auto simp add: iso_def)
+ fix a b
+ assume "embed r r' f"
+ hence "compat r r' f" using embed_compat[of r] by auto
+ moreover assume "(a,b) \<in> r"
+ ultimately show "(f a, f b) \<in> r'" using compat_def[of r] by auto
+next
+ let ?f' = "inv_into (Field r) f"
+ assume "embed r r' f" and 1: "bij_betw f (Field r) (Field r')"
+ hence "embed r' r ?f'" using assms
+ by (auto simp add: inv_into_Field_embed_bij_betw)
+ hence 2: "compat r' r ?f'" using embed_compat[of r'] by auto
+ fix a b assume *: "a \<in> Field r" "b \<in> Field r" and **: "(f a,f b) \<in> r'"
+ hence "?f'(f a) = a \<and> ?f'(f b) = b" using 1
+ by (auto simp add: bij_betw_inv_into_left)
+ thus "(a,b) \<in> r" using ** 2 compat_def[of r' r ?f'] by fastforce
+next
+ assume *: "bij_betw f (Field r) (Field r')" and
+ **: "\<forall>a\<in>Field r. \<forall>b\<in>Field r. ((a, b) \<in> r) = ((f a, f b) \<in> r')"
+ have 1: "\<And> a. rel.under r a \<le> Field r \<and> rel.under r' (f a) \<le> Field r'"
+ by (auto simp add: rel.under_Field)
+ have 2: "inj_on f (Field r)" using * by (auto simp add: bij_betw_def)
+ {fix a assume ***: "a \<in> Field r"
+ have "bij_betw f (rel.under r a) (rel.under r' (f a))"
+ proof(unfold bij_betw_def, auto)
+ show "inj_on f (rel.under r a)"
+ using 1 2 by (auto simp add: subset_inj_on)
+ next
+ fix b assume "b \<in> rel.under r a"
+ hence "a \<in> Field r \<and> b \<in> Field r \<and> (b,a) \<in> r"
+ unfolding rel.under_def by (auto simp add: Field_def Range_def Domain_def)
+ with 1 ** show "f b \<in> rel.under r' (f a)"
+ unfolding rel.under_def by auto
+ next
+ fix b' assume "b' \<in> rel.under r' (f a)"
+ hence 3: "(b',f a) \<in> r'" unfolding rel.under_def by simp
+ hence "b' \<in> Field r'" unfolding Field_def by auto
+ with * obtain b where "b \<in> Field r \<and> f b = b'"
+ unfolding bij_betw_def by force
+ with 3 ** ***
+ show "b' \<in> f ` (rel.under r a)" unfolding rel.under_def by blast
+ qed
+ }
+ thus "embed r r' f" unfolding embed_def using * by auto
+qed
+
+
+lemma iso_iff3:
+assumes WELL: "Well_order r" and WELL': "Well_order r'"
+shows "iso r r' f = (bij_betw f (Field r) (Field r') \<and> compat r r' f)"
+proof
+ assume "iso r r' f"
+ thus "bij_betw f (Field r) (Field r') \<and> compat r r' f"
+ unfolding compat_def using WELL by (auto simp add: iso_iff2 Field_def)
+next
+ have Well: "wo_rel r \<and> wo_rel r'" using WELL WELL'
+ by (auto simp add: wo_rel_def)
+ assume *: "bij_betw f (Field r) (Field r') \<and> compat r r' f"
+ thus "iso r r' f"
+ unfolding "compat_def" using assms
+ proof(auto simp add: iso_iff2)
+ fix a b assume **: "a \<in> Field r" "b \<in> Field r" and
+ ***: "(f a, f b) \<in> r'"
+ {assume "(b,a) \<in> r \<or> b = a"
+ hence "(b,a): r"using Well ** wo_rel.REFL[of r] refl_on_def[of _ r] by blast
+ hence "(f b, f a) \<in> r'" using * unfolding compat_def by auto
+ hence "f a = f b"
+ using Well *** wo_rel.ANTISYM[of r'] antisym_def[of r'] by blast
+ hence "a = b" using * ** unfolding bij_betw_def inj_on_def by auto
+ hence "(a,b) \<in> r" using Well ** wo_rel.REFL[of r] refl_on_def[of _ r] by blast
+ }
+ thus "(a,b) \<in> r"
+ using Well ** wo_rel.TOTAL[of r] total_on_def[of _ r] by blast
+ qed
+qed
+
+
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Ordinals_and_Cardinals/Wellorder_Relation.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,511 @@
+(* Title: HOL/Ordinals_and_Cardinals/Wellorder_Relation.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Well-order relations.
+*)
+
+header {* Well-Order Relations *}
+
+theory Wellorder_Relation
+imports Wellorder_Relation_Base Wellfounded_More
+begin
+
+
+subsection {* Auxiliaries *}
+
+lemma (in wo_rel) PREORD: "Preorder r"
+using WELL order_on_defs[of _ r] by auto
+
+lemma (in wo_rel) PARORD: "Partial_order r"
+using WELL order_on_defs[of _ r] by auto
+
+lemma (in wo_rel) cases_Total2:
+"\<And> phi a b. \<lbrakk>{a,b} \<le> Field r; ((a,b) \<in> r - Id \<Longrightarrow> phi a b);
+ ((b,a) \<in> r - Id \<Longrightarrow> phi a b); (a = b \<Longrightarrow> phi a b)\<rbrakk>
+ \<Longrightarrow> phi a b"
+using TOTALS by auto
+
+
+subsection {* Well-founded induction and recursion adapted to non-strict well-order relations *}
+
+lemma (in wo_rel) worec_unique_fixpoint:
+assumes ADM: "adm_wo H" and fp: "f = H f"
+shows "f = worec H"
+proof-
+ have "adm_wf (r - Id) H"
+ unfolding adm_wf_def
+ using ADM adm_wo_def[of H] underS_def by auto
+ hence "f = wfrec (r - Id) H"
+ using fp WF wfrec_unique_fixpoint[of "r - Id" H] by simp
+ thus ?thesis unfolding worec_def .
+qed
+
+
+subsubsection {* Properties of max2 *}
+
+lemma (in wo_rel) max2_iff:
+assumes "a \<in> Field r" and "b \<in> Field r"
+shows "((max2 a b, c) \<in> r) = ((a,c) \<in> r \<and> (b,c) \<in> r)"
+proof
+ assume "(max2 a b, c) \<in> r"
+ thus "(a,c) \<in> r \<and> (b,c) \<in> r"
+ using assms max2_greater[of a b] TRANS trans_def[of r] by blast
+next
+ assume "(a,c) \<in> r \<and> (b,c) \<in> r"
+ thus "(max2 a b, c) \<in> r"
+ using assms max2_among[of a b] by auto
+qed
+
+
+subsubsection{* Properties of minim *}
+
+lemma (in wo_rel) minim_Under:
+"\<lbrakk>B \<le> Field r; B \<noteq> {}\<rbrakk> \<Longrightarrow> minim B \<in> Under B"
+by(auto simp add: Under_def
+minim_in
+minim_inField
+minim_least
+under_ofilter
+underS_ofilter
+Field_ofilter
+ofilter_Under
+ofilter_UnderS
+ofilter_Un
+)
+
+lemma (in wo_rel) equals_minim_Under:
+"\<lbrakk>B \<le> Field r; a \<in> B; a \<in> Under B\<rbrakk>
+ \<Longrightarrow> a = minim B"
+by(auto simp add: Under_def equals_minim)
+
+lemma (in wo_rel) minim_iff_In_Under:
+assumes SUB: "B \<le> Field r" and NE: "B \<noteq> {}"
+shows "(a = minim B) = (a \<in> B \<and> a \<in> Under B)"
+proof
+ assume "a = minim B"
+ thus "a \<in> B \<and> a \<in> Under B"
+ using assms minim_in minim_Under by simp
+next
+ assume "a \<in> B \<and> a \<in> Under B"
+ thus "a = minim B"
+ using assms equals_minim_Under by simp
+qed
+
+lemma (in wo_rel) minim_Under_under:
+assumes NE: "A \<noteq> {}" and SUB: "A \<le> Field r"
+shows "Under A = under (minim A)"
+proof-
+ (* Preliminary facts *)
+ have 1: "minim A \<in> A"
+ using assms minim_in by auto
+ have 2: "\<forall>x \<in> A. (minim A, x) \<in> r"
+ using assms minim_least by auto
+ (* Main proof *)
+ have "Under A \<le> under (minim A)"
+ proof
+ fix x assume "x \<in> Under A"
+ with 1 Under_def have "(x,minim A) \<in> r" by auto
+ thus "x \<in> under(minim A)" unfolding under_def by simp
+ qed
+ (* *)
+ moreover
+ (* *)
+ have "under (minim A) \<le> Under A"
+ proof
+ fix x assume "x \<in> under(minim A)"
+ hence 11: "(x,minim A) \<in> r" unfolding under_def by simp
+ hence "x \<in> Field r" unfolding Field_def by auto
+ moreover
+ {fix a assume "a \<in> A"
+ with 2 have "(minim A, a) \<in> r" by simp
+ with 11 have "(x,a) \<in> r"
+ using TRANS trans_def[of r] by blast
+ }
+ ultimately show "x \<in> Under A" by (unfold Under_def, auto)
+ qed
+ (* *)
+ ultimately show ?thesis by blast
+qed
+
+lemma (in wo_rel) minim_UnderS_underS:
+assumes NE: "A \<noteq> {}" and SUB: "A \<le> Field r"
+shows "UnderS A = underS (minim A)"
+proof-
+ (* Preliminary facts *)
+ have 1: "minim A \<in> A"
+ using assms minim_in by auto
+ have 2: "\<forall>x \<in> A. (minim A, x) \<in> r"
+ using assms minim_least by auto
+ (* Main proof *)
+ have "UnderS A \<le> underS (minim A)"
+ proof
+ fix x assume "x \<in> UnderS A"
+ with 1 UnderS_def have "x \<noteq> minim A \<and> (x,minim A) \<in> r" by auto
+ thus "x \<in> underS(minim A)" unfolding underS_def by simp
+ qed
+ (* *)
+ moreover
+ (* *)
+ have "underS (minim A) \<le> UnderS A"
+ proof
+ fix x assume "x \<in> underS(minim A)"
+ hence 11: "x \<noteq> minim A \<and> (x,minim A) \<in> r" unfolding underS_def by simp
+ hence "x \<in> Field r" unfolding Field_def by auto
+ moreover
+ {fix a assume "a \<in> A"
+ with 2 have 3: "(minim A, a) \<in> r" by simp
+ with 11 have "(x,a) \<in> r"
+ using TRANS trans_def[of r] by blast
+ moreover
+ have "x \<noteq> a"
+ proof
+ assume "x = a"
+ with 11 3 ANTISYM antisym_def[of r]
+ show False by auto
+ qed
+ ultimately
+ have "x \<noteq> a \<and> (x,a) \<in> r" by simp
+ }
+ ultimately show "x \<in> UnderS A" by (unfold UnderS_def, auto)
+ qed
+ (* *)
+ ultimately show ?thesis by blast
+qed
+
+
+subsubsection{* Properties of supr *}
+
+lemma (in wo_rel) supr_Above:
+assumes SUB: "B \<le> Field r" and ABOVE: "Above B \<noteq> {}"
+shows "supr B \<in> Above B"
+proof(unfold supr_def)
+ have "Above B \<le> Field r"
+ using Above_Field by auto
+ thus "minim (Above B) \<in> Above B"
+ using assms by (simp add: minim_in)
+qed
+
+lemma (in wo_rel) supr_greater:
+assumes SUB: "B \<le> Field r" and ABOVE: "Above B \<noteq> {}" and
+ IN: "b \<in> B"
+shows "(b, supr B) \<in> r"
+proof-
+ from assms supr_Above
+ have "supr B \<in> Above B" by simp
+ with IN Above_def show ?thesis by simp
+qed
+
+lemma (in wo_rel) supr_least_Above:
+assumes SUB: "B \<le> Field r" and
+ ABOVE: "a \<in> Above B"
+shows "(supr B, a) \<in> r"
+proof(unfold supr_def)
+ have "Above B \<le> Field r"
+ using Above_Field by auto
+ thus "(minim (Above B), a) \<in> r"
+ using assms minim_least
+ by simp
+qed
+
+lemma (in wo_rel) supr_least:
+"\<lbrakk>B \<le> Field r; a \<in> Field r; (\<And> b. b \<in> B \<Longrightarrow> (b,a) \<in> r)\<rbrakk>
+ \<Longrightarrow> (supr B, a) \<in> r"
+by(auto simp add: supr_least_Above Above_def)
+
+lemma (in wo_rel) equals_supr_Above:
+assumes SUB: "B \<le> Field r" and ABV: "a \<in> Above B" and
+ MINIM: "\<And> a'. a' \<in> Above B \<Longrightarrow> (a,a') \<in> r"
+shows "a = supr B"
+proof(unfold supr_def)
+ have "Above B \<le> Field r"
+ using Above_Field by auto
+ thus "a = minim (Above B)"
+ using assms equals_minim by simp
+qed
+
+lemma (in wo_rel) equals_supr:
+assumes SUB: "B \<le> Field r" and IN: "a \<in> Field r" and
+ ABV: "\<And> b. b \<in> B \<Longrightarrow> (b,a) \<in> r" and
+ MINIM: "\<And> a'. \<lbrakk> a' \<in> Field r; \<And> b. b \<in> B \<Longrightarrow> (b,a') \<in> r\<rbrakk> \<Longrightarrow> (a,a') \<in> r"
+shows "a = supr B"
+proof-
+ have "a \<in> Above B"
+ unfolding Above_def using ABV IN by simp
+ moreover
+ have "\<And> a'. a' \<in> Above B \<Longrightarrow> (a,a') \<in> r"
+ unfolding Above_def using MINIM by simp
+ ultimately show ?thesis
+ using equals_supr_Above SUB by auto
+qed
+
+lemma (in wo_rel) supr_inField:
+assumes "B \<le> Field r" and "Above B \<noteq> {}"
+shows "supr B \<in> Field r"
+proof-
+ have "supr B \<in> Above B" using supr_Above assms by simp
+ thus ?thesis using assms Above_Field by auto
+qed
+
+lemma (in wo_rel) supr_above_Above:
+assumes SUB: "B \<le> Field r" and ABOVE: "Above B \<noteq> {}"
+shows "Above B = above (supr B)"
+proof(unfold Above_def above_def, auto)
+ fix a assume "a \<in> Field r" "\<forall>b \<in> B. (b,a) \<in> r"
+ with supr_least assms
+ show "(supr B, a) \<in> r" by auto
+next
+ fix b assume "(supr B, b) \<in> r"
+ thus "b \<in> Field r"
+ using REFL refl_on_def[of _ r] by auto
+next
+ fix a b
+ assume 1: "(supr B, b) \<in> r" and 2: "a \<in> B"
+ with assms supr_greater
+ have "(a,supr B) \<in> r" by auto
+ thus "(a,b) \<in> r"
+ using 1 TRANS trans_def[of r] by blast
+qed
+
+lemma (in wo_rel) supr_under:
+assumes IN: "a \<in> Field r"
+shows "a = supr (under a)"
+proof-
+ have "under a \<le> Field r"
+ using under_Field by auto
+ moreover
+ have "under a \<noteq> {}"
+ using IN Refl_under_in REFL by auto
+ moreover
+ have "a \<in> Above (under a)"
+ using in_Above_under IN by auto
+ moreover
+ have "\<forall>a' \<in> Above (under a). (a,a') \<in> r"
+ proof(unfold Above_def under_def, auto)
+ fix a'
+ assume "\<forall>aa. (aa, a) \<in> r \<longrightarrow> (aa, a') \<in> r"
+ hence "(a,a) \<in> r \<longrightarrow> (a,a') \<in> r" by blast
+ moreover have "(a,a) \<in> r"
+ using REFL IN by (auto simp add: refl_on_def)
+ ultimately
+ show "(a, a') \<in> r" by (rule mp)
+ qed
+ ultimately show ?thesis
+ using equals_supr_Above by auto
+qed
+
+
+subsubsection{* Properties of successor *}
+
+lemma (in wo_rel) suc_least:
+"\<lbrakk>B \<le> Field r; a \<in> Field r; (\<And> b. b \<in> B \<Longrightarrow> a \<noteq> b \<and> (b,a) \<in> r)\<rbrakk>
+ \<Longrightarrow> (suc B, a) \<in> r"
+by(auto simp add: suc_least_AboveS AboveS_def)
+
+lemma (in wo_rel) equals_suc:
+assumes SUB: "B \<le> Field r" and IN: "a \<in> Field r" and
+ ABVS: "\<And> b. b \<in> B \<Longrightarrow> a \<noteq> b \<and> (b,a) \<in> r" and
+ MINIM: "\<And> a'. \<lbrakk>a' \<in> Field r; \<And> b. b \<in> B \<Longrightarrow> a' \<noteq> b \<and> (b,a') \<in> r\<rbrakk> \<Longrightarrow> (a,a') \<in> r"
+shows "a = suc B"
+proof-
+ have "a \<in> AboveS B"
+ unfolding AboveS_def using ABVS IN by simp
+ moreover
+ have "\<And> a'. a' \<in> AboveS B \<Longrightarrow> (a,a') \<in> r"
+ unfolding AboveS_def using MINIM by simp
+ ultimately show ?thesis
+ using equals_suc_AboveS SUB by auto
+qed
+
+lemma (in wo_rel) suc_above_AboveS:
+assumes SUB: "B \<le> Field r" and
+ ABOVE: "AboveS B \<noteq> {}"
+shows "AboveS B = above (suc B)"
+proof(unfold AboveS_def above_def, auto)
+ fix a assume "a \<in> Field r" "\<forall>b \<in> B. a \<noteq> b \<and> (b,a) \<in> r"
+ with suc_least assms
+ show "(suc B,a) \<in> r" by auto
+next
+ fix b assume "(suc B, b) \<in> r"
+ thus "b \<in> Field r"
+ using REFL refl_on_def[of _ r] by auto
+next
+ fix a b
+ assume 1: "(suc B, b) \<in> r" and 2: "a \<in> B"
+ with assms suc_greater[of B a]
+ have "(a,suc B) \<in> r" by auto
+ thus "(a,b) \<in> r"
+ using 1 TRANS trans_def[of r] by blast
+next
+ fix a
+ assume 1: "(suc B, a) \<in> r" and 2: "a \<in> B"
+ with assms suc_greater[of B a]
+ have "(a,suc B) \<in> r" by auto
+ moreover have "suc B \<in> Field r"
+ using assms suc_inField by simp
+ ultimately have "a = suc B"
+ using 1 2 SUB ANTISYM antisym_def[of r] by auto
+ thus False
+ using assms suc_greater[of B a] 2 by auto
+qed
+
+lemma (in wo_rel) suc_singl_pred:
+assumes IN: "a \<in> Field r" and ABOVE_NE: "aboveS a \<noteq> {}" and
+ REL: "(a',suc {a}) \<in> r" and DIFF: "a' \<noteq> suc {a}"
+shows "a' = a \<or> (a',a) \<in> r"
+proof-
+ have *: "suc {a} \<in> Field r \<and> a' \<in> Field r"
+ using WELL REL well_order_on_domain by auto
+ {assume **: "a' \<noteq> a"
+ hence "(a,a') \<in> r \<or> (a',a) \<in> r"
+ using TOTAL IN * by (auto simp add: total_on_def)
+ moreover
+ {assume "(a,a') \<in> r"
+ with ** * assms WELL suc_least[of "{a}" a']
+ have "(suc {a},a') \<in> r" by auto
+ with REL DIFF * ANTISYM antisym_def[of r]
+ have False by simp
+ }
+ ultimately have "(a',a) \<in> r"
+ by blast
+ }
+ thus ?thesis by blast
+qed
+
+lemma (in wo_rel) under_underS_suc:
+assumes IN: "a \<in> Field r" and ABV: "aboveS a \<noteq> {}"
+shows "underS (suc {a}) = under a"
+proof-
+ have 1: "AboveS {a} \<noteq> {}"
+ using ABV aboveS_AboveS_singl by auto
+ have 2: "a \<noteq> suc {a} \<and> (a,suc {a}) \<in> r"
+ using suc_greater[of "{a}" a] IN 1 by auto
+ (* *)
+ have "underS (suc {a}) \<le> under a"
+ proof(unfold underS_def under_def, auto)
+ fix x assume *: "x \<noteq> suc {a}" and **: "(x,suc {a}) \<in> r"
+ with suc_singl_pred[of a x] IN ABV
+ have "x = a \<or> (x,a) \<in> r" by auto
+ with REFL refl_on_def[of _ r] IN
+ show "(x,a) \<in> r" by auto
+ qed
+ (* *)
+ moreover
+ (* *)
+ have "under a \<le> underS (suc {a})"
+ proof(unfold underS_def under_def, auto)
+ assume "(suc {a}, a) \<in> r"
+ with 2 ANTISYM antisym_def[of r]
+ show False by auto
+ next
+ fix x assume *: "(x,a) \<in> r"
+ with 2 TRANS trans_def[of r]
+ show "(x,suc {a}) \<in> r" by blast
+ (* blast is often better than auto/auto for transitivity-like properties *)
+ qed
+ (* *)
+ ultimately show ?thesis by blast
+qed
+
+
+subsubsection {* Properties of order filters *}
+
+lemma (in wo_rel) ofilter_INTER:
+"\<lbrakk>I \<noteq> {}; \<And> i. i \<in> I \<Longrightarrow> ofilter(A i)\<rbrakk> \<Longrightarrow> ofilter (\<Inter> i \<in> I. A i)"
+unfolding ofilter_def by blast
+
+lemma (in wo_rel) ofilter_Inter:
+"\<lbrakk>S \<noteq> {}; \<And> A. A \<in> S \<Longrightarrow> ofilter A\<rbrakk> \<Longrightarrow> ofilter (Inter S)"
+unfolding ofilter_def by blast
+
+lemma (in wo_rel) ofilter_Union:
+"(\<And> A. A \<in> S \<Longrightarrow> ofilter A) \<Longrightarrow> ofilter (Union S)"
+unfolding ofilter_def by blast
+
+lemma (in wo_rel) ofilter_under_Union:
+"ofilter A \<Longrightarrow> A = Union {under a| a. a \<in> A}"
+using ofilter_under_UNION[of A]
+by(unfold Union_eq, auto)
+
+
+subsubsection{* Other properties *}
+
+lemma (in wo_rel) Trans_Under_regressive:
+assumes NE: "A \<noteq> {}" and SUB: "A \<le> Field r"
+shows "Under(Under A) \<le> Under A"
+proof
+ let ?a = "minim A"
+ (* Preliminary facts *)
+ have 1: "minim A \<in> Under A"
+ using assms minim_Under by auto
+ have 2: "\<forall>y \<in> A. (minim A, y) \<in> r"
+ using assms minim_least by auto
+ (* Main proof *)
+ fix x assume "x \<in> Under(Under A)"
+ with 1 have 1: "(x,minim A) \<in> r"
+ using Under_def by auto
+ with Field_def have "x \<in> Field r" by fastforce
+ moreover
+ {fix y assume *: "y \<in> A"
+ hence "(x,y) \<in> r"
+ using 1 2 TRANS trans_def[of r] by blast
+ with Field_def have "(x,y) \<in> r" by auto
+ }
+ ultimately
+ show "x \<in> Under A" unfolding Under_def by auto
+qed
+
+lemma (in wo_rel) ofilter_suc_Field:
+assumes OF: "ofilter A" and NE: "A \<noteq> Field r"
+shows "ofilter (A \<union> {suc A})"
+proof-
+ (* Preliminary facts *)
+ have 1: "A \<le> Field r" using OF ofilter_def by auto
+ hence 2: "AboveS A \<noteq> {}"
+ using ofilter_AboveS_Field NE OF by blast
+ from 1 2 suc_inField
+ have 3: "suc A \<in> Field r" by auto
+ (* Main proof *)
+ show ?thesis
+ proof(unfold ofilter_def, auto simp add: 1 3)
+ fix a x
+ assume "a \<in> A" "x \<in> under a" "x \<notin> A"
+ with OF ofilter_def have False by auto
+ thus "x = suc A" by simp
+ next
+ fix x assume *: "x \<in> under (suc A)" and **: "x \<notin> A"
+ hence "x \<in> Field r" using under_def Field_def by fastforce
+ with ** have "x \<in> AboveS A"
+ using ofilter_AboveS_Field[of A] OF by auto
+ hence "(suc A,x) \<in> r"
+ using suc_least_AboveS by auto
+ moreover
+ have "(x,suc A) \<in> r" using * under_def by auto
+ ultimately show "x = suc A"
+ using ANTISYM antisym_def[of r] by auto
+ qed
+qed
+
+(* FIXME: needed? *)
+declare (in wo_rel)
+ minim_in[simp]
+ minim_inField[simp]
+ minim_least[simp]
+ under_ofilter[simp]
+ underS_ofilter[simp]
+ Field_ofilter[simp]
+ ofilter_Under[simp]
+ ofilter_UnderS[simp]
+ ofilter_Int[simp]
+ ofilter_Un[simp]
+
+abbreviation "worec \<equiv> wo_rel.worec"
+abbreviation "adm_wo \<equiv> wo_rel.adm_wo"
+abbreviation "isMinim \<equiv> wo_rel.isMinim"
+abbreviation "minim \<equiv> wo_rel.minim"
+abbreviation "max2 \<equiv> wo_rel.max2"
+abbreviation "supr \<equiv> wo_rel.supr"
+abbreviation "suc \<equiv> wo_rel.suc"
+abbreviation "ofilter \<equiv> wo_rel.ofilter"
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Ordinals_and_Cardinals/Wellorder_Relation_Base.thy Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,669 @@
+(* Title: HOL/Ordinals_and_Cardinals/Wellorder_Relation_Base.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Well-order relations (base).
+*)
+
+header {* Well-Order Relations (Base) *}
+
+theory Wellorder_Relation_Base
+imports Wellfounded_More_Base
+begin
+
+
+text{* In this section, we develop basic concepts and results pertaining
+to well-order relations. Note that we consider well-order relations
+as {\em non-strict relations},
+i.e., as containing the diagonals of their fields. *}
+
+
+locale wo_rel = rel + assumes WELL: "Well_order r"
+begin
+
+text{* The following context encompasses all this section. In other words,
+for the whole section, we consider a fixed well-order relation @{term "r"}. *}
+
+(* context wo_rel *)
+
+
+subsection {* Auxiliaries *}
+
+
+lemma REFL: "Refl r"
+using WELL order_on_defs[of _ r] by auto
+
+
+lemma TRANS: "trans r"
+using WELL order_on_defs[of _ r] by auto
+
+
+lemma ANTISYM: "antisym r"
+using WELL order_on_defs[of _ r] by auto
+
+
+lemma TOTAL: "Total r"
+using WELL order_on_defs[of _ r] by auto
+
+
+lemma TOTALS: "\<forall>a \<in> Field r. \<forall>b \<in> Field r. (a,b) \<in> r \<or> (b,a) \<in> r"
+using REFL TOTAL refl_on_def[of _ r] total_on_def[of _ r] by force
+
+
+lemma LIN: "Linear_order r"
+using WELL well_order_on_def[of _ r] by auto
+
+
+lemma WF: "wf (r - Id)"
+using WELL well_order_on_def[of _ r] by auto
+
+
+lemma cases_Total:
+"\<And> phi a b. \<lbrakk>{a,b} <= Field r; ((a,b) \<in> r \<Longrightarrow> phi a b); ((b,a) \<in> r \<Longrightarrow> phi a b)\<rbrakk>
+ \<Longrightarrow> phi a b"
+using TOTALS by auto
+
+
+lemma cases_Total3:
+"\<And> phi a b. \<lbrakk>{a,b} \<le> Field r; ((a,b) \<in> r - Id \<or> (b,a) \<in> r - Id \<Longrightarrow> phi a b);
+ (a = b \<Longrightarrow> phi a b)\<rbrakk> \<Longrightarrow> phi a b"
+using TOTALS by auto
+
+
+subsection {* Well-founded induction and recursion adapted to non-strict well-order relations *}
+
+
+text{* Here we provide induction and recursion principles specific to {\em non-strict}
+well-order relations.
+Although minor variations of those for well-founded relations, they will be useful
+for doing away with the tediousness of
+having to take out the diagonal each time in order to switch to a well-founded relation. *}
+
+
+lemma well_order_induct:
+assumes IND: "\<And>x. \<forall>y. y \<noteq> x \<and> (y, x) \<in> r \<longrightarrow> P y \<Longrightarrow> P x"
+shows "P a"
+proof-
+ have "\<And>x. \<forall>y. (y, x) \<in> r - Id \<longrightarrow> P y \<Longrightarrow> P x"
+ using IND by blast
+ thus "P a" using WF wf_induct[of "r - Id" P a] by blast
+qed
+
+
+definition
+worec :: "(('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
+where
+"worec F \<equiv> wfrec (r - Id) F"
+
+
+definition
+adm_wo :: "(('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> bool"
+where
+"adm_wo H \<equiv> \<forall>f g x. (\<forall>y \<in> underS x. f y = g y) \<longrightarrow> H f x = H g x"
+
+
+lemma worec_fixpoint:
+assumes ADM: "adm_wo H"
+shows "worec H = H (worec H)"
+proof-
+ let ?rS = "r - Id"
+ have "adm_wf (r - Id) H"
+ unfolding adm_wf_def
+ using ADM adm_wo_def[of H] underS_def by auto
+ hence "wfrec ?rS H = H (wfrec ?rS H)"
+ using WF wfrec_fixpoint[of ?rS H] by simp
+ thus ?thesis unfolding worec_def .
+qed
+
+
+subsection {* The notions of maximum, minimum, supremum, successor and order filter *}
+
+
+text{*
+We define the successor {\em of a set}, and not of an element (the latter is of course
+a particular case). Also, we define the maximum {\em of two elements}, @{text "max2"},
+and the minimum {\em of a set}, @{text "minim"} -- we chose these variants since we
+consider them the most useful for well-orders. The minimum is defined in terms of the
+auxiliary relational operator @{text "isMinim"}. Then, supremum and successor are
+defined in terms of minimum as expected.
+The minimum is only meaningful for non-empty sets, and the successor is only
+meaningful for sets for which strict upper bounds exist.
+Order filters for well-orders are also known as ``initial segments". *}
+
+
+definition max2 :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
+where "max2 a b \<equiv> if (a,b) \<in> r then b else a"
+
+
+definition isMinim :: "'a set \<Rightarrow> 'a \<Rightarrow> bool"
+where "isMinim A b \<equiv> b \<in> A \<and> (\<forall>a \<in> A. (b,a) \<in> r)"
+
+definition minim :: "'a set \<Rightarrow> 'a"
+where "minim A \<equiv> THE b. isMinim A b"
+
+
+definition supr :: "'a set \<Rightarrow> 'a"
+where "supr A \<equiv> minim (Above A)"
+
+definition suc :: "'a set \<Rightarrow> 'a"
+where "suc A \<equiv> minim (AboveS A)"
+
+definition ofilter :: "'a set \<Rightarrow> bool"
+where
+"ofilter A \<equiv> (A \<le> Field r) \<and> (\<forall>a \<in> A. under a \<le> A)"
+
+
+subsubsection {* Properties of max2 *}
+
+
+lemma max2_greater_among:
+assumes "a \<in> Field r" and "b \<in> Field r"
+shows "(a, max2 a b) \<in> r \<and> (b, max2 a b) \<in> r \<and> max2 a b \<in> {a,b}"
+proof-
+ {assume "(a,b) \<in> r"
+ hence ?thesis using max2_def assms REFL refl_on_def
+ by (auto simp add: refl_on_def)
+ }
+ moreover
+ {assume "a = b"
+ hence "(a,b) \<in> r" using REFL assms
+ by (auto simp add: refl_on_def)
+ }
+ moreover
+ {assume *: "a \<noteq> b \<and> (b,a) \<in> r"
+ hence "(a,b) \<notin> r" using ANTISYM
+ by (auto simp add: antisym_def)
+ hence ?thesis using * max2_def assms REFL refl_on_def
+ by (auto simp add: refl_on_def)
+ }
+ ultimately show ?thesis using assms TOTAL
+ total_on_def[of "Field r" r] by blast
+qed
+
+
+lemma max2_greater:
+assumes "a \<in> Field r" and "b \<in> Field r"
+shows "(a, max2 a b) \<in> r \<and> (b, max2 a b) \<in> r"
+using assms by (auto simp add: max2_greater_among)
+
+
+lemma max2_among:
+assumes "a \<in> Field r" and "b \<in> Field r"
+shows "max2 a b \<in> {a, b}"
+using assms max2_greater_among[of a b] by simp
+
+
+lemma max2_equals1:
+assumes "a \<in> Field r" and "b \<in> Field r"
+shows "(max2 a b = a) = ((b,a) \<in> r)"
+using assms ANTISYM unfolding antisym_def using TOTALS
+by(auto simp add: max2_def max2_among)
+
+
+lemma max2_equals2:
+assumes "a \<in> Field r" and "b \<in> Field r"
+shows "(max2 a b = b) = ((a,b) \<in> r)"
+using assms ANTISYM unfolding antisym_def using TOTALS
+unfolding max2_def by auto
+
+
+subsubsection {* Existence and uniqueness for isMinim and well-definedness of minim *}
+
+
+lemma isMinim_unique:
+assumes MINIM: "isMinim B a" and MINIM': "isMinim B a'"
+shows "a = a'"
+proof-
+ {have "a \<in> B"
+ using MINIM isMinim_def by simp
+ hence "(a',a) \<in> r"
+ using MINIM' isMinim_def by simp
+ }
+ moreover
+ {have "a' \<in> B"
+ using MINIM' isMinim_def by simp
+ hence "(a,a') \<in> r"
+ using MINIM isMinim_def by simp
+ }
+ ultimately
+ show ?thesis using ANTISYM antisym_def[of r] by blast
+qed
+
+
+lemma Well_order_isMinim_exists:
+assumes SUB: "B \<le> Field r" and NE: "B \<noteq> {}"
+shows "\<exists>b. isMinim B b"
+proof-
+ from WF wf_eq_minimal[of "r - Id"] NE Id_def obtain b where
+ *: "b \<in> B \<and> (\<forall>b'. b' \<noteq> b \<and> (b',b) \<in> r \<longrightarrow> b' \<notin> B)" by force
+ show ?thesis
+ proof(simp add: isMinim_def, rule exI[of _ b], auto)
+ show "b \<in> B" using * by simp
+ next
+ fix b' assume As: "b' \<in> B"
+ hence **: "b \<in> Field r \<and> b' \<in> Field r" using As SUB * by auto
+ (* *)
+ from As * have "b' = b \<or> (b',b) \<notin> r" by auto
+ moreover
+ {assume "b' = b"
+ hence "(b,b') \<in> r"
+ using ** REFL by (auto simp add: refl_on_def)
+ }
+ moreover
+ {assume "b' \<noteq> b \<and> (b',b) \<notin> r"
+ hence "(b,b') \<in> r"
+ using ** TOTAL by (auto simp add: total_on_def)
+ }
+ ultimately show "(b,b') \<in> r" by blast
+ qed
+qed
+
+
+lemma minim_isMinim:
+assumes SUB: "B \<le> Field r" and NE: "B \<noteq> {}"
+shows "isMinim B (minim B)"
+proof-
+ let ?phi = "(\<lambda> b. isMinim B b)"
+ from assms Well_order_isMinim_exists
+ obtain b where *: "?phi b" by blast
+ moreover
+ have "\<And> b'. ?phi b' \<Longrightarrow> b' = b"
+ using isMinim_unique * by auto
+ ultimately show ?thesis
+ unfolding minim_def using theI[of ?phi b] by blast
+qed
+
+
+subsubsection{* Properties of minim *}
+
+
+lemma minim_in:
+assumes "B \<le> Field r" and "B \<noteq> {}"
+shows "minim B \<in> B"
+proof-
+ from minim_isMinim[of B] assms
+ have "isMinim B (minim B)" by simp
+ thus ?thesis by (simp add: isMinim_def)
+qed
+
+
+lemma minim_inField:
+assumes "B \<le> Field r" and "B \<noteq> {}"
+shows "minim B \<in> Field r"
+proof-
+ have "minim B \<in> B" using assms by (simp add: minim_in)
+ thus ?thesis using assms by blast
+qed
+
+
+lemma minim_least:
+assumes SUB: "B \<le> Field r" and IN: "b \<in> B"
+shows "(minim B, b) \<in> r"
+proof-
+ from minim_isMinim[of B] assms
+ have "isMinim B (minim B)" by auto
+ thus ?thesis by (auto simp add: isMinim_def IN)
+qed
+
+
+lemma equals_minim:
+assumes SUB: "B \<le> Field r" and IN: "a \<in> B" and
+ LEAST: "\<And> b. b \<in> B \<Longrightarrow> (a,b) \<in> r"
+shows "a = minim B"
+proof-
+ from minim_isMinim[of B] assms
+ have "isMinim B (minim B)" by auto
+ moreover have "isMinim B a" using IN LEAST isMinim_def by auto
+ ultimately show ?thesis
+ using isMinim_unique by auto
+qed
+
+
+subsubsection{* Properties of successor *}
+
+
+lemma suc_AboveS:
+assumes SUB: "B \<le> Field r" and ABOVES: "AboveS B \<noteq> {}"
+shows "suc B \<in> AboveS B"
+proof(unfold suc_def)
+ have "AboveS B \<le> Field r"
+ using AboveS_Field by auto
+ thus "minim (AboveS B) \<in> AboveS B"
+ using assms by (simp add: minim_in)
+qed
+
+
+lemma suc_greater:
+assumes SUB: "B \<le> Field r" and ABOVES: "AboveS B \<noteq> {}" and
+ IN: "b \<in> B"
+shows "suc B \<noteq> b \<and> (b,suc B) \<in> r"
+proof-
+ from assms suc_AboveS
+ have "suc B \<in> AboveS B" by simp
+ with IN AboveS_def show ?thesis by simp
+qed
+
+
+lemma suc_least_AboveS:
+assumes ABOVES: "a \<in> AboveS B"
+shows "(suc B,a) \<in> r"
+proof(unfold suc_def)
+ have "AboveS B \<le> Field r"
+ using AboveS_Field by auto
+ thus "(minim (AboveS B),a) \<in> r"
+ using assms minim_least by simp
+qed
+
+
+lemma suc_inField:
+assumes "B \<le> Field r" and "AboveS B \<noteq> {}"
+shows "suc B \<in> Field r"
+proof-
+ have "suc B \<in> AboveS B" using suc_AboveS assms by simp
+ thus ?thesis
+ using assms AboveS_Field by auto
+qed
+
+
+lemma equals_suc_AboveS:
+assumes SUB: "B \<le> Field r" and ABV: "a \<in> AboveS B" and
+ MINIM: "\<And> a'. a' \<in> AboveS B \<Longrightarrow> (a,a') \<in> r"
+shows "a = suc B"
+proof(unfold suc_def)
+ have "AboveS B \<le> Field r"
+ using AboveS_Field[of B] by auto
+ thus "a = minim (AboveS B)"
+ using assms equals_minim
+ by simp
+qed
+
+
+lemma suc_underS:
+assumes IN: "a \<in> Field r"
+shows "a = suc (underS a)"
+proof-
+ have "underS a \<le> Field r"
+ using underS_Field by auto
+ moreover
+ have "a \<in> AboveS (underS a)"
+ using in_AboveS_underS IN by auto
+ moreover
+ have "\<forall>a' \<in> AboveS (underS a). (a,a') \<in> r"
+ proof(clarify)
+ fix a'
+ assume *: "a' \<in> AboveS (underS a)"
+ hence **: "a' \<in> Field r"
+ using AboveS_Field by auto
+ {assume "(a,a') \<notin> r"
+ hence "a' = a \<or> (a',a) \<in> r"
+ using TOTAL IN ** by (auto simp add: total_on_def)
+ moreover
+ {assume "a' = a"
+ hence "(a,a') \<in> r"
+ using REFL IN ** by (auto simp add: refl_on_def)
+ }
+ moreover
+ {assume "a' \<noteq> a \<and> (a',a) \<in> r"
+ hence "a' \<in> underS a"
+ unfolding underS_def by simp
+ hence "a' \<notin> AboveS (underS a)"
+ using AboveS_disjoint by blast
+ with * have False by simp
+ }
+ ultimately have "(a,a') \<in> r" by blast
+ }
+ thus "(a, a') \<in> r" by blast
+ qed
+ ultimately show ?thesis
+ using equals_suc_AboveS by auto
+qed
+
+
+subsubsection {* Properties of order filters *}
+
+
+lemma under_ofilter:
+"ofilter (under a)"
+proof(unfold ofilter_def under_def, auto simp add: Field_def)
+ fix aa x
+ assume "(aa,a) \<in> r" "(x,aa) \<in> r"
+ thus "(x,a) \<in> r"
+ using TRANS trans_def[of r] by blast
+qed
+
+
+lemma underS_ofilter:
+"ofilter (underS a)"
+proof(unfold ofilter_def underS_def under_def, auto simp add: Field_def)
+ fix aa assume "(a, aa) \<in> r" "(aa, a) \<in> r" and DIFF: "aa \<noteq> a"
+ thus False
+ using ANTISYM antisym_def[of r] by blast
+next
+ fix aa x
+ assume "(aa,a) \<in> r" "aa \<noteq> a" "(x,aa) \<in> r"
+ thus "(x,a) \<in> r"
+ using TRANS trans_def[of r] by blast
+qed
+
+
+lemma Field_ofilter:
+"ofilter (Field r)"
+by(unfold ofilter_def under_def, auto simp add: Field_def)
+
+
+lemma ofilter_underS_Field:
+"ofilter A = ((\<exists>a \<in> Field r. A = underS a) \<or> (A = Field r))"
+proof
+ assume "(\<exists>a\<in>Field r. A = underS a) \<or> A = Field r"
+ thus "ofilter A"
+ by (auto simp: underS_ofilter Field_ofilter)
+next
+ assume *: "ofilter A"
+ let ?One = "(\<exists>a\<in>Field r. A = underS a)"
+ let ?Two = "(A = Field r)"
+ show "?One \<or> ?Two"
+ proof(cases ?Two, simp)
+ let ?B = "(Field r) - A"
+ let ?a = "minim ?B"
+ assume "A \<noteq> Field r"
+ moreover have "A \<le> Field r" using * ofilter_def by simp
+ ultimately have 1: "?B \<noteq> {}" by blast
+ hence 2: "?a \<in> Field r" using minim_inField[of ?B] by blast
+ have 3: "?a \<in> ?B" using minim_in[of ?B] 1 by blast
+ hence 4: "?a \<notin> A" by blast
+ have 5: "A \<le> Field r" using * ofilter_def[of A] by auto
+ (* *)
+ moreover
+ have "A = underS ?a"
+ proof
+ show "A \<le> underS ?a"
+ proof(unfold underS_def, auto simp add: 4)
+ fix x assume **: "x \<in> A"
+ hence 11: "x \<in> Field r" using 5 by auto
+ have 12: "x \<noteq> ?a" using 4 ** by auto
+ have 13: "under x \<le> A" using * ofilter_def ** by auto
+ {assume "(x,?a) \<notin> r"
+ hence "(?a,x) \<in> r"
+ using TOTAL total_on_def[of "Field r" r]
+ 2 4 11 12 by auto
+ hence "?a \<in> under x" using under_def by auto
+ hence "?a \<in> A" using ** 13 by blast
+ with 4 have False by simp
+ }
+ thus "(x,?a) \<in> r" by blast
+ qed
+ next
+ show "underS ?a \<le> A"
+ proof(unfold underS_def, auto)
+ fix x
+ assume **: "x \<noteq> ?a" and ***: "(x,?a) \<in> r"
+ hence 11: "x \<in> Field r" using Field_def by fastforce
+ {assume "x \<notin> A"
+ hence "x \<in> ?B" using 11 by auto
+ hence "(?a,x) \<in> r" using 3 minim_least[of ?B x] by blast
+ hence False
+ using ANTISYM antisym_def[of r] ** *** by auto
+ }
+ thus "x \<in> A" by blast
+ qed
+ qed
+ ultimately have ?One using 2 by blast
+ thus ?thesis by simp
+ qed
+qed
+
+
+lemma ofilter_Under:
+assumes "A \<le> Field r"
+shows "ofilter(Under A)"
+proof(unfold ofilter_def, auto)
+ fix x assume "x \<in> Under A"
+ thus "x \<in> Field r"
+ using Under_Field assms by auto
+next
+ fix a x
+ assume "a \<in> Under A" and "x \<in> under a"
+ thus "x \<in> Under A"
+ using TRANS under_Under_trans by auto
+qed
+
+
+lemma ofilter_UnderS:
+assumes "A \<le> Field r"
+shows "ofilter(UnderS A)"
+proof(unfold ofilter_def, auto)
+ fix x assume "x \<in> UnderS A"
+ thus "x \<in> Field r"
+ using UnderS_Field assms by auto
+next
+ fix a x
+ assume "a \<in> UnderS A" and "x \<in> under a"
+ thus "x \<in> UnderS A"
+ using TRANS ANTISYM under_UnderS_trans by auto
+qed
+
+
+lemma ofilter_Int: "\<lbrakk>ofilter A; ofilter B\<rbrakk> \<Longrightarrow> ofilter(A Int B)"
+unfolding ofilter_def by blast
+
+
+lemma ofilter_Un: "\<lbrakk>ofilter A; ofilter B\<rbrakk> \<Longrightarrow> ofilter(A \<union> B)"
+unfolding ofilter_def by blast
+
+
+lemma ofilter_UNION:
+"(\<And> i. i \<in> I \<Longrightarrow> ofilter(A i)) \<Longrightarrow> ofilter (\<Union> i \<in> I. A i)"
+unfolding ofilter_def by blast
+
+
+lemma ofilter_under_UNION:
+assumes "ofilter A"
+shows "A = (\<Union> a \<in> A. under a)"
+proof
+ have "\<forall>a \<in> A. under a \<le> A"
+ using assms ofilter_def by auto
+ thus "(\<Union> a \<in> A. under a) \<le> A" by blast
+next
+ have "\<forall>a \<in> A. a \<in> under a"
+ using REFL Refl_under_in assms ofilter_def by blast
+ thus "A \<le> (\<Union> a \<in> A. under a)" by blast
+qed
+
+
+subsubsection{* Other properties *}
+
+
+lemma ofilter_linord:
+assumes OF1: "ofilter A" and OF2: "ofilter B"
+shows "A \<le> B \<or> B \<le> A"
+proof(cases "A = Field r")
+ assume Case1: "A = Field r"
+ hence "B \<le> A" using OF2 ofilter_def by auto
+ thus ?thesis by simp
+next
+ assume Case2: "A \<noteq> Field r"
+ with ofilter_underS_Field OF1 obtain a where
+ 1: "a \<in> Field r \<and> A = underS a" by auto
+ show ?thesis
+ proof(cases "B = Field r")
+ assume Case21: "B = Field r"
+ hence "A \<le> B" using OF1 ofilter_def by auto
+ thus ?thesis by simp
+ next
+ assume Case22: "B \<noteq> Field r"
+ with ofilter_underS_Field OF2 obtain b where
+ 2: "b \<in> Field r \<and> B = underS b" by auto
+ have "a = b \<or> (a,b) \<in> r \<or> (b,a) \<in> r"
+ using 1 2 TOTAL total_on_def[of _ r] by auto
+ moreover
+ {assume "a = b" with 1 2 have ?thesis by auto
+ }
+ moreover
+ {assume "(a,b) \<in> r"
+ with underS_incr TRANS ANTISYM 1 2
+ have "A \<le> B" by auto
+ hence ?thesis by auto
+ }
+ moreover
+ {assume "(b,a) \<in> r"
+ with underS_incr TRANS ANTISYM 1 2
+ have "B \<le> A" by auto
+ hence ?thesis by auto
+ }
+ ultimately show ?thesis by blast
+ qed
+qed
+
+
+lemma ofilter_AboveS_Field:
+assumes "ofilter A"
+shows "A \<union> (AboveS A) = Field r"
+proof
+ show "A \<union> (AboveS A) \<le> Field r"
+ using assms ofilter_def AboveS_Field by auto
+next
+ {fix x assume *: "x \<in> Field r" and **: "x \<notin> A"
+ {fix y assume ***: "y \<in> A"
+ with ** have 1: "y \<noteq> x" by auto
+ {assume "(y,x) \<notin> r"
+ moreover
+ have "y \<in> Field r" using assms ofilter_def *** by auto
+ ultimately have "(x,y) \<in> r"
+ using 1 * TOTAL total_on_def[of _ r] by auto
+ with *** assms ofilter_def under_def have "x \<in> A" by auto
+ with ** have False by contradiction
+ }
+ hence "(y,x) \<in> r" by blast
+ with 1 have "y \<noteq> x \<and> (y,x) \<in> r" by auto
+ }
+ with * have "x \<in> AboveS A" unfolding AboveS_def by auto
+ }
+ thus "Field r \<le> A \<union> (AboveS A)" by blast
+qed
+
+
+lemma suc_ofilter_in:
+assumes OF: "ofilter A" and ABOVE_NE: "AboveS A \<noteq> {}" and
+ REL: "(b,suc A) \<in> r" and DIFF: "b \<noteq> suc A"
+shows "b \<in> A"
+proof-
+ have *: "suc A \<in> Field r \<and> b \<in> Field r"
+ using WELL REL well_order_on_domain by auto
+ {assume **: "b \<notin> A"
+ hence "b \<in> AboveS A"
+ using OF * ofilter_AboveS_Field by auto
+ hence "(suc A, b) \<in> r"
+ using suc_least_AboveS by auto
+ hence False using REL DIFF ANTISYM *
+ by (auto simp add: antisym_def)
+ }
+ thus ?thesis by blast
+qed
+
+
+
+end (* context wo_rel *)
+
+
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Ordinals_and_Cardinals/document/intro.tex Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,123 @@
+\newcommand{\eqo}{\mbox{$=\!\!o$}}
+\newcommand{\leqo}{\mbox{$\leq\!\!o$}}
+\newcommand{\lesso}{\mbox{$<\!\!o$}}
+
+
+\begin{abstract}
+We develop a basic theory of ordinals and cardinals in Isabelle/HOL, up to the
+point where some cardinality facts relevant for the ``working mathematician" become available.
+Unlike in set theory, here we do not have at hand canonical notions of ordinal and cardinal.
+Therefore, here an ordinal is merely a well-order relation and a cardinal is an
+ordinal minim w.r.t. order embedding on its field.
+\end{abstract}
+
+
+
+\section{Introduction}
+
+In set theory (under formalizations such as Zermelo-Fraenkel or Von Neumann-Bernays-G\"{o}del), an
+{\em ordinal} is a special kind of well-order, namely one
+whose strict version is the restriction of the membership relation to a set. In particular,
+the field of a set-theoretic ordinal is a transitive set, and the non-strict version
+of an ordinal relation is set inclusion. Set-theoretic ordinals enjoy the nice properties
+of membership on transitive sets, while at the same time forming a complete class of
+representatives for well-orders (since any well-order turns out isomorphic to an ordinal).
+Moreover, the class of ordinals is itself transitive and well-ordered by membership as the strict relation
+and inclusion as the non-strict relation.
+Also knowing that any set can be well-ordered (in the presence of the axiom of choice), one then defines
+the {\em cardinal} of a set to be the smallest ordinal isomorphic to a well-order on that set.
+This makes the class of cardinals a complete set of representatives for the intuitive notion
+of set cardinality.\footnote{The ``intuitive" cardinality of a set $A$ is the class of all
+sets equipollent to $A$, i.e., being in bijection with $A$.}
+The ability to produce {\em canonical well-orders} from the membership relation (having the aforementioned
+convenient properties)
+allows for a harmonious development of the theory of cardinals in set-theoretic settings.
+Non-trivial cardinality results, such as $A$ being equipollent to $A \times A$ for any infinite $A$,
+follow rather quickly within this theory.
+
+However, a canonical notion of well-order is {\em not} available in HOL.
+Here, one has to do with well-order ``as is", but otherwise has all the necessary infrastructure
+(including Hilbert choice) to ``climb" well-orders recursively and to well-oder arbitrary sets.
+
+The current work, formalized in Isabelle/HOL, develops the basic theory of ordinals and cardinals
+up to the point where there are inferred a collection of non-trivial cardinality facts useful
+for the ``working mathematician", among which:
+\begin{itemize}
+\item Given any two sets (on any two given types)\footnote{Recall that, in HOL, a set
+on a type $\alpha$ is modeled, just like a predicate, as a function from $\alpha$ to \textsf{bool}.},
+one is injectable in the other.
+\item If at least one of two sets is infinite, then their sum and their Cartesian product are equipollent
+to the larger of the two.
+\item The set of lists (and also the set of finite sets) with element from an
+infinite set is equipollent with that set.
+\end{itemize}
+
+Our development emulates the standard one from set-theory, but keeps everything
+{\em up to order isomorphism}.
+An (HOL) ordinal is merely a well-order. An {\em ordinal embedding} is an
+injective and order-compatible function which maps its source into an initial segment
+(i.e., order filter) of its target. Now, a {\em cardinal} (called in this work a {\em cardinal order})
+is an ordinal minim w.r.t. the existence of embeddings among all
+well-orders on its field. After showing the existence of cardinals on any given set,
+we define the cardinal of a set $A$, denoted $|A|$, to be {\em some} cardinal order
+on $A$. This concept is unique only up to order isomorphism (denoted $\eqo$), but meets
+its purpose: any two sets $A$ and $B$ (laying at potentially distinct types)
+are in bijection if and only if $|A|\;\eqo\;|B|$. Moreover, we also show that numeric cardinals
+assigned to finite sets\footnote{Numeric cardinals of finite sets are already formalized in
+Isabelle/HOL.}
+are {\em conservatively extended} by our general (order-theoretic) notion of
+cardinal. We study the interaction of cardinals with standard set-theoretic
+constructions such as powersets, products, sums and lists. These constructions are shown
+to preserve the ``cardinal identity" $\eqo$ and also to be monotonic w.r.t. $\leqo$, the ordinal
+embedding relation. By studying the interaction between these constructions, infinite sets and
+cardinals, we obtain the
+aforementioned results for ``working mathematicians".
+
+For this development, we did not follow closely any particular textbook, and in fact are not
+aware of such basic theory of cardinals previously
+developed in HOL.\footnote{After writing this formalization, we became aware of
+Paul Taylor's membership-free development of the theory of ordinals \cite{taylor-ordinals}.} On
+the other hand,
+the set-theoretic versions of the facts proved here are folklore in set theory, and can be found,
+e.g., in the textbook \cite{card-book}. Beyond taking care of some locality aspects
+concerning the spreading of our concepts throughout types, we have not departed
+much from the techniques used in set theory for establishing these facts -- for instance,
+in the proof of one of our major theorems,
+\textit{Card-order-Times-same-infinite} from Section 8.4,\footnote{This theorem states that, for any
+infinite cardinal $r$ on a set $A$, $|A\times A|$ is not larger than $r$.}
+we have essentially applied the technique described, e.g., in the proof of
+theorem 1.5.11 from \cite{card-book}, page 47.
+
+Here is the structure of the rest of this document.
+
+The next three sections, 2-4, develop some
+mathematical prerequisites.
+In Section 2, a large collection of simple facts about
+injections, bijections, inverses, (in)finite sets and numeric cardinals are proved,
+making life easier
+for later, when proving less trivial facts.
+Section 3 introduces upper and lower
+bounds operators for order-like relations and studies their basic properties.
+Section 4 states some useful variations of well-founded recursion and induction principles.
+
+Then come the major sections, 5-8.
+Section 5 defines and studies, in the context of a well-order relation,
+the notions of minimum (of a set), maximum (of two elements), supremum, successor (of a set),
+and order filter (i.e., initial segment, i.e., downward-closed set).
+Section 6 defines and studies (well-order) embeddings, strict embeddings, isomorphisms, and
+compatible functions.
+Section 7 deals with various constructions on well-orders, and with the relations
+$\leqo$, $\lesso$ and $\eqo$ of well-order embedding, strict embedding, and isomorphism, respectively.
+Section 8 defines and studies cardinal order relations, the cardinal of a set, the connection
+of cardinals with set-theoretic constructs,
+the canonical cardinal of natural numbers and finite cardinals, the successor
+of a cardinal, as well as regular cardinals. (The latter play a crucial role in the development of
+a new (co)datatype package in HOL.)
+
+Finally, section 9 provides an abstraction of the previous developments on
+cardinals, to provide a simpler user interface to cardinals, which in most of
+the cases allows to forget that cardinals are represented by orders and use them
+through defined arithmetic operators.
+
+More informal details are provided at the beginning of each section, and also at the
+beginning of some of the subsections.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Ordinals_and_Cardinals/document/root.bib Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,17 @@
+ @book{card-book,
+ title = {Introduction to {C}ardinal {A}rithmetic},
+ author = {M. Holz and K. Steffens and E. Weitz},
+ publisher = "Birkh{\"{a}}user",
+ year = 1999,
+ }
+
+ @article{taylor-ordinals,
+ author = {Paul Taylor},
+ title = {Intuitionistic Sets and Ordinals},
+ journal = {J. Symb. Log.},
+ volume = {61},
+ number = {3},
+ year = {1996},
+ pages = {705-744},
+ bibsource = {DBLP, http://dblp.uni-trier.de}
+}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Ordinals_and_Cardinals/document/root.tex Tue Aug 28 17:53:08 2012 +0200
@@ -0,0 +1,68 @@
+\documentclass[11pt,a4paper]{article}
+\usepackage{isabelle,isabellesym}
+
+% further packages required for unusual symbols (see also
+% isabellesym.sty), use only when needed
+
+%\usepackage{amssymb}
+ %for \<leadsto>, \<box>, \<diamond>, \<sqsupset>, \<mho>, \<Join>,
+ %\<lhd>, \<lesssim>, \<greatersim>, \<lessapprox>, \<greaterapprox>,
+ %\<triangleq>, \<yen>, \<lozenge>
+
+%\usepackage[greek,english]{babel}
+ %option greek for \<euro>
+ %option english (default language) for \<guillemotleft>, \<guillemotright>
+
+%\usepackage[latin1]{inputenc}
+ %for \<onesuperior>, \<onequarter>, \<twosuperior>, \<onehalf>,
+ %\<threesuperior>, \<threequarters>, \<degree>
+
+%\usepackage[only,bigsqcap]{stmaryrd}
+ %for \<Sqinter>
+
+%\usepackage{eufrak}
+ %for \<AA> ... \<ZZ>, \<aa> ... \<zz> (also included in amssymb)
+
+%\usepackage{textcomp}
+ %for \<cent>, \<currency>
+
+% this should be the last package used
+\usepackage{pdfsetup}
+
+% urls in roman style, theory text in math-similar italics
+\urlstyle{rm}
+\isabellestyle{it}
+
+% for uniform font size
+%\renewcommand{\isastyle}{\isastyleminor}
+
+\bibliographystyle{plain}
+\begin{document}
+
+\title{Ordinals and cardinals in HOL}
+\author{Andrei Popescu \& Dmitriy Traytel}
+\date{}
+\maketitle
+
+\tableofcontents
+
+% sane default for proof documents
+\parindent 0pt\parskip 0.5ex
+
+\input{intro.tex}
+
+\bibliography{root}
+
+% generated text of all theories
+\input{session}
+
+% optional bibliography
+%\bibliographystyle{abbrv}
+%\bibliography{root}
+
+\end{document}
+
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: t
+%%% End:
--- a/src/HOL/ROOT Tue Aug 28 17:51:43 2012 +0200
+++ b/src/HOL/ROOT Tue Aug 28 17:53:08 2012 +0200
@@ -597,6 +597,41 @@
theories Nominal_Examples
theories [quick_and_dirty] VC_Condition
+session "HOL-Ordinals_and_Cardinals-Base" in Ordinals_and_Cardinals = HOL +
+ description {* Ordinals and Cardinals, Base Theories *}
+ options [document = false]
+ theories Cardinal_Arithmetic
+
+session "HOL-Ordinals_and_Cardinals" in Ordinals_and_Cardinals =
+ "HOL-Ordinals_and_Cardinals-Base" +
+ description {* Ordinals and Cardinals, Full Theories *}
+ options [document = pdf, document_output = "."]
+ theories Cardinal_Order_Relation
+ files "document/intro.tex" "document/root.tex" "document/root.bib"
+
+session "HOL-Codatatype" in Codatatype = "HOL-Ordinals_and_Cardinals-Base" +
+ description {* New (Co)datatype Package *}
+ options [document = false]
+ theories Codatatype
+
+session "HOL-Codatatype-Examples" in "Codatatype/Examples" = "HOL-Codatatype" +
+ description {* Examples for the New (Co)datatype Package *}
+ options [document = false]
+ theories
+ HFset
+ Lambda_Term
+ Process
+ TreeFsetI
+ (* FIXME: shouldn't require "parallel_proofs = 0";
+ with parallel proofs enabled, the build of this session takes 10 times longer *)
+ theories [parallel_proofs = 0]
+ "Infinite_Derivation_Trees/Gram_Lang"
+ "Infinite_Derivation_Trees/Parallel"
+ Stream
+ theories [parallel_proofs = 0, condition = ISABELLE_FULL_TEST]
+ Misc_Codata
+ Misc_Data
+
session "HOL-Word" in Word = HOL +
options [document_graph]
theories Word