--- a/NEWS Wed Nov 20 16:15:54 2013 +0100
+++ b/NEWS Wed Nov 20 16:43:09 2013 +0100
@@ -18,6 +18,14 @@
even_zero_(nat|int) ~> even_zero
INCOMPATIBILITY.
+* Abolished neg_numeral.
+ * Canonical representation for minus one is "- 1".
+ * Canonical representation for other negative numbers is "- (numeral _)".
+ * When devising rule sets for number calculation, consider the
+ following cases: 0, 1, numeral _, - 1, - numeral _.
+ * Syntax for negative numerals is mere input syntax.
+INCOMPATBILITY.
+
* Elimination of fact duplicates:
equals_zero_I ~> minus_unique
diff_eq_0_iff_eq ~> right_minus_eq
--- a/src/Doc/Datatypes/Datatypes.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/Doc/Datatypes/Datatypes.thy Wed Nov 20 16:43:09 2013 +0100
@@ -350,7 +350,7 @@
custom names. In the example below, the familiar names @{text null}, @{text hd},
@{text tl}, @{text set}, @{text map}, and @{text list_all2}, override the
default names @{text is_Nil}, @{text un_Cons1}, @{text un_Cons2},
-@{text list_set}, @{text list_map}, and @{text list_rel}:
+@{text set_list}, @{text map_list}, and @{text rel_list}:
*}
(*<*)
@@ -363,7 +363,7 @@
Cons (infixr "#" 65)
hide_type list
- hide_const Nil Cons hd tl set map list_all2 list_case list_rec
+ hide_const Nil Cons hd tl set map list_all2
context early begin
(*>*)
@@ -501,7 +501,7 @@
reference manual \cite{isabelle-isar-ref}.
The optional names preceding the type variables allow to override the default
-names of the set functions (@{text t_set1}, \ldots, @{text t_setM}).
+names of the set functions (@{text set1_t}, \ldots, @{text setM_t}).
Inside a mutually recursive specification, all defined datatypes must
mention exactly the same type variables in the same order.
@@ -626,7 +626,7 @@
\begin{itemize}
\setlength{\itemsep}{0pt}
-\item \relax{Case combinator}: @{text t_case} (rendered using the familiar
+\item \relax{Case combinator}: @{text t.case_t} (rendered using the familiar
@{text case}--@{text of} syntax)
\item \relax{Discriminators}: @{text "t.is_C\<^sub>1"}, \ldots,
@@ -638,22 +638,22 @@
\phantom{\relax{Selectors:}} @{text t.un_C\<^sub>n1}$, \ldots, @{text t.un_C\<^sub>nk\<^sub>n}.
\item \relax{Set functions} (or \relax{natural transformations}):
-@{text t_set1}, \ldots, @{text t_setm}
-
-\item \relax{Map function} (or \relax{functorial action}): @{text t_map}
-
-\item \relax{Relator}: @{text t_rel}
-
-\item \relax{Iterator}: @{text t_fold}
-
-\item \relax{Recursor}: @{text t_rec}
+@{text set1_t}, \ldots, @{text t.setm_t}
+
+\item \relax{Map function} (or \relax{functorial action}): @{text t.map_t}
+
+\item \relax{Relator}: @{text t.rel_t}
+
+\item \relax{Iterator}: @{text t.fold_t}
+
+\item \relax{Recursor}: @{text t.rec_t}
\end{itemize}
\noindent
The case combinator, discriminators, and selectors are collectively called
\emph{destructors}. The prefix ``@{text "t."}'' is an optional component of the
-name and is normally hidden.
+names and is normally hidden.
*}
@@ -810,8 +810,8 @@
\item[@{text "t."}\hthm{sel\_split\_asm}\rm:] ~ \\
@{thm list.sel_split_asm[no_vars]}
-\item[@{text "t."}\hthm{case\_if}\rm:] ~ \\
-@{thm list.case_if[no_vars]}
+\item[@{text "t."}\hthm{case\_eq\_if}\rm:] ~ \\
+@{thm list.case_eq_if[no_vars]}
\end{description}
\end{indentblock}
@@ -914,7 +914,10 @@
is recommended to use @{command datatype_new_compat} or \keyw{rep\_datatype}
to register new-style datatypes as old-style datatypes.
-\item \emph{The recursor @{text "t_rec"} has a different signature for nested
+\item \emph{The constants @{text "t_case"} and @{text "t_rec"} are now called
+@{text "case_t"} and @{text "rec_t"}.
+
+\item \emph{The recursor @{text "rec_t"} has a different signature for nested
recursive datatypes.} In the old package, nested recursion through non-functions
was internally reduced to mutual recursion. This reduction was visible in the
type of the recursor, used by \keyw{primrec}. Recursion through functions was
@@ -1150,13 +1153,13 @@
\noindent
The next example features recursion through the @{text option} type. Although
@{text option} is not a new-style datatype, it is registered as a BNF with the
-map function @{const option_map}:
+map function @{const map_option}:
*}
primrec_new (*<*)(in early) (*>*)sum_btree :: "('a\<Colon>{zero,plus}) btree \<Rightarrow> 'a" where
"sum_btree (BNode a lt rt) =
- a + the_default 0 (option_map sum_btree lt) +
- the_default 0 (option_map sum_btree rt)"
+ a + the_default 0 (map_option sum_btree lt) +
+ the_default 0 (map_option sum_btree rt)"
text {*
\noindent
@@ -1552,9 +1555,9 @@
\begin{itemize}
\setlength{\itemsep}{0pt}
-\item \relax{Coiterator}: @{text t_unfold}
-
-\item \relax{Corecursor}: @{text t_corec}
+\item \relax{Coiterator}: @{text unfold_t}
+
+\item \relax{Corecursor}: @{text corec_t}
\end{itemize}
*}
--- a/src/Doc/ProgProve/Bool_nat_list.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/Doc/ProgProve/Bool_nat_list.thy Wed Nov 20 16:43:09 2013 +0100
@@ -99,10 +99,10 @@
For example, given the goal @{text"x + 0 = x"}, there is nothing to indicate
that you are talking about natural numbers. Hence Isabelle can only infer
that @{term x} is of some arbitrary type where @{text 0} and @{text"+"}
- exist. As a consequence, you will be unable to prove the
- goal. To alert you to such pitfalls, Isabelle flags numerals without a
- fixed type in its output: @{prop"x+0 = x"}. In this particular example,
- you need to include
+ exist. As a consequence, you will be unable to prove the goal.
+% To alert you to such pitfalls, Isabelle flags numerals without a
+% fixed type in its output: @ {prop"x+0 = x"}.
+ In this particular example, you need to include
an explicit type constraint, for example @{text"x+0 = (x::nat)"}. If there
is enough contextual information this may not be necessary: @{prop"Suc x =
x"} automatically implies @{text"x::nat"} because @{term Suc} is not
--- a/src/Doc/ProgProve/document/intro-isabelle.tex Wed Nov 20 16:15:54 2013 +0100
+++ b/src/Doc/ProgProve/document/intro-isabelle.tex Wed Nov 20 16:43:09 2013 +0100
@@ -16,7 +16,7 @@
of recursive functions.
\ifsem
\autoref{sec:CaseStudyExp} contains a
-little case study: arithmetic and boolean expressions, their evaluation,
+small case study: arithmetic and boolean expressions, their evaluation,
optimization and compilation.
\fi
\autoref{ch:Logic} introduces the rest of HOL: the
@@ -35,8 +35,8 @@
% in the intersection of computation and logic.
This introduction to the core of Isabelle is intentionally concrete and
-example-based: we concentrate on examples that illustrate the typical cases;
-we do not explain the general case if it can be inferred from the examples.
+example-based: we concentrate on examples that illustrate the typical cases
+without explaining the general case if it can be inferred from the examples.
We cover the essentials (from a functional programming point of view) as
quickly and compactly as possible.
\ifsem
@@ -46,7 +46,7 @@
For a comprehensive treatment of all things Isabelle we recommend the
\emph{Isabelle/Isar Reference Manual}~\cite{IsarRef}, which comes with the
Isabelle distribution.
-The tutorial by Nipkow, Paulson and Wenzel~\cite{LNCS2283} (in its updated version that comes with the Isabelle distribution) is still recommended for the wealth of examples and material, but its proof style is outdated. In particular it fails to cover the structured proof language Isar.
+The tutorial by Nipkow, Paulson and Wenzel~\cite{LNCS2283} (in its updated version that comes with the Isabelle distribution) is still recommended for the wealth of examples and material, but its proof style is outdated. In particular it does not cover the structured proof language Isar.
%This introduction to Isabelle has grown out of many years of teaching
%Isabelle courses.
--- a/src/HOL/Archimedean_Field.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Archimedean_Field.thy Wed Nov 20 16:43:09 2013 +0100
@@ -204,8 +204,8 @@
lemma floor_numeral [simp]: "floor (numeral v) = numeral v"
using floor_of_int [of "numeral v"] by simp
-lemma floor_neg_numeral [simp]: "floor (neg_numeral v) = neg_numeral v"
- using floor_of_int [of "neg_numeral v"] by simp
+lemma floor_neg_numeral [simp]: "floor (- numeral v) = - numeral v"
+ using floor_of_int [of "- numeral v"] by simp
lemma zero_le_floor [simp]: "0 \<le> floor x \<longleftrightarrow> 0 \<le> x"
by (simp add: le_floor_iff)
@@ -218,7 +218,7 @@
by (simp add: le_floor_iff)
lemma neg_numeral_le_floor [simp]:
- "neg_numeral v \<le> floor x \<longleftrightarrow> neg_numeral v \<le> x"
+ "- numeral v \<le> floor x \<longleftrightarrow> - numeral v \<le> x"
by (simp add: le_floor_iff)
lemma zero_less_floor [simp]: "0 < floor x \<longleftrightarrow> 1 \<le> x"
@@ -232,7 +232,7 @@
by (simp add: less_floor_iff)
lemma neg_numeral_less_floor [simp]:
- "neg_numeral v < floor x \<longleftrightarrow> neg_numeral v + 1 \<le> x"
+ "- numeral v < floor x \<longleftrightarrow> - numeral v + 1 \<le> x"
by (simp add: less_floor_iff)
lemma floor_le_zero [simp]: "floor x \<le> 0 \<longleftrightarrow> x < 1"
@@ -246,7 +246,7 @@
by (simp add: floor_le_iff)
lemma floor_le_neg_numeral [simp]:
- "floor x \<le> neg_numeral v \<longleftrightarrow> x < neg_numeral v + 1"
+ "floor x \<le> - numeral v \<longleftrightarrow> x < - numeral v + 1"
by (simp add: floor_le_iff)
lemma floor_less_zero [simp]: "floor x < 0 \<longleftrightarrow> x < 0"
@@ -260,7 +260,7 @@
by (simp add: floor_less_iff)
lemma floor_less_neg_numeral [simp]:
- "floor x < neg_numeral v \<longleftrightarrow> x < neg_numeral v"
+ "floor x < - numeral v \<longleftrightarrow> x < - numeral v"
by (simp add: floor_less_iff)
text {* Addition and subtraction of integers *}
@@ -272,10 +272,6 @@
"floor (x + numeral v) = floor x + numeral v"
using floor_add_of_int [of x "numeral v"] by simp
-lemma floor_add_neg_numeral [simp]:
- "floor (x + neg_numeral v) = floor x + neg_numeral v"
- using floor_add_of_int [of x "neg_numeral v"] by simp
-
lemma floor_add_one [simp]: "floor (x + 1) = floor x + 1"
using floor_add_of_int [of x 1] by simp
@@ -286,10 +282,6 @@
"floor (x - numeral v) = floor x - numeral v"
using floor_diff_of_int [of x "numeral v"] by simp
-lemma floor_diff_neg_numeral [simp]:
- "floor (x - neg_numeral v) = floor x - neg_numeral v"
- using floor_diff_of_int [of x "neg_numeral v"] by simp
-
lemma floor_diff_one [simp]: "floor (x - 1) = floor x - 1"
using floor_diff_of_int [of x 1] by simp
@@ -353,8 +345,8 @@
lemma ceiling_numeral [simp]: "ceiling (numeral v) = numeral v"
using ceiling_of_int [of "numeral v"] by simp
-lemma ceiling_neg_numeral [simp]: "ceiling (neg_numeral v) = neg_numeral v"
- using ceiling_of_int [of "neg_numeral v"] by simp
+lemma ceiling_neg_numeral [simp]: "ceiling (- numeral v) = - numeral v"
+ using ceiling_of_int [of "- numeral v"] by simp
lemma ceiling_le_zero [simp]: "ceiling x \<le> 0 \<longleftrightarrow> x \<le> 0"
by (simp add: ceiling_le_iff)
@@ -367,7 +359,7 @@
by (simp add: ceiling_le_iff)
lemma ceiling_le_neg_numeral [simp]:
- "ceiling x \<le> neg_numeral v \<longleftrightarrow> x \<le> neg_numeral v"
+ "ceiling x \<le> - numeral v \<longleftrightarrow> x \<le> - numeral v"
by (simp add: ceiling_le_iff)
lemma ceiling_less_zero [simp]: "ceiling x < 0 \<longleftrightarrow> x \<le> -1"
@@ -381,7 +373,7 @@
by (simp add: ceiling_less_iff)
lemma ceiling_less_neg_numeral [simp]:
- "ceiling x < neg_numeral v \<longleftrightarrow> x \<le> neg_numeral v - 1"
+ "ceiling x < - numeral v \<longleftrightarrow> x \<le> - numeral v - 1"
by (simp add: ceiling_less_iff)
lemma zero_le_ceiling [simp]: "0 \<le> ceiling x \<longleftrightarrow> -1 < x"
@@ -395,7 +387,7 @@
by (simp add: le_ceiling_iff)
lemma neg_numeral_le_ceiling [simp]:
- "neg_numeral v \<le> ceiling x \<longleftrightarrow> neg_numeral v - 1 < x"
+ "- numeral v \<le> ceiling x \<longleftrightarrow> - numeral v - 1 < x"
by (simp add: le_ceiling_iff)
lemma zero_less_ceiling [simp]: "0 < ceiling x \<longleftrightarrow> 0 < x"
@@ -409,7 +401,7 @@
by (simp add: less_ceiling_iff)
lemma neg_numeral_less_ceiling [simp]:
- "neg_numeral v < ceiling x \<longleftrightarrow> neg_numeral v < x"
+ "- numeral v < ceiling x \<longleftrightarrow> - numeral v < x"
by (simp add: less_ceiling_iff)
text {* Addition and subtraction of integers *}
@@ -421,10 +413,6 @@
"ceiling (x + numeral v) = ceiling x + numeral v"
using ceiling_add_of_int [of x "numeral v"] by simp
-lemma ceiling_add_neg_numeral [simp]:
- "ceiling (x + neg_numeral v) = ceiling x + neg_numeral v"
- using ceiling_add_of_int [of x "neg_numeral v"] by simp
-
lemma ceiling_add_one [simp]: "ceiling (x + 1) = ceiling x + 1"
using ceiling_add_of_int [of x 1] by simp
@@ -435,10 +423,6 @@
"ceiling (x - numeral v) = ceiling x - numeral v"
using ceiling_diff_of_int [of x "numeral v"] by simp
-lemma ceiling_diff_neg_numeral [simp]:
- "ceiling (x - neg_numeral v) = ceiling x - neg_numeral v"
- using ceiling_diff_of_int [of x "neg_numeral v"] by simp
-
lemma ceiling_diff_one [simp]: "ceiling (x - 1) = ceiling x - 1"
using ceiling_diff_of_int [of x 1] by simp
--- a/src/HOL/BNF/BNF_Comp.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/BNF/BNF_Comp.thy Wed Nov 20 16:43:09 2013 +0100
@@ -11,6 +11,9 @@
imports Basic_BNFs
begin
+lemma wpull_id: "wpull UNIV B1 B2 id id id id"
+unfolding wpull_def by simp
+
lemma empty_natural: "(\<lambda>_. {}) o f = image g o (\<lambda>_. {})"
by (rule ext) simp
--- a/src/HOL/BNF/BNF_Def.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/BNF/BNF_Def.thy Wed Nov 20 16:43:09 2013 +0100
@@ -190,9 +190,6 @@
lemma vimage2pI: "R (f x) (g y) \<Longrightarrow> vimage2p f g R x y"
unfolding vimage2p_def by -
-lemma vimage2pD: "vimage2p f g R x y \<Longrightarrow> R (f x) (g y)"
- unfolding vimage2p_def by -
-
lemma fun_rel_iff_leq_vimage2p: "(fun_rel R S) f g = (R \<le> vimage2p f g S)"
unfolding fun_rel_def vimage2p_def by auto
--- a/src/HOL/BNF/BNF_FP_Base.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/BNF/BNF_FP_Base.thy Wed Nov 20 16:43:09 2013 +0100
@@ -13,12 +13,6 @@
imports BNF_Comp Ctr_Sugar
begin
-lemma not_TrueE: "\<not> True \<Longrightarrow> P"
-by (erule notE, rule TrueI)
-
-lemma neq_eq_eq_contradict: "\<lbrakk>t \<noteq> u; s = t; s = u\<rbrakk> \<Longrightarrow> P"
-by fast
-
lemma mp_conj: "(P \<longrightarrow> Q) \<and> R \<Longrightarrow> P \<Longrightarrow> R \<and> Q"
by auto
--- a/src/HOL/BNF/BNF_GFP.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/BNF/BNF_GFP.thy Wed Nov 20 16:43:09 2013 +0100
@@ -15,14 +15,22 @@
"primcorec" :: thy_decl
begin
+lemma not_TrueE: "\<not> True \<Longrightarrow> P"
+by (erule notE, rule TrueI)
+
+lemma neq_eq_eq_contradict: "\<lbrakk>t \<noteq> u; s = t; s = u\<rbrakk> \<Longrightarrow> P"
+by fast
+
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 sum_case_expand_Inr': "f o Inl = g \<Longrightarrow> h = f o Inr \<longleftrightarrow> sum_case g h = f"
-by (metis sum_case_o_inj(1,2) surjective_sum)
+apply rule
+ apply (rule ext, force split: sum.split)
+by (rule ext, metis sum_case_o_inj(2))
lemma converse_Times: "(A \<times> B) ^-1 = B \<times> A"
-by auto
+by fast
lemma equiv_proj:
assumes e: "equiv A R" and "z \<in> R"
@@ -37,7 +45,6 @@
(* Operators: *)
definition image2 where "image2 A f g = {(f a, g a) | a. a \<in> A}"
-
lemma Id_onD: "(a, b) \<in> Id_on A \<Longrightarrow> a = b"
unfolding Id_on_def by simp
@@ -56,9 +63,6 @@
lemma Id_on_Gr: "Id_on A = Gr A id"
unfolding Id_on_def Gr_def by auto
-lemma Id_on_UNIV_I: "x = y \<Longrightarrow> (x, y) \<in> Id_on UNIV"
-unfolding Id_on_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
@@ -77,6 +81,12 @@
lemma Gr_incl: "Gr A f \<subseteq> A <*> B \<longleftrightarrow> f ` A \<subseteq> B"
unfolding Gr_def by auto
+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 in_rel_Collect_split_eq: "in_rel (Collect (split X)) = X"
unfolding fun_eq_iff by auto
@@ -130,9 +140,6 @@
"R \<subseteq> relInvImage UNIV (relImage R f) f"
unfolding relInvImage_def relImage_def by auto
-lemma equiv_Image: "equiv A R \<Longrightarrow> (\<And>a b. (a, b) \<in> R \<Longrightarrow> a \<in> A \<and> b \<in> A \<and> R `` {a} = R `` {b})"
-unfolding equiv_def refl_on_def Image_def by (auto intro: transD symD)
-
lemma relImage_proj:
assumes "equiv A R"
shows "relImage R (proj R) \<subseteq> Id_on (A//R)"
@@ -143,7 +150,7 @@
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
+using assms unfolding relImage_def relInvImage_def by fast
lemma subst_Pair: "P x y \<Longrightarrow> a = (x, y) \<Longrightarrow> P (fst a) (snd a)"
by simp
@@ -255,13 +262,18 @@
shows "\<exists> a. a \<in> A \<and> p1 a = b1 \<and> p2 a = b2"
using assms unfolding wpull_def by blast
-lemma pickWP:
+lemma pickWP_raw:
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"
-unfolding pickWP_def using assms someI_ex[OF pickWP_pred] by fastforce+
+shows "pickWP A p1 p2 b1 b2 \<in> A
+ \<and> p1 (pickWP A p1 p2 b1 b2) = b1
+ \<and> p2 (pickWP A p1 p2 b1 b2) = b2"
+unfolding pickWP_def using assms someI_ex[OF pickWP_pred] by fastforce
+
+lemmas pickWP =
+ pickWP_raw[THEN conjunct1]
+ pickWP_raw[THEN conjunct2, THEN conjunct1]
+ pickWP_raw[THEN conjunct2, THEN conjunct2]
lemma Inl_Field_csum: "a \<in> Field r \<Longrightarrow> Inl a \<in> Field (r +c s)"
unfolding Field_card_of csum_def by auto
@@ -293,18 +305,12 @@
lemma image2pI: "R x y \<Longrightarrow> (image2p f g R) (f x) (g y)"
unfolding image2p_def by blast
-lemma image2p_eqI: "\<lbrakk>fx = f x; gy = g y; R x y\<rbrakk> \<Longrightarrow> (image2p f g R) fx gy"
- unfolding image2p_def by blast
-
lemma image2pE: "\<lbrakk>(image2p f g R) fx gy; (\<And>x y. fx = f x \<Longrightarrow> gy = g y \<Longrightarrow> R x y \<Longrightarrow> P)\<rbrakk> \<Longrightarrow> P"
unfolding image2p_def by blast
lemma fun_rel_iff_geq_image2p: "(fun_rel R S) f g = (image2p f g R \<le> S)"
unfolding fun_rel_def image2p_def by auto
-lemma convol_image_image2p: "<f o fst, g o snd> ` Collect (split R) \<subseteq> Collect (split (image2p f g R))"
- unfolding convol_def image2p_def by fastforce
-
lemma fun_rel_image2p: "(fun_rel R (image2p f g R)) f g"
unfolding fun_rel_def image2p_def by auto
--- a/src/HOL/BNF/BNF_Util.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/BNF/BNF_Util.thy Wed Nov 20 16:43:09 2013 +0100
@@ -9,15 +9,9 @@
header {* Library for Bounded Natural Functors *}
theory BNF_Util
-imports "../Cardinals/Cardinal_Arithmetic"
+imports "../Cardinals/Cardinal_Arithmetic_FP"
begin
-lemma subset_Collect_iff: "B \<subseteq> A \<Longrightarrow> (B \<subseteq> {x \<in> A. P x}) = (\<forall>x \<in> B. P x)"
-by blast
-
-lemma subset_CollectI: "B \<subseteq> A \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> Q x \<Longrightarrow> P x) \<Longrightarrow> ({x \<in> B. Q x} \<subseteq> {x \<in> A. P x})"
-by blast
-
definition collect where
"collect F x = (\<Union>f \<in> F. f x)"
--- a/src/HOL/BNF/Basic_BNFs.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/BNF/Basic_BNFs.thy Wed Nov 20 16:43:09 2013 +0100
@@ -13,14 +13,8 @@
imports BNF_Def
begin
-lemma wpull_id: "wpull UNIV B1 B2 id id id id"
-unfolding wpull_def by simp
-
lemmas natLeq_card_order = natLeq_Card_order[unfolded Field_natLeq]
-lemma ctwo_card_order: "card_order ctwo"
-using Card_order_ctwo by (unfold ctwo_def Field_card_of)
-
lemma natLeq_cinfinite: "cinfinite natLeq"
unfolding cinfinite_def Field_natLeq by (rule nat_infinite)
@@ -62,11 +56,11 @@
proof -
show "sum_map id id = id" by (rule sum_map.id)
next
- fix f1 f2 g1 g2
+ fix f1 :: "'o \<Rightarrow> 's" and f2 :: "'p \<Rightarrow> 't" and g1 :: "'s \<Rightarrow> 'q" and g2 :: "'t \<Rightarrow> 'r"
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
+ fix x and f1 :: "'o \<Rightarrow> 'q" and f2 :: "'p \<Rightarrow> 'r" and g1 g2
assume a1: "\<And>z. z \<in> setl x \<Longrightarrow> f1 z = g1 z" and
a2: "\<And>z. z \<in> setr x \<Longrightarrow> f2 z = g2 z"
thus "sum_map f1 f2 x = sum_map g1 g2 x"
@@ -76,11 +70,11 @@
case Inr thus ?thesis using a2 by (clarsimp simp: setr_def)
qed
next
- fix f1 f2
+ fix f1 :: "'o \<Rightarrow> 'q" and f2 :: "'p \<Rightarrow> 'r"
show "setl o sum_map f1 f2 = image f1 o setl"
by (rule ext, unfold o_apply) (simp add: setl_def split: sum.split)
next
- fix f1 f2
+ fix f1 :: "'o \<Rightarrow> 'q" and f2 :: "'p \<Rightarrow> 'r"
show "setr o sum_map f1 f2 = image f2 o setr"
by (rule ext, unfold o_apply) (simp add: setr_def split: sum.split)
next
@@ -88,13 +82,13 @@
next
show "cinfinite natLeq" by (rule natLeq_cinfinite)
next
- fix x
+ fix x :: "'o + 'p"
show "|setl x| \<le>o natLeq"
apply (rule ordLess_imp_ordLeq)
apply (rule finite_iff_ordLess_natLeq[THEN iffD1])
by (simp add: setl_def split: sum.split)
next
- fix x
+ fix x :: "'o + 'p"
show "|setr x| \<le>o natLeq"
apply (rule ordLess_imp_ordLeq)
apply (rule finite_iff_ordLess_natLeq[THEN iffD1])
@@ -229,22 +223,6 @@
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 f x else undefined"
- 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 "'a \<Rightarrow> 'b"
map: "op \<circ>"
sets: range
@@ -275,7 +253,7 @@
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)
+ 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 A B1 B2 f1 f2 p1 p2 assume p: "wpull A B1 B2 f1 f2 p1 p2"
@@ -294,7 +272,7 @@
(Grp {x. range x \<subseteq> Collect (split R)} (op \<circ> fst))\<inverse>\<inverse> OO
Grp {x. range x \<subseteq> Collect (split R)} (op \<circ> snd)"
unfolding fun_rel_def Grp_def fun_eq_iff relcompp.simps conversep.simps subset_iff image_iff
- by auto (force, metis pair_collapse)
+ by auto (force, metis (no_types) pair_collapse)
qed
end
--- a/src/HOL/BNF/Equiv_Relations_More.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/BNF/Equiv_Relations_More.thy Wed Nov 20 16:43:09 2013 +0100
@@ -59,7 +59,7 @@
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
+using quotient_eq_iff[THEN iffD1] 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"
--- a/src/HOL/BNF/Examples/Koenig.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/BNF/Examples/Koenig.thy Wed Nov 20 16:43:09 2013 +0100
@@ -12,44 +12,33 @@
imports TreeFI Stream
begin
-(* selectors for streams *)
-lemma shd_def': "shd as = fst (stream_dtor as)"
-apply (case_tac as)
-apply (auto simp add: shd_def)
-by (simp add: Stream_def stream.dtor_ctor)
-
-lemma stl_def': "stl as = snd (stream_dtor as)"
-apply (case_tac as)
-apply (auto simp add: stl_def)
-by (simp add: Stream_def stream.dtor_ctor)
-
(* infinite trees: *)
coinductive infiniteTr where
-"\<lbrakk>tr' \<in> listF_set (sub tr); infiniteTr tr'\<rbrakk> \<Longrightarrow> infiniteTr tr"
+"\<lbrakk>tr' \<in> set_listF (sub tr); infiniteTr tr'\<rbrakk> \<Longrightarrow> infiniteTr tr"
lemma infiniteTr_strong_coind[consumes 1, case_names sub]:
assumes *: "phi tr" and
-**: "\<And> tr. phi tr \<Longrightarrow> \<exists> tr' \<in> listF_set (sub tr). phi tr' \<or> infiniteTr tr'"
+**: "\<And> tr. phi tr \<Longrightarrow> \<exists> tr' \<in> set_listF (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'"
+**: "\<And> tr. phi tr \<Longrightarrow> \<exists> tr' \<in> set_listF (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')"
+"infiniteTr tr \<Longrightarrow> (\<exists> tr' \<in> set_listF (sub tr). infiniteTr tr')"
by (erule infiniteTr.cases) blast
primcorec konigPath where
"shd (konigPath t) = lab t"
-| "stl (konigPath t) = konigPath (SOME tr. tr \<in> listF_set (sub t) \<and> infiniteTr tr)"
+| "stl (konigPath t) = konigPath (SOME tr. tr \<in> set_listF (sub t) \<and> infiniteTr tr)"
(* proper paths in trees: *)
coinductive properPath where
-"\<lbrakk>shd as = lab tr; tr' \<in> listF_set (sub tr); properPath (stl as) tr'\<rbrakk> \<Longrightarrow>
+"\<lbrakk>shd as = lab tr; tr' \<in> set_listF (sub tr); properPath (stl as) tr'\<rbrakk> \<Longrightarrow>
properPath as tr"
lemma properPath_strong_coind[consumes 1, case_names shd_lab sub]:
@@ -57,7 +46,7 @@
**: "\<And> as tr. phi as tr \<Longrightarrow> shd as = lab tr" and
***: "\<And> as tr.
phi as tr \<Longrightarrow>
- \<exists> tr' \<in> listF_set (sub tr). phi (stl as) tr' \<or> properPath (stl as) tr'"
+ \<exists> tr' \<in> set_listF (sub tr). phi (stl as) tr' \<or> properPath (stl as) tr'"
shows "properPath as tr"
using assms by (elim properPath.coinduct) blast
@@ -66,7 +55,7 @@
**: "\<And> as tr. phi as tr \<Longrightarrow> shd as = lab tr" and
***: "\<And> as tr.
phi as tr \<Longrightarrow>
- \<exists> tr' \<in> listF_set (sub tr). phi (stl as) tr'"
+ \<exists> tr' \<in> set_listF (sub tr). phi (stl as) tr'"
shows "properPath as tr"
using properPath_strong_coind[of phi, OF * **] *** by blast
@@ -76,7 +65,7 @@
lemma properPath_sub:
"properPath as tr \<Longrightarrow>
- \<exists> tr' \<in> listF_set (sub tr). phi (stl as) tr' \<or> properPath (stl as) tr'"
+ \<exists> tr' \<in> set_listF (sub tr). phi (stl as) tr' \<or> properPath (stl as) tr'"
by (erule properPath.cases) blast
(* prove the following by coinduction *)
@@ -88,10 +77,10 @@
assume "infiniteTr tr \<and> as = konigPath tr" hence "properPath as tr"
proof (coinduction arbitrary: tr as rule: properPath_coind)
case (sub tr as)
- let ?t = "SOME t'. t' \<in> listF_set (sub tr) \<and> infiniteTr t'"
- from sub have "\<exists>t' \<in> listF_set (sub tr). infiniteTr t'" by simp
- then have "\<exists>t'. t' \<in> listF_set (sub tr) \<and> infiniteTr t'" by blast
- then have "?t \<in> listF_set (sub tr) \<and> infiniteTr ?t" by (rule someI_ex)
+ let ?t = "SOME t'. t' \<in> set_listF (sub tr) \<and> infiniteTr t'"
+ from sub have "\<exists>t' \<in> set_listF (sub tr). infiniteTr t'" by simp
+ then have "\<exists>t'. t' \<in> set_listF (sub tr) \<and> infiniteTr t'" by blast
+ then have "?t \<in> set_listF (sub tr) \<and> infiniteTr ?t" by (rule someI_ex)
moreover have "stl (konigPath tr) = konigPath ?t" by simp
ultimately show ?case using sub by blast
qed simp
--- a/src/HOL/BNF/Examples/ListF.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/BNF/Examples/ListF.thy Wed Nov 20 16:43:09 2013 +0100
@@ -62,7 +62,7 @@
"i < lengthh xs \<Longrightarrow> nthh (mapF 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"
+lemma nthh_listF_set[simp]: "i < lengthh xs \<Longrightarrow> nthh xs i \<in> set_listF xs"
by (induct rule: nthh.induct) auto
lemma NilF_iff[iff]: "(lengthh xs = 0) = (xs = NilF)"
@@ -105,7 +105,7 @@
qed simp
lemma list_set_nthh[simp]:
- "(x \<in> listF_set xs) \<Longrightarrow> (\<exists>i < lengthh xs. nthh xs i = x)"
+ "(x \<in> set_listF xs) \<Longrightarrow> (\<exists>i < lengthh xs. nthh xs i = x)"
by (induct xs) (auto, induct rule: nthh.induct, auto)
end
--- a/src/HOL/BNF/Examples/Process.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/BNF/Examples/Process.thy Wed Nov 20 16:43:09 2013 +0100
@@ -22,7 +22,7 @@
subsection {* Basic properties *}
declare
- pre_process_rel_def[simp]
+ rel_pre_process_def[simp]
sum_rel_def[simp]
prod_rel_def[simp]
--- a/src/HOL/BNF/Examples/Stream.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/BNF/Examples/Stream.thy Wed Nov 20 16:43:09 2013 +0100
@@ -18,7 +18,7 @@
code_datatype Stream
lemma stream_case_cert:
- assumes "CASE \<equiv> stream_case c"
+ assumes "CASE \<equiv> case_stream c"
shows "CASE (a ## s) \<equiv> c a s"
using assms by simp_all
@@ -87,10 +87,10 @@
by (induct xs) auto
-subsection {* set of streams with elements in some fixes set *}
+subsection {* set of streams with elements in some fixed set *}
coinductive_set
- streams :: "'a set => 'a stream set"
+ streams :: "'a set \<Rightarrow> 'a stream set"
for A :: "'a set"
where
Stream[intro!, simp, no_atp]: "\<lbrakk>a \<in> A; s \<in> streams A\<rbrakk> \<Longrightarrow> a ## s \<in> streams A"
@@ -98,6 +98,15 @@
lemma shift_streams: "\<lbrakk>w \<in> lists A; s \<in> streams A\<rbrakk> \<Longrightarrow> w @- s \<in> streams A"
by (induct w) auto
+lemma streams_Stream: "x ## s \<in> streams A \<longleftrightarrow> x \<in> A \<and> s \<in> streams A"
+ by (auto elim: streams.cases)
+
+lemma streams_stl: "s \<in> streams A \<Longrightarrow> stl s \<in> streams A"
+ by (cases s) (auto simp: streams_Stream)
+
+lemma streams_shd: "s \<in> streams A \<Longrightarrow> shd s \<in> A"
+ by (cases s) (auto simp: streams_Stream)
+
lemma sset_streams:
assumes "sset s \<subseteq> A"
shows "s \<in> streams A"
@@ -105,6 +114,28 @@
case streams then show ?case by (cases s) simp
qed
+lemma streams_sset:
+ assumes "s \<in> streams A"
+ shows "sset s \<subseteq> A"
+proof
+ fix x assume "x \<in> sset s" from this `s \<in> streams A` show "x \<in> A"
+ by (induct s) (auto intro: streams_shd streams_stl)
+qed
+
+lemma streams_iff_sset: "s \<in> streams A \<longleftrightarrow> sset s \<subseteq> A"
+ by (metis sset_streams streams_sset)
+
+lemma streams_mono: "s \<in> streams A \<Longrightarrow> A \<subseteq> B \<Longrightarrow> s \<in> streams B"
+ unfolding streams_iff_sset by auto
+
+lemma smap_streams: "s \<in> streams A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<in> B) \<Longrightarrow> smap f s \<in> streams B"
+ unfolding streams_iff_sset stream.set_map by auto
+
+lemma streams_empty: "streams {} = {}"
+ by (auto elim: streams.cases)
+
+lemma streams_UNIV[simp]: "streams UNIV = UNIV"
+ by (auto simp: streams_iff_sset)
subsection {* nth, take, drop for streams *}
@@ -234,6 +265,9 @@
lemma stream_all_shift[simp]: "stream_all P (xs @- s) = (list_all P xs \<and> stream_all P s)"
unfolding stream_all_iff list_all_iff by auto
+lemma stream_all_Stream: "stream_all P (x ## X) \<longleftrightarrow> P x \<and> stream_all P X"
+ by simp
+
subsection {* recurring stream out of a list *}
@@ -285,59 +319,60 @@
by (induct n arbitrary: u) (auto simp: rotate1_rotate_swap rotate1_hd_tl rotate_conv_mod[symmetric])
+subsection {* iterated application of a function *}
+
+primcorec siterate where
+ "shd (siterate f x) = x"
+| "stl (siterate f x) = siterate f (f x)"
+
+lemma stake_Suc: "stake (Suc n) s = stake n s @ [s !! n]"
+ by (induct n arbitrary: s) auto
+
+lemma snth_siterate[simp]: "siterate f x !! n = (f^^n) x"
+ by (induct n arbitrary: x) (auto simp: funpow_swap1)
+
+lemma sdrop_siterate[simp]: "sdrop n (siterate f x) = siterate f ((f^^n) x)"
+ by (induct n arbitrary: x) (auto simp: funpow_swap1)
+
+lemma stake_siterate[simp]: "stake n (siterate f x) = map (\<lambda>n. (f^^n) x) [0 ..< n]"
+ by (induct n arbitrary: x) (auto simp del: stake.simps(2) simp: stake_Suc)
+
+lemma sset_siterate: "sset (siterate f x) = {(f^^n) x | n. True}"
+ by (auto simp: sset_range)
+
+lemma smap_siterate: "smap f (siterate f x) = siterate f (f x)"
+ by (coinduction arbitrary: x) auto
+
+
subsection {* stream repeating a single element *}
-primcorec same where
- "shd (same x) = x"
-| "stl (same x) = same x"
+abbreviation "sconst \<equiv> siterate id"
-lemma snth_same[simp]: "same x !! n = x"
- unfolding same_def by (induct n) auto
+lemma shift_replicate_sconst[simp]: "replicate n x @- sconst x = sconst x"
+ by (subst (3) stake_sdrop[symmetric]) (simp add: map_replicate_trivial)
-lemma stake_same[simp]: "stake n (same x) = replicate n x"
- unfolding same_def by (induct n) (auto simp: upt_rec)
+lemma stream_all_same[simp]: "sset (sconst x) = {x}"
+ by (simp add: sset_siterate)
-lemma sdrop_same[simp]: "sdrop n (same x) = same x"
- unfolding same_def by (induct n) auto
-
-lemma shift_replicate_same[simp]: "replicate n x @- same x = same x"
- by (metis sdrop_same stake_same stake_sdrop)
+lemma same_cycle: "sconst x = cycle [x]"
+ by coinduction auto
-lemma stream_all_same[simp]: "stream_all P (same x) \<longleftrightarrow> P x"
- unfolding stream_all_def by auto
+lemma smap_sconst: "smap f (sconst x) = sconst (f x)"
+ by coinduction auto
-lemma same_cycle: "same x = cycle [x]"
- by coinduction auto
+lemma sconst_streams: "x \<in> A \<Longrightarrow> sconst x \<in> streams A"
+ by (simp add: streams_iff_sset)
subsection {* stream of natural numbers *}
-primcorec fromN :: "nat \<Rightarrow> nat stream" where
- "fromN n = n ## fromN (n + 1)"
-
-lemma snth_fromN[simp]: "fromN n !! m = n + m"
- unfolding fromN_def by (induct m arbitrary: n) auto
-
-lemma stake_fromN[simp]: "stake m (fromN n) = [n ..< m + n]"
- unfolding fromN_def by (induct m arbitrary: n) (auto simp: upt_rec)
-
-lemma sdrop_fromN[simp]: "sdrop m (fromN n) = fromN (n + m)"
- unfolding fromN_def by (induct m arbitrary: n) auto
-
-lemma sset_fromN[simp]: "sset (fromN n) = {n ..}" (is "?L = ?R")
-proof safe
- fix m assume "m \<in> ?L"
- moreover
- { fix s assume "m \<in> sset s" "\<exists>n'\<ge>n. s = fromN n'"
- hence "n \<le> m" by (induct arbitrary: n rule: sset_induct1) fastforce+
- }
- ultimately show "n \<le> m" by auto
-next
- fix m assume "n \<le> m" thus "m \<in> ?L" by (metis le_iff_add snth_fromN snth_sset)
-qed
+abbreviation "fromN \<equiv> siterate Suc"
abbreviation "nats \<equiv> fromN 0"
+lemma sset_fromN[simp]: "sset (fromN n) = {n ..}"
+ by (auto simp add: sset_siterate) arith
+
subsection {* flatten a stream of lists *}
@@ -498,26 +533,4 @@
"smap2 f s1 s2 = smap (split f) (szip s1 s2)"
by (coinduction arbitrary: s1 s2) auto
-
-subsection {* iterated application of a function *}
-
-primcorec siterate where
- "shd (siterate f x) = x"
-| "stl (siterate f x) = siterate f (f x)"
-
-lemma stake_Suc: "stake (Suc n) s = stake n s @ [s !! n]"
- by (induct n arbitrary: s) auto
-
-lemma snth_siterate[simp]: "siterate f x !! n = (f^^n) x"
- by (induct n arbitrary: x) (auto simp: funpow_swap1)
-
-lemma sdrop_siterate[simp]: "sdrop n (siterate f x) = siterate f ((f^^n) x)"
- by (induct n arbitrary: x) (auto simp: funpow_swap1)
-
-lemma stake_siterate[simp]: "stake n (siterate f x) = map (\<lambda>n. (f^^n) x) [0 ..< n]"
- by (induct n arbitrary: x) (auto simp del: stake.simps(2) simp: stake_Suc)
-
-lemma sset_siterate: "sset (siterate f x) = {(f^^n) x | n. True}"
- by (auto simp: sset_range)
-
end
--- a/src/HOL/BNF/More_BNFs.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/BNF/More_BNFs.thy Wed Nov 20 16:43:09 2013 +0100
@@ -909,18 +909,18 @@
by (auto simp add: mmap_id0 mmap_comp set_of_mmap natLeq_card_order natLeq_cinfinite set_of_bd
intro: mmap_cong wpull_mmap)
-inductive multiset_rel' where
-Zero: "multiset_rel' R {#} {#}"
+inductive rel_multiset' where
+Zero: "rel_multiset' R {#} {#}"
|
-Plus: "\<lbrakk>R a b; multiset_rel' R M N\<rbrakk> \<Longrightarrow> multiset_rel' R (M + {#a#}) (N + {#b#})"
+Plus: "\<lbrakk>R a b; rel_multiset' R M N\<rbrakk> \<Longrightarrow> rel_multiset' R (M + {#a#}) (N + {#b#})"
-lemma multiset_map_Zero_iff[simp]: "mmap f M = {#} \<longleftrightarrow> M = {#}"
+lemma map_multiset_Zero_iff[simp]: "mmap f M = {#} \<longleftrightarrow> M = {#}"
by (metis image_is_empty multiset.set_map set_of_eq_empty_iff)
-lemma multiset_map_Zero[simp]: "mmap f {#} = {#}" by simp
+lemma map_multiset_Zero[simp]: "mmap f {#} = {#}" by simp
-lemma multiset_rel_Zero: "multiset_rel R {#} {#}"
-unfolding multiset_rel_def Grp_def by auto
+lemma rel_multiset_Zero: "rel_multiset R {#} {#}"
+unfolding rel_multiset_def Grp_def by auto
declare multiset.count[simp]
declare Abs_multiset_inverse[simp]
@@ -928,7 +928,7 @@
declare union_preserves_multiset[simp]
-lemma multiset_map_Plus[simp]: "mmap f (M1 + M2) = mmap f M1 + mmap f M2"
+lemma map_multiset_Plus[simp]: "mmap f (M1 + M2) = mmap f M1 + mmap f M2"
proof (intro multiset_eqI, transfer fixing: f)
fix x :: 'a and M1 M2 :: "'b \<Rightarrow> nat"
assume "M1 \<in> multiset" "M2 \<in> multiset"
@@ -941,12 +941,12 @@
by (auto simp: setsum.distrib[symmetric])
qed
-lemma multiset_map_singl[simp]: "mmap f {#a#} = {#f a#}"
+lemma map_multiset_singl[simp]: "mmap f {#a#} = {#f a#}"
by transfer auto
-lemma multiset_rel_Plus:
-assumes ab: "R a b" and MN: "multiset_rel R M N"
-shows "multiset_rel R (M + {#a#}) (N + {#b#})"
+lemma rel_multiset_Plus:
+assumes ab: "R a b" and MN: "rel_multiset R M N"
+shows "rel_multiset R (M + {#a#}) (N + {#b#})"
proof-
{fix y assume "R a b" and "set_of y \<subseteq> {(x, y). R x y}"
hence "\<exists>ya. mmap fst y + {#a#} = mmap fst ya \<and>
@@ -956,13 +956,13 @@
}
thus ?thesis
using assms
- unfolding multiset_rel_def Grp_def by force
+ unfolding rel_multiset_def Grp_def by force
qed
-lemma multiset_rel'_imp_multiset_rel:
-"multiset_rel' R M N \<Longrightarrow> multiset_rel R M N"
-apply(induct rule: multiset_rel'.induct)
-using multiset_rel_Zero multiset_rel_Plus by auto
+lemma rel_multiset'_imp_rel_multiset:
+"rel_multiset' R M N \<Longrightarrow> rel_multiset R M N"
+apply(induct rule: rel_multiset'.induct)
+using rel_multiset_Zero rel_multiset_Plus by auto
lemma mcard_mmap[simp]: "mcard (mmap f M) = mcard M"
proof -
@@ -973,8 +973,7 @@
using finite_Collect_mem .
ultimately have fin: "finite {b. \<exists>a. f a = b \<and> a \<in># M}" by(rule finite_subset)
have i: "inj_on A ?B" unfolding inj_on_def A_def apply clarsimp
- by (metis (lifting, mono_tags) mem_Collect_eq rel_simps(54)
- setsum_gt_0_iff setsum_infinite)
+ by (metis (lifting, full_types) mem_Collect_eq neq0_conv setsum.neutral)
have 0: "\<And> b. 0 < setsum (count M) (A b) \<longleftrightarrow> (\<exists> a \<in> A b. count M a > 0)"
apply safe
apply (metis less_not_refl setsum_gt_0_iff setsum_infinite)
@@ -995,10 +994,10 @@
then show ?thesis unfolding mcard_unfold_setsum A_def by transfer
qed
-lemma multiset_rel_mcard:
-assumes "multiset_rel R M N"
+lemma rel_multiset_mcard:
+assumes "rel_multiset R M N"
shows "mcard M = mcard N"
-using assms unfolding multiset_rel_def Grp_def by auto
+using assms unfolding rel_multiset_def Grp_def by auto
lemma multiset_induct2[case_names empty addL addR]:
assumes empty: "P {#} {#}"
@@ -1053,67 +1052,67 @@
qed
lemma msed_rel_invL:
-assumes "multiset_rel R (M + {#a#}) N"
-shows "\<exists> N1 b. N = N1 + {#b#} \<and> R a b \<and> multiset_rel R M N1"
+assumes "rel_multiset R (M + {#a#}) N"
+shows "\<exists> N1 b. N = N1 + {#b#} \<and> R a b \<and> rel_multiset R M N1"
proof-
obtain K where KM: "mmap fst K = M + {#a#}"
and KN: "mmap snd K = N" and sK: "set_of K \<subseteq> {(a, b). R a b}"
using assms
- unfolding multiset_rel_def Grp_def by auto
+ unfolding rel_multiset_def Grp_def by auto
obtain K1 ab where K: "K = K1 + {#ab#}" and a: "fst ab = a"
and K1M: "mmap fst K1 = M" using msed_map_invR[OF KM] by auto
obtain N1 where N: "N = N1 + {#snd ab#}" and K1N1: "mmap snd K1 = N1"
using msed_map_invL[OF KN[unfolded K]] by auto
have Rab: "R a (snd ab)" using sK a unfolding K by auto
- have "multiset_rel R M N1" using sK K1M K1N1
- unfolding K multiset_rel_def Grp_def by auto
+ have "rel_multiset R M N1" using sK K1M K1N1
+ unfolding K rel_multiset_def Grp_def by auto
thus ?thesis using N Rab by auto
qed
lemma msed_rel_invR:
-assumes "multiset_rel R M (N + {#b#})"
-shows "\<exists> M1 a. M = M1 + {#a#} \<and> R a b \<and> multiset_rel R M1 N"
+assumes "rel_multiset R M (N + {#b#})"
+shows "\<exists> M1 a. M = M1 + {#a#} \<and> R a b \<and> rel_multiset R M1 N"
proof-
obtain K where KN: "mmap snd K = N + {#b#}"
and KM: "mmap fst K = M" and sK: "set_of K \<subseteq> {(a, b). R a b}"
using assms
- unfolding multiset_rel_def Grp_def by auto
+ unfolding rel_multiset_def Grp_def by auto
obtain K1 ab where K: "K = K1 + {#ab#}" and b: "snd ab = b"
and K1N: "mmap snd K1 = N" using msed_map_invR[OF KN] by auto
obtain M1 where M: "M = M1 + {#fst ab#}" and K1M1: "mmap fst K1 = M1"
using msed_map_invL[OF KM[unfolded K]] by auto
have Rab: "R (fst ab) b" using sK b unfolding K by auto
- have "multiset_rel R M1 N" using sK K1N K1M1
- unfolding K multiset_rel_def Grp_def by auto
+ have "rel_multiset R M1 N" using sK K1N K1M1
+ unfolding K rel_multiset_def Grp_def by auto
thus ?thesis using M Rab by auto
qed
-lemma multiset_rel_imp_multiset_rel':
-assumes "multiset_rel R M N"
-shows "multiset_rel' R M N"
+lemma rel_multiset_imp_rel_multiset':
+assumes "rel_multiset R M N"
+shows "rel_multiset' R M N"
using assms proof(induct M arbitrary: N rule: measure_induct_rule[of mcard])
case (less M)
- have c: "mcard M = mcard N" using multiset_rel_mcard[OF less.prems] .
+ have c: "mcard M = mcard N" using rel_multiset_mcard[OF less.prems] .
show ?case
proof(cases "M = {#}")
case True hence "N = {#}" using c by simp
- thus ?thesis using True multiset_rel'.Zero by auto
+ thus ?thesis using True rel_multiset'.Zero by auto
next
case False then obtain M1 a where M: "M = M1 + {#a#}" by (metis multi_nonempty_split)
- obtain N1 b where N: "N = N1 + {#b#}" and R: "R a b" and ms: "multiset_rel R M1 N1"
+ obtain N1 b where N: "N = N1 + {#b#}" and R: "R a b" and ms: "rel_multiset R M1 N1"
using msed_rel_invL[OF less.prems[unfolded M]] by auto
- have "multiset_rel' R M1 N1" using less.hyps[of M1 N1] ms unfolding M by simp
- thus ?thesis using multiset_rel'.Plus[of R a b, OF R] unfolding M N by simp
+ have "rel_multiset' R M1 N1" using less.hyps[of M1 N1] ms unfolding M by simp
+ thus ?thesis using rel_multiset'.Plus[of R a b, OF R] unfolding M N by simp
qed
qed
-lemma multiset_rel_multiset_rel':
-"multiset_rel R M N = multiset_rel' R M N"
-using multiset_rel_imp_multiset_rel' multiset_rel'_imp_multiset_rel by auto
+lemma rel_multiset_rel_multiset':
+"rel_multiset R M N = rel_multiset' R M N"
+using rel_multiset_imp_rel_multiset' rel_multiset'_imp_rel_multiset by auto
-(* The main end product for multiset_rel: inductive characterization *)
-theorems multiset_rel_induct[case_names empty add, induct pred: multiset_rel] =
- multiset_rel'.induct[unfolded multiset_rel_multiset_rel'[symmetric]]
+(* The main end product for rel_multiset: inductive characterization *)
+theorems rel_multiset_induct[case_names empty add, induct pred: rel_multiset] =
+ rel_multiset'.induct[unfolded rel_multiset_rel_multiset'[symmetric]]
@@ -1184,5 +1183,4 @@
qed
qed
-
end
--- a/src/HOL/BNF/Tools/bnf_def.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/BNF/Tools/bnf_def.ML Wed Nov 20 16:43:09 2013 +0100
@@ -678,7 +678,7 @@
val def_qualify = Binding.conceal o Binding.qualify false (Binding.name_of bnf_b);
- fun mk_suffix_binding suf = Binding.suffix_name ("_" ^ suf) bnf_b;
+ fun mk_prefix_binding pre = Binding.prefix_name (pre ^ "_") bnf_b;
fun maybe_define user_specified (b, rhs) lthy =
let
@@ -703,7 +703,7 @@
lthy |> not (pointer_eq (lthy_old, lthy)) ? Local_Theory.restore;
val map_bind_def =
- (fn () => def_qualify (if Binding.is_empty map_b then mk_suffix_binding mapN else map_b),
+ (fn () => def_qualify (if Binding.is_empty map_b then mk_prefix_binding mapN else map_b),
map_rhs);
val set_binds_defs =
let
@@ -711,10 +711,10 @@
(case try (nth set_bs) (i - 1) of
SOME b => if Binding.is_empty b then get_b else K b
| NONE => get_b) #> def_qualify;
- val bs = if live = 1 then [set_name 1 (fn () => mk_suffix_binding setN)]
- else map (fn i => set_name i (fn () => mk_suffix_binding (mk_setN i))) (1 upto live);
+ val bs = if live = 1 then [set_name 1 (fn () => mk_prefix_binding setN)]
+ else map (fn i => set_name i (fn () => mk_prefix_binding (mk_setN i))) (1 upto live);
in bs ~~ set_rhss end;
- val bd_bind_def = (fn () => def_qualify (mk_suffix_binding bdN), bd_rhs);
+ val bd_bind_def = (fn () => def_qualify (mk_prefix_binding bdN), bd_rhs);
val ((((bnf_map_term, raw_map_def),
(bnf_set_terms, raw_set_defs)),
@@ -861,7 +861,7 @@
| SOME raw_rel => prep_term no_defs_lthy raw_rel);
val rel_bind_def =
- (fn () => def_qualify (if Binding.is_empty rel_b then mk_suffix_binding relN else rel_b),
+ (fn () => def_qualify (if Binding.is_empty rel_b then mk_prefix_binding relN else rel_b),
rel_rhs);
val wit_rhss =
@@ -873,8 +873,8 @@
val nwits = length wit_rhss;
val wit_binds_defs =
let
- val bs = if nwits = 1 then [fn () => def_qualify (mk_suffix_binding witN)]
- else map (fn i => fn () => def_qualify (mk_suffix_binding (mk_witN i))) (1 upto nwits);
+ val bs = if nwits = 1 then [fn () => def_qualify (mk_prefix_binding witN)]
+ else map (fn i => fn () => def_qualify (mk_prefix_binding (mk_witN i))) (1 upto nwits);
in bs ~~ wit_rhss end;
val (((bnf_rel_term, raw_rel_def), (bnf_wit_terms, raw_wit_defs)), (lthy, lthy_old)) =
--- a/src/HOL/BNF/Tools/bnf_fp_def_sugar.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/BNF/Tools/bnf_fp_def_sugar.ML Wed Nov 20 16:43:09 2013 +0100
@@ -544,10 +544,10 @@
val fpT_to_C as Type (_, [fpT, _]) = snd (strip_typeN nn (fastype_of (hd ctor_iters)));
- fun generate_iter suf (f_Tss, _, fss, xssss) ctor_iter =
+ fun generate_iter pre (f_Tss, _, fss, xssss) ctor_iter =
let
val res_T = fold_rev (curry (op --->)) f_Tss fpT_to_C;
- val b = mk_binding suf;
+ val b = mk_binding pre;
val spec =
mk_Trueprop_eq (lists_bmoc fss (Free (Binding.name_of b, res_T)),
mk_iter_body ctor_iter fss xssss);
@@ -563,10 +563,10 @@
val C_to_fpT as Type (_, [_, fpT]) = snd (strip_typeN nn (fastype_of (hd dtor_coiters)));
- fun generate_coiter suf ((pfss, cqfsss), (f_sum_prod_Ts, pf_Tss)) dtor_coiter =
+ fun generate_coiter pre ((pfss, cqfsss), (f_sum_prod_Ts, pf_Tss)) dtor_coiter =
let
val res_T = fold_rev (curry (op --->)) pf_Tss C_to_fpT;
- val b = mk_binding suf;
+ val b = mk_binding pre;
val spec =
mk_Trueprop_eq (lists_bmoc pfss (Free (Binding.name_of b, res_T)),
mk_coiter_body cs cpss f_sum_prod_Ts cqfsss dtor_coiter);
@@ -1356,7 +1356,7 @@
lthy |> Local_Theory.notes (anonymous_notes @ notes) |> snd)
end;
- fun mk_binding suf = qualify false fp_b_name (Binding.suffix_name ("_" ^ suf) fp_b);
+ fun mk_binding pre = qualify false fp_b_name (Binding.prefix_name (pre ^ "_") fp_b);
fun massage_res (((maps_sets_rels, ctr_sugar), co_iter_res), lthy) =
(((maps_sets_rels, (ctrs, xss, ctr_defs, ctr_sugar)), co_iter_res), lthy);
--- a/src/HOL/BNF/Tools/bnf_gfp.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/BNF/Tools/bnf_gfp.ML Wed Nov 20 16:43:09 2013 +0100
@@ -74,7 +74,7 @@
val mk_internal_b = Binding.name #> Binding.prefix true b_name #> Binding.conceal;
fun mk_internal_bs name =
map (fn b =>
- Binding.prefix true b_name (Binding.suffix_name ("_" ^ name) b) |> Binding.conceal) bs;
+ Binding.prefix true b_name (Binding.prefix_name (name ^ "_") b) |> Binding.conceal) bs;
val external_bs = map2 (Binding.prefix false) b_names bs
|> note_all = false ? map Binding.conceal;
@@ -1695,7 +1695,7 @@
||>> mk_Frees "s" corec_sTs
||>> mk_Frees "P" (map2 mk_pred2T Ts Ts);
- fun dtor_bind i = nth external_bs (i - 1) |> Binding.suffix_name ("_" ^ dtorN);
+ fun dtor_bind i = nth external_bs (i - 1) |> Binding.prefix_name (dtorN ^ "_");
val dtor_name = Binding.name_of o dtor_bind;
val dtor_def_bind = rpair [] o Binding.conceal o Thm.def_binding o dtor_bind;
@@ -1747,7 +1747,7 @@
val timer = time (timer "dtor definitions & thms");
- fun unfold_bind i = nth external_bs (i - 1) |> Binding.suffix_name ("_" ^ dtor_unfoldN);
+ fun unfold_bind i = nth external_bs (i - 1) |> Binding.prefix_name (dtor_unfoldN ^ "_");
val unfold_name = Binding.name_of o unfold_bind;
val unfold_def_bind = rpair [] o Binding.conceal o Thm.def_binding o unfold_bind;
@@ -1868,7 +1868,7 @@
Term.list_comb (mk_map_of_bnf Ds (passiveAs @ Ts) (passiveAs @ FTs) bnf,
map HOLogic.id_const passiveAs @ dtors)) Dss bnfs;
- fun ctor_bind i = nth external_bs (i - 1) |> Binding.suffix_name ("_" ^ ctorN);
+ fun ctor_bind i = nth external_bs (i - 1) |> Binding.prefix_name (ctorN ^ "_");
val ctor_name = Binding.name_of o ctor_bind;
val ctor_def_bind = rpair [] o Binding.conceal o Thm.def_binding o ctor_bind;
@@ -1939,7 +1939,7 @@
trans OF [mor RS unique, unfold_dtor]) unfold_unique_mor_thms unfold_dtor_thms
end;
- fun corec_bind i = nth external_bs (i - 1) |> Binding.suffix_name ("_" ^ dtor_corecN);
+ fun corec_bind i = nth external_bs (i - 1) |> Binding.prefix_name (dtor_corecN ^ "_");
val corec_name = Binding.name_of o corec_bind;
val corec_def_bind = rpair [] o Binding.conceal o Thm.def_binding o corec_bind;
--- a/src/HOL/BNF/Tools/bnf_lfp.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/BNF/Tools/bnf_lfp.ML Wed Nov 20 16:43:09 2013 +0100
@@ -44,7 +44,7 @@
val mk_internal_b = Binding.name #> Binding.prefix true b_name #> Binding.conceal;
fun mk_internal_bs name =
map (fn b =>
- Binding.prefix true b_name (Binding.suffix_name ("_" ^ name) b) |> Binding.conceal) bs;
+ Binding.prefix true b_name (Binding.prefix_name (name ^ "_") b) |> Binding.conceal) bs;
val external_bs = map2 (Binding.prefix false) b_names bs
|> note_all = false ? map Binding.conceal;
@@ -1021,7 +1021,7 @@
val phis = map2 retype_free (map mk_pred1T Ts) init_phis;
val phi2s = map2 retype_free (map2 mk_pred2T Ts Ts') init_phis;
- fun ctor_bind i = nth external_bs (i - 1) |> Binding.suffix_name ("_" ^ ctorN);
+ fun ctor_bind i = nth external_bs (i - 1) |> Binding.prefix_name (ctorN ^ "_");
val ctor_name = Binding.name_of o ctor_bind;
val ctor_def_bind = rpair [] o Binding.conceal o Thm.def_binding o ctor_bind;
@@ -1080,7 +1080,7 @@
(mk_mor UNIVs ctors active_UNIVs ss (map (mk_nthN n fold_f) ks));
val foldx = HOLogic.choice_const foldT $ fold_fun;
- fun fold_bind i = nth external_bs (i - 1) |> Binding.suffix_name ("_" ^ ctor_foldN);
+ fun fold_bind i = nth external_bs (i - 1) |> Binding.prefix_name (ctor_foldN ^ "_");
val fold_name = Binding.name_of o fold_bind;
val fold_def_bind = rpair [] o Binding.conceal o Thm.def_binding o fold_bind;
@@ -1170,7 +1170,7 @@
Term.list_comb (mk_map_of_bnf Ds (passiveAs @ FTs) (passiveAs @ Ts) bnf,
map HOLogic.id_const passiveAs @ ctors)) Dss bnfs;
- fun dtor_bind i = nth external_bs (i - 1) |> Binding.suffix_name ("_" ^ dtorN);
+ fun dtor_bind i = nth external_bs (i - 1) |> Binding.prefix_name (dtorN ^ "_");
val dtor_name = Binding.name_of o dtor_bind;
val dtor_def_bind = rpair [] o Binding.conceal o Thm.def_binding o dtor_bind;
@@ -1243,7 +1243,7 @@
trans OF [mor RS unique, fold_ctor]) fold_unique_mor_thms fold_ctor_thms
end;
- fun rec_bind i = nth external_bs (i - 1) |> Binding.suffix_name ("_" ^ ctor_recN);
+ fun rec_bind i = nth external_bs (i - 1) |> Binding.prefix_name (ctor_recN ^ "_");
val rec_name = Binding.name_of o rec_bind;
val rec_def_bind = rpair [] o Binding.conceal o Thm.def_binding o rec_bind;
@@ -1354,7 +1354,7 @@
val cTs = map (SOME o certifyT lthy o TFree) induct_params;
val weak_ctor_induct_thms =
- let fun insts i = (replicate (i - 1) TrueI) @ (@{thm asm_rl} :: replicate (n - i) TrueI);
+ let fun insts i = (replicate (i - 1) TrueI) @ (asm_rl :: replicate (n - i) TrueI);
in map (fn i => (ctor_induct_thm OF insts i) RS mk_conjunctN n i) ks end;
val (ctor_induct2_thm, induct2_params) =
--- a/src/HOL/Cardinals/Cardinal_Arithmetic.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Cardinals/Cardinal_Arithmetic.thy Wed Nov 20 16:43:09 2013 +0100
@@ -8,270 +8,17 @@
header {* Cardinal Arithmetic *}
theory Cardinal_Arithmetic
-imports Cardinal_Order_Relation_Base
+imports Cardinal_Arithmetic_FP Cardinal_Order_Relation
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 Field_card_order: "card_order r \<Longrightarrow> Field r = UNIV"
-using card_order_on_Card_order[of UNIV r] by simp
-
-(*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 Func_Times_Range:
- "|Func A (B <*> C)| =o |Func A B <*> Func A C|" (is "|?LHS| =o |?RHS|")
-proof -
- let ?F = "\<lambda>fg. (\<lambda>x. if x \<in> A then fst (fg x) else undefined,
- \<lambda>x. if x \<in> A then snd (fg x) else undefined)"
- let ?G = "\<lambda>(f, g) x. if x \<in> A then (f x, g x) else undefined"
- have "bij_betw ?F ?LHS ?RHS" unfolding bij_betw_def inj_on_def
- proof safe
- fix f g assume "f \<in> Func A B" "g \<in> Func A C"
- thus "(f, g) \<in> ?F ` Func A (B \<times> C)"
- by (intro image_eqI[of _ _ "?G (f, g)"]) (auto simp: Func_def)
- qed (auto simp: Func_def fun_eq_iff, metis pair_collapse)
- thus ?thesis using card_of_ordIso by blast
-qed
-
-
-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 :: 'b rel) \<longleftrightarrow> A = ({} :: 'a set)"
-unfolding czero_def by (rule iffI[OF card_of_empty2]) (auto simp: card_of_refl card_of_empty_ordIso)
-
-(* 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 {* (In)finite cardinals *}
-
-definition cinfinite where
- "cinfinite r = infinite (Field r)"
-
-abbreviation Cinfinite where
- "Cinfinite r \<equiv> cinfinite r \<and> Card_order r"
-
-definition cfinite where
- "cfinite r = finite (Field r)"
-
-abbreviation Cfinite where
- "Cfinite r \<equiv> cfinite r \<and> Card_order r"
-
-lemma Cfinite_ordLess_Cinfinite: "\<lbrakk>Cfinite r; Cinfinite s\<rbrakk> \<Longrightarrow> r <o s"
- unfolding cfinite_def cinfinite_def
- by (metis card_order_on_well_order_on finite_ordLess_infinite)
-
-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 Field_csum: "Field (r +c s) = Inl ` Field r \<union> Inr ` Field s"
- unfolding csum_def Field_card_of by auto
-
-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 Cfinite_csum: "\<lbrakk>Cfinite r; Cfinite s\<rbrakk> \<Longrightarrow> Cfinite (r +c s)"
- unfolding cfinite_def csum_def Field_card_of using card_of_card_order_on by simp
-
-lemma csum_csum: "(r1 +c r2) +c (r3 +c r4) =o (r1 +c r3) +c (r2 +c r4)"
-proof -
- have "(r1 +c r2) +c (r3 +c r4) =o r1 +c r2 +c (r3 +c r4)"
- by (metis csum_assoc)
- also have "r1 +c r2 +c (r3 +c r4) =o r1 +c (r2 +c r3) +c r4"
- by (metis csum_assoc csum_cong2 ordIso_symmetric)
- also have "r1 +c (r2 +c r3) +c r4 =o r1 +c (r3 +c r2) +c r4"
- by (metis csum_com csum_cong1 csum_cong2)
- also have "r1 +c (r3 +c r2) +c r4 =o r1 +c r3 +c r2 +c r4"
- by (metis csum_assoc csum_cong2 ordIso_symmetric)
- also have "r1 +c r3 +c r2 +c r4 =o (r1 +c r3) +c (r2 +c r4)"
- by (metis csum_assoc ordIso_symmetric)
- finally show ?thesis .
-qed
-
-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 Cfinite_cone: "Cfinite cone"
- unfolding cfinite_def by (simp add: Card_order_cone)
-
lemma single_cone:
"|{x}| =o cone"
proof -
@@ -280,349 +27,37 @@
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" 90) where
- "r1 ^c r2 \<equiv> |Func (Field r2) (Field r1)|"
-
-lemma card_order_cexp:
- 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 cexp_def Func_def by (simp add: card_of_card_order_on)
-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 n: "Field p2 = {} \<Longrightarrow> Field r2 = {}"
- shows "p1 ^c p2 \<le>o r1 ^c r2"
-proof(cases "Field p1 = {}")
- case True
- hence "|Field |Func (Field p2) (Field p1)|| \<le>o cone"
- unfolding cone_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 "|Func (Field p2) (Field p1)| \<le>o cone" by (simp add: Field_card_of cexp_def)
- hence "p1 ^c p2 \<le>o cone" unfolding cexp_def .
- thus ?thesis
- proof (cases "Field p2 = {}")
- case True
- with n have "Field r2 = {}" .
- hence "cone \<le>o r1 ^c r2" unfolding cone_def cexp_def Func_def by (auto intro: card_of_ordLeqI)
- thus ?thesis using `p1 ^c p2 \<le>o cone` ordLeq_transitive by auto
- next
- case False with True have "|Field (p1 ^c p2)| =o czero"
- unfolding card_of_ordIso_czero_iff_empty cexp_def Field_card_of Func_def by auto
- thus ?thesis unfolding cexp_def card_of_ordIso_czero_iff_empty Field_card_of
- by (simp add: card_of_empty)
- qed
-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 n, 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 n: "p2 =o czero \<Longrightarrow> r2 =o czero" and card: "Card_order p2"
- shows "p1 ^c p2 \<le>o r1 ^c r2"
- by (metis (full_types) "1" "2" card cexp_mono' czeroE czeroI n)
-
-lemma cexp_mono1:
- assumes 1: "p1 \<le>o r1" 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]) (auto simp: q)
-
-lemma cexp_mono2':
- assumes 2: "p2 \<le>o r2" and q: "Card_order q"
- and n: "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 n]) auto
-
-lemma cexp_mono2:
- assumes 2: "p2 \<le>o r2" and q: "Card_order q"
- and n: "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 n card]) auto
-
-lemma cexp_mono2_Cnotzero:
- assumes "p2 \<le>o r2" "Card_order q" "Cnotzero p2"
- shows "q ^c p2 \<le>o q ^c r2"
-by (metis assms cexp_mono2' czeroI)
-
-lemma cexp_cong:
- assumes 1: "p1 =o r1" and 2: "p2 =o r2"
- and Cr: "Card_order r2"
- 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 _ _ _ Cp] cexp_mono[OF _ _ _ Cr] by blast
-qed
-
-lemma cexp_cong1:
- assumes 1: "p1 =o r1" and q: "Card_order q"
- shows "p1 ^c q =o r1 ^c q"
-by (rule cexp_cong[OF 1 _ q q]) (rule ordIso_refl[OF q])
-
-lemma cexp_cong2:
- assumes 2: "p2 =o r2" and q: "Card_order 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. f ()"]) auto
- also have "|Field r| =o r" by (rule card_of_Field_ordIso[OF assms])
- finally show ?thesis .
-qed
-
-lemma cexp_cprod:
- assumes r1: "Card_order 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: "Card_order 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 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 assms(2))
- done
-qed
-
lemma Cnotzero_cexp:
assumes "Cnotzero q" "Card_order r"
shows "Cnotzero (q ^c r)"
@@ -664,41 +99,7 @@
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_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]])
+by (rule conjI, rule infinite_Func, auto)
lemma cone_ordLeq_iff_Field:
assumes "cone \<le>o r"
@@ -731,22 +132,6 @@
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)
@@ -782,131 +167,30 @@
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 Inl a else 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 Cfinite_cprod_Cinfinite: "\<lbrakk>Cfinite r; Cinfinite s\<rbrakk> \<Longrightarrow> r *c s \<le>o s"
-by (intro cprod_cinfinite_bound)
- (auto intro: ordLeq_refl ordLess_imp_ordLeq[OF Cfinite_ordLess_Cinfinite])
-
-lemma cprod_cexp: "(r *c s) ^c t =o r ^c t *c s ^c t"
- unfolding cprod_def cexp_def Field_card_of by (rule Func_Times_Range)
-
-lemma cprod_cexp_csum_cexp_Cinfinite:
- assumes t: "Cinfinite t"
- shows "(r *c s) ^c t \<le>o (r +c s) ^c t"
-proof -
- have "(r *c s) ^c t \<le>o ((r +c s) ^c ctwo) ^c t"
- by (rule cexp_mono1[OF cprod_csum_cexp conjunct2[OF t]])
- also have "((r +c s) ^c ctwo) ^c t =o (r +c s) ^c (ctwo *c t)"
- by (rule cexp_cprod[OF Card_order_csum])
- also have "(r +c s) ^c (ctwo *c t) =o (r +c s) ^c (t *c ctwo)"
- by (rule cexp_cong2[OF cprod_com Card_order_csum Card_order_cprod])
- also have "(r +c s) ^c (t *c ctwo) =o ((r +c s) ^c t) ^c ctwo"
- by (rule ordIso_symmetric[OF cexp_cprod[OF Card_order_csum]])
- also have "((r +c s) ^c t) ^c ctwo =o (r +c s) ^c t"
- by (rule cexp_cprod_ordLeq[OF Card_order_csum t ctwo_Cnotzero ctwo_ordLeq_Cinfinite[OF t]])
- finally show ?thesis .
-qed
-
-lemma Cfinite_cexp_Cinfinite:
- assumes s: "Cfinite s" and t: "Cinfinite t"
- shows "s ^c t \<le>o ctwo ^c t"
-proof (cases "s \<le>o ctwo")
- case True thus ?thesis using t by (blast intro: cexp_mono1)
-next
- case False
- hence "ctwo \<le>o s" by (metis card_order_on_well_order_on ctwo_Cnotzero ordLeq_total s)
- hence "Cnotzero s" by (metis Cnotzero_mono ctwo_Cnotzero s)
- hence st: "Cnotzero (s *c t)" by (metis Cinfinite_cprod2 cinfinite_not_czero t)
- have "s ^c t \<le>o (ctwo ^c s) ^c t"
- using assms by (blast intro: cexp_mono1 ordLess_imp_ordLeq[OF ordLess_ctwo_cexp])
- also have "(ctwo ^c s) ^c t =o ctwo ^c (s *c t)"
- by (blast intro: Card_order_ctwo cexp_cprod)
- also have "ctwo ^c (s *c t) \<le>o ctwo ^c t"
- using assms st by (intro cexp_mono2_Cnotzero Cfinite_cprod_Cinfinite Card_order_ctwo)
- finally show ?thesis .
-qed
-
-lemma csum_Cfinite_cexp_Cinfinite:
- assumes r: "Card_order r" and s: "Cfinite s" and t: "Cinfinite t"
- shows "(r +c s) ^c t \<le>o (r +c ctwo) ^c t"
-proof (cases "Cinfinite r")
- case True
- hence "r +c s =o r" by (intro csum_absorb1 ordLess_imp_ordLeq[OF Cfinite_ordLess_Cinfinite] s)
- hence "(r +c s) ^c t =o r ^c t" using t by (blast intro: cexp_cong1)
- also have "r ^c t \<le>o (r +c ctwo) ^c t" using t by (blast intro: cexp_mono1 ordLeq_csum1 r)
- finally show ?thesis .
-next
- case False
- with r have "Cfinite r" unfolding cinfinite_def cfinite_def by auto
- hence "Cfinite (r +c s)" by (intro Cfinite_csum s)
- hence "(r +c s) ^c t \<le>o ctwo ^c t" by (intro Cfinite_cexp_Cinfinite t)
- also have "ctwo ^c t \<le>o (r +c ctwo) ^c t" using t
- by (blast intro: cexp_mono1 ordLeq_csum2 Card_order_ctwo)
- finally show ?thesis .
-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)|"
+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.
+*}
lemma clists_Cinfinite: "Cinfinite r \<Longrightarrow> clists r =o r"
unfolding cinfinite_def clists_def by (blast intro: Card_order_lists_infinite)
@@ -915,6 +199,6 @@
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)
+by (simp add: clists_def card_of_ordIso_czero_iff_empty lists_not_empty)
end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Cardinals/Cardinal_Arithmetic_FP.thy Wed Nov 20 16:43:09 2013 +0100
@@ -0,0 +1,747 @@
+(* Title: HOL/Cardinals/Cardinal_Arithmetic_FP.thy
+ Author: Dmitriy Traytel, TU Muenchen
+ Copyright 2012
+
+Cardinal arithmetic (FP).
+*)
+
+header {* Cardinal Arithmetic (FP) *}
+
+theory Cardinal_Arithmetic_FP
+imports Cardinal_Order_Relation_FP
+begin
+
+(*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 fast
+
+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 Field_card_order: "card_order r \<Longrightarrow> Field r = UNIV"
+using card_order_on_Card_order[of UNIV r] by simp
+
+(*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 Func_Times_Range:
+ "|Func A (B <*> C)| =o |Func A B <*> Func A C|" (is "|?LHS| =o |?RHS|")
+proof -
+ let ?F = "\<lambda>fg. (\<lambda>x. if x \<in> A then fst (fg x) else undefined,
+ \<lambda>x. if x \<in> A then snd (fg x) else undefined)"
+ let ?G = "\<lambda>(f, g) x. if x \<in> A then (f x, g x) else undefined"
+ have "bij_betw ?F ?LHS ?RHS" unfolding bij_betw_def inj_on_def
+ apply safe
+ apply (simp add: Func_def fun_eq_iff)
+ apply (metis (no_types) pair_collapse)
+ apply (auto simp: Func_def fun_eq_iff)[2]
+ proof -
+ fix f g assume "f \<in> Func A B" "g \<in> Func A C"
+ thus "(f, g) \<in> ?F ` Func A (B \<times> C)"
+ by (intro image_eqI[of _ _ "?G (f, g)"]) (auto simp: Func_def)
+ qed
+ thus ?thesis using card_of_ordIso by blast
+qed
+
+
+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 :: 'b rel) \<longleftrightarrow> A = ({} :: 'a set)"
+unfolding czero_def by (rule iffI[OF card_of_empty2]) (auto simp: card_of_refl card_of_empty_ordIso)
+
+(* 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 {* (In)finite cardinals *}
+
+definition cinfinite where
+ "cinfinite r = infinite (Field r)"
+
+abbreviation Cinfinite where
+ "Cinfinite r \<equiv> cinfinite r \<and> Card_order r"
+
+definition cfinite where
+ "cfinite r = finite (Field r)"
+
+abbreviation Cfinite where
+ "Cfinite r \<equiv> cfinite r \<and> Card_order r"
+
+lemma Cfinite_ordLess_Cinfinite: "\<lbrakk>Cfinite r; Cinfinite s\<rbrakk> \<Longrightarrow> r <o s"
+ unfolding cfinite_def cinfinite_def
+ by (metis card_order_on_well_order_on finite_ordLess_infinite)
+
+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 Field_csum: "Field (r +c s) = Inl ` Field r \<union> Inr ` Field s"
+ unfolding csum_def Field_card_of by auto
+
+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 Plus_eq_empty_conv card_of_Card_order card_of_ordIso_czero_iff_empty)
+
+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_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 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 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 Cfinite_csum: "\<lbrakk>Cfinite r; Cfinite s\<rbrakk> \<Longrightarrow> Cfinite (r +c s)"
+ unfolding cfinite_def csum_def Field_card_of using card_of_card_order_on by simp
+
+lemma csum_csum: "(r1 +c r2) +c (r3 +c r4) =o (r1 +c r3) +c (r2 +c r4)"
+proof -
+ have "(r1 +c r2) +c (r3 +c r4) =o r1 +c r2 +c (r3 +c r4)"
+ by (metis csum_assoc)
+ also have "r1 +c r2 +c (r3 +c r4) =o r1 +c (r2 +c r3) +c r4"
+ by (metis csum_assoc csum_cong2 ordIso_symmetric)
+ also have "r1 +c (r2 +c r3) +c r4 =o r1 +c (r3 +c r2) +c r4"
+ by (metis csum_com csum_cong1 csum_cong2)
+ also have "r1 +c (r3 +c r2) +c r4 =o r1 +c r3 +c r2 +c r4"
+ by (metis csum_assoc csum_cong2 ordIso_symmetric)
+ also have "r1 +c r3 +c r2 +c r4 =o (r1 +c r3) +c (r2 +c r4)"
+ by (metis csum_assoc ordIso_symmetric)
+ finally show ?thesis .
+qed
+
+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 Cfinite_cone: "Cfinite cone"
+ unfolding cfinite_def by (simp add: Card_order_cone)
+
+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_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 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 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_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_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
+
+
+subsection {* Exponentiation *}
+
+definition cexp (infixr "^c" 90) where
+ "r1 ^c r2 \<equiv> |Func (Field r2) (Field r1)|"
+
+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 n: "Field p2 = {} \<Longrightarrow> Field r2 = {}"
+ shows "p1 ^c p2 \<le>o r1 ^c r2"
+proof(cases "Field p1 = {}")
+ case True
+ hence "|Field |Func (Field p2) (Field p1)|| \<le>o cone"
+ unfolding cone_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 "|Func (Field p2) (Field p1)| \<le>o cone" by (simp add: Field_card_of cexp_def)
+ hence "p1 ^c p2 \<le>o cone" unfolding cexp_def .
+ thus ?thesis
+ proof (cases "Field p2 = {}")
+ case True
+ with n have "Field r2 = {}" .
+ hence "cone \<le>o r1 ^c r2" unfolding cone_def cexp_def Func_def by (auto intro: card_of_ordLeqI)
+ thus ?thesis using `p1 ^c p2 \<le>o cone` ordLeq_transitive by auto
+ next
+ case False with True have "|Field (p1 ^c p2)| =o czero"
+ unfolding card_of_ordIso_czero_iff_empty cexp_def Field_card_of Func_def by auto
+ thus ?thesis unfolding cexp_def card_of_ordIso_czero_iff_empty Field_card_of
+ by (simp add: card_of_empty)
+ qed
+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 n, 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 n: "p2 =o czero \<Longrightarrow> r2 =o czero" and card: "Card_order p2"
+ shows "p1 ^c p2 \<le>o r1 ^c r2"
+ by (metis (full_types) "1" "2" card cexp_mono' czeroE czeroI n)
+
+lemma cexp_mono1:
+ assumes 1: "p1 \<le>o r1" 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]) (auto simp: q)
+
+lemma cexp_mono2':
+ assumes 2: "p2 \<le>o r2" and q: "Card_order q"
+ and n: "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 n]) auto
+
+lemma cexp_mono2:
+ assumes 2: "p2 \<le>o r2" and q: "Card_order q"
+ and n: "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 n card]) auto
+
+lemma cexp_mono2_Cnotzero:
+ assumes "p2 \<le>o r2" "Card_order q" "Cnotzero p2"
+ shows "q ^c p2 \<le>o q ^c r2"
+by (metis assms cexp_mono2' czeroI)
+
+lemma cexp_cong:
+ assumes 1: "p1 =o r1" and 2: "p2 =o r2"
+ and Cr: "Card_order r2"
+ 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 _ _ _ Cp] cexp_mono[OF _ _ _ Cr] by metis
+qed
+
+lemma cexp_cong1:
+ assumes 1: "p1 =o r1" and q: "Card_order q"
+ shows "p1 ^c q =o r1 ^c q"
+by (rule cexp_cong[OF 1 _ q q]) (rule ordIso_refl[OF q])
+
+lemma cexp_cong2:
+ assumes 2: "p2 =o r2" and q: "Card_order 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_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. f ()"]) auto
+ also have "|Field r| =o r" by (rule card_of_Field_ordIso[OF assms])
+ finally show ?thesis .
+qed
+
+lemma cexp_cprod:
+ assumes r1: "Card_order 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 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 cexp_cprod_ordLeq:
+ assumes r1: "Card_order 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 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 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 assms(2))
+ done
+qed
+
+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_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 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 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 Inl a else 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 Cfinite_cprod_Cinfinite: "\<lbrakk>Cfinite r; Cinfinite s\<rbrakk> \<Longrightarrow> r *c s \<le>o s"
+by (intro cprod_cinfinite_bound)
+ (auto intro: ordLeq_refl ordLess_imp_ordLeq[OF Cfinite_ordLess_Cinfinite])
+
+lemma cprod_cexp: "(r *c s) ^c t =o r ^c t *c s ^c t"
+ unfolding cprod_def cexp_def Field_card_of by (rule Func_Times_Range)
+
+lemma cprod_cexp_csum_cexp_Cinfinite:
+ assumes t: "Cinfinite t"
+ shows "(r *c s) ^c t \<le>o (r +c s) ^c t"
+proof -
+ have "(r *c s) ^c t \<le>o ((r +c s) ^c ctwo) ^c t"
+ by (rule cexp_mono1[OF cprod_csum_cexp conjunct2[OF t]])
+ also have "((r +c s) ^c ctwo) ^c t =o (r +c s) ^c (ctwo *c t)"
+ by (rule cexp_cprod[OF Card_order_csum])
+ also have "(r +c s) ^c (ctwo *c t) =o (r +c s) ^c (t *c ctwo)"
+ by (rule cexp_cong2[OF cprod_com Card_order_csum Card_order_cprod])
+ also have "(r +c s) ^c (t *c ctwo) =o ((r +c s) ^c t) ^c ctwo"
+ by (rule ordIso_symmetric[OF cexp_cprod[OF Card_order_csum]])
+ also have "((r +c s) ^c t) ^c ctwo =o (r +c s) ^c t"
+ by (rule cexp_cprod_ordLeq[OF Card_order_csum t ctwo_Cnotzero ctwo_ordLeq_Cinfinite[OF t]])
+ finally show ?thesis .
+qed
+
+lemma Cfinite_cexp_Cinfinite:
+ assumes s: "Cfinite s" and t: "Cinfinite t"
+ shows "s ^c t \<le>o ctwo ^c t"
+proof (cases "s \<le>o ctwo")
+ case True thus ?thesis using t by (blast intro: cexp_mono1)
+next
+ case False
+ hence "ctwo \<le>o s" by (metis card_order_on_well_order_on ctwo_Cnotzero ordLeq_total s)
+ hence "Cnotzero s" by (metis Cnotzero_mono ctwo_Cnotzero s)
+ hence st: "Cnotzero (s *c t)" by (metis Cinfinite_cprod2 cinfinite_not_czero t)
+ have "s ^c t \<le>o (ctwo ^c s) ^c t"
+ using assms by (blast intro: cexp_mono1 ordLess_imp_ordLeq[OF ordLess_ctwo_cexp])
+ also have "(ctwo ^c s) ^c t =o ctwo ^c (s *c t)"
+ by (blast intro: Card_order_ctwo cexp_cprod)
+ also have "ctwo ^c (s *c t) \<le>o ctwo ^c t"
+ using assms st by (intro cexp_mono2_Cnotzero Cfinite_cprod_Cinfinite Card_order_ctwo)
+ finally show ?thesis .
+qed
+
+lemma csum_Cfinite_cexp_Cinfinite:
+ assumes r: "Card_order r" and s: "Cfinite s" and t: "Cinfinite t"
+ shows "(r +c s) ^c t \<le>o (r +c ctwo) ^c t"
+proof (cases "Cinfinite r")
+ case True
+ hence "r +c s =o r" by (intro csum_absorb1 ordLess_imp_ordLeq[OF Cfinite_ordLess_Cinfinite] s)
+ hence "(r +c s) ^c t =o r ^c t" using t by (blast intro: cexp_cong1)
+ also have "r ^c t \<le>o (r +c ctwo) ^c t" using t by (blast intro: cexp_mono1 ordLeq_csum1 r)
+ finally show ?thesis .
+next
+ case False
+ with r have "Cfinite r" unfolding cinfinite_def cfinite_def by auto
+ hence "Cfinite (r +c s)" by (intro Cfinite_csum s)
+ hence "(r +c s) ^c t \<le>o ctwo ^c t" by (intro Cfinite_cexp_Cinfinite t)
+ also have "ctwo ^c t \<le>o (r +c ctwo) ^c t" using t
+ by (blast intro: cexp_mono1 ordLeq_csum2 Card_order_ctwo)
+ finally show ?thesis .
+qed
+
+lemma card_order_cexp:
+ 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 cexp_def Func_def by (simp add: card_of_card_order_on)
+qed
+
+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]])
+
+(* 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 Cinfinite_cpow: "Cinfinite r \<Longrightarrow> Cinfinite (cpow r)"
+unfolding cpow_def cinfinite_def by (metis Field_card_of card_of_Card_order infinite_Pow)
+
+subsection {* Lists *}
+
+definition clists where "clists r = |lists (Field r)|"
+
+end
--- a/src/HOL/Cardinals/Cardinal_Order_Relation.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Cardinals/Cardinal_Order_Relation.thy Wed Nov 20 16:43:09 2013 +0100
@@ -8,7 +8,7 @@
header {* Cardinal-Order Relations *}
theory Cardinal_Order_Relation
-imports Cardinal_Order_Relation_Base Constructions_on_Wellorders
+imports Cardinal_Order_Relation_FP Constructions_on_Wellorders
begin
declare
@@ -34,7 +34,6 @@
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]
@@ -44,25 +43,19 @@
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]
@@ -143,6 +136,17 @@
subsection {* Cardinals versus set operations on arbitrary sets *}
+lemma card_of_set_type[simp]: "|UNIV::'a set| <o |UNIV::'a set set|"
+using card_of_Pow[of "UNIV::'a set"] by simp
+
+lemma card_of_Un1[simp]:
+shows "|A| \<le>o |A \<union> B| "
+using inj_on_id[of A] card_of_ordLeq[of A _] by fastforce
+
+lemma card_of_diff[simp]:
+shows "|A - B| \<le>o |A|"
+using inj_on_id[of "A - B"] card_of_ordLeq[of "A - B" _] by fastforce
+
lemma subset_ordLeq_strict:
assumes "A \<le> B" and "|A| <o |B|"
shows "A < B"
@@ -307,6 +311,16 @@
using card_of_Times3 card_of_Field_ordIso
ordIso_ordLeq_trans ordIso_symmetric by blast
+lemma card_of_Times_cong1[simp]:
+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[simp]:
+assumes "|A| =o |B|"
+shows "|C \<times> A| =o |C \<times> B|"
+using assms by (simp add: ordIso_iff_ordLeq card_of_Times_mono2)
+
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|"
@@ -323,6 +337,11 @@
shows "|(Field r) \<times> C| =o |(Field r') \<times> C|"
using assms card_of_cong card_of_Times_cong1 by blast
+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_Times_cong[simp]:
assumes "|A| =o |B|" and "|C| =o |D|"
shows "|A \<times> C| =o |B \<times> D|"
@@ -501,11 +520,55 @@
using assms Plus_infinite_bij_betw[of "UNIV::'a set" g "UNIV::'b set"]
by auto
+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_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
+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
+
corollary Card_order_Un_infinite:
assumes INF: "infinite(Field r)" and CARD: "Card_order r" and
LEQ: "p \<le>o r"
@@ -597,6 +660,33 @@
thus ?thesis using 1 ordLess_ordIso_trans by blast
qed
+
+subsection {* Cardinals versus set operations involving infinite sets *}
+
+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[simp]:
+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
+
+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
+
lemma card_of_Un_singl_ordLess_infinite1:
assumes "infinite B" and "|A| <o |B|"
shows "|{a} Un A| <o |B|"
@@ -616,7 +706,83 @@
qed
-subsection {* Cardinals versus lists *}
+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_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
@@ -690,6 +856,22 @@
thus ?thesis using card_of_ordIso[of "lists A"] by auto
qed
+lemma card_of_lists_infinite[simp]:
+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
+
lemma ordIso_lists_cong:
assumes "r =o r'"
shows "|lists(Field r)| =o |lists(Field r')|"
@@ -827,13 +1009,22 @@
lemma Field_natLess: "Field natLess = (UNIV::nat set)"
by(unfold Field_def, auto)
+lemma 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_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)
+ 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)
+ simp add: Field_natLeq, unfold rel.under_def, auto)
+
+lemma natLeq_UNIV_ofilter: "wo_rel.ofilter natLeq UNIV"
+using natLeq_wo_rel Field_natLeq wo_rel.Field_ofilter[of natLeq] by auto
lemma natLeq_ofilter_iff:
"ofilter natLeq A = (A = UNIV \<or> (\<exists>n. A = {0 ..< n}))"
@@ -900,7 +1091,7 @@
qed
-subsubsection {* "Backwards compatibility" with the numeric cardinal operator for finite sets *}
+subsubsection {* "Backward compatibility" with the numeric cardinal operator for finite sets *}
lemma finite_card_of_iff_card:
assumes FIN: "finite A" and FIN': "finite B"
@@ -993,6 +1184,11 @@
shows "relChain r (\<lambda> i. under r i)"
using assms unfolding relChain_def by auto
+lemma card_of_infinite_diff_finite:
+assumes "infinite A" and "finite B"
+shows "|A - B| =o |A|"
+by (metis assms card_of_Un_diff_infinite finite_ordLess_infinite2)
+
lemma infinite_card_of_diff_singl:
assumes "infinite A"
shows "|A - {a}| =o |A|"
@@ -1110,6 +1306,30 @@
thus "f \<in> Pfunc A B" unfolding Func_option_def Pfunc_def by auto
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) a. if a \<in> A2 then (if a \<in> A1 then f1 a else bb)
+ else undefined"
+ 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_Func_option_mono:
fixes A1 A2 :: "'a set" and B :: "'b set"
assumes A12: "A1 \<subseteq> A2" and B: "B \<noteq> {}"
@@ -1178,4 +1398,18 @@
"|Func (UNIV::'a set) (UNIV::'b set)| =o |UNIV::('a \<Rightarrow> 'b) set|"
using card_of_Func_UNIV[of "UNIV::'b set"] by auto
+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 b1 else b2) else undefined"
+ 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 by auto
+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)
+
end
--- a/src/HOL/Cardinals/Cardinal_Order_Relation_Base.thy Wed Nov 20 16:15:54 2013 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,2438 +0,0 @@
-(* Title: HOL/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 force
- 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 metis
- 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}"
- unfolding ordLeq_def using
- embed_inj_on[OF natLeq_on_Well_order[of m], of "natLeq_on n", unfolded Field_natLeq_on]
- embed_Field[OF natLeq_on_Well_order[of m], of "natLeq_on n", unfolded Field_natLeq_on] by blast
- 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 *}
-
-lemma card_of_infinite_diff_finite:
-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 = {f . (\<forall> a \<in> A. f a \<in> B) \<and> (\<forall> a. a \<notin> A \<longrightarrow> f a = undefined)}"
-
-lemma Func_empty:
-"Func {} B = {\<lambda>x. undefined}"
-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 = b"
-using assms unfolding Func_def by (cases "g a = undefined") auto
-
-definition curr where
-"curr A f \<equiv> \<lambda> a. if a \<in> A then \<lambda>b. f (a,b) else undefined"
-
-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 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 by(elim allE[of _ a]) simp
- 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. if fst ab \<in> A \<and> snd ab \<in> B then g (fst ab) (snd ab) else undefined"
- 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 = undefined" 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 = g1"
- using assms using Func_elim[OF assms True] by blast
- thus ?thesis using True unfolding Func_def curr_def by auto
- qed
- qed
- show "?f \<in> Func (A <*> B) C" using assms unfolding Func_def mem_Collect_eq by auto
-qed
-
-lemma bij_betw_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_betw_curr by blast
-
-definition Func_map where
-"Func_map B2 f1 f2 g b2 \<equiv> if b2 \<in> B2 then f1 (g (f2 b2)) else undefined"
-
-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"
-using assms unfolding Func_def Func_map_def mem_Collect_eq by 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 b else undefined) \<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 = undefined", 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_def)
-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 = undefined")
- case True
- hence b1: "h b2 \<in> f1 ` A1" using h `b2 \<in> B2` unfolding B1 Func_def by auto
- show ?thesis using A2 f_inv_into_f[OF b1]
- unfolding True g_def Func_map_def j1_def j2[OF `b2 \<in> B2`] by auto
- qed(insert A2 True j2[OF True] h B1, unfold j1_def g_def Func_def Func_map_def,
- auto intro: f_inv_into_f)
- 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 fast+
- 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
-
-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 True else False)
- else undefined"
- 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 \<in> Pow A" "A2 \<in> Pow A" "F A1 = F A2"
- thus "A1 = A2" unfolding F_def Pow_def fun_eq_iff by (auto split: split_if_asm)
- 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 = True}"
- show "f = F ?A1" unfolding F_def apply(rule ext)
- using f unfolding Func_def mem_Collect_eq by auto
- 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) a. if a \<in> A2 then (if a \<in> A1 then f1 a else bb)
- else undefined"
- 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 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 b1 else b2) else undefined"
- 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 by auto
-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)
-
-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). ((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" assume h: "h \<in> Func UNIV B"
- hence "\<forall> a. \<exists> b. h a = b" unfolding Func_def by auto
- then obtain f where f: "\<forall> a. h a = f a" by metis
- hence "range f \<subseteq> B" using h unfolding Func_def by auto
- thus "h \<in> (\<lambda>f a. 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/Cardinals/Cardinal_Order_Relation_FP.thy Wed Nov 20 16:43:09 2013 +0100
@@ -0,0 +1,2174 @@
+(* Title: HOL/Cardinals/Cardinal_Order_Relation_FP.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Cardinal-order relations (FP).
+*)
+
+header {* Cardinal-Order Relations (FP) *}
+
+theory Cardinal_Order_Relation_FP
+imports Constructions_on_Wellorders_FP
+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 fast
+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
+
+
+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
+
+
+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
+
+
+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 fastforce
+ 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 g_def by force
+ }
+ 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_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
+
+
+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_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
+
+
+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_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 metis
+ 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)
+
+
+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 metis
+ 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 fast
+ 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_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_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)
+
+
+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
+
+
+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
+
+
+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 force
+
+
+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}"
+apply (auto simp add: natLeq_on_wo_rel wo_rel.ofilter_def)
+apply (simp add: Field_natLeq_on)
+by (auto simp add: rel.under_def)
+
+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 auto
+
+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}"
+ unfolding ordLeq_def using
+ embed_inj_on[OF natLeq_on_Well_order[of m], of "natLeq_on n", unfolded Field_natLeq_on]
+ embed_Field[OF natLeq_on_Well_order[of m], of "natLeq_on n", unfolded Field_natLeq_on] by blast
+ 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 fastforce
+
+
+
+subsubsection {* "Backward 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_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 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 fast
+
+
+
+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 *}
+
+(* function space *)
+definition Func where
+"Func A B = {f . (\<forall> a \<in> A. f a \<in> B) \<and> (\<forall> a. a \<notin> A \<longrightarrow> f a = undefined)}"
+
+lemma Func_empty:
+"Func {} B = {\<lambda>x. undefined}"
+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 = b"
+using assms unfolding Func_def by (cases "g a = undefined") auto
+
+definition curr where
+"curr A f \<equiv> \<lambda> a. if a \<in> A then \<lambda>b. f (a,b) else undefined"
+
+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 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 by(elim allE[of _ a]) simp
+ 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. if fst ab \<in> A \<and> snd ab \<in> B then g (fst ab) (snd ab) else undefined"
+ 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 = undefined" 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 = g1"
+ using assms using Func_elim[OF assms True] by blast
+ thus ?thesis using True unfolding Func_def curr_def by auto
+ qed
+ qed
+ show "?f \<in> Func (A <*> B) C" using assms unfolding Func_def mem_Collect_eq by auto
+qed
+
+lemma bij_betw_curr:
+"bij_betw (curr A) (Func (A <*> B) C) (Func A (Func B C))"
+unfolding bij_betw_def inj_on_def image_def
+apply (intro impI conjI ballI)
+apply (erule curr_inj[THEN iffD1], assumption+)
+apply auto
+apply (erule curr_in)
+using 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_betw_curr by blast
+
+definition Func_map where
+"Func_map B2 f1 f2 g b2 \<equiv> if b2 \<in> B2 then f1 (g (f2 b2)) else undefined"
+
+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"
+using assms unfolding Func_def Func_map_def mem_Collect_eq by 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 b else undefined) \<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 blast
+ with R have False by blast
+ }
+ 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_def)
+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 = undefined")
+ case True
+ hence b1: "h b2 \<in> f1 ` A1" using h `b2 \<in> B2` unfolding B1 Func_def by auto
+ show ?thesis using A2 f_inv_into_f[OF b1]
+ unfolding True g_def Func_map_def j1_def j2[OF `b2 \<in> B2`] by auto
+ qed(insert A2 True j2[OF True] h B1, unfold j1_def g_def Func_def Func_map_def,
+ auto intro: f_inv_into_f)
+ 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 fast+
+ 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
+
+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 True else False)
+ else undefined"
+ 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 \<in> Pow A" "A2 \<in> Pow A" "F A1 = F A2"
+ thus "A1 = A2" unfolding F_def Pow_def fun_eq_iff by (auto split: split_if_asm)
+ 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 = True}"
+ show "f = F ?A1" unfolding F_def apply(rule ext)
+ using f unfolding Func_def mem_Collect_eq by auto
+ 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_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). ((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" assume h: "h \<in> Func UNIV B"
+ hence "\<forall> a. \<exists> b. h a = b" unfolding Func_def by auto
+ then obtain f where f: "\<forall> a. h a = f a" by metis
+ hence "range f \<subseteq> B" using h unfolding Func_def by auto
+ thus "h \<in> (\<lambda>f a. f a) ` {f. range f \<subseteq> B}" using f unfolding image_def by auto
+ qed(unfold Func_def fun_eq_iff, auto)
+qed
+
+end
--- a/src/HOL/Cardinals/Constructions_on_Wellorders.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Cardinals/Constructions_on_Wellorders.thy Wed Nov 20 16:43:09 2013 +0100
@@ -8,13 +8,11 @@
header {* Constructions on Wellorders *}
theory Constructions_on_Wellorders
-imports Constructions_on_Wellorders_Base Wellorder_Embedding
+imports Constructions_on_Wellorders_FP 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]
@@ -88,7 +86,7 @@
by (auto simp add: ofilter_subset_embedS_iso)
-subsection {* Ordering the well-orders by existence of embeddings *}
+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
@@ -113,6 +111,16 @@
corollary ordIso_equiv: "equiv {r. Well_order r} ordIso"
by (auto simp add: equiv_def ordIso_sym ordIso_refl_on ordIso_trans)
+lemma ordLess_Well_order_simp[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[simp]:
+assumes "r =o r'"
+shows "Well_order r \<and> Well_order r'"
+using assms unfolding ordIso_def by simp
+
lemma ordLess_irrefl: "irrefl ordLess"
by(unfold irrefl_def, auto simp add: ordLess_irreflexive)
--- a/src/HOL/Cardinals/Constructions_on_Wellorders_Base.thy Wed Nov 20 16:15:54 2013 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1633 +0,0 @@
-(* Title: HOL/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 metis
- }
- 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 metis
- 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 force
- }
- 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/Cardinals/Constructions_on_Wellorders_FP.thy Wed Nov 20 16:43:09 2013 +0100
@@ -0,0 +1,1621 @@
+(* Title: HOL/Cardinals/Constructions_on_Wellorders_FP.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Constructions on wellorders (FP).
+*)
+
+header {* Constructions on Wellorders (FP) *}
+
+theory Constructions_on_Wellorders_FP
+imports Wellorder_Embedding_FP
+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"
+unfolding refl_on_def Field_def by blast
+
+
+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
+
+
+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 fast
+ 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 metis
+ }
+ 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 metis
+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 metis
+ 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 force
+ }
+ 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
--- a/src/HOL/Cardinals/Fun_More.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Cardinals/Fun_More.thy Wed Nov 20 16:43:09 2013 +0100
@@ -8,7 +8,7 @@
header {* More on Injections, Bijections and Inverses *}
theory Fun_More
-imports Fun_More_Base
+imports Fun_More_FP
begin
@@ -132,6 +132,18 @@
subsection {* Properties involving Hilbert choice *}
+(*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
+
+
subsection {* Other facts *}
(*3*)lemma atLeastLessThan_injective:
--- a/src/HOL/Cardinals/Fun_More_Base.thy Wed Nov 20 16:15:54 2013 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,252 +0,0 @@
-(* Title: HOL/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, safe)
- 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/Cardinals/Fun_More_FP.thy Wed Nov 20 16:43:09 2013 +0100
@@ -0,0 +1,239 @@
+(* Title: HOL/Cardinals/Fun_More_FP.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+More on injections, bijections and inverses (FP).
+*)
+
+header {* More on Injections, Bijections and Inverses (FP) *}
+
+theory Fun_More_FP
+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, safe)
+ 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_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 fastforce
+
+
+
+end
--- a/src/HOL/Cardinals/Order_Relation_More.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Cardinals/Order_Relation_More.thy Wed Nov 20 16:43:09 2013 +0100
@@ -8,64 +8,70 @@
header {* Basics on Order-Like Relations *}
theory Order_Relation_More
-imports Order_Relation_More_Base
+imports Order_Relation_More_FP
begin
subsection {* The upper and lower bounds operators *}
-lemma (in rel) aboveS_subset_above: "aboveS a \<le> above a"
+context rel
+begin
+
+lemma 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"
+lemma 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) = {}"
+lemma UnderS_disjoint: "A Int (UnderS A) = {}"
by(auto simp add: UnderS_def)
-lemma (in rel) aboveS_notIn: "a \<notin> aboveS a"
+lemma 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"
+lemma 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)"
+lemma 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)"
+lemma 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)"
+lemma 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)"
+lemma UnderS_subset_Under: "UnderS A \<le> Under A"
+by(auto simp add: UnderS_def Under_def)
+
+lemma 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)"
+lemma 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)"
+lemma 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)"
+lemma 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:
+lemma 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:
+lemma 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:
+lemma 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:
+lemma 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)
@@ -88,7 +94,7 @@
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:
+lemma 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)
@@ -111,19 +117,25 @@
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 = {}"
+lemma 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"
+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 above_Field: "above a \<le> Field r"
by(unfold above_def Field_def, auto)
-lemma (in rel) aboveS_Field: "aboveS a \<le> Field r"
+lemma aboveS_Field: "aboveS a \<le> Field r"
by(unfold aboveS_def Field_def, auto)
-lemma (in rel) Above_Field: "Above A \<le> Field r"
+lemma Above_Field: "Above A \<le> Field r"
by(unfold Above_def Field_def, auto)
-lemma (in rel) Refl_under_Under:
+lemma Refl_under_Under:
assumes REFL: "Refl r" and NE: "A \<noteq> {}"
shows "Under A = (\<Inter> a \<in> A. under a)"
proof
@@ -147,7 +159,7 @@
qed
qed
-lemma (in rel) Refl_underS_UnderS:
+lemma Refl_underS_UnderS:
assumes REFL: "Refl r" and NE: "A \<noteq> {}"
shows "UnderS A = (\<Inter> a \<in> A. underS a)"
proof
@@ -171,7 +183,7 @@
qed
qed
-lemma (in rel) Refl_above_Above:
+lemma Refl_above_Above:
assumes REFL: "Refl r" and NE: "A \<noteq> {}"
shows "Above A = (\<Inter> a \<in> A. above a)"
proof
@@ -195,7 +207,7 @@
qed
qed
-lemma (in rel) Refl_aboveS_AboveS:
+lemma Refl_aboveS_AboveS:
assumes REFL: "Refl r" and NE: "A \<noteq> {}"
shows "AboveS A = (\<Inter> a \<in> A. aboveS a)"
proof
@@ -219,31 +231,31 @@
qed
qed
-lemma (in rel) under_Under_singl: "under a = Under {a}"
+lemma 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}"
+lemma 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}"
+lemma 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}"
+lemma 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"
+lemma 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"
+lemma 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"
+lemma 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"
+lemma AboveS_decr: "A \<le> B \<Longrightarrow> AboveS B \<le> AboveS A"
by(unfold AboveS_def, auto)
-lemma (in rel) under_incl_iff:
+lemma 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
@@ -259,7 +271,7 @@
by (auto simp add: under_def)
qed
-lemma (in rel) above_decr:
+lemma above_decr:
assumes TRANS: "trans r" and REL: "(a,b) \<in> r"
shows "above b \<le> above a"
proof(unfold above_def, auto)
@@ -268,7 +280,7 @@
show "(a,x) \<in> r" by blast
qed
-lemma (in rel) aboveS_decr:
+lemma aboveS_decr:
assumes TRANS: "trans r" and ANTISYM: "antisym r" and
REL: "(a,b) \<in> r"
shows "aboveS b \<le> aboveS a"
@@ -282,7 +294,7 @@
show "(a,x) \<in> r" by blast
qed
-lemma (in rel) under_trans:
+lemma under_trans:
assumes TRANS: "trans r" and
IN1: "a \<in> under b" and IN2: "b \<in> under c"
shows "a \<in> under c"
@@ -294,7 +306,7 @@
thus ?thesis unfolding under_def by simp
qed
-lemma (in rel) under_underS_trans:
+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"
@@ -312,7 +324,7 @@
from 1 3 show ?thesis unfolding underS_def by simp
qed
-lemma (in rel) underS_under_trans:
+lemma 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"
@@ -330,7 +342,7 @@
from 1 3 show ?thesis unfolding underS_def by simp
qed
-lemma (in rel) underS_underS_trans:
+lemma 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"
@@ -340,7 +352,7 @@
with assms under_underS_trans show ?thesis by auto
qed
-lemma (in rel) above_trans:
+lemma above_trans:
assumes TRANS: "trans r" and
IN1: "b \<in> above a" and IN2: "c \<in> above b"
shows "c \<in> above a"
@@ -352,7 +364,7 @@
thus ?thesis unfolding above_def by simp
qed
-lemma (in rel) above_aboveS_trans:
+lemma 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"
@@ -370,7 +382,7 @@
from 1 3 show ?thesis unfolding aboveS_def by simp
qed
-lemma (in rel) aboveS_above_trans:
+lemma 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"
@@ -388,7 +400,7 @@
from 1 3 show ?thesis unfolding aboveS_def by simp
qed
-lemma (in rel) aboveS_aboveS_trans:
+lemma 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"
@@ -398,7 +410,22 @@
with assms above_aboveS_trans show ?thesis by auto
qed
-lemma (in rel) underS_Under_trans:
+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 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"
@@ -426,7 +453,7 @@
using Under_def by auto
qed
-lemma (in rel) underS_UnderS_trans:
+lemma 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"
@@ -437,7 +464,7 @@
show ?thesis by auto
qed
-lemma (in rel) above_Above_trans:
+lemma above_Above_trans:
assumes TRANS: "trans r" and
IN1: "a \<in> above b" and IN2: "b \<in> Above C"
shows "a \<in> Above C"
@@ -452,7 +479,7 @@
show ?thesis unfolding Above_def by auto
qed
-lemma (in rel) aboveS_Above_trans:
+lemma 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"
@@ -480,7 +507,7 @@
using Above_def by auto
qed
-lemma (in rel) above_AboveS_trans:
+lemma 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"
@@ -508,7 +535,7 @@
using Above_def by auto
qed
-lemma (in rel) aboveS_AboveS_trans:
+lemma 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"
@@ -519,6 +546,35 @@
show ?thesis by auto
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 *)
+
subsection {* Properties depending on more than one relation *}
--- a/src/HOL/Cardinals/Order_Relation_More_Base.thy Wed Nov 20 16:15:54 2013 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,286 +0,0 @@
-(* Title: HOL/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_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/Cardinals/Order_Relation_More_FP.thy Wed Nov 20 16:43:09 2013 +0100
@@ -0,0 +1,230 @@
+(* Title: HOL/Cardinals/Order_Relation_More_FP.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Basics on order-like relations (FP).
+*)
+
+header {* Basics on Order-Like Relations (FP) *}
+
+theory Order_Relation_More_FP
+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_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_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 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
+
+
+end (* context rel *)
+
+end
--- a/src/HOL/Cardinals/README.txt Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Cardinals/README.txt Wed Nov 20 16:43:09 2013 +0100
@@ -89,15 +89,16 @@
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.
+1. Most of the definitions and theorems are proved in files suffixed with "_FP".
+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.
+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,
@@ -148,7 +149,7 @@
Notes for anyone who would like to enrich these theories in the future
--------------------------------------------------------------------------------------
-Theory Fun_More (and Fun_More_Base):
+Theory Fun_More (and Fun_More_FP):
- 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".
@@ -166,7 +167,7 @@
- In subsection "Other facts":
-- Does the lemma "atLeastLessThan_injective" already exist anywhere?
-Theory Order_Relation_More (and Order_Relation_More_Base):
+Theory Order_Relation_More (and Order_Relation_More_FP):
- 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.
@@ -181,16 +182,16 @@
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):
+Theory Wellorder_Relation (and Wellorder_Relation_FP):
- 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):
+Theory Wellfounded_More (and Wellfounded_More_FP):
Recall the lemmas "wfrec" and "wf_induct".
-Theory Wellorder_Embedding (and Wellorder_Embedding_Base):
+Theory Wellorder_Embedding (and Wellorder_Embedding_FP):
- 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?
@@ -200,7 +201,7 @@
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):
+Theory Constructions_on_Wellorders (and Constructions_on_Wellorders_FP):
- 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,
@@ -208,7 +209,7 @@
- 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):
+Theory Cardinal_Order_Relation (and Cardinal_Order_Relation_FP):
- 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 --
--- a/src/HOL/Cardinals/Wellfounded_More.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Cardinals/Wellfounded_More.thy Wed Nov 20 16:43:09 2013 +0100
@@ -8,7 +8,7 @@
header {* More on Well-Founded Relations *}
theory Wellfounded_More
-imports Wellfounded_More_Base Order_Relation_More
+imports Wellfounded_More_FP Order_Relation_More
begin
--- a/src/HOL/Cardinals/Wellfounded_More_Base.thy Wed Nov 20 16:15:54 2013 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,194 +0,0 @@
-(* Title: HOL/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 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 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/Cardinals/Wellfounded_More_FP.thy Wed Nov 20 16:43:09 2013 +0100
@@ -0,0 +1,194 @@
+(* Title: HOL/Cardinals/Wellfounded_More_FP.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+More on well-founded relations (FP).
+*)
+
+header {* More on Well-Founded Relations (FP) *}
+
+theory Wellfounded_More_FP
+imports Order_Relation_More_FP "~~/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 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
--- a/src/HOL/Cardinals/Wellorder_Embedding.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Cardinals/Wellorder_Embedding.thy Wed Nov 20 16:43:09 2013 +0100
@@ -8,7 +8,7 @@
header {* Well-Order Embeddings *}
theory Wellorder_Embedding
-imports Wellorder_Embedding_Base Fun_More Wellorder_Relation
+imports Wellorder_Embedding_FP Fun_More Wellorder_Relation
begin
--- a/src/HOL/Cardinals/Wellorder_Embedding_Base.thy Wed Nov 20 16:15:54 2013 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1146 +0,0 @@
-(* Title: HOL/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 (subst (asm) wo_rel.ofilter_def) fast+
- 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 del: not_False_in_image_Ball)
- 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/Cardinals/Wellorder_Embedding_FP.thy Wed Nov 20 16:43:09 2013 +0100
@@ -0,0 +1,1145 @@
+(* Title: HOL/Cardinals/Wellorder_Embedding_FP.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Well-order embeddings (FP).
+*)
+
+header {* Well-Order Embeddings (FP) *}
+
+theory Wellorder_Embedding_FP
+imports "~~/src/HOL/Library/Zorn" Fun_More_FP Wellorder_Relation_FP
+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 (metis (no_types) 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 (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 (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 force
+ 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 (subst (asm) wo_rel.ofilter_def) fast+
+ 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 del: not_False_in_image_Ball)
+ 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 fast
+ 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 (metis 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
--- a/src/HOL/Cardinals/Wellorder_Relation.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Cardinals/Wellorder_Relation.thy Wed Nov 20 16:43:09 2013 +0100
@@ -8,7 +8,7 @@
header {* Well-Order Relations *}
theory Wellorder_Relation
-imports Wellorder_Relation_Base Wellfounded_More
+imports Wellorder_Relation_FP Wellfounded_More
begin
context wo_rel
@@ -64,17 +64,7 @@
lemma 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
-)
+by(auto simp add: Under_def minim_inField minim_least)
lemma equals_minim_Under:
"\<lbrakk>B \<le> Field r; a \<in> B; a \<in> Under B\<rbrakk>
@@ -410,7 +400,41 @@
qed
-subsubsection {* Properties of order filters *}
+subsubsection {* Properties of order filters *}
+
+lemma ofilter_Under[simp]:
+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[simp]:
+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[simp]: "\<lbrakk>ofilter A; ofilter B\<rbrakk> \<Longrightarrow> ofilter(A Int B)"
+unfolding ofilter_def by blast
+
+lemma ofilter_Un[simp]: "\<lbrakk>ofilter A; ofilter B\<rbrakk> \<Longrightarrow> ofilter(A \<union> B)"
+unfolding ofilter_def by blast
lemma ofilter_INTER:
"\<lbrakk>I \<noteq> {}; \<And> i. i \<in> I \<Longrightarrow> ofilter(A i)\<rbrakk> \<Longrightarrow> ofilter (\<Inter> i \<in> I. A i)"
@@ -496,10 +520,6 @@
under_ofilter[simp]
underS_ofilter[simp]
Field_ofilter[simp]
- ofilter_Under[simp]
- ofilter_UnderS[simp]
- ofilter_Int[simp]
- ofilter_Un[simp]
end
--- a/src/HOL/Cardinals/Wellorder_Relation_Base.thy Wed Nov 20 16:15:54 2013 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,669 +0,0 @@
-(* Title: HOL/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 spec[OF WF[unfolded wf_eq_minimal[of "r - Id"]], of B] NE obtain b where
- *: "b \<in> B \<and> (\<forall>b'. b' \<noteq> b \<and> (b',b) \<in> r \<longrightarrow> b' \<notin> B)" by auto
- 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/Cardinals/Wellorder_Relation_FP.thy Wed Nov 20 16:43:09 2013 +0100
@@ -0,0 +1,631 @@
+(* Title: HOL/Cardinals/Wellorder_Relation_FP.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Well-order relations (FP).
+*)
+
+header {* Well-Order Relations (FP) *}
+
+theory Wellorder_Relation_FP
+imports Wellfounded_More_FP
+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 spec[OF WF[unfolded wf_eq_minimal[of "r - Id"]], of B] NE obtain b where
+ *: "b \<in> B \<and> (\<forall>b'. b' \<noteq> b \<and> (b',b) \<in> r \<longrightarrow> b' \<notin> B)" by auto
+ 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_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
--- a/src/HOL/Code_Numeral.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Code_Numeral.thy Wed Nov 20 16:43:09 2013 +0100
@@ -96,10 +96,6 @@
qed
lemma [transfer_rule]:
- "fun_rel HOL.eq pcr_integer (neg_numeral :: num \<Rightarrow> int) (neg_numeral :: num \<Rightarrow> integer)"
- by (unfold neg_numeral_def [abs_def]) transfer_prover
-
-lemma [transfer_rule]:
"fun_rel HOL.eq (fun_rel HOL.eq pcr_integer) (Num.sub :: _ \<Rightarrow> _ \<Rightarrow> int) (Num.sub :: _ \<Rightarrow> _ \<Rightarrow> integer)"
by (unfold Num.sub_def [abs_def]) transfer_prover
@@ -147,10 +143,6 @@
"int_of_integer (numeral k) = numeral k"
by transfer rule
-lemma int_of_integer_neg_numeral [simp]:
- "int_of_integer (neg_numeral k) = neg_numeral k"
- by transfer rule
-
lemma int_of_integer_sub [simp]:
"int_of_integer (Num.sub k l) = Num.sub k l"
by transfer rule
@@ -253,11 +245,11 @@
definition Neg :: "num \<Rightarrow> integer"
where
- [simp, code_abbrev]: "Neg = neg_numeral"
+ [simp, code_abbrev]: "Neg n = - Pos n"
lemma [transfer_rule]:
- "fun_rel HOL.eq pcr_integer neg_numeral Neg"
- by simp transfer_prover
+ "fun_rel HOL.eq pcr_integer (\<lambda>n. - numeral n) Neg"
+ by (simp add: Neg_def [abs_def]) transfer_prover
code_datatype "0::integer" Pos Neg
@@ -272,7 +264,7 @@
"dup 0 = 0"
"dup (Pos n) = Pos (Num.Bit0 n)"
"dup (Neg n) = Neg (Num.Bit0 n)"
- by (transfer, simp only: neg_numeral_def numeral_Bit0 minus_add_distrib)+
+ by (transfer, simp only: numeral_Bit0 minus_add_distrib)+
lift_definition sub :: "num \<Rightarrow> num \<Rightarrow> integer"
is "\<lambda>m n. numeral m - numeral n :: int"
--- a/src/HOL/Complex.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Complex.thy Wed Nov 20 16:43:09 2013 +0100
@@ -108,7 +108,12 @@
definition complex_divide_def:
"x / (y\<Colon>complex) = x * inverse y"
-lemma Complex_eq_1 [simp]: "(Complex a b = 1) = (a = 1 \<and> b = 0)"
+lemma Complex_eq_1 [simp]:
+ "Complex a b = 1 \<longleftrightarrow> a = 1 \<and> b = 0"
+ by (simp add: complex_one_def)
+
+lemma Complex_eq_neg_1 [simp]:
+ "Complex a b = - 1 \<longleftrightarrow> a = - 1 \<and> b = 0"
by (simp add: complex_one_def)
lemma complex_Re_one [simp]: "Re 1 = 1"
@@ -166,21 +171,21 @@
lemma complex_Re_numeral [simp]: "Re (numeral v) = numeral v"
using complex_Re_of_int [of "numeral v"] by simp
-lemma complex_Re_neg_numeral [simp]: "Re (neg_numeral v) = neg_numeral v"
- using complex_Re_of_int [of "neg_numeral v"] by simp
+lemma complex_Re_neg_numeral [simp]: "Re (- numeral v) = - numeral v"
+ using complex_Re_of_int [of "- numeral v"] by simp
lemma complex_Im_numeral [simp]: "Im (numeral v) = 0"
using complex_Im_of_int [of "numeral v"] by simp
-lemma complex_Im_neg_numeral [simp]: "Im (neg_numeral v) = 0"
- using complex_Im_of_int [of "neg_numeral v"] by simp
+lemma complex_Im_neg_numeral [simp]: "Im (- numeral v) = 0"
+ using complex_Im_of_int [of "- numeral v"] by simp
lemma Complex_eq_numeral [simp]:
- "(Complex a b = numeral w) = (a = numeral w \<and> b = 0)"
+ "Complex a b = numeral w \<longleftrightarrow> a = numeral w \<and> b = 0"
by (simp add: complex_eq_iff)
lemma Complex_eq_neg_numeral [simp]:
- "(Complex a b = neg_numeral w) = (a = neg_numeral w \<and> b = 0)"
+ "Complex a b = - numeral w \<longleftrightarrow> a = - numeral w \<and> b = 0"
by (simp add: complex_eq_iff)
@@ -421,7 +426,7 @@
lemma complex_i_not_numeral [simp]: "ii \<noteq> numeral w"
by (simp add: complex_eq_iff)
-lemma complex_i_not_neg_numeral [simp]: "ii \<noteq> neg_numeral w"
+lemma complex_i_not_neg_numeral [simp]: "ii \<noteq> - numeral w"
by (simp add: complex_eq_iff)
lemma i_mult_Complex [simp]: "ii * Complex a b = Complex (- b) a"
@@ -508,7 +513,7 @@
lemma complex_cnj_numeral [simp]: "cnj (numeral w) = numeral w"
by (simp add: complex_eq_iff)
-lemma complex_cnj_neg_numeral [simp]: "cnj (neg_numeral w) = neg_numeral w"
+lemma complex_cnj_neg_numeral [simp]: "cnj (- numeral w) = - numeral w"
by (simp add: complex_eq_iff)
lemma complex_cnj_scaleR: "cnj (scaleR r x) = scaleR r (cnj x)"
--- a/src/HOL/Decision_Procs/Approximation.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Decision_Procs/Approximation.thy Wed Nov 20 16:43:09 2013 +0100
@@ -11,8 +11,10 @@
"~~/src/HOL/Library/Code_Target_Numeral"
begin
-declare powr_numeral[simp]
-declare powr_neg_numeral[simp]
+declare powr_one [simp]
+declare powr_numeral [simp]
+declare powr_neg_one [simp]
+declare powr_neg_numeral [simp]
section "Horner Scheme"
@@ -1261,8 +1263,8 @@
unfolding cos_periodic_int ..
also have "\<dots> \<le> cos ((?ux - 2 * ?lpi))"
using cos_monotone_minus_pi_0'[OF pi_x x_le_ux ux_0]
- by (simp only: minus_float.rep_eq real_of_int_minus real_of_one minus_one[symmetric]
- mult_minus_left mult_1_left) simp
+ by (simp only: minus_float.rep_eq real_of_int_minus real_of_one
+ mult_minus_left mult_1_left) simp
also have "\<dots> = cos ((- (?ux - 2 * ?lpi)))"
unfolding uminus_float.rep_eq cos_minus ..
also have "\<dots> \<le> (ub_cos prec (- (?ux - 2 * ?lpi)))"
@@ -1306,7 +1308,7 @@
also have "\<dots> \<le> cos ((?lx + 2 * ?lpi))"
using cos_monotone_0_pi'[OF lx_0 lx_le_x pi_x]
by (simp only: minus_float.rep_eq real_of_int_minus real_of_one
- minus_one[symmetric] mult_minus_left mult_1_left) simp
+ mult_minus_left mult_1_left) simp
also have "\<dots> \<le> (ub_cos prec (?lx + 2 * ?lpi))"
using lb_cos[OF lx_0 pi_lx] by simp
finally show ?thesis unfolding u by (simp add: real_of_float_max)
@@ -2104,8 +2106,9 @@
lemma interpret_floatarith_num:
shows "interpret_floatarith (Num (Float 0 0)) vs = 0"
and "interpret_floatarith (Num (Float 1 0)) vs = 1"
+ and "interpret_floatarith (Num (Float (- 1) 0)) vs = - 1"
and "interpret_floatarith (Num (Float (numeral a) 0)) vs = numeral a"
- and "interpret_floatarith (Num (Float (neg_numeral a) 0)) vs = neg_numeral a" by auto
+ and "interpret_floatarith (Num (Float (- numeral a) 0)) vs = - numeral a" by auto
subsection "Implement approximation function"
--- a/src/HOL/Decision_Procs/Cooper.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Decision_Procs/Cooper.thy Wed Nov 20 16:43:09 2013 +0100
@@ -2006,9 +2006,10 @@
| SOME n => @{code Bound} (@{code nat_of_integer} n))
| num_of_term vs @{term "0::int"} = @{code C} (@{code int_of_integer} 0)
| num_of_term vs @{term "1::int"} = @{code C} (@{code int_of_integer} 1)
+ | num_of_term vs @{term "- 1::int"} = @{code C} (@{code int_of_integer} (~ 1))
| num_of_term vs (@{term "numeral :: _ \<Rightarrow> int"} $ t) =
@{code C} (@{code int_of_integer} (HOLogic.dest_num t))
- | num_of_term vs (@{term "neg_numeral :: _ \<Rightarrow> int"} $ t) =
+ | num_of_term vs (@{term "- numeral :: _ \<Rightarrow> int"} $ t) =
@{code C} (@{code int_of_integer} (~(HOLogic.dest_num t)))
| num_of_term vs (Bound i) = @{code Bound} (@{code nat_of_integer} i)
| num_of_term vs (@{term "uminus :: int \<Rightarrow> int"} $ t') = @{code Neg} (num_of_term vs t')
--- a/src/HOL/Decision_Procs/MIR.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Decision_Procs/MIR.thy Wed Nov 20 16:43:09 2013 +0100
@@ -3154,7 +3154,7 @@
hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = (- ?N i e + real j1) - 1"
by (simp only: algebra_simps)
hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = - 1 - ?N i e + real j1"
- by (simp add: algebra_simps minus_one [symmetric] del: minus_one)
+ by (simp add: algebra_simps)
with nob have ?case by blast }
ultimately show ?case by blast
next
@@ -5549,6 +5549,7 @@
| num_of_term vs @{term "real (1::int)"} = mk_C 1
| num_of_term vs @{term "0::real"} = mk_C 0
| num_of_term vs @{term "1::real"} = mk_C 1
+ | num_of_term vs @{term "- 1::real"} = mk_C (~ 1)
| num_of_term vs (Bound i) = mk_Bound i
| num_of_term vs (@{term "uminus :: real \<Rightarrow> real"} $ t') = @{code Neg} (num_of_term vs t')
| num_of_term vs (@{term "op + :: real \<Rightarrow> real \<Rightarrow> real"} $ t1 $ t2) =
@@ -5561,7 +5562,7 @@
| _ => error "num_of_term: unsupported Multiplication")
| num_of_term vs (@{term "real :: int \<Rightarrow> real"} $ (@{term "numeral :: _ \<Rightarrow> int"} $ t')) =
mk_C (HOLogic.dest_num t')
- | num_of_term vs (@{term "real :: int \<Rightarrow> real"} $ (@{term "neg_numeral :: _ \<Rightarrow> int"} $ t')) =
+ | num_of_term vs (@{term "real :: int \<Rightarrow> real"} $ (@{term "- numeral :: _ \<Rightarrow> int"} $ t')) =
mk_C (~ (HOLogic.dest_num t'))
| num_of_term vs (@{term "real :: int \<Rightarrow> real"} $ (@{term "floor :: real \<Rightarrow> int"} $ t')) =
@{code Floor} (num_of_term vs t')
@@ -5569,7 +5570,7 @@
@{code Neg} (@{code Floor} (@{code Neg} (num_of_term vs t')))
| num_of_term vs (@{term "numeral :: _ \<Rightarrow> real"} $ t') =
mk_C (HOLogic.dest_num t')
- | num_of_term vs (@{term "neg_numeral :: _ \<Rightarrow> real"} $ t') =
+ | num_of_term vs (@{term "- numeral :: _ \<Rightarrow> real"} $ t') =
mk_C (~ (HOLogic.dest_num t'))
| num_of_term vs t = error ("num_of_term: unknown term " ^ Syntax.string_of_term @{context} t);
@@ -5583,7 +5584,7 @@
@{code Eq} (@{code Sub} (num_of_term vs t1, num_of_term vs t2))
| fm_of_term vs (@{term "op rdvd"} $ (@{term "real :: int \<Rightarrow> real"} $ (@{term "numeral :: _ \<Rightarrow> int"} $ t1)) $ t2) =
mk_Dvd (HOLogic.dest_num t1, num_of_term vs t2)
- | fm_of_term vs (@{term "op rdvd"} $ (@{term "real :: int \<Rightarrow> real"} $ (@{term "neg_numeral :: _ \<Rightarrow> int"} $ t1)) $ t2) =
+ | fm_of_term vs (@{term "op rdvd"} $ (@{term "real :: int \<Rightarrow> real"} $ (@{term "- numeral :: _ \<Rightarrow> int"} $ t1)) $ t2) =
mk_Dvd (~ (HOLogic.dest_num t1), num_of_term vs t2)
| fm_of_term vs (@{term "op = :: bool \<Rightarrow> bool \<Rightarrow> bool"} $ t1 $ t2) =
@{code Iff} (fm_of_term vs t1, fm_of_term vs t2)
--- a/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy Wed Nov 20 16:43:09 2013 +0100
@@ -1256,12 +1256,10 @@
apply (case_tac n', simp, simp)
apply (case_tac n, simp, simp)
apply (case_tac n, case_tac n', simp add: Let_def)
- apply (case_tac "pa +\<^sub>p p' = 0\<^sub>p")
- apply (auto simp add: polyadd_eq_const_degree)
+ apply (auto simp add: polyadd_eq_const_degree)[2]
apply (metis head_nz)
apply (metis head_nz)
apply (metis degree.simps(9) gr0_conv_Suc head.simps(1) less_Suc0 not_less_eq)
- apply (metis degree.simps(9) gr0_conv_Suc nat_less_le order_le_less_trans)
done
lemma polymul_head_polyeq:
--- a/src/HOL/Divides.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Divides.thy Wed Nov 20 16:43:09 2013 +0100
@@ -1919,10 +1919,9 @@
val zero = @{term "0 :: int"}
val less = @{term "op < :: int \<Rightarrow> int \<Rightarrow> bool"}
val le = @{term "op \<le> :: int \<Rightarrow> int \<Rightarrow> bool"}
- val simps = @{thms arith_simps} @ @{thms rel_simps} @
- map (fn th => th RS sym) [@{thm numeral_1_eq_1}]
- fun prove ctxt goal = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop goal)
- (K (ALLGOALS (full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps simps))));
+ val simps = @{thms arith_simps} @ @{thms rel_simps} @ [@{thm numeral_1_eq_1 [symmetric]}]
+ fun prove ctxt goal = (writeln "prove"; Goal.prove ctxt [] [] (HOLogic.mk_Trueprop goal)
+ (K (ALLGOALS (full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps simps)))));
fun binary_proc proc ctxt ct =
(case Thm.term_of ct of
_ $ t $ u =>
@@ -1945,23 +1944,23 @@
simproc_setup binary_int_div
("numeral m div numeral n :: int" |
- "numeral m div neg_numeral n :: int" |
- "neg_numeral m div numeral n :: int" |
- "neg_numeral m div neg_numeral n :: int") =
+ "numeral m div - numeral n :: int" |
+ "- numeral m div numeral n :: int" |
+ "- numeral m div - numeral n :: int") =
{* K (divmod_proc @{thm int_div_pos_eq} @{thm int_div_neg_eq}) *}
simproc_setup binary_int_mod
("numeral m mod numeral n :: int" |
- "numeral m mod neg_numeral n :: int" |
- "neg_numeral m mod numeral n :: int" |
- "neg_numeral m mod neg_numeral n :: int") =
+ "numeral m mod - numeral n :: int" |
+ "- numeral m mod numeral n :: int" |
+ "- numeral m mod - numeral n :: int") =
{* K (divmod_proc @{thm int_mod_pos_eq} @{thm int_mod_neg_eq}) *}
lemmas posDivAlg_eqn_numeral [simp] =
posDivAlg_eqn [of "numeral v" "numeral w", OF zero_less_numeral] for v w
lemmas negDivAlg_eqn_numeral [simp] =
- negDivAlg_eqn [of "numeral v" "neg_numeral w", OF zero_less_numeral] for v w
+ negDivAlg_eqn [of "numeral v" "- numeral w", OF zero_less_numeral] for v w
text{*Special-case simplification *}
@@ -1973,14 +1972,14 @@
div_pos_pos [OF zero_less_one, of "numeral w", OF zero_le_numeral] for w
lemmas div_pos_neg_1_numeral [simp] =
- div_pos_neg [OF zero_less_one, of "neg_numeral w",
+ div_pos_neg [OF zero_less_one, of "- numeral w",
OF neg_numeral_less_zero] for w
lemmas mod_pos_pos_1_numeral [simp] =
mod_pos_pos [OF zero_less_one, of "numeral w", OF zero_le_numeral] for w
lemmas mod_pos_neg_1_numeral [simp] =
- mod_pos_neg [OF zero_less_one, of "neg_numeral w",
+ mod_pos_neg [OF zero_less_one, of "- numeral w",
OF neg_numeral_less_zero] for w
lemmas posDivAlg_eqn_1_numeral [simp] =
@@ -2290,6 +2289,8 @@
shows "divmod_int_rel (1 + 2*a) (2*b) (q, 1 + 2*r)"
using assms unfolding divmod_int_rel_def by auto
+declaration {* K (Lin_Arith.add_simps @{thms uminus_numeral_One}) *}
+
lemma neg_divmod_int_rel_mult_2:
assumes "b \<le> 0"
assumes "divmod_int_rel (a + 1) b (q, r)"
@@ -2427,13 +2428,13 @@
lemma dvd_neg_numeral_left [simp]:
fixes y :: "'a::comm_ring_1"
- shows "(neg_numeral k) dvd y \<longleftrightarrow> (numeral k) dvd y"
- unfolding neg_numeral_def minus_dvd_iff ..
+ shows "(- numeral k) dvd y \<longleftrightarrow> (numeral k) dvd y"
+ by (fact minus_dvd_iff)
lemma dvd_neg_numeral_right [simp]:
fixes x :: "'a::comm_ring_1"
- shows "x dvd (neg_numeral k) \<longleftrightarrow> x dvd (numeral k)"
- unfolding neg_numeral_def dvd_minus_iff ..
+ shows "x dvd (- numeral k) \<longleftrightarrow> x dvd (numeral k)"
+ by (fact dvd_minus_iff)
lemmas dvd_eq_mod_eq_0_numeral [simp] =
dvd_eq_mod_eq_0 [of "numeral x" "numeral y"] for x y
--- a/src/HOL/GCD.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/GCD.thy Wed Nov 20 16:43:09 2013 +0100
@@ -134,6 +134,14 @@
lemma gcd_neg2_int [simp]: "gcd (x::int) (-y) = gcd x y"
by (simp add: gcd_int_def)
+lemma gcd_neg_numeral_1_int [simp]:
+ "gcd (- numeral n :: int) x = gcd (numeral n) x"
+ by (fact gcd_neg1_int)
+
+lemma gcd_neg_numeral_2_int [simp]:
+ "gcd x (- numeral n :: int) = gcd x (numeral n)"
+ by (fact gcd_neg2_int)
+
lemma abs_gcd_int[simp]: "abs(gcd (x::int) y) = gcd x y"
by(simp add: gcd_int_def)
--- a/src/HOL/IMP/Hoare_Examples.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/IMP/Hoare_Examples.thy Wed Nov 20 16:43:09 2013 +0100
@@ -2,17 +2,6 @@
theory Hoare_Examples imports Hoare begin
-text{* Improves proof automation for negative numerals: *}
-
-lemma add_neg1R[simp]:
- "x + -1 = x - (1 :: int)"
-by arith
-
-lemma add_neg_numeralR[simp]:
- "x + neg_numeral n = (x::'a::neg_numeral) - numeral(n)"
-by (simp only: diff_minus_eq_add[symmetric] minus_neg_numeral)
-
-
text{* Summing up the first @{text x} natural numbers in variable @{text y}. *}
fun sum :: "int \<Rightarrow> int" where
--- a/src/HOL/Int.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Int.thy Wed Nov 20 16:43:09 2013 +0100
@@ -232,9 +232,8 @@
lemma of_int_numeral [simp, code_post]: "of_int (numeral k) = numeral k"
by (simp add: of_nat_numeral [symmetric] of_int_of_nat_eq [symmetric])
-lemma of_int_neg_numeral [simp, code_post]: "of_int (neg_numeral k) = neg_numeral k"
- unfolding neg_numeral_def neg_numeral_class.neg_numeral_def
- by (simp only: of_int_minus of_int_numeral)
+lemma of_int_neg_numeral [code_post]: "of_int (- numeral k) = - numeral k"
+ by simp
lemma of_int_power:
"of_int (z ^ n) = of_int z ^ n"
@@ -370,7 +369,7 @@
by (simp add: nat_eq_iff)
lemma nat_neg_numeral [simp]:
- "nat (neg_numeral k) = 0"
+ "nat (- numeral k) = 0"
by simp
lemma nat_2: "nat 2 = Suc (Suc 0)"
@@ -511,13 +510,13 @@
lemma nonneg_int_cases:
assumes "0 \<le> k" obtains n where "k = int n"
- using assms by (cases k, simp, simp del: of_nat_Suc)
+ using assms by (rule nonneg_eq_int)
lemma Let_numeral [simp]: "Let (numeral v) f = f (numeral v)"
-- {* Unfold all @{text let}s involving constants *}
unfolding Let_def ..
-lemma Let_neg_numeral [simp]: "Let (neg_numeral v) f = f (neg_numeral v)"
+lemma Let_neg_numeral [simp]: "Let (- numeral v) f = f (- numeral v)"
-- {* Unfold all @{text let}s involving constants *}
unfolding Let_def ..
@@ -525,15 +524,15 @@
lemmas max_number_of [simp] =
max_def [of "numeral u" "numeral v"]
- max_def [of "numeral u" "neg_numeral v"]
- max_def [of "neg_numeral u" "numeral v"]
- max_def [of "neg_numeral u" "neg_numeral v"] for u v
+ max_def [of "numeral u" "- numeral v"]
+ max_def [of "- numeral u" "numeral v"]
+ max_def [of "- numeral u" "- numeral v"] for u v
lemmas min_number_of [simp] =
min_def [of "numeral u" "numeral v"]
- min_def [of "numeral u" "neg_numeral v"]
- min_def [of "neg_numeral u" "numeral v"]
- min_def [of "neg_numeral u" "neg_numeral v"] for u v
+ min_def [of "numeral u" "- numeral v"]
+ min_def [of "- numeral u" "numeral v"]
+ min_def [of "- numeral u" "- numeral v"] for u v
subsubsection {* Binary comparisons *}
@@ -1070,8 +1069,6 @@
by auto
qed
-ML_val {* @{const_name neg_numeral} *}
-
lemma pos_zmult_eq_1_iff_lemma: "(m * n = 1) ==> m = (1::int) | m = -1"
by (insert abs_zmult_eq_1 [of m n], arith)
@@ -1127,62 +1124,30 @@
inverse_eq_divide [of "numeral w"] for w
lemmas inverse_eq_divide_neg_numeral [simp] =
- inverse_eq_divide [of "neg_numeral w"] for w
+ inverse_eq_divide [of "- numeral w"] for w
text {*These laws simplify inequalities, moving unary minus from a term
into the literal.*}
-lemmas le_minus_iff_numeral [simp, no_atp] =
- le_minus_iff [of "numeral v"]
- le_minus_iff [of "neg_numeral v"] for v
-
-lemmas equation_minus_iff_numeral [simp, no_atp] =
- equation_minus_iff [of "numeral v"]
- equation_minus_iff [of "neg_numeral v"] for v
+lemmas equation_minus_iff_numeral [no_atp] =
+ equation_minus_iff [of "numeral v"] for v
-lemmas minus_less_iff_numeral [simp, no_atp] =
- minus_less_iff [of _ "numeral v"]
- minus_less_iff [of _ "neg_numeral v"] for v
+lemmas minus_equation_iff_numeral [no_atp] =
+ minus_equation_iff [of _ "numeral v"] for v
-lemmas minus_le_iff_numeral [simp, no_atp] =
- minus_le_iff [of _ "numeral v"]
- minus_le_iff [of _ "neg_numeral v"] for v
-
-lemmas minus_equation_iff_numeral [simp, no_atp] =
- minus_equation_iff [of _ "numeral v"]
- minus_equation_iff [of _ "neg_numeral v"] for v
-
-text{*To Simplify Inequalities Where One Side is the Constant 1*}
+lemmas le_minus_iff_numeral [no_atp] =
+ le_minus_iff [of "numeral v"] for v
-lemma less_minus_iff_1 [simp]:
- fixes b::"'b::linordered_idom"
- shows "(1 < - b) = (b < -1)"
-by auto
-
-lemma le_minus_iff_1 [simp]:
- fixes b::"'b::linordered_idom"
- shows "(1 \<le> - b) = (b \<le> -1)"
-by auto
-
-lemma equation_minus_iff_1 [simp]:
- fixes b::"'b::ring_1"
- shows "(1 = - b) = (b = -1)"
-by (subst equation_minus_iff, auto)
+lemmas minus_le_iff_numeral [no_atp] =
+ minus_le_iff [of _ "numeral v"] for v
-lemma minus_less_iff_1 [simp]:
- fixes a::"'b::linordered_idom"
- shows "(- a < 1) = (-1 < a)"
-by auto
+lemmas less_minus_iff_numeral [no_atp] =
+ less_minus_iff [of "numeral v"] for v
-lemma minus_le_iff_1 [simp]:
- fixes a::"'b::linordered_idom"
- shows "(- a \<le> 1) = (-1 \<le> a)"
-by auto
+lemmas minus_less_iff_numeral [no_atp] =
+ minus_less_iff [of _ "numeral v"] for v
-lemma minus_equation_iff_1 [simp]:
- fixes a::"'b::ring_1"
- shows "(- a = 1) = (a = -1)"
-by (subst minus_equation_iff, auto)
+-- {* FIXME maybe simproc *}
text {*Cancellation of constant factors in comparisons (@{text "<"} and @{text "\<le>"}) *}
@@ -1197,27 +1162,28 @@
lemmas le_divide_eq_numeral1 [simp] =
pos_le_divide_eq [of "numeral w", OF zero_less_numeral]
- neg_le_divide_eq [of "neg_numeral w", OF neg_numeral_less_zero] for w
+ neg_le_divide_eq [of "- numeral w", OF neg_numeral_less_zero] for w
lemmas divide_le_eq_numeral1 [simp] =
pos_divide_le_eq [of "numeral w", OF zero_less_numeral]
- neg_divide_le_eq [of "neg_numeral w", OF neg_numeral_less_zero] for w
+ neg_divide_le_eq [of "- numeral w", OF neg_numeral_less_zero] for w
lemmas less_divide_eq_numeral1 [simp] =
pos_less_divide_eq [of "numeral w", OF zero_less_numeral]
- neg_less_divide_eq [of "neg_numeral w", OF neg_numeral_less_zero] for w
+ neg_less_divide_eq [of "- numeral w", OF neg_numeral_less_zero] for w
lemmas divide_less_eq_numeral1 [simp] =
pos_divide_less_eq [of "numeral w", OF zero_less_numeral]
- neg_divide_less_eq [of "neg_numeral w", OF neg_numeral_less_zero] for w
+ neg_divide_less_eq [of "- numeral w", OF neg_numeral_less_zero] for w
lemmas eq_divide_eq_numeral1 [simp] =
eq_divide_eq [of _ _ "numeral w"]
- eq_divide_eq [of _ _ "neg_numeral w"] for w
+ eq_divide_eq [of _ _ "- numeral w"] for w
lemmas divide_eq_eq_numeral1 [simp] =
divide_eq_eq [of _ "numeral w"]
- divide_eq_eq [of _ "neg_numeral w"] for w
+ divide_eq_eq [of _ "- numeral w"] for w
+
subsubsection{*Optional Simplification Rules Involving Constants*}
@@ -1225,27 +1191,27 @@
lemmas le_divide_eq_numeral =
le_divide_eq [of "numeral w"]
- le_divide_eq [of "neg_numeral w"] for w
+ le_divide_eq [of "- numeral w"] for w
lemmas divide_le_eq_numeral =
divide_le_eq [of _ _ "numeral w"]
- divide_le_eq [of _ _ "neg_numeral w"] for w
+ divide_le_eq [of _ _ "- numeral w"] for w
lemmas less_divide_eq_numeral =
less_divide_eq [of "numeral w"]
- less_divide_eq [of "neg_numeral w"] for w
+ less_divide_eq [of "- numeral w"] for w
lemmas divide_less_eq_numeral =
divide_less_eq [of _ _ "numeral w"]
- divide_less_eq [of _ _ "neg_numeral w"] for w
+ divide_less_eq [of _ _ "- numeral w"] for w
lemmas eq_divide_eq_numeral =
eq_divide_eq [of "numeral w"]
- eq_divide_eq [of "neg_numeral w"] for w
+ eq_divide_eq [of "- numeral w"] for w
lemmas divide_eq_eq_numeral =
divide_eq_eq [of _ _ "numeral w"]
- divide_eq_eq [of _ _ "neg_numeral w"] for w
+ divide_eq_eq [of _ _ "- numeral w"] for w
text{*Not good as automatic simprules because they cause case splits.*}
@@ -1257,21 +1223,20 @@
text{*Division By @{text "-1"}*}
lemma divide_minus1 [simp]: "(x::'a::field) / -1 = - x"
- unfolding minus_one [symmetric]
unfolding nonzero_minus_divide_right [OF one_neq_zero, symmetric]
by simp
lemma minus1_divide [simp]: "-1 / (x::'a::field) = - (1 / x)"
- unfolding minus_one [symmetric] by (rule divide_minus_left)
+ by (fact divide_minus_left)
lemma half_gt_zero_iff:
- "(0 < r/2) = (0 < (r::'a::linordered_field_inverse_zero))"
-by auto
+ "(0 < r/2) = (0 < (r::'a::linordered_field_inverse_zero))"
+ by auto
lemmas half_gt_zero [simp] = half_gt_zero_iff [THEN iffD2]
lemma divide_Numeral1: "(x::'a::field) / Numeral1 = x"
- by simp
+ by (fact divide_numeral_1)
subsection {* The divides relation *}
@@ -1475,7 +1440,7 @@
[simp, code_abbrev]: "Pos = numeral"
definition Neg :: "num \<Rightarrow> int" where
- [simp, code_abbrev]: "Neg = neg_numeral"
+ [simp, code_abbrev]: "Neg n = - (Pos n)"
code_datatype "0::int" Pos Neg
@@ -1489,7 +1454,7 @@
"dup 0 = 0"
"dup (Pos n) = Pos (Num.Bit0 n)"
"dup (Neg n) = Neg (Num.Bit0 n)"
- unfolding Pos_def Neg_def neg_numeral_def
+ unfolding Pos_def Neg_def
by (simp_all add: numeral_Bit0)
definition sub :: "num \<Rightarrow> num \<Rightarrow> int" where
@@ -1505,12 +1470,10 @@
"sub (Num.Bit1 m) (Num.Bit1 n) = dup (sub m n)"
"sub (Num.Bit1 m) (Num.Bit0 n) = dup (sub m n) + 1"
"sub (Num.Bit0 m) (Num.Bit1 n) = dup (sub m n) - 1"
- apply (simp_all only: sub_def dup_def numeral.simps Pos_def Neg_def
- neg_numeral_def numeral_BitM)
+ apply (simp_all only: sub_def dup_def numeral.simps Pos_def Neg_def numeral_BitM)
apply (simp_all only: algebra_simps minus_diff_eq)
apply (simp_all only: add.commute [of _ "- (numeral n + numeral n)"])
apply (simp_all only: minus_add add.assoc left_minus)
- apply (simp_all only: algebra_simps right_minus)
done
text {* Implementations *}
@@ -1606,10 +1569,10 @@
"nat (Int.Neg k) = 0"
"nat 0 = 0"
"nat (Int.Pos k) = nat_of_num k"
- by (simp_all add: nat_of_num_numeral nat_numeral)
+ by (simp_all add: nat_of_num_numeral)
lemma (in ring_1) of_int_code [code]:
- "of_int (Int.Neg k) = neg_numeral k"
+ "of_int (Int.Neg k) = - numeral k"
"of_int 0 = 0"
"of_int (Int.Pos k) = numeral k"
by simp_all
@@ -1653,7 +1616,7 @@
lemma int_power:
"int (m ^ n) = int m ^ n"
- by (rule of_nat_power)
+ by (fact of_nat_power)
lemmas zpower_int = int_power [symmetric]
--- a/src/HOL/Library/Binomial.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Library/Binomial.thy Wed Nov 20 16:43:09 2013 +0100
@@ -370,7 +370,7 @@
by auto
from False show ?thesis
by (simp add: pochhammer_def gbinomial_def field_simps
- eq setprod_timesf[symmetric] del: minus_one)
+ eq setprod_timesf[symmetric])
qed
lemma binomial_fact_lemma: "k \<le> n \<Longrightarrow> fact k * fact (n - k) * (n choose k) = fact n"
@@ -441,9 +441,9 @@
from eq[symmetric]
have ?thesis using kn
apply (simp add: binomial_fact[OF kn, where ?'a = 'a]
- gbinomial_pochhammer field_simps pochhammer_Suc_setprod del: minus_one)
+ gbinomial_pochhammer field_simps pochhammer_Suc_setprod)
apply (simp add: pochhammer_Suc_setprod fact_altdef_nat h
- of_nat_setprod setprod_timesf[symmetric] eq' del: One_nat_def power_Suc del: minus_one)
+ of_nat_setprod setprod_timesf[symmetric] eq' del: One_nat_def power_Suc)
unfolding setprod_Un_disjoint[OF th0, unfolded eq3, of "of_nat:: nat \<Rightarrow> 'a"] eq[unfolded h]
unfolding mult_assoc[symmetric]
unfolding setprod_timesf[symmetric]
--- a/src/HOL/Library/Bit.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Library/Bit.thy Wed Nov 20 16:43:09 2013 +0100
@@ -147,11 +147,11 @@
text {* All numerals reduce to either 0 or 1. *}
-lemma bit_minus1 [simp]: "-1 = (1 :: bit)"
- by (simp only: minus_one [symmetric] uminus_bit_def)
+lemma bit_minus1 [simp]: "- 1 = (1 :: bit)"
+ by (simp only: uminus_bit_def)
-lemma bit_neg_numeral [simp]: "(neg_numeral w :: bit) = numeral w"
- by (simp only: neg_numeral_def uminus_bit_def)
+lemma bit_neg_numeral [simp]: "(- numeral w :: bit) = numeral w"
+ by (simp only: uminus_bit_def)
lemma bit_numeral_even [simp]: "numeral (Num.Bit0 w) = (0 :: bit)"
by (simp only: numeral_Bit0 bit_add_self)
--- a/src/HOL/Library/Code_Prolog.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Library/Code_Prolog.thy Wed Nov 20 16:43:09 2013 +0100
@@ -12,10 +12,8 @@
section {* Setup for Numerals *}
-setup {* Predicate_Compile_Data.ignore_consts
- [@{const_name numeral}, @{const_name neg_numeral}] *}
+setup {* Predicate_Compile_Data.ignore_consts [@{const_name numeral}] *}
-setup {* Predicate_Compile_Data.keep_functions
- [@{const_name numeral}, @{const_name neg_numeral}] *}
+setup {* Predicate_Compile_Data.keep_functions [@{const_name numeral}] *}
end
--- a/src/HOL/Library/Code_Real_Approx_By_Float.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Library/Code_Real_Approx_By_Float.thy Wed Nov 20 16:43:09 2013 +0100
@@ -169,7 +169,7 @@
by simp
lemma [code_unfold del]:
- "neg_numeral k \<equiv> (of_rat (neg_numeral k) :: real)"
+ "- numeral k \<equiv> (of_rat (- numeral k) :: real)"
by simp
hide_const (open) real_of_int
--- a/src/HOL/Library/Code_Target_Int.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Library/Code_Target_Int.thy Wed Nov 20 16:43:09 2013 +0100
@@ -30,7 +30,7 @@
by transfer simp
lemma [code_abbrev]:
- "int_of_integer (neg_numeral k) = Int.Neg k"
+ "int_of_integer (- numeral k) = Int.Neg k"
by transfer simp
lemma [code, symmetric, code_post]:
--- a/src/HOL/Library/Extended.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Library/Extended.thy Wed Nov 20 16:43:09 2013 +0100
@@ -161,8 +161,8 @@
apply (simp only: numeral_inc one_extended_def plus_extended.simps(1)[symmetric])
done
-lemma Fin_neg_numeral: "Fin(neg_numeral w) = - numeral w"
-by (simp only: Fin_numeral minus_numeral[symmetric] uminus_extended.simps[symmetric])
+lemma Fin_neg_numeral: "Fin (- numeral w) = - numeral w"
+by (simp only: Fin_numeral uminus_extended.simps[symmetric])
instantiation extended :: (lattice)bounded_lattice
--- a/src/HOL/Library/Float.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Library/Float.thy Wed Nov 20 16:43:09 2013 +0100
@@ -45,7 +45,7 @@
lemma zero_float[simp]: "0 \<in> float" by (auto simp: float_def)
lemma one_float[simp]: "1 \<in> float" by (intro floatI[of 1 0]) simp
lemma numeral_float[simp]: "numeral i \<in> float" by (intro floatI[of "numeral i" 0]) simp
-lemma neg_numeral_float[simp]: "neg_numeral i \<in> float" by (intro floatI[of "neg_numeral i" 0]) simp
+lemma neg_numeral_float[simp]: "- numeral i \<in> float" by (intro floatI[of "- numeral i" 0]) simp
lemma real_of_int_float[simp]: "real (x :: int) \<in> float" by (intro floatI[of x 0]) simp
lemma real_of_nat_float[simp]: "real (x :: nat) \<in> float" by (intro floatI[of x 0]) simp
lemma two_powr_int_float[simp]: "2 powr (real (i::int)) \<in> float" by (intro floatI[of 1 i]) simp
@@ -53,7 +53,7 @@
lemma two_powr_minus_int_float[simp]: "2 powr - (real (i::int)) \<in> float" by (intro floatI[of 1 "-i"]) simp
lemma two_powr_minus_nat_float[simp]: "2 powr - (real (i::nat)) \<in> float" by (intro floatI[of 1 "-i"]) simp
lemma two_powr_numeral_float[simp]: "2 powr numeral i \<in> float" by (intro floatI[of 1 "numeral i"]) simp
-lemma two_powr_neg_numeral_float[simp]: "2 powr neg_numeral i \<in> float" by (intro floatI[of 1 "neg_numeral i"]) simp
+lemma two_powr_neg_numeral_float[simp]: "2 powr - numeral i \<in> float" by (intro floatI[of 1 "- numeral i"]) simp
lemma two_pow_float[simp]: "2 ^ n \<in> float" by (intro floatI[of 1 "n"]) (simp add: powr_realpow)
lemma real_of_float_float[simp]: "real (f::float) \<in> float" by (cases f) simp
@@ -121,11 +121,11 @@
qed
lemma div_neg_numeral_Bit0_float[simp]:
- assumes x: "x / numeral n \<in> float" shows "x / (neg_numeral (Num.Bit0 n)) \<in> float"
+ assumes x: "x / numeral n \<in> float" shows "x / (- numeral (Num.Bit0 n)) \<in> float"
proof -
have "- (x / numeral (Num.Bit0 n)) \<in> float" using x by simp
- also have "- (x / numeral (Num.Bit0 n)) = x / neg_numeral (Num.Bit0 n)"
- unfolding neg_numeral_def by (simp del: minus_numeral)
+ also have "- (x / numeral (Num.Bit0 n)) = x / - numeral (Num.Bit0 n)"
+ by simp
finally show ?thesis .
qed
@@ -197,7 +197,7 @@
then show "\<exists>c. a < c \<and> c < b"
apply (intro exI[of _ "(a + b) * Float 1 -1"])
apply transfer
- apply (simp add: powr_neg_numeral)
+ apply (simp add: powr_minus)
done
qed
@@ -226,16 +226,16 @@
"fun_rel (op =) pcr_float (numeral :: _ \<Rightarrow> real) (numeral :: _ \<Rightarrow> float)"
unfolding fun_rel_def float.pcr_cr_eq cr_float_def by simp
-lemma float_neg_numeral[simp]: "real (neg_numeral x :: float) = neg_numeral x"
- by (simp add: minus_numeral[symmetric] del: minus_numeral)
+lemma float_neg_numeral[simp]: "real (- numeral x :: float) = - numeral x"
+ by simp
lemma transfer_neg_numeral [transfer_rule]:
- "fun_rel (op =) pcr_float (neg_numeral :: _ \<Rightarrow> real) (neg_numeral :: _ \<Rightarrow> float)"
+ "fun_rel (op =) pcr_float (- numeral :: _ \<Rightarrow> real) (- numeral :: _ \<Rightarrow> float)"
unfolding fun_rel_def float.pcr_cr_eq cr_float_def by simp
lemma
shows float_of_numeral[simp]: "numeral k = float_of (numeral k)"
- and float_of_neg_numeral[simp]: "neg_numeral k = float_of (neg_numeral k)"
+ and float_of_neg_numeral[simp]: "- numeral k = float_of (- numeral k)"
unfolding real_of_float_eq by simp_all
subsection {* Represent floats as unique mantissa and exponent *}
@@ -439,7 +439,7 @@
by transfer simp
hide_fact (open) compute_float_numeral
-lemma compute_float_neg_numeral[code_abbrev]: "Float (neg_numeral k) 0 = neg_numeral k"
+lemma compute_float_neg_numeral[code_abbrev]: "Float (- numeral k) 0 = - numeral k"
by transfer simp
hide_fact (open) compute_float_neg_numeral
@@ -960,7 +960,7 @@
also have "... < (1 / 2) * 2 powr real (rat_precision n (int x) (int y))"
apply (rule mult_strict_right_mono) by (insert assms) auto
also have "\<dots> = 2 powr real (rat_precision n (int x) (int y) - 1)"
- using powr_add [of 2 _ "- 1", simplified add_uminus_conv_diff] by (simp add: powr_neg_numeral)
+ using powr_add [of 2 _ "- 1", simplified add_uminus_conv_diff] by (simp add: powr_minus)
also have "\<dots> = 2 ^ nat (rat_precision n (int x) (int y) - 1)"
using rat_precision_pos[of x y n] assms by (simp add: powr_realpow[symmetric])
also have "\<dots> \<le> 2 ^ nat (rat_precision n (int x) (int y)) - 1"
--- a/src/HOL/Library/Formal_Power_Series.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Library/Formal_Power_Series.thy Wed Nov 20 16:43:09 2013 +0100
@@ -384,8 +384,8 @@
by (induct k) (simp_all only: numeral.simps fps_const_1_eq_1
fps_const_add [symmetric])
-lemma neg_numeral_fps_const: "neg_numeral k = fps_const (neg_numeral k)"
- by (simp only: neg_numeral_def numeral_fps_const fps_const_neg)
+lemma neg_numeral_fps_const: "- numeral k = fps_const (- numeral k)"
+ by (simp only: numeral_fps_const fps_const_neg)
subsection{* The eXtractor series X*}
@@ -1202,7 +1202,7 @@
have eq: "(1 + X) * ?r = 1"
unfolding minus_one_power_iff
by (auto simp add: field_simps fps_eq_iff)
- show ?thesis by (auto simp add: eq intro: fps_inverse_unique simp del: minus_one)
+ show ?thesis by (auto simp add: eq intro: fps_inverse_unique)
qed
@@ -1245,7 +1245,7 @@
lemma numeral_compose[simp]: "(numeral k::('a::{comm_ring_1}) fps) oo b = numeral k"
unfolding numeral_fps_const by simp
-lemma neg_numeral_compose[simp]: "(neg_numeral k::('a::{comm_ring_1}) fps) oo b = neg_numeral k"
+lemma neg_numeral_compose[simp]: "(- numeral k::('a::{comm_ring_1}) fps) oo b = - numeral k"
unfolding neg_numeral_fps_const by simp
lemma X_fps_compose_startby0[simp]: "a$0 = 0 \<Longrightarrow> X oo a = (a :: ('a :: comm_ring_1) fps)"
@@ -2363,7 +2363,7 @@
next
case (Suc n1)
have "?i $ n = setsum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1} + fps_inv a $ Suc n1 * (a $ 1)^ Suc n1"
- by (simp add: fps_compose_nth Suc startsby_zero_power_nth_same[OF a0] del: power_Suc)
+ by (simp only: fps_compose_nth) (simp add: Suc startsby_zero_power_nth_same [OF a0] del: power_Suc)
also have "\<dots> = setsum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1} +
(X$ Suc n1 - setsum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1})"
using a0 a1 Suc by (simp add: fps_inv_def)
@@ -2404,7 +2404,7 @@
next
case (Suc n1)
have "?i $ n = setsum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1} + fps_ginv b a $ Suc n1 * (a $ 1)^ Suc n1"
- by (simp add: fps_compose_nth Suc startsby_zero_power_nth_same[OF a0] del: power_Suc)
+ by (simp only: fps_compose_nth) (simp add: Suc startsby_zero_power_nth_same [OF a0] del: power_Suc)
also have "\<dots> = setsum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1} +
(b$ Suc n1 - setsum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1})"
using a0 a1 Suc by (simp add: fps_ginv_def)
@@ -2564,9 +2564,9 @@
lemma fps_compose_mult_distrib:
- assumes c0: "c$0 = (0::'a::idom)"
- shows "(a * b) oo c = (a oo c) * (b oo c)" (is "?l = ?r")
- apply (simp add: fps_eq_iff fps_compose_mult_distrib_lemma[OF c0])
+ assumes c0: "c $ 0 = (0::'a::idom)"
+ shows "(a * b) oo c = (a oo c) * (b oo c)"
+ apply (simp add: fps_eq_iff fps_compose_mult_distrib_lemma [OF c0])
apply (simp add: fps_compose_nth fps_mult_nth setsum_left_distrib)
done
@@ -2941,7 +2941,7 @@
(is "inverse ?l = ?r")
proof -
have th: "?l * ?r = 1"
- by (auto simp add: field_simps fps_eq_iff minus_one_power_iff simp del: minus_one)
+ by (auto simp add: field_simps fps_eq_iff minus_one_power_iff)
have th': "?l $ 0 \<noteq> 0" by (simp add: )
from fps_inverse_unique[OF th' th] show ?thesis .
qed
@@ -3165,7 +3165,7 @@
have th: "?r$0 \<noteq> 0" by simp
have th': "fps_deriv (inverse ?r) = fps_const (- 1) * inverse ?r / (1 + X)"
by (simp add: fps_inverse_deriv[OF th] fps_divide_def
- power2_eq_square mult_commute fps_const_neg[symmetric] del: fps_const_neg minus_one)
+ power2_eq_square mult_commute fps_const_neg[symmetric] del: fps_const_neg)
have eq: "inverse ?r $ 0 = 1"
by (simp add: fps_inverse_def)
from iffD1[OF fps_binomial_ODE_unique[of "inverse (1 + X)" "- 1"] th'] eq
@@ -3276,7 +3276,7 @@
have th1: "(?m1 k * ?p (of_nat n) k) / ?f n = 1 / of_nat(fact (n - k))"
unfolding m1nk
unfolding m h pochhammer_Suc_setprod
- apply (simp add: field_simps del: fact_Suc minus_one)
+ apply (simp add: field_simps del: fact_Suc)
unfolding fact_altdef_nat id_def
unfolding of_nat_setprod
unfolding setprod_timesf[symmetric]
@@ -3593,7 +3593,7 @@
unfolding even_mult_two_ex by blast
have "?l $n = ?r$n"
- by (simp add: m fps_sin_def fps_cos_def power_mult_distrib power_mult power_minus)
+ by (simp add: m fps_sin_def fps_cos_def power_mult_distrib power_mult power_minus [of "c ^ 2"])
}
moreover
{
@@ -3602,7 +3602,7 @@
unfolding odd_nat_equiv_def2 by (auto simp add: mult_2)
have "?l $n = ?r$n"
by (simp add: m fps_sin_def fps_cos_def power_mult_distrib
- power_mult power_minus)
+ power_mult power_minus [of "c ^ 2"])
}
ultimately have "?l $n = ?r$n" by blast
} then show ?thesis by (simp add: fps_eq_iff)
--- a/src/HOL/Library/Fundamental_Theorem_Algebra.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Library/Fundamental_Theorem_Algebra.thy Wed Nov 20 16:43:09 2013 +0100
@@ -207,12 +207,14 @@
from unimodular_reduce_norm[OF th0] o
have "\<exists>v. cmod (complex_of_real (cmod b) / b + v^n) < 1"
apply (cases "cmod (complex_of_real (cmod b) / b + 1) < 1", rule_tac x="1" in exI, simp)
- apply (cases "cmod (complex_of_real (cmod b) / b - 1) < 1", rule_tac x="-1" in exI, simp del: minus_one add: minus_one [symmetric])
+ apply (cases "cmod (complex_of_real (cmod b) / b - 1) < 1", rule_tac x="-1" in exI, simp)
apply (cases "cmod (complex_of_real (cmod b) / b + ii) < 1")
apply (cases "even m", rule_tac x="ii" in exI, simp add: m power_mult)
apply (rule_tac x="- ii" in exI, simp add: m power_mult)
apply (cases "even m", rule_tac x="- ii" in exI, simp add: m power_mult)
- apply (rule_tac x="ii" in exI, simp add: m power_mult)
+ apply (auto simp add: m power_mult)
+ apply (rule_tac x="ii" in exI)
+ apply (auto simp add: m power_mult)
done
then obtain v where v: "cmod (complex_of_real (cmod b) / b + v^n) < 1" by blast
let ?w = "v / complex_of_real (root n (cmod b))"
--- a/src/HOL/Library/Order_Relation.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Library/Order_Relation.thy Wed Nov 20 16:43:09 2013 +0100
@@ -93,7 +93,7 @@
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
+ then obtain b and c where "b \<noteq> c \<and> (b,c) \<in> r" using NID by auto
hence 1: "b \<noteq> c \<and> {b,c} \<le> Field r" by (auto simp: Field_def)
(* *)
fix a assume *: "a \<in> Field r"
--- a/src/HOL/Library/Order_Union.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Library/Order_Union.thy Wed Nov 20 16:43:09 2013 +0100
@@ -7,7 +7,7 @@
header {* Order Union *}
theory Order_Union
-imports "~~/src/HOL/Cardinals/Wellfounded_More_Base"
+imports "~~/src/HOL/Cardinals/Wellfounded_More_FP"
begin
definition Osum :: "'a rel \<Rightarrow> 'a rel \<Rightarrow> 'a rel" (infix "Osum" 60) where
@@ -31,7 +31,7 @@
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
+ using WF unfolding wf_eq_minimal2 by metis
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'"
@@ -59,7 +59,7 @@
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
+ using WF' unfolding wf_eq_minimal2 by metis
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'"
@@ -299,7 +299,7 @@
using assms Total_Id_Field by blast
hence ?thesis unfolding Osum_def by auto
}
- ultimately show ?thesis using * unfolding Osum_def by blast
+ ultimately show ?thesis using * unfolding Osum_def by fast
qed
}
thus ?thesis by(auto simp add: Osum_def)
@@ -308,12 +308,7 @@
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
+unfolding wf_def using assms by blast
lemma Osum_wf_Id:
assumes TOT: "Total r" and TOT': "Total r'" and
@@ -343,7 +338,7 @@
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)
+ ultimately have ?thesis by (metis wf_subset)
}
moreover
{assume Case22: "r' \<le> Id"
@@ -356,7 +351,7 @@
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 have ?thesis by (metis wf_subset)
}
ultimately show ?thesis by blast
qed
--- a/src/HOL/Library/Polynomial.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Library/Polynomial.thy Wed Nov 20 16:43:09 2013 +0100
@@ -1575,12 +1575,12 @@
lemma poly_div_minus_left [simp]:
fixes x y :: "'a::field poly"
shows "(- x) div y = - (x div y)"
- using div_smult_left [of "- 1::'a"] by (simp del: minus_one) (* FIXME *)
+ using div_smult_left [of "- 1::'a"] by simp
lemma poly_mod_minus_left [simp]:
fixes x y :: "'a::field poly"
shows "(- x) mod y = - (x mod y)"
- using mod_smult_left [of "- 1::'a"] by (simp del: minus_one) (* FIXME *)
+ using mod_smult_left [of "- 1::'a"] by simp
lemma pdivmod_rel_smult_right:
"\<lbrakk>a \<noteq> 0; pdivmod_rel x y q r\<rbrakk>
@@ -1597,13 +1597,12 @@
lemma poly_div_minus_right [simp]:
fixes x y :: "'a::field poly"
shows "x div (- y) = - (x div y)"
- using div_smult_right [of "- 1::'a"]
- by (simp add: nonzero_inverse_minus_eq del: minus_one) (* FIXME *)
+ using div_smult_right [of "- 1::'a"] by (simp add: nonzero_inverse_minus_eq)
lemma poly_mod_minus_right [simp]:
fixes x y :: "'a::field poly"
shows "x mod (- y) = x mod y"
- using mod_smult_right [of "- 1::'a"] by (simp del: minus_one) (* FIXME *)
+ using mod_smult_right [of "- 1::'a"] by simp
lemma pdivmod_rel_mult:
"\<lbrakk>pdivmod_rel x y q r; pdivmod_rel q z q' r'\<rbrakk>
--- a/src/HOL/Library/Predicate_Compile_Alternative_Defs.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Library/Predicate_Compile_Alternative_Defs.thy Wed Nov 20 16:43:09 2013 +0100
@@ -45,8 +45,8 @@
section {* Setup for Numerals *}
-setup {* Predicate_Compile_Data.ignore_consts [@{const_name numeral}, @{const_name neg_numeral}] *}
-setup {* Predicate_Compile_Data.keep_functions [@{const_name numeral}, @{const_name neg_numeral}] *}
+setup {* Predicate_Compile_Data.ignore_consts [@{const_name numeral}] *}
+setup {* Predicate_Compile_Data.keep_functions [@{const_name numeral}] *}
setup {* Predicate_Compile_Data.ignore_consts [@{const_name div}, @{const_name mod}, @{const_name times}] *}
--- a/src/HOL/Library/Sublist.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Library/Sublist.thy Wed Nov 20 16:43:09 2013 +0100
@@ -107,16 +107,22 @@
lemma append_one_prefixeq:
"prefixeq xs ys \<Longrightarrow> length xs < length ys \<Longrightarrow> prefixeq (xs @ [ys ! length xs]) ys"
- unfolding prefixeq_def
- by (metis Cons_eq_appendI append_eq_appendI append_eq_conv_conj
- eq_Nil_appendI nth_drop')
+ proof (unfold prefixeq_def)
+ assume a1: "\<exists>zs. ys = xs @ zs"
+ then obtain sk :: "'a list" where sk: "ys = xs @ sk" by fastforce
+ assume a2: "length xs < length ys"
+ have f1: "\<And>v. ([]\<Colon>'a list) @ v = v" using append_Nil2 by simp
+ have "[] \<noteq> sk" using a1 a2 sk less_not_refl by force
+ hence "\<exists>v. xs @ hd sk # v = ys" using sk by (metis hd_Cons_tl)
+ thus "\<exists>zs. ys = (xs @ [ys ! length xs]) @ zs" using f1 by fastforce
+ qed
theorem prefixeq_length_le: "prefixeq xs ys \<Longrightarrow> length xs \<le> length ys"
by (auto simp add: prefixeq_def)
lemma prefixeq_same_cases:
"prefixeq (xs\<^sub>1::'a list) ys \<Longrightarrow> prefixeq xs\<^sub>2 ys \<Longrightarrow> prefixeq xs\<^sub>1 xs\<^sub>2 \<or> prefixeq xs\<^sub>2 xs\<^sub>1"
- unfolding prefixeq_def by (metis append_eq_append_conv2)
+ unfolding prefixeq_def by (force simp: append_eq_append_conv2)
lemma set_mono_prefixeq: "prefixeq xs ys \<Longrightarrow> set xs \<subseteq> set ys"
by (auto simp add: prefixeq_def)
@@ -224,9 +230,9 @@
then show ?thesis by (metis append_Nil2 parallelE prefixeqI snoc.prems ys)
next
fix c cs assume ys': "ys' = c # cs"
- then show ?thesis
- by (metis Cons_eq_appendI eq_Nil_appendI parallelE prefixeqI
- same_prefixeq_prefixeq snoc.prems ys)
+ have "x \<noteq> c" using snoc.prems ys ys' by fastforce
+ thus "\<exists>as b bs c cs. b \<noteq> c \<and> xs @ [x] = as @ b # bs \<and> ys = as @ c # cs"
+ using ys ys' by blast
qed
next
assume "prefix ys xs"
@@ -464,7 +470,7 @@
then show ?case by (metis append_Cons)
next
case (list_hembeq_Cons2 x y xs ys)
- then show ?case by (cases xs) (auto, blast+)
+ then show ?case by blast
qed
lemma list_hembeq_appendD:
@@ -475,9 +481,14 @@
case Nil then show ?case by auto
next
case (Cons x xs)
- then obtain us v vs where "zs = us @ v # vs"
- and "P\<^sup>=\<^sup>= x v" and "list_hembeq P (xs @ ys) vs" by (auto dest: list_hembeq_ConsD)
- with Cons show ?case by (metis append_Cons append_assoc list_hembeq_Cons2 list_hembeq_append2)
+ then obtain us v vs where
+ zs: "zs = us @ v # vs" and p: "P\<^sup>=\<^sup>= x v" and lh: "list_hembeq P (xs @ ys) vs"
+ by (auto dest: list_hembeq_ConsD)
+ obtain sk\<^sub>0 :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" and sk\<^sub>1 :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
+ sk: "\<forall>x\<^sub>0 x\<^sub>1. \<not> list_hembeq P (xs @ x\<^sub>0) x\<^sub>1 \<or> sk\<^sub>0 x\<^sub>0 x\<^sub>1 @ sk\<^sub>1 x\<^sub>0 x\<^sub>1 = x\<^sub>1 \<and> list_hembeq P xs (sk\<^sub>0 x\<^sub>0 x\<^sub>1) \<and> list_hembeq P x\<^sub>0 (sk\<^sub>1 x\<^sub>0 x\<^sub>1)"
+ using Cons(1) by (metis (no_types))
+ hence "\<forall>x\<^sub>2. list_hembeq P (x # xs) (x\<^sub>2 @ v # sk\<^sub>0 ys vs)" using p lh by auto
+ thus ?case using lh zs sk by (metis (no_types) append_Cons append_assoc)
qed
lemma list_hembeq_suffix:
@@ -550,7 +561,7 @@
by (simp_all)
lemma sublisteq_Cons': "sublisteq (x#xs) ys \<Longrightarrow> sublisteq xs ys"
- by (induct xs) (auto dest: list_hembeq_ConsD)
+ by (induct xs, simp, blast dest: list_hembeq_ConsD)
lemma sublisteq_Cons2':
assumes "sublisteq (x#xs) (x#ys)" shows "sublisteq xs ys"
@@ -579,11 +590,11 @@
from list_hembeq_Nil2 [OF this] show ?case by simp
next
case list_hembeq_Cons2
- then show ?case by simp
+ thus ?case by simp
next
case list_hembeq_Cons
- then show ?case
- by (metis sublisteq_Cons' list_hembeq_length Suc_length_conv Suc_n_not_le_n)
+ hence False using sublisteq_Cons' by fastforce
+ thus ?case ..
qed
lemma sublisteq_trans: "sublisteq xs ys \<Longrightarrow> sublisteq ys zs \<Longrightarrow> sublisteq xs zs"
@@ -650,7 +661,7 @@
lemma sublisteq_filter [simp]:
assumes "sublisteq xs ys" shows "sublisteq (filter P xs) (filter P ys)"
- using assms by (induct) auto
+ using assms by induct auto
lemma "sublisteq xs ys \<longleftrightarrow> (\<exists>N. xs = sublist ys N)" (is "?L = ?R")
proof
--- a/src/HOL/Library/Sum_of_Squares/sum_of_squares.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Library/Sum_of_Squares/sum_of_squares.ML Wed Nov 20 16:43:09 2013 +0100
@@ -875,7 +875,6 @@
@{term "0::real"}, @{term "1::real"},
@{term "numeral :: num => nat"},
@{term "numeral :: num => real"},
- @{term "neg_numeral :: num => real"},
@{term "Num.Bit0"}, @{term "Num.Bit1"}, @{term "Num.One"}];
fun check_sos kcts ct =
--- a/src/HOL/Library/Wfrec.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Library/Wfrec.thy Wed Nov 20 16:43:09 2013 +0100
@@ -48,7 +48,7 @@
apply (fast dest!: theI')
apply (erule wfrec_rel.cases, simp)
apply (erule allE, erule allE, erule allE, erule mp)
-apply (fast intro: the_equality [symmetric])
+apply (blast intro: the_equality [symmetric])
done
lemma adm_lemma: "adm_wf R (%f x. F (cut f R x) x)"
--- a/src/HOL/Library/Zorn.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Library/Zorn.thy Wed Nov 20 16:43:09 2013 +0100
@@ -71,7 +71,7 @@
lemma suc_not_equals:
"chain C \<Longrightarrow> \<not> maxchain C \<Longrightarrow> suc C \<noteq> C"
- by (auto simp: suc_def) (metis less_irrefl not_maxchain_Some)
+ by (auto simp: suc_def) (metis (no_types) less_irrefl not_maxchain_Some)
lemma subset_suc:
assumes "X \<subseteq> Y" shows "X \<subseteq> suc Y"
@@ -258,7 +258,7 @@
shows "chain X"
using assms
proof (induct)
- case (suc X) then show ?case by (simp add: suc_def) (metis not_maxchain_Some)
+ case (suc X) then show ?case by (simp add: suc_def) (metis (no_types) not_maxchain_Some)
next
case (Union X)
then have "\<Union>X \<subseteq> A" by (auto dest: suc_Union_closed_in_carrier)
@@ -378,7 +378,7 @@
using `subset.maxchain A M` by (auto simp: subset.maxchain_def)
qed
qed
- ultimately show ?thesis by blast
+ ultimately show ?thesis by metis
qed
text{*Alternative version of Zorn's lemma for the subset relation.*}
@@ -423,7 +423,7 @@
unfolding Chains_def by blast
lemma chain_subset_alt_def: "chain\<^sub>\<subseteq> C = subset.chain UNIV C"
- by (auto simp add: chain_subset_def subset.chain_def)
+ unfolding chain_subset_def subset.chain_def by fast
lemma chains_alt_def: "chains A = {C. subset.chain A C}"
by (simp add: chains_def chain_subset_alt_def subset.chain_def)
@@ -487,7 +487,7 @@
fix a B assume aB: "B \<in> C" "a \<in> B"
with 1 obtain x where "x \<in> Field r" and "B = r\<inverse> `` {x}" by auto
thus "(a, u) \<in> r" using uA and aB and `Preorder r`
- by (auto simp add: preorder_on_def refl_on_def) (metis transD)
+ unfolding preorder_on_def refl_on_def by simp (fast dest: transD)
qed
then have "\<exists>u\<in>Field r. ?P u" using `u \<in> Field r` by blast
}
@@ -524,8 +524,7 @@
lemma trans_init_seg_of:
"r initial_segment_of s \<Longrightarrow> s initial_segment_of t \<Longrightarrow> r initial_segment_of t"
- by (simp (no_asm_use) add: init_seg_of_def)
- (metis UnCI Un_absorb2 subset_trans)
+ by (simp (no_asm_use) add: init_seg_of_def) blast
lemma antisym_init_seg_of:
"r initial_segment_of s \<Longrightarrow> s initial_segment_of r \<Longrightarrow> r = s"
@@ -539,14 +538,13 @@
"chain\<^sub>\<subseteq> R \<Longrightarrow> \<forall>r\<in>R. trans r \<Longrightarrow> trans (\<Union>R)"
apply (auto simp add: chain_subset_def)
apply (simp (no_asm_use) add: trans_def)
-apply (metis subsetD)
-done
+by (metis subsetD)
lemma chain_subset_antisym_Union:
"chain\<^sub>\<subseteq> R \<Longrightarrow> \<forall>r\<in>R. antisym r \<Longrightarrow> antisym (\<Union>R)"
-apply (auto simp add: chain_subset_def antisym_def)
-apply (metis subsetD)
-done
+unfolding chain_subset_def antisym_def
+apply simp
+by (metis (no_types) subsetD)
lemma chain_subset_Total_Union:
assumes "chain\<^sub>\<subseteq> R" and "\<forall>r\<in>R. Total r"
@@ -558,11 +556,11 @@
thus "(\<exists>r\<in>R. (a, b) \<in> r) \<or> (\<exists>r\<in>R. (b, a) \<in> r)"
proof
assume "r \<subseteq> s" hence "(a, b) \<in> s \<or> (b, a) \<in> s" using assms(2) A
- by (simp add: total_on_def) (metis mono_Field subsetD)
+ by (simp add: total_on_def) (metis (no_types) mono_Field subsetD)
thus ?thesis using `s \<in> R` by blast
next
assume "s \<subseteq> r" hence "(a, b) \<in> r \<or> (b, a) \<in> r" using assms(2) A
- by (simp add: total_on_def) (metis mono_Field subsetD)
+ by (simp add: total_on_def) (metis (no_types) mono_Field subsetD)
thus ?thesis using `r \<in> R` by blast
qed
qed
@@ -604,7 +602,7 @@
def I \<equiv> "init_seg_of \<inter> ?WO \<times> ?WO"
have I_init: "I \<subseteq> init_seg_of" by (auto simp: I_def)
hence subch: "\<And>R. R \<in> Chains I \<Longrightarrow> chain\<^sub>\<subseteq> R"
- by (auto simp: init_seg_of_def chain_subset_def Chains_def)
+ unfolding init_seg_of_def chain_subset_def Chains_def by blast
have Chains_wo: "\<And>R r. R \<in> Chains I \<Longrightarrow> r \<in> R \<Longrightarrow> Well_order r"
by (simp add: Chains_def I_def) blast
have FI: "Field I = ?WO" by (auto simp add: I_def init_seg_of_def Field_def)
@@ -619,7 +617,7 @@
have "\<forall>r\<in>R. Refl r" and "\<forall>r\<in>R. trans r" and "\<forall>r\<in>R. antisym r"
and "\<forall>r\<in>R. Total r" and "\<forall>r\<in>R. wf (r - Id)"
using Chains_wo [OF `R \<in> Chains I`] by (simp_all add: order_on_defs)
- have "Refl (\<Union>R)" using `\<forall>r\<in>R. Refl r` by (auto simp: refl_on_def)
+ have "Refl (\<Union>R)" using `\<forall>r\<in>R. Refl r` unfolding refl_on_def by fastforce
moreover have "trans (\<Union>R)"
by (rule chain_subset_trans_Union [OF subch `\<forall>r\<in>R. trans r`])
moreover have "antisym (\<Union>R)"
@@ -630,7 +628,7 @@
proof -
have "(\<Union>R) - Id = \<Union>{r - Id | r. r \<in> R}" by blast
with `\<forall>r\<in>R. wf (r - Id)` and wf_Union_wf_init_segs [OF Chains_inits_DiffI [OF Ris]]
- show ?thesis by (simp (no_asm_simp)) blast
+ show ?thesis by fastforce
qed
ultimately have "Well_order (\<Union>R)" by(simp add:order_on_defs)
moreover have "\<forall>r \<in> R. r initial_segment_of \<Union>R" using Ris
@@ -643,7 +641,7 @@
--{*Zorn's Lemma yields a maximal well-order m:*}
then obtain m::"'a rel" where "Well_order m" and
max: "\<forall>r. Well_order r \<and> (m, r) \<in> I \<longrightarrow> r = m"
- using Zorns_po_lemma[OF 0 1] by (auto simp:FI)
+ using Zorns_po_lemma[OF 0 1] unfolding FI by fastforce
--{*Now show by contradiction that m covers the whole type:*}
{ fix x::'a assume "x \<notin> Field m"
--{*We assume that x is not covered and extend m at the top with x*}
@@ -666,7 +664,7 @@
have "Refl m" and "trans m" and "antisym m" and "Total m" and "wf (m - Id)"
using `Well_order m` by (simp_all add: order_on_defs)
--{*We show that the extension is a well-order*}
- have "Refl ?m" using `Refl m` Fm by (auto simp: refl_on_def)
+ have "Refl ?m" using `Refl m` Fm unfolding refl_on_def by blast
moreover have "trans ?m" using `trans m` and `x \<notin> Field m`
unfolding trans_def Field_def by blast
moreover have "antisym ?m" using `antisym m` and `x \<notin> Field m`
@@ -678,7 +676,7 @@
by (auto simp add: wf_eq_minimal Field_def) metis
thus ?thesis using `wf (m - Id)` and `x \<notin> Field m`
wf_subset [OF `wf ?s` Diff_subset]
- by (fastforce intro!: wf_Un simp add: Un_Diff Field_def)
+ unfolding Un_Diff Field_def by (auto intro: wf_Un)
qed
ultimately have "Well_order ?m" by (simp add: order_on_defs)
--{*We show that the extension is above m*}
@@ -709,7 +707,7 @@
moreover have "Total ?r" using `Total r` by (simp add:total_on_def 1 univ)
moreover have "wf (?r - Id)" by (rule wf_subset [OF `wf (r - Id)`]) blast
ultimately have "Well_order ?r" by (simp add: order_on_defs)
- with 1 show ?thesis by metis
+ with 1 show ?thesis by auto
qed
subsection {* Extending Well-founded Relations to Well-Orders *}
@@ -732,7 +730,7 @@
lemma downset_onD:
"downset_on A r \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> y \<in> A \<Longrightarrow> x \<in> A"
- by (auto simp: downset_on_def)
+ unfolding downset_on_def by blast
text {*Extensions of relations w.r.t.\ a given set.*}
definition extension_on where
@@ -755,7 +753,8 @@
assumes "chain\<^sub>\<subseteq> R" and "\<And>r. r \<in> R \<Longrightarrow> extension_on (Field r) r p"
shows "extension_on (Field (\<Union>R)) (\<Union>R) p"
using assms
- by (simp add: chain_subset_def extension_on_def) (metis mono_Field set_mp)
+ by (simp add: chain_subset_def extension_on_def)
+ (metis (no_types) mono_Field set_mp)
lemma downset_on_empty [simp]: "downset_on {} p"
by (auto simp: downset_on_def)
@@ -789,7 +788,7 @@
"\<And>r. r \<in> R \<Longrightarrow> downset_on (Field r) p" and
"\<And>r. r \<in> R \<Longrightarrow> extension_on (Field r) r p"
using Chains_wo [OF `R \<in> Chains I`] by (simp_all add: order_on_defs)
- have "Refl (\<Union>R)" using `\<forall>r\<in>R. Refl r` by (auto simp: refl_on_def)
+ have "Refl (\<Union>R)" using `\<forall>r\<in>R. Refl r` unfolding refl_on_def by fastforce
moreover have "trans (\<Union>R)"
by (rule chain_subset_trans_Union [OF subch `\<forall>r\<in>R. trans r`])
moreover have "antisym (\<Union>R)"
@@ -800,7 +799,7 @@
proof -
have "(\<Union>R) - Id = \<Union>{r - Id | r. r \<in> R}" by blast
with `\<forall>r\<in>R. wf (r - Id)` wf_Union_wf_init_segs [OF Chains_inits_DiffI [OF Ris]]
- show ?thesis by (simp (no_asm_simp)) blast
+ show ?thesis by fastforce
qed
ultimately have "Well_order (\<Union>R)" by (simp add: order_on_defs)
moreover have "\<forall>r\<in>R. r initial_segment_of \<Union>R" using Ris
@@ -853,8 +852,9 @@
using `extension_on (Field m) m p` `downset_on (Field m) p`
by (subst Fm) (auto simp: extension_on_def dest: downset_onD)
moreover have "downset_on (Field ?m) p"
+ apply (subst Fm)
using `downset_on (Field m) p` and min
- by (subst Fm, simp add: downset_on_def Field_def) (metis Domain_iff)
+ unfolding downset_on_def Field_def by blast
moreover have "(m, ?m) \<in> I"
using `Well_order m` and `Well_order ?m` and
`downset_on (Field m) p` and `downset_on (Field ?m) p` and
@@ -867,7 +867,7 @@
qed
have "p \<subseteq> m"
using `Field p \<subseteq> Field m` and `extension_on (Field m) m p`
- by (force simp: Field_def extension_on_def)
+ unfolding Field_def extension_on_def by auto fast
with `Well_order m` show ?thesis by blast
qed
--- a/src/HOL/List.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/List.thy Wed Nov 20 16:43:09 2013 +0100
@@ -2988,6 +2988,9 @@
lemma map_Suc_upt: "map Suc [m..<n] = [Suc m..<Suc n]"
by (induct n) auto
+lemma map_add_upt: "map (\<lambda>i. i + n) [0..<m] = [n..<m + n]"
+ by (induct m) simp_all
+
lemma nth_map_upt: "i < n-m ==> (map f [m..<n]) ! i = f(m+i)"
apply (induct n m arbitrary: i rule: diff_induct)
prefer 3 apply (subst map_Suc_upt[symmetric])
@@ -3072,9 +3075,9 @@
lemmas upto_rec_numeral [simp] =
upto.simps[of "numeral m" "numeral n"]
- upto.simps[of "numeral m" "neg_numeral n"]
- upto.simps[of "neg_numeral m" "numeral n"]
- upto.simps[of "neg_numeral m" "neg_numeral n"] for m n
+ upto.simps[of "numeral m" "- numeral n"]
+ upto.simps[of "- numeral m" "numeral n"]
+ upto.simps[of "- numeral m" "- numeral n"] for m n
lemma upto_empty[simp]: "j < i \<Longrightarrow> [i..j] = []"
by(simp add: upto.simps)
--- a/src/HOL/Matrix_LP/ComputeFloat.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Matrix_LP/ComputeFloat.thy Wed Nov 20 16:43:09 2013 +0100
@@ -79,8 +79,8 @@
lemma real_is_int_numeral[simp]: "real_is_int (numeral x)"
by (auto simp: real_is_int_def intro!: exI[of _ "numeral x"])
-lemma real_is_int_neg_numeral[simp]: "real_is_int (neg_numeral x)"
- by (auto simp: real_is_int_def intro!: exI[of _ "neg_numeral x"])
+lemma real_is_int_neg_numeral[simp]: "real_is_int (- numeral x)"
+ by (auto simp: real_is_int_def intro!: exI[of _ "- numeral x"])
lemma int_of_real_0[simp]: "int_of_real (0::real) = (0::int)"
by (simp add: int_of_real_def)
@@ -96,7 +96,7 @@
by (intro some_equality)
(auto simp add: real_of_int_inject[symmetric] simp del: real_of_int_inject)
-lemma int_of_real_neg_numeral[simp]: "int_of_real (neg_numeral b) = neg_numeral b"
+lemma int_of_real_neg_numeral[simp]: "int_of_real (- numeral b) = - numeral b"
unfolding int_of_real_def
by (intro some_equality)
(auto simp add: real_of_int_inject[symmetric] simp del: real_of_int_inject)
--- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML Wed Nov 20 16:43:09 2013 +0100
@@ -374,8 +374,7 @@
fun get_prover ctxt name params goal all_facts =
let
- fun learn prover =
- Sledgehammer_MaSh.mash_learn_proof ctxt params prover (prop_of goal) all_facts
+ val learn = Sledgehammer_MaSh.mash_learn_proof ctxt params (prop_of goal) all_facts
in
Sledgehammer_Minimize.get_minimizing_prover ctxt Sledgehammer_Provers.Normal
learn name
@@ -609,8 +608,8 @@
|> max_new_mono_instancesLST
|> max_mono_itersLST)
val minimize =
- Sledgehammer_Minimize.minimize_facts (K (K ())) prover_name params
- true 1 (Sledgehammer_Util.subgoal_count st)
+ Sledgehammer_Minimize.minimize_facts (K ()) prover_name params true 1
+ (Sledgehammer_Util.subgoal_count st)
val _ = log separator
val (used_facts, (preplay, message, message_tail)) =
minimize st goal NONE (these (!named_thms))
@@ -665,7 +664,7 @@
SMT_Solver.smt_tac ctxt thms
else if full then
Metis_Tactic.metis_tac [ATP_Proof_Reconstruct.full_typesN]
- ATP_Proof_Reconstruct.metis_default_lam_trans ctxt thms
+ ATP_Proof_Reconstruct.default_metis_lam_trans ctxt thms
else if String.isPrefix "metis (" (!reconstructor) then
let
val (type_encs, lam_trans) =
@@ -674,10 +673,10 @@
|> filter Token.is_proper |> tl
|> Metis_Tactic.parse_metis_options |> fst
|>> the_default [ATP_Proof_Reconstruct.partial_typesN]
- ||> the_default ATP_Proof_Reconstruct.metis_default_lam_trans
+ ||> the_default ATP_Proof_Reconstruct.default_metis_lam_trans
in Metis_Tactic.metis_tac type_encs lam_trans ctxt thms end
else if !reconstructor = "metis" then
- Metis_Tactic.metis_tac [] ATP_Proof_Reconstruct.metis_default_lam_trans ctxt
+ Metis_Tactic.metis_tac [] ATP_Proof_Reconstruct.default_metis_lam_trans ctxt
thms
else
K all_tac
--- a/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Wed Nov 20 16:43:09 2013 +0100
@@ -228,7 +228,10 @@
then show ?case by vector
qed
-lemma one_index[simp]: "(1 :: 'a::one ^'n)$i = 1"
+lemma one_index [simp]: "(1 :: 'a :: one ^ 'n) $ i = 1"
+ by vector
+
+lemma neg_one_index [simp]: "(- 1 :: 'a :: {one, uminus} ^ 'n) $ i = - 1"
by vector
instance vec :: (semiring_char_0, finite) semiring_char_0
@@ -244,8 +247,8 @@
lemma numeral_index [simp]: "numeral w $ i = numeral w"
by (induct w) (simp_all only: numeral.simps vector_add_component one_index)
-lemma neg_numeral_index [simp]: "neg_numeral w $ i = neg_numeral w"
- by (simp only: neg_numeral_def vector_uminus_component numeral_index)
+lemma neg_numeral_index [simp]: "- numeral w $ i = - numeral w"
+ by (simp only: vector_uminus_component numeral_index)
instance vec :: (comm_ring_1, finite) comm_ring_1 ..
instance vec :: (ring_char_0, finite) ring_char_0 ..
--- a/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Wed Nov 20 16:43:09 2013 +0100
@@ -337,10 +337,10 @@
by (simp add: bilinear_def linear_iff)
lemma bilinear_lneg: "bilinear h \<Longrightarrow> h (- x) y = - h x y"
- by (simp only: scaleR_minus1_left [symmetric] bilinear_lmul)
+ by (drule bilinear_lmul [of _ "- 1"]) simp
lemma bilinear_rneg: "bilinear h \<Longrightarrow> h x (- y) = - h x y"
- by (simp only: scaleR_minus1_left [symmetric] bilinear_rmul)
+ by (drule bilinear_rmul [of _ _ "- 1"]) simp
lemma (in ab_group_add) eq_add_iff: "x = x + y \<longleftrightarrow> y = 0"
using add_imp_eq[of x y 0] by auto
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Wed Nov 20 16:43:09 2013 +0100
@@ -5163,9 +5163,8 @@
lemma open_negations:
fixes s :: "'a::real_normed_vector set"
- shows "open s \<Longrightarrow> open ((\<lambda> x. -x) ` s)"
- unfolding scaleR_minus1_left [symmetric]
- by (rule open_scaling, auto)
+ shows "open s \<Longrightarrow> open ((\<lambda>x. - x) ` s)"
+ using open_scaling [of "- 1" s] by simp
lemma open_translation:
fixes s :: "'a::real_normed_vector set"
--- a/src/HOL/NSA/HTranscendental.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/NSA/HTranscendental.thy Wed Nov 20 16:43:09 2013 +0100
@@ -258,7 +258,7 @@
simp add: mult_assoc)
apply (rule approx_add_right_cancel [where d="-1"])
apply (rule approx_sym [THEN [2] approx_trans2])
-apply (auto simp add: mem_infmal_iff minus_one [symmetric] simp del: minus_one)
+apply (auto simp add: mem_infmal_iff)
done
lemma STAR_exp_epsilon [simp]: "( *f* exp) epsilon @= 1"
@@ -450,7 +450,7 @@
apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]
simp add: mult_assoc)
apply (rule approx_add_right_cancel [where d = "-1"])
-apply (simp add: minus_one [symmetric] del: minus_one)
+apply simp
done
lemma STAR_tan_zero [simp]: "( *f* tan) 0 = 0"
--- a/src/HOL/NSA/HyperDef.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/NSA/HyperDef.thy Wed Nov 20 16:43:09 2013 +0100
@@ -425,7 +425,7 @@
declare power_hypreal_of_real_numeral [of _ "numeral w", simp] for w
lemma power_hypreal_of_real_neg_numeral:
- "(neg_numeral v :: hypreal) ^ n = hypreal_of_real ((neg_numeral v) ^ n)"
+ "(- numeral v :: hypreal) ^ n = hypreal_of_real ((- numeral v) ^ n)"
by simp
declare power_hypreal_of_real_neg_numeral [of _ "numeral w", simp] for w
(*
--- a/src/HOL/NSA/NSA.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/NSA/NSA.thy Wed Nov 20 16:43:09 2013 +0100
@@ -654,7 +654,7 @@
(*reorientation simplification procedure: reorients (polymorphic)
0 = x, 1 = x, nnn = x provided x isn't 0, 1 or a numeral.*)
simproc_setup approx_reorient_simproc
- ("0 @= x" | "1 @= y" | "numeral w @= z" | "neg_numeral w @= r") =
+ ("0 @= x" | "1 @= y" | "numeral w @= z" | "- 1 @= y" | "- numeral w @= r") =
{*
let val rule = @{thm approx_reorient} RS eq_reflection
fun proc phi ss ct = case term_of ct of
@@ -1849,8 +1849,12 @@
lemma st_numeral [simp]: "st (numeral w) = numeral w"
by (rule Reals_numeral [THEN st_SReal_eq])
-lemma st_neg_numeral [simp]: "st (neg_numeral w) = neg_numeral w"
-by (rule Reals_neg_numeral [THEN st_SReal_eq])
+lemma st_neg_numeral [simp]: "st (- numeral w) = - numeral w"
+proof -
+ from Reals_numeral have "numeral w \<in> \<real>" .
+ then have "- numeral w \<in> \<real>" by simp
+ with st_SReal_eq show ?thesis .
+qed
lemma st_0 [simp]: "st 0 = 0"
by (simp add: st_SReal_eq)
@@ -1858,6 +1862,9 @@
lemma st_1 [simp]: "st 1 = 1"
by (simp add: st_SReal_eq)
+lemma st_neg_1 [simp]: "st (- 1) = - 1"
+by (simp add: st_SReal_eq)
+
lemma st_minus: "x \<in> HFinite \<Longrightarrow> st (- x) = - st x"
by (simp add: st_unique st_SReal st_approx_self approx_minus)
--- a/src/HOL/NSA/NSComplex.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/NSA/NSComplex.thy Wed Nov 20 16:43:09 2013 +0100
@@ -635,7 +635,7 @@
by transfer (rule of_real_numeral [symmetric])
lemma hcomplex_hypreal_neg_numeral:
- "hcomplex_of_complex (neg_numeral w) = hcomplex_of_hypreal(neg_numeral w)"
+ "hcomplex_of_complex (- numeral w) = hcomplex_of_hypreal(- numeral w)"
by transfer (rule of_real_neg_numeral [symmetric])
lemma hcomplex_numeral_hcnj [simp]:
@@ -647,7 +647,7 @@
by transfer (rule norm_numeral)
lemma hcomplex_neg_numeral_hcmod [simp]:
- "hcmod(neg_numeral v :: hcomplex) = (numeral v :: hypreal)"
+ "hcmod(- numeral v :: hcomplex) = (numeral v :: hypreal)"
by transfer (rule norm_neg_numeral)
lemma hcomplex_numeral_hRe [simp]:
--- a/src/HOL/NSA/StarDef.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/NSA/StarDef.thy Wed Nov 20 16:43:09 2013 +0100
@@ -968,13 +968,13 @@
by transfer (rule refl)
lemma star_neg_numeral_def [transfer_unfold]:
- "neg_numeral k = star_of (neg_numeral k)"
-by (simp only: neg_numeral_def star_of_minus star_of_numeral)
+ "- numeral k = star_of (- numeral k)"
+by (simp only: star_of_minus star_of_numeral)
-lemma Standard_neg_numeral [simp]: "neg_numeral k \<in> Standard"
-by (simp add: star_neg_numeral_def)
+lemma Standard_neg_numeral [simp]: "- numeral k \<in> Standard"
+ using star_neg_numeral_def [of k] by simp
-lemma star_of_neg_numeral [simp]: "star_of (neg_numeral k) = neg_numeral k"
+lemma star_of_neg_numeral [simp]: "star_of (- numeral k) = - numeral k"
by transfer (rule refl)
lemma star_of_nat_def [transfer_unfold]: "of_nat n = star_of (of_nat n)"
@@ -987,12 +987,12 @@
star_of_less [of _ "numeral k", simplified star_of_numeral]
star_of_le [of _ "numeral k", simplified star_of_numeral]
star_of_eq [of _ "numeral k", simplified star_of_numeral]
- star_of_less [of "neg_numeral k", simplified star_of_numeral]
- star_of_le [of "neg_numeral k", simplified star_of_numeral]
- star_of_eq [of "neg_numeral k", simplified star_of_numeral]
- star_of_less [of _ "neg_numeral k", simplified star_of_numeral]
- star_of_le [of _ "neg_numeral k", simplified star_of_numeral]
- star_of_eq [of _ "neg_numeral k", simplified star_of_numeral] for k
+ star_of_less [of "- numeral k", simplified star_of_numeral]
+ star_of_le [of "- numeral k", simplified star_of_numeral]
+ star_of_eq [of "- numeral k", simplified star_of_numeral]
+ star_of_less [of _ "- numeral k", simplified star_of_numeral]
+ star_of_le [of _ "- numeral k", simplified star_of_numeral]
+ star_of_eq [of _ "- numeral k", simplified star_of_numeral] for k
lemma Standard_of_nat [simp]: "of_nat n \<in> Standard"
by (simp add: star_of_nat_def)
--- a/src/HOL/Nat.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Nat.thy Wed Nov 20 16:43:09 2013 +0100
@@ -1315,6 +1315,11 @@
shows "comp f ^^ n = comp (f ^^ n)"
by (induct n) simp_all
+lemma Suc_funpow[simp]: "Suc ^^ n = (op + n)"
+ by (induct n) simp_all
+
+lemma id_funpow[simp]: "id ^^ n = id"
+ by (induct n) simp_all
subsection {* Kleene iteration *}
--- a/src/HOL/Nominal/Nominal.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Nominal/Nominal.thy Wed Nov 20 16:43:09 2013 +0100
@@ -3517,7 +3517,7 @@
by (simp add: perm_int_def perm_int_def)
lemma neg_numeral_int_eqvt:
- shows "pi\<bullet>((neg_numeral n)::int) = neg_numeral n"
+ shows "pi\<bullet>((- numeral n)::int) = - numeral n"
by (simp add: perm_int_def perm_int_def)
lemma max_int_eqvt:
--- a/src/HOL/Num.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Num.thy Wed Nov 20 16:43:09 2013 +0100
@@ -275,16 +275,6 @@
end
-text {* Negative numerals. *}
-
-class neg_numeral = numeral + group_add
-begin
-
-definition neg_numeral :: "num \<Rightarrow> 'a" where
- "neg_numeral k = - numeral k"
-
-end
-
text {* Numeral syntax. *}
syntax
@@ -299,8 +289,8 @@
| (n, 0) => Syntax.const @{const_name Bit0} $ num_of_int n
| (n, 1) => Syntax.const @{const_name Bit1} $ num_of_int n)
else raise Match
- val pos = Syntax.const @{const_name numeral}
- val neg = Syntax.const @{const_name neg_numeral}
+ val numeral = Syntax.const @{const_name numeral}
+ val uminus = Syntax.const @{const_name uminus}
val one = Syntax.const @{const_name Groups.one}
val zero = Syntax.const @{const_name Groups.zero}
fun numeral_tr [(c as Const (@{syntax_const "_constrain"}, _)) $ t $ u] =
@@ -311,8 +301,9 @@
in
if value = 0 then zero else
if value > 0
- then pos $ num_of_int value
- else neg $ num_of_int (~value)
+ then numeral $ num_of_int value
+ else if value = ~1 then uminus $ one
+ else uminus $ (numeral $ num_of_int (~value))
end
| numeral_tr ts = raise TERM ("numeral_tr", ts);
in [("_Numeral", K numeral_tr)] end
@@ -323,12 +314,12 @@
fun dest_num (Const (@{const_syntax Bit0}, _) $ n) = 2 * dest_num n
| dest_num (Const (@{const_syntax Bit1}, _) $ n) = 2 * dest_num n + 1
| dest_num (Const (@{const_syntax One}, _)) = 1;
- fun num_tr' sign ctxt T [n] =
+ fun num_tr' ctxt T [n] =
let
val k = dest_num n;
val t' =
Syntax.const @{syntax_const "_Numeral"} $
- Syntax.free (sign ^ string_of_int k);
+ Syntax.free (string_of_int k);
in
(case T of
Type (@{type_name fun}, [_, T']) =>
@@ -339,8 +330,7 @@
| _ => if T = dummyT then t' else raise Match)
end;
in
- [(@{const_syntax numeral}, num_tr' ""),
- (@{const_syntax neg_numeral}, num_tr' "-")]
+ [(@{const_syntax numeral}, num_tr')]
end
*}
@@ -383,9 +373,13 @@
Structures with negation: class @{text neg_numeral}
*}
-context neg_numeral
+class neg_numeral = numeral + group_add
begin
+lemma uminus_numeral_One:
+ "- Numeral1 = - 1"
+ by (simp add: numeral_One)
+
text {* Numerals form an abelian subgroup. *}
inductive is_num :: "'a \<Rightarrow> bool" where
@@ -403,7 +397,7 @@
apply simp
apply (rule_tac a=x in add_left_imp_eq)
apply (rule_tac a=x in add_right_imp_eq)
- apply (simp add: add_assoc minus_add_cancel)
+ apply (simp add: add_assoc)
apply (simp add: add_assoc [symmetric], simp add: add_assoc)
apply (rule_tac a=x in add_left_imp_eq)
apply (rule_tac a=x in add_right_imp_eq)
@@ -431,77 +425,85 @@
by (simp only: BitM_plus_one [symmetric] numeral_add numeral_One eq_diff_eq)
lemma dbl_simps [simp]:
- "dbl (neg_numeral k) = neg_numeral (Bit0 k)"
+ "dbl (- numeral k) = - dbl (numeral k)"
"dbl 0 = 0"
"dbl 1 = 2"
+ "dbl (- 1) = - 2"
"dbl (numeral k) = numeral (Bit0 k)"
- by (simp_all add: dbl_def neg_numeral_def numeral.simps minus_add)
+ by (simp_all add: dbl_def numeral.simps minus_add)
lemma dbl_inc_simps [simp]:
- "dbl_inc (neg_numeral k) = neg_numeral (BitM k)"
+ "dbl_inc (- numeral k) = - dbl_dec (numeral k)"
"dbl_inc 0 = 1"
"dbl_inc 1 = 3"
+ "dbl_inc (- 1) = - 1"
"dbl_inc (numeral k) = numeral (Bit1 k)"
- by (simp_all add: dbl_inc_def neg_numeral_def numeral.simps numeral_BitM is_num_normalize algebra_simps del: add_uminus_conv_diff)
+ by (simp_all add: dbl_inc_def dbl_dec_def numeral.simps numeral_BitM is_num_normalize algebra_simps del: add_uminus_conv_diff)
lemma dbl_dec_simps [simp]:
- "dbl_dec (neg_numeral k) = neg_numeral (Bit1 k)"
- "dbl_dec 0 = -1"
+ "dbl_dec (- numeral k) = - dbl_inc (numeral k)"
+ "dbl_dec 0 = - 1"
"dbl_dec 1 = 1"
+ "dbl_dec (- 1) = - 3"
"dbl_dec (numeral k) = numeral (BitM k)"
- by (simp_all add: dbl_dec_def neg_numeral_def numeral.simps numeral_BitM is_num_normalize)
+ by (simp_all add: dbl_dec_def dbl_inc_def numeral.simps numeral_BitM is_num_normalize)
lemma sub_num_simps [simp]:
"sub One One = 0"
- "sub One (Bit0 l) = neg_numeral (BitM l)"
- "sub One (Bit1 l) = neg_numeral (Bit0 l)"
+ "sub One (Bit0 l) = - numeral (BitM l)"
+ "sub One (Bit1 l) = - numeral (Bit0 l)"
"sub (Bit0 k) One = numeral (BitM k)"
"sub (Bit1 k) One = numeral (Bit0 k)"
"sub (Bit0 k) (Bit0 l) = dbl (sub k l)"
"sub (Bit0 k) (Bit1 l) = dbl_dec (sub k l)"
"sub (Bit1 k) (Bit0 l) = dbl_inc (sub k l)"
"sub (Bit1 k) (Bit1 l) = dbl (sub k l)"
- by (simp_all add: dbl_def dbl_dec_def dbl_inc_def sub_def neg_numeral_def numeral.simps
+ by (simp_all add: dbl_def dbl_dec_def dbl_inc_def sub_def numeral.simps
numeral_BitM is_num_normalize del: add_uminus_conv_diff add: diff_conv_add_uminus)
lemma add_neg_numeral_simps:
- "numeral m + neg_numeral n = sub m n"
- "neg_numeral m + numeral n = sub n m"
- "neg_numeral m + neg_numeral n = neg_numeral (m + n)"
- by (simp_all add: sub_def neg_numeral_def numeral_add numeral.simps is_num_normalize
+ "numeral m + - numeral n = sub m n"
+ "- numeral m + numeral n = sub n m"
+ "- numeral m + - numeral n = - (numeral m + numeral n)"
+ by (simp_all add: sub_def numeral_add numeral.simps is_num_normalize
del: add_uminus_conv_diff add: diff_conv_add_uminus)
lemma add_neg_numeral_special:
- "1 + neg_numeral m = sub One m"
- "neg_numeral m + 1 = sub One m"
- by (simp_all add: sub_def neg_numeral_def numeral_add numeral.simps is_num_normalize)
+ "1 + - numeral m = sub One m"
+ "- numeral m + 1 = sub One m"
+ "numeral m + - 1 = sub m One"
+ "- 1 + numeral n = sub n One"
+ "- 1 + - numeral n = - numeral (inc n)"
+ "- numeral m + - 1 = - numeral (inc m)"
+ "1 + - 1 = 0"
+ "- 1 + 1 = 0"
+ "- 1 + - 1 = - 2"
+ by (simp_all add: sub_def numeral_add numeral.simps is_num_normalize right_minus numeral_inc
+ del: add_uminus_conv_diff add: diff_conv_add_uminus)
lemma diff_numeral_simps:
"numeral m - numeral n = sub m n"
- "numeral m - neg_numeral n = numeral (m + n)"
- "neg_numeral m - numeral n = neg_numeral (m + n)"
- "neg_numeral m - neg_numeral n = sub n m"
- by (simp_all add: neg_numeral_def sub_def numeral_add numeral.simps is_num_normalize
+ "numeral m - - numeral n = numeral (m + n)"
+ "- numeral m - numeral n = - numeral (m + n)"
+ "- numeral m - - numeral n = sub n m"
+ by (simp_all add: sub_def numeral_add numeral.simps is_num_normalize
del: add_uminus_conv_diff add: diff_conv_add_uminus)
lemma diff_numeral_special:
"1 - numeral n = sub One n"
- "1 - neg_numeral n = numeral (One + n)"
"numeral m - 1 = sub m One"
- "neg_numeral m - 1 = neg_numeral (m + One)"
- by (simp_all add: neg_numeral_def sub_def numeral_add numeral.simps add: is_num_normalize)
-
-lemma minus_one: "- 1 = -1"
- unfolding neg_numeral_def numeral.simps ..
-
-lemma minus_numeral: "- numeral n = neg_numeral n"
- unfolding neg_numeral_def ..
-
-lemma minus_neg_numeral: "- neg_numeral n = numeral n"
- unfolding neg_numeral_def by simp
-
-lemmas minus_numeral_simps [simp] =
- minus_one minus_numeral minus_neg_numeral
+ "1 - - numeral n = numeral (One + n)"
+ "- numeral m - 1 = - numeral (m + One)"
+ "- 1 - numeral n = - numeral (inc n)"
+ "numeral m - - 1 = numeral (inc m)"
+ "- 1 - - numeral n = sub n One"
+ "- numeral m - - 1 = sub One m"
+ "1 - 1 = 0"
+ "- 1 - 1 = - 2"
+ "1 - - 1 = 2"
+ "- 1 - - 1 = 0"
+ by (simp_all add: sub_def numeral_add numeral.simps is_num_normalize numeral_inc
+ del: add_uminus_conv_diff add: diff_conv_add_uminus)
end
@@ -675,17 +677,17 @@
subclass neg_numeral ..
lemma mult_neg_numeral_simps:
- "neg_numeral m * neg_numeral n = numeral (m * n)"
- "neg_numeral m * numeral n = neg_numeral (m * n)"
- "numeral m * neg_numeral n = neg_numeral (m * n)"
- unfolding neg_numeral_def mult_minus_left mult_minus_right
+ "- numeral m * - numeral n = numeral (m * n)"
+ "- numeral m * numeral n = - numeral (m * n)"
+ "numeral m * - numeral n = - numeral (m * n)"
+ unfolding mult_minus_left mult_minus_right
by (simp_all only: minus_minus numeral_mult)
-lemma mult_minus1 [simp]: "-1 * z = - z"
- unfolding neg_numeral_def numeral.simps mult_minus_left by simp
+lemma mult_minus1 [simp]: "- 1 * z = - z"
+ unfolding numeral.simps mult_minus_left by simp
-lemma mult_minus1_right [simp]: "z * -1 = - z"
- unfolding neg_numeral_def numeral.simps mult_minus_right by simp
+lemma mult_minus1_right [simp]: "z * - 1 = - z"
+ unfolding numeral.simps mult_minus_right by simp
end
@@ -708,9 +710,15 @@
lemma not_iszero_Numeral1: "\<not> iszero Numeral1"
by (simp add: numeral_One)
+lemma not_iszero_neg_1 [simp]: "\<not> iszero (- 1)"
+ by (simp add: iszero_def)
+
+lemma not_iszero_neg_Numeral1: "\<not> iszero (- Numeral1)"
+ by (simp add: numeral_One)
+
lemma iszero_neg_numeral [simp]:
- "iszero (neg_numeral w) \<longleftrightarrow> iszero (numeral w)"
- unfolding iszero_def neg_numeral_def
+ "iszero (- numeral w) \<longleftrightarrow> iszero (numeral w)"
+ unfolding iszero_def
by (rule neg_equal_0_iff_equal)
lemma eq_iff_iszero_diff: "x = y \<longleftrightarrow> iszero (x - y)"
@@ -730,17 +738,17 @@
lemma eq_numeral_iff_iszero:
"numeral x = numeral y \<longleftrightarrow> iszero (sub x y)"
- "numeral x = neg_numeral y \<longleftrightarrow> iszero (numeral (x + y))"
- "neg_numeral x = numeral y \<longleftrightarrow> iszero (numeral (x + y))"
- "neg_numeral x = neg_numeral y \<longleftrightarrow> iszero (sub y x)"
+ "numeral x = - numeral y \<longleftrightarrow> iszero (numeral (x + y))"
+ "- numeral x = numeral y \<longleftrightarrow> iszero (numeral (x + y))"
+ "- numeral x = - numeral y \<longleftrightarrow> iszero (sub y x)"
"numeral x = 1 \<longleftrightarrow> iszero (sub x One)"
"1 = numeral y \<longleftrightarrow> iszero (sub One y)"
- "neg_numeral x = 1 \<longleftrightarrow> iszero (numeral (x + One))"
- "1 = neg_numeral y \<longleftrightarrow> iszero (numeral (One + y))"
+ "- numeral x = 1 \<longleftrightarrow> iszero (numeral (x + One))"
+ "1 = - numeral y \<longleftrightarrow> iszero (numeral (One + y))"
"numeral x = 0 \<longleftrightarrow> iszero (numeral x)"
"0 = numeral y \<longleftrightarrow> iszero (numeral y)"
- "neg_numeral x = 0 \<longleftrightarrow> iszero (numeral x)"
- "0 = neg_numeral y \<longleftrightarrow> iszero (numeral y)"
+ "- numeral x = 0 \<longleftrightarrow> iszero (numeral x)"
+ "0 = - numeral y \<longleftrightarrow> iszero (numeral y)"
unfolding eq_iff_iszero_diff diff_numeral_simps diff_numeral_special
by simp_all
@@ -756,33 +764,69 @@
lemma not_iszero_numeral [simp]: "\<not> iszero (numeral w)"
by (simp add: iszero_def)
-lemma neg_numeral_eq_iff: "neg_numeral m = neg_numeral n \<longleftrightarrow> m = n"
- by (simp only: neg_numeral_def neg_equal_iff_equal numeral_eq_iff)
+lemma neg_numeral_eq_iff: "- numeral m = - numeral n \<longleftrightarrow> m = n"
+ by simp
-lemma numeral_neq_neg_numeral: "numeral m \<noteq> neg_numeral n"
- unfolding neg_numeral_def eq_neg_iff_add_eq_0
+lemma numeral_neq_neg_numeral: "numeral m \<noteq> - numeral n"
+ unfolding eq_neg_iff_add_eq_0
by (simp add: numeral_plus_numeral)
-lemma neg_numeral_neq_numeral: "neg_numeral m \<noteq> numeral n"
+lemma neg_numeral_neq_numeral: "- numeral m \<noteq> numeral n"
by (rule numeral_neq_neg_numeral [symmetric])
-lemma zero_neq_neg_numeral: "0 \<noteq> neg_numeral n"
- unfolding neg_numeral_def neg_0_equal_iff_equal by simp
+lemma zero_neq_neg_numeral: "0 \<noteq> - numeral n"
+ unfolding neg_0_equal_iff_equal by simp
-lemma neg_numeral_neq_zero: "neg_numeral n \<noteq> 0"
- unfolding neg_numeral_def neg_equal_0_iff_equal by simp
+lemma neg_numeral_neq_zero: "- numeral n \<noteq> 0"
+ unfolding neg_equal_0_iff_equal by simp
-lemma one_neq_neg_numeral: "1 \<noteq> neg_numeral n"
+lemma one_neq_neg_numeral: "1 \<noteq> - numeral n"
using numeral_neq_neg_numeral [of One n] by (simp add: numeral_One)
-lemma neg_numeral_neq_one: "neg_numeral n \<noteq> 1"
+lemma neg_numeral_neq_one: "- numeral n \<noteq> 1"
using neg_numeral_neq_numeral [of n One] by (simp add: numeral_One)
+lemma neg_one_neq_numeral:
+ "- 1 \<noteq> numeral n"
+ using neg_numeral_neq_numeral [of One n] by (simp add: numeral_One)
+
+lemma numeral_neq_neg_one:
+ "numeral n \<noteq> - 1"
+ using numeral_neq_neg_numeral [of n One] by (simp add: numeral_One)
+
+lemma neg_one_eq_numeral_iff:
+ "- 1 = - numeral n \<longleftrightarrow> n = One"
+ using neg_numeral_eq_iff [of One n] by (auto simp add: numeral_One)
+
+lemma numeral_eq_neg_one_iff:
+ "- numeral n = - 1 \<longleftrightarrow> n = One"
+ using neg_numeral_eq_iff [of n One] by (auto simp add: numeral_One)
+
+lemma neg_one_neq_zero:
+ "- 1 \<noteq> 0"
+ by simp
+
+lemma zero_neq_neg_one:
+ "0 \<noteq> - 1"
+ by simp
+
+lemma neg_one_neq_one:
+ "- 1 \<noteq> 1"
+ using neg_numeral_neq_numeral [of One One] by (simp only: numeral_One not_False_eq_True)
+
+lemma one_neq_neg_one:
+ "1 \<noteq> - 1"
+ using numeral_neq_neg_numeral [of One One] by (simp only: numeral_One not_False_eq_True)
+
lemmas eq_neg_numeral_simps [simp] =
neg_numeral_eq_iff
numeral_neq_neg_numeral neg_numeral_neq_numeral
one_neq_neg_numeral neg_numeral_neq_one
zero_neq_neg_numeral neg_numeral_neq_zero
+ neg_one_neq_numeral numeral_neq_neg_one
+ neg_one_eq_numeral_iff numeral_eq_neg_one_iff
+ neg_one_neq_zero zero_neq_neg_one
+ neg_one_neq_one one_neq_neg_one
end
@@ -795,48 +839,72 @@
subclass ring_char_0 ..
-lemma neg_numeral_le_iff: "neg_numeral m \<le> neg_numeral n \<longleftrightarrow> n \<le> m"
- by (simp only: neg_numeral_def neg_le_iff_le numeral_le_iff)
+lemma neg_numeral_le_iff: "- numeral m \<le> - numeral n \<longleftrightarrow> n \<le> m"
+ by (simp only: neg_le_iff_le numeral_le_iff)
-lemma neg_numeral_less_iff: "neg_numeral m < neg_numeral n \<longleftrightarrow> n < m"
- by (simp only: neg_numeral_def neg_less_iff_less numeral_less_iff)
+lemma neg_numeral_less_iff: "- numeral m < - numeral n \<longleftrightarrow> n < m"
+ by (simp only: neg_less_iff_less numeral_less_iff)
-lemma neg_numeral_less_zero: "neg_numeral n < 0"
- by (simp only: neg_numeral_def neg_less_0_iff_less zero_less_numeral)
+lemma neg_numeral_less_zero: "- numeral n < 0"
+ by (simp only: neg_less_0_iff_less zero_less_numeral)
-lemma neg_numeral_le_zero: "neg_numeral n \<le> 0"
- by (simp only: neg_numeral_def neg_le_0_iff_le zero_le_numeral)
+lemma neg_numeral_le_zero: "- numeral n \<le> 0"
+ by (simp only: neg_le_0_iff_le zero_le_numeral)
-lemma not_zero_less_neg_numeral: "\<not> 0 < neg_numeral n"
+lemma not_zero_less_neg_numeral: "\<not> 0 < - numeral n"
by (simp only: not_less neg_numeral_le_zero)
-lemma not_zero_le_neg_numeral: "\<not> 0 \<le> neg_numeral n"
+lemma not_zero_le_neg_numeral: "\<not> 0 \<le> - numeral n"
by (simp only: not_le neg_numeral_less_zero)
-lemma neg_numeral_less_numeral: "neg_numeral m < numeral n"
+lemma neg_numeral_less_numeral: "- numeral m < numeral n"
using neg_numeral_less_zero zero_less_numeral by (rule less_trans)
-lemma neg_numeral_le_numeral: "neg_numeral m \<le> numeral n"
+lemma neg_numeral_le_numeral: "- numeral m \<le> numeral n"
by (simp only: less_imp_le neg_numeral_less_numeral)
-lemma not_numeral_less_neg_numeral: "\<not> numeral m < neg_numeral n"
+lemma not_numeral_less_neg_numeral: "\<not> numeral m < - numeral n"
by (simp only: not_less neg_numeral_le_numeral)
-lemma not_numeral_le_neg_numeral: "\<not> numeral m \<le> neg_numeral n"
+lemma not_numeral_le_neg_numeral: "\<not> numeral m \<le> - numeral n"
by (simp only: not_le neg_numeral_less_numeral)
-lemma neg_numeral_less_one: "neg_numeral m < 1"
+lemma neg_numeral_less_one: "- numeral m < 1"
by (rule neg_numeral_less_numeral [of m One, unfolded numeral_One])
-lemma neg_numeral_le_one: "neg_numeral m \<le> 1"
+lemma neg_numeral_le_one: "- numeral m \<le> 1"
by (rule neg_numeral_le_numeral [of m One, unfolded numeral_One])
-lemma not_one_less_neg_numeral: "\<not> 1 < neg_numeral m"
+lemma not_one_less_neg_numeral: "\<not> 1 < - numeral m"
by (simp only: not_less neg_numeral_le_one)
-lemma not_one_le_neg_numeral: "\<not> 1 \<le> neg_numeral m"
+lemma not_one_le_neg_numeral: "\<not> 1 \<le> - numeral m"
by (simp only: not_le neg_numeral_less_one)
+lemma not_numeral_less_neg_one: "\<not> numeral m < - 1"
+ using not_numeral_less_neg_numeral [of m One] by (simp add: numeral_One)
+
+lemma not_numeral_le_neg_one: "\<not> numeral m \<le> - 1"
+ using not_numeral_le_neg_numeral [of m One] by (simp add: numeral_One)
+
+lemma neg_one_less_numeral: "- 1 < numeral m"
+ using neg_numeral_less_numeral [of One m] by (simp add: numeral_One)
+
+lemma neg_one_le_numeral: "- 1 \<le> numeral m"
+ using neg_numeral_le_numeral [of One m] by (simp add: numeral_One)
+
+lemma neg_numeral_less_neg_one_iff: "- numeral m < - 1 \<longleftrightarrow> m \<noteq> One"
+ by (cases m) simp_all
+
+lemma neg_numeral_le_neg_one: "- numeral m \<le> - 1"
+ by simp
+
+lemma not_neg_one_less_neg_numeral: "\<not> - 1 < - numeral m"
+ by simp
+
+lemma not_neg_one_le_neg_numeral_iff: "\<not> - 1 \<le> - numeral m \<longleftrightarrow> m \<noteq> One"
+ by (cases m) simp_all
+
lemma sub_non_negative:
"sub n m \<ge> 0 \<longleftrightarrow> n \<ge> m"
by (simp only: sub_def le_diff_eq) simp
@@ -858,18 +926,40 @@
neg_numeral_le_numeral not_numeral_le_neg_numeral
neg_numeral_le_zero not_zero_le_neg_numeral
neg_numeral_le_one not_one_le_neg_numeral
+ neg_one_le_numeral not_numeral_le_neg_one
+ neg_numeral_le_neg_one not_neg_one_le_neg_numeral_iff
+
+lemma le_minus_one_simps [simp]:
+ "- 1 \<le> 0"
+ "- 1 \<le> 1"
+ "\<not> 0 \<le> - 1"
+ "\<not> 1 \<le> - 1"
+ by simp_all
lemmas less_neg_numeral_simps [simp] =
neg_numeral_less_iff
neg_numeral_less_numeral not_numeral_less_neg_numeral
neg_numeral_less_zero not_zero_less_neg_numeral
neg_numeral_less_one not_one_less_neg_numeral
+ neg_one_less_numeral not_numeral_less_neg_one
+ neg_numeral_less_neg_one_iff not_neg_one_less_neg_numeral
+
+lemma less_minus_one_simps [simp]:
+ "- 1 < 0"
+ "- 1 < 1"
+ "\<not> 0 < - 1"
+ "\<not> 1 < - 1"
+ by (simp_all add: less_le)
lemma abs_numeral [simp]: "abs (numeral n) = numeral n"
by simp
-lemma abs_neg_numeral [simp]: "abs (neg_numeral n) = numeral n"
- by (simp only: neg_numeral_def abs_minus_cancel abs_numeral)
+lemma abs_neg_numeral [simp]: "abs (- numeral n) = numeral n"
+ by (simp only: abs_minus_cancel abs_numeral)
+
+lemma abs_neg_one [simp]:
+ "abs (- 1) = 1"
+ by simp
end
@@ -1032,31 +1122,36 @@
text{*Theorem lists for the cancellation simprocs. The use of a binary
numeral for 1 reduces the number of special cases.*}
-lemmas mult_1s =
- mult_numeral_1 mult_numeral_1_right
- mult_minus1 mult_minus1_right
+lemma mult_1s:
+ fixes a :: "'a::semiring_numeral"
+ and b :: "'b::ring_1"
+ shows "Numeral1 * a = a"
+ "a * Numeral1 = a"
+ "- Numeral1 * b = - b"
+ "b * - Numeral1 = - b"
+ by simp_all
setup {*
Reorient_Proc.add
(fn Const (@{const_name numeral}, _) $ _ => true
- | Const (@{const_name neg_numeral}, _) $ _ => true
+ | Const (@{const_name uminus}, _) $ (Const (@{const_name numeral}, _) $ _) => true
| _ => false)
*}
simproc_setup reorient_numeral
- ("numeral w = x" | "neg_numeral w = y") = Reorient_Proc.proc
+ ("numeral w = x" | "- numeral w = y") = Reorient_Proc.proc
subsubsection {* Simplification of arithmetic operations on integer constants. *}
lemmas arith_special = (* already declared simp above *)
add_numeral_special add_neg_numeral_special
- diff_numeral_special minus_one
+ diff_numeral_special
(* rules already in simpset *)
lemmas arith_extra_simps =
numeral_plus_numeral add_neg_numeral_simps add_0_left add_0_right
- minus_numeral minus_neg_numeral minus_zero minus_one
+ minus_zero
diff_numeral_simps diff_0 diff_0_right
numeral_times_numeral mult_neg_numeral_simps
mult_zero_left mult_zero_right
@@ -1089,15 +1184,15 @@
lemmas rel_simps =
le_num_simps less_num_simps eq_num_simps
- le_numeral_simps le_neg_numeral_simps le_numeral_extra
- less_numeral_simps less_neg_numeral_simps less_numeral_extra
+ le_numeral_simps le_neg_numeral_simps le_minus_one_simps le_numeral_extra
+ less_numeral_simps less_neg_numeral_simps less_minus_one_simps less_numeral_extra
eq_numeral_simps eq_neg_numeral_simps eq_numeral_extra
lemma Let_numeral [simp]: "Let (numeral v) f = f (numeral v)"
-- {* Unfold all @{text let}s involving constants *}
unfolding Let_def ..
-lemma Let_neg_numeral [simp]: "Let (neg_numeral v) f = f (neg_numeral v)"
+lemma Let_neg_numeral [simp]: "Let (- numeral v) f = f (- numeral v)"
-- {* Unfold all @{text let}s involving constants *}
unfolding Let_def ..
@@ -1133,16 +1228,16 @@
by (simp_all add: add_assoc [symmetric])
lemma add_neg_numeral_left [simp]:
- "numeral v + (neg_numeral w + y) = (sub v w + y)"
- "neg_numeral v + (numeral w + y) = (sub w v + y)"
- "neg_numeral v + (neg_numeral w + y) = (neg_numeral(v + w) + y)"
+ "numeral v + (- numeral w + y) = (sub v w + y)"
+ "- numeral v + (numeral w + y) = (sub w v + y)"
+ "- numeral v + (- numeral w + y) = (- numeral(v + w) + y)"
by (simp_all add: add_assoc [symmetric])
lemma mult_numeral_left [simp]:
"numeral v * (numeral w * z) = (numeral(v * w) * z :: 'a::semiring_numeral)"
- "neg_numeral v * (numeral w * y) = (neg_numeral(v * w) * y :: 'b::ring_1)"
- "numeral v * (neg_numeral w * y) = (neg_numeral(v * w) * y :: 'b::ring_1)"
- "neg_numeral v * (neg_numeral w * y) = (numeral(v * w) * y :: 'b::ring_1)"
+ "- numeral v * (numeral w * y) = (- numeral(v * w) * y :: 'b::ring_1)"
+ "numeral v * (- numeral w * y) = (- numeral(v * w) * y :: 'b::ring_1)"
+ "- numeral v * (- numeral w * y) = (numeral(v * w) * y :: 'b::ring_1)"
by (simp_all add: mult_assoc [symmetric])
hide_const (open) One Bit0 Bit1 BitM inc pow sqr sub dbl dbl_inc dbl_dec
--- a/src/HOL/Number_Theory/Cong.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Number_Theory/Cong.thy Wed Nov 20 16:43:09 2013 +0100
@@ -323,8 +323,6 @@
\<Longrightarrow> [a = 1] (mod p) \<or> [a = - 1] (mod p)"
apply (simp only: cong_altdef_int)
apply (subst prime_dvd_mult_eq_int [symmetric], assumption)
- (* any way around this? *)
- apply (subgoal_tac "a * a - 1 = (a - 1) * (a - -1)")
apply (auto simp add: field_simps)
done
@@ -665,7 +663,6 @@
apply auto
apply (cases "n \<ge> 0")
apply auto
- apply (subst cong_int_def, auto)
apply (frule cong_solve_int [of a n])
apply auto
done
--- a/src/HOL/Number_Theory/Residues.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Number_Theory/Residues.thy Wed Nov 20 16:43:09 2013 +0100
@@ -455,6 +455,7 @@
apply (subst fact_altdef_int, simp)
apply (subst cong_int_def)
apply simp
+ apply arith
apply (rule residues_prime.wilson_theorem1)
apply (rule residues_prime.intro)
apply auto
--- a/src/HOL/Numeral_Simprocs.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Numeral_Simprocs.thy Wed Nov 20 16:43:09 2013 +0100
@@ -17,7 +17,7 @@
if_False if_True
add_0 add_Suc add_numeral_left
add_neg_numeral_left mult_numeral_left
- numeral_1_eq_1 [symmetric] Suc_eq_plus1
+ numeral_One [symmetric] uminus_numeral_One [symmetric] Suc_eq_plus1
eq_numeral_iff_iszero not_iszero_Numeral1
declare split_div [of _ _ "numeral k", arith_split] for k
@@ -85,18 +85,19 @@
text {* For @{text cancel_factor} *}
-lemma nat_mult_le_cancel_disj: "(k*m <= k*n) = ((0::nat) < k --> m<=n)"
-by auto
+lemmas nat_mult_le_cancel_disj = mult_le_cancel1
-lemma nat_mult_less_cancel_disj: "(k*m < k*n) = ((0::nat) < k & m<n)"
-by auto
+lemmas nat_mult_less_cancel_disj = mult_less_cancel1
-lemma nat_mult_eq_cancel_disj: "(k*m = k*n) = (k = (0::nat) | m=n)"
-by auto
+lemma nat_mult_eq_cancel_disj:
+ fixes k m n :: nat
+ shows "k * m = k * n \<longleftrightarrow> k = 0 \<or> m = n"
+ by auto
-lemma nat_mult_div_cancel_disj[simp]:
- "(k*m) div (k*n) = (if k = (0::nat) then 0 else m div n)"
-by (simp add: nat_mult_div_cancel1)
+lemma nat_mult_div_cancel_disj [simp]:
+ fixes k m n :: nat
+ shows "(k * m) div (k * n) = (if k = 0 then 0 else m div n)"
+ by (fact div_mult_mult1_if)
ML_file "Tools/numeral_simprocs.ML"
--- a/src/HOL/Parity.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Parity.thy Wed Nov 20 16:43:09 2013 +0100
@@ -78,7 +78,7 @@
unfolding even_def by simp
(* TODO: proper simp rules for Num.Bit0, Num.Bit1 *)
-declare even_def [of "neg_numeral v", simp] for v
+declare even_def [of "- numeral v", simp] for v
lemma even_numeral_nat [simp]: "even (numeral (Num.Bit0 k) :: nat)"
unfolding even_nat_def by simp
@@ -190,14 +190,9 @@
by (induct n) simp_all
lemma (in comm_ring_1)
- shows minus_one_even_power [simp]: "even n \<Longrightarrow> (- 1) ^ n = 1"
- and minus_one_odd_power [simp]: "odd n \<Longrightarrow> (- 1) ^ n = - 1"
- by (simp_all add: neg_power_if del: minus_one)
-
-lemma (in comm_ring_1)
- shows neg_one_even_power [simp]: "even n \<Longrightarrow> (-1) ^ n = 1"
- and neg_one_odd_power [simp]: "odd n \<Longrightarrow> (-1) ^ n = - 1"
- by (simp_all add: minus_one [symmetric] del: minus_one)
+ shows neg_one_even_power [simp]: "even n \<Longrightarrow> (- 1) ^ n = 1"
+ and neg_one_odd_power [simp]: "odd n \<Longrightarrow> (- 1) ^ n = - 1"
+ by (simp_all add: neg_power_if)
lemma zero_le_even_power: "even n ==>
0 <= (x::'a::{linordered_ring,monoid_mult}) ^ n"
--- a/src/HOL/Power.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Power.thy Wed Nov 20 16:43:09 2013 +0100
@@ -209,14 +209,6 @@
"(- x) ^ numeral (Num.Bit1 k) = - (x ^ numeral (Num.Bit1 k))"
by (simp only: eval_nat_numeral(3) power_Suc power_minus_Bit0 mult_minus_left)
-lemma power_neg_numeral_Bit0 [simp]:
- "neg_numeral k ^ numeral (Num.Bit0 l) = numeral (Num.pow k (Num.Bit0 l))"
- by (simp only: neg_numeral_def power_minus_Bit0 power_numeral)
-
-lemma power_neg_numeral_Bit1 [simp]:
- "neg_numeral k ^ numeral (Num.Bit1 l) = neg_numeral (Num.pow k (Num.Bit1 l))"
- by (simp only: neg_numeral_def power_minus_Bit1 power_numeral pow.simps)
-
lemma power2_minus [simp]:
"(- a)\<^sup>2 = a\<^sup>2"
by (rule power_minus_Bit0)
--- a/src/HOL/ROOT Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/ROOT Wed Nov 20 16:43:09 2013 +0100
@@ -687,14 +687,14 @@
theories Nominal_Examples
theories [quick_and_dirty] VC_Condition
-session "HOL-Cardinals-Base" in Cardinals = HOL +
+session "HOL-Cardinals-FP" in Cardinals = HOL +
description {*
- Ordinals and Cardinals, Base Theories.
+ Ordinals and Cardinals, Theories Needed for BNF FP Constructions.
*}
options [document = false]
- theories Cardinal_Arithmetic
+ theories Cardinal_Arithmetic_FP
-session "HOL-Cardinals" in Cardinals = "HOL-Cardinals-Base" +
+session "HOL-Cardinals" in Cardinals = "HOL-Cardinals-FP" +
description {*
Ordinals and Cardinals, Full Theories.
*}
@@ -705,16 +705,16 @@
"document/root.tex"
"document/root.bib"
-session "HOL-BNF-LFP" in BNF = "HOL-Cardinals-Base" +
+session "HOL-BNF-FP" in BNF = "HOL-Cardinals-FP" +
description {*
- Bounded Natural Functors for Datatypes.
+ Bounded Natural Functors for (Co)datatypes.
*}
options [document = false]
- theories BNF_LFP
+ theories BNF_LFP BNF_GFP
-session "HOL-BNF" in BNF = "HOL-Cardinals" +
+session "HOL-BNF" in BNF = "HOL-BNF-FP" +
description {*
- Bounded Natural Functors for (Co)datatypes.
+ Bounded Natural Functors for (Co)datatypes, Including More BNFs.
*}
options [document = false]
theories BNF
--- a/src/HOL/Rat.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Rat.thy Wed Nov 20 16:43:09 2013 +0100
@@ -215,17 +215,19 @@
"Fract 0 k = 0"
"Fract 1 1 = 1"
"Fract (numeral w) 1 = numeral w"
- "Fract (neg_numeral w) 1 = neg_numeral w"
+ "Fract (- numeral w) 1 = - numeral w"
+ "Fract (- 1) 1 = - 1"
"Fract k 0 = 0"
using Fract_of_int_eq [of "numeral w"]
- using Fract_of_int_eq [of "neg_numeral w"]
+ using Fract_of_int_eq [of "- numeral w"]
by (simp_all add: Zero_rat_def One_rat_def eq_rat)
lemma rat_number_expand:
"0 = Fract 0 1"
"1 = Fract 1 1"
"numeral k = Fract (numeral k) 1"
- "neg_numeral k = Fract (neg_numeral k) 1"
+ "- 1 = Fract (- 1) 1"
+ "- numeral k = Fract (- numeral k) 1"
by (simp_all add: rat_number_collapse)
lemma Rat_cases_nonzero [case_names Fract 0]:
@@ -356,7 +358,8 @@
"quotient_of 0 = (0, 1)"
"quotient_of 1 = (1, 1)"
"quotient_of (numeral k) = (numeral k, 1)"
- "quotient_of (neg_numeral k) = (neg_numeral k, 1)"
+ "quotient_of (- 1) = (- 1, 1)"
+ "quotient_of (- numeral k) = (- numeral k, 1)"
by (simp_all add: rat_number_expand quotient_of_Fract)
lemma quotient_of_eq: "quotient_of (Fract a b) = (p, q) \<Longrightarrow> Fract p q = Fract a b"
@@ -620,7 +623,7 @@
#> Lin_Arith.add_simps [@{thm neg_less_iff_less},
@{thm True_implies_equals},
read_instantiate @{context} [(("a", 0), "(numeral ?v)")] @{thm distrib_left},
- read_instantiate @{context} [(("a", 0), "(neg_numeral ?v)")] @{thm distrib_left},
+ read_instantiate @{context} [(("a", 0), "(- numeral ?v)")] @{thm distrib_left},
@{thm divide_1}, @{thm divide_zero_left},
@{thm times_divide_eq_right}, @{thm times_divide_eq_left},
@{thm minus_divide_left} RS sym, @{thm minus_divide_right} RS sym,
@@ -664,6 +667,10 @@
lemma of_rat_minus: "of_rat (- a) = - of_rat a"
by transfer simp
+lemma of_rat_neg_one [simp]:
+ "of_rat (- 1) = - 1"
+ by (simp add: of_rat_minus)
+
lemma of_rat_diff: "of_rat (a - b) = of_rat a - of_rat b"
using of_rat_add [of a "- b"] by (simp add: of_rat_minus)
@@ -778,8 +785,8 @@
using of_rat_of_int_eq [of "numeral w"] by simp
lemma of_rat_neg_numeral_eq [simp]:
- "of_rat (neg_numeral w) = neg_numeral w"
-using of_rat_of_int_eq [of "neg_numeral w"] by simp
+ "of_rat (- numeral w) = - numeral w"
+using of_rat_of_int_eq [of "- numeral w"] by simp
lemmas zero_rat = Zero_rat_def
lemmas one_rat = One_rat_def
@@ -820,9 +827,6 @@
lemma Rats_number_of [simp]: "numeral w \<in> Rats"
by (subst of_rat_numeral_eq [symmetric], rule Rats_of_rat)
-lemma Rats_neg_number_of [simp]: "neg_numeral w \<in> Rats"
-by (subst of_rat_neg_numeral_eq [symmetric], rule Rats_of_rat)
-
lemma Rats_0 [simp]: "0 \<in> Rats"
apply (unfold Rats_def)
apply (rule range_eqI)
@@ -943,7 +947,7 @@
by (simp add: Rat.of_int_def)
lemma [code_unfold]:
- "neg_numeral k = Rat.of_int (neg_numeral k)"
+ "- numeral k = Rat.of_int (- numeral k)"
by (simp add: Rat.of_int_def)
lemma Frct_code_post [code_post]:
@@ -951,13 +955,13 @@
"Frct (a, 0) = 0"
"Frct (1, 1) = 1"
"Frct (numeral k, 1) = numeral k"
- "Frct (neg_numeral k, 1) = neg_numeral k"
+ "Frct (- numeral k, 1) = - numeral k"
"Frct (1, numeral k) = 1 / numeral k"
- "Frct (1, neg_numeral k) = 1 / neg_numeral k"
+ "Frct (1, - numeral k) = 1 / - numeral k"
"Frct (numeral k, numeral l) = numeral k / numeral l"
- "Frct (numeral k, neg_numeral l) = numeral k / neg_numeral l"
- "Frct (neg_numeral k, numeral l) = neg_numeral k / numeral l"
- "Frct (neg_numeral k, neg_numeral l) = neg_numeral k / neg_numeral l"
+ "Frct (numeral k, - numeral l) = numeral k / - numeral l"
+ "Frct (- numeral k, numeral l) = - numeral k / numeral l"
+ "Frct (- numeral k, - numeral l) = - numeral k / - numeral l"
by (simp_all add: Fract_of_int_quotient)
@@ -1156,7 +1160,7 @@
in
if i = 0 then Syntax.const @{const_syntax Groups.zero}
else if i > 0 then Syntax.const @{const_syntax Num.numeral} $ mk i
- else Syntax.const @{const_syntax Num.neg_numeral} $ mk (~i)
+ else Syntax.const @{const_syntax Groups.uminus} $ (Syntax.const @{const_syntax Num.numeral} $ mk (~i))
end;
fun mk_frac str =
--- a/src/HOL/Real.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Real.thy Wed Nov 20 16:43:09 2013 +0100
@@ -1447,13 +1447,13 @@
lemma [code_abbrev]:
"real_of_int (numeral k) = numeral k"
- "real_of_int (neg_numeral k) = neg_numeral k"
+ "real_of_int (- numeral k) = - numeral k"
by simp_all
-text{*Collapse applications of @{term real} to @{term number_of}*}
+text{*Collapse applications of @{const real} to @{const numeral}*}
lemma real_numeral [simp]:
"real (numeral v :: int) = numeral v"
- "real (neg_numeral v :: int) = neg_numeral v"
+ "real (- numeral v :: int) = - numeral v"
by (simp_all add: real_of_int_def)
lemma real_of_nat_numeral [simp]:
@@ -1559,11 +1559,11 @@
unfolding real_of_int_le_iff[symmetric] by simp
lemma neg_numeral_power_le_real_of_int_cancel_iff[simp]:
- "(neg_numeral x::real) ^ n \<le> real a \<longleftrightarrow> (neg_numeral x::int) ^ n \<le> a"
+ "(- numeral x::real) ^ n \<le> real a \<longleftrightarrow> (- numeral x::int) ^ n \<le> a"
unfolding real_of_int_le_iff[symmetric] by simp
lemma real_of_int_le_neg_numeral_power_cancel_iff[simp]:
- "real a \<le> (neg_numeral x::real) ^ n \<longleftrightarrow> a \<le> (neg_numeral x::int) ^ n"
+ "real a \<le> (- numeral x::real) ^ n \<longleftrightarrow> a \<le> (- numeral x::int) ^ n"
unfolding real_of_int_le_iff[symmetric] by simp
subsection{*Density of the Reals*}
@@ -2051,7 +2051,7 @@
by simp
lemma [code_abbrev]:
- "(of_rat (neg_numeral k) :: real) = neg_numeral k"
+ "(of_rat (- numeral k) :: real) = - numeral k"
by simp
lemma [code_post]:
@@ -2059,14 +2059,14 @@
"(of_rat (r / 0) :: real) = 0"
"(of_rat (1 / 1) :: real) = 1"
"(of_rat (numeral k / 1) :: real) = numeral k"
- "(of_rat (neg_numeral k / 1) :: real) = neg_numeral k"
+ "(of_rat (- numeral k / 1) :: real) = - numeral k"
"(of_rat (1 / numeral k) :: real) = 1 / numeral k"
- "(of_rat (1 / neg_numeral k) :: real) = 1 / neg_numeral k"
+ "(of_rat (1 / - numeral k) :: real) = 1 / - numeral k"
"(of_rat (numeral k / numeral l) :: real) = numeral k / numeral l"
- "(of_rat (numeral k / neg_numeral l) :: real) = numeral k / neg_numeral l"
- "(of_rat (neg_numeral k / numeral l) :: real) = neg_numeral k / numeral l"
- "(of_rat (neg_numeral k / neg_numeral l) :: real) = neg_numeral k / neg_numeral l"
- by (simp_all add: of_rat_divide)
+ "(of_rat (numeral k / - numeral l) :: real) = numeral k / - numeral l"
+ "(of_rat (- numeral k / numeral l) :: real) = - numeral k / numeral l"
+ "(of_rat (- numeral k / - numeral l) :: real) = - numeral k / - numeral l"
+ by (simp_all add: of_rat_divide of_rat_minus)
text {* Operations *}
--- a/src/HOL/Real_Vector_Spaces.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Real_Vector_Spaces.thy Wed Nov 20 16:43:09 2013 +0100
@@ -307,8 +307,8 @@
lemma of_real_numeral: "of_real (numeral w) = numeral w"
using of_real_of_int_eq [of "numeral w"] by simp
-lemma of_real_neg_numeral: "of_real (neg_numeral w) = neg_numeral w"
-using of_real_of_int_eq [of "neg_numeral w"] by simp
+lemma of_real_neg_numeral: "of_real (- numeral w) = - numeral w"
+using of_real_of_int_eq [of "- numeral w"] by simp
text{*Every real algebra has characteristic zero*}
@@ -341,9 +341,6 @@
lemma Reals_numeral [simp]: "numeral w \<in> Reals"
by (subst of_real_numeral [symmetric], rule Reals_of_real)
-lemma Reals_neg_numeral [simp]: "neg_numeral w \<in> Reals"
-by (subst of_real_neg_numeral [symmetric], rule Reals_of_real)
-
lemma Reals_0 [simp]: "0 \<in> Reals"
apply (unfold Reals_def)
apply (rule range_eqI)
@@ -602,7 +599,7 @@
by (subst of_real_numeral [symmetric], subst norm_of_real, simp)
lemma norm_neg_numeral [simp]:
- "norm (neg_numeral w::'a::real_normed_algebra_1) = numeral w"
+ "norm (- numeral w::'a::real_normed_algebra_1) = numeral w"
by (subst of_real_neg_numeral [symmetric], subst norm_of_real, simp)
lemma norm_of_int [simp]:
--- a/src/HOL/Rings.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Rings.thy Wed Nov 20 16:43:09 2013 +0100
@@ -1058,6 +1058,34 @@
"\<bar>l\<bar> = \<bar>k\<bar> \<Longrightarrow> l dvd k"
by(subst abs_dvd_iff[symmetric]) simp
+text {* The following lemmas can be proven in more generale structures, but
+are dangerous as simp rules in absence of @{thm neg_equal_zero},
+@{thm neg_less_pos}, @{thm neg_less_eq_nonneg}. *}
+
+lemma equation_minus_iff_1 [simp, no_atp]:
+ "1 = - a \<longleftrightarrow> a = - 1"
+ by (fact equation_minus_iff)
+
+lemma minus_equation_iff_1 [simp, no_atp]:
+ "- a = 1 \<longleftrightarrow> a = - 1"
+ by (subst minus_equation_iff, auto)
+
+lemma le_minus_iff_1 [simp, no_atp]:
+ "1 \<le> - b \<longleftrightarrow> b \<le> - 1"
+ by (fact le_minus_iff)
+
+lemma minus_le_iff_1 [simp, no_atp]:
+ "- a \<le> 1 \<longleftrightarrow> - 1 \<le> a"
+ by (fact minus_le_iff)
+
+lemma less_minus_iff_1 [simp, no_atp]:
+ "1 < - b \<longleftrightarrow> b < - 1"
+ by (fact less_minus_iff)
+
+lemma minus_less_iff_1 [simp, no_atp]:
+ "- a < 1 \<longleftrightarrow> - 1 < a"
+ by (fact minus_less_iff)
+
end
text {* Simprules for comparisons where common factors can be cancelled. *}
--- a/src/HOL/SMT_Examples/SMT_Tests.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/SMT_Examples/SMT_Tests.thy Wed Nov 20 16:43:09 2013 +0100
@@ -374,7 +374,6 @@
lemma
"(0::int) = 0"
- "(0::int) = -0"
"(0::int) = (- 0)"
"(1::int) = 1"
"\<not>(-1 = (1::int))"
--- a/src/HOL/SMT_Examples/SMT_Word_Examples.certs Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/SMT_Examples/SMT_Word_Examples.certs Wed Nov 20 16:43:09 2013 +0100
@@ -54,3 +54,5 @@
unsat
e5c27ae0a583eeafeaa4ef3c59b1b4ec53e06b0f 1 0
unsat
+7d3ef49480d3ed3a7e5f2d7a12e7108cf7fc7819 1 0
+unsat
--- a/src/HOL/TPTP/atp_problem_import.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/TPTP/atp_problem_import.ML Wed Nov 20 16:43:09 2013 +0100
@@ -20,7 +20,7 @@
val sledgehammer_tptp_file : theory -> int -> string -> unit
val isabelle_tptp_file : theory -> int -> string -> unit
val isabelle_hot_tptp_file : theory -> int -> string -> unit
- val translate_tptp_file : theory -> string -> string -> string -> unit
+ val translate_tptp_file : theory -> string -> string -> unit
end;
structure ATP_Problem_Import : ATP_PROBLEM_IMPORT =
@@ -301,9 +301,9 @@
(** Translator between TPTP(-like) file formats **)
-fun translate_tptp_file thy format_str in_file_name out_file_name =
+fun translate_tptp_file thy format_str file_name =
let
- val (conjs, (defs, nondefs), ctxt) = read_tptp_file thy I in_file_name
+ val (conjs, (defs, nondefs), ctxt) = read_tptp_file thy I file_name
val conj = make_conj ([], []) (map snd conjs)
val (format, type_enc, lam_trans) =
@@ -327,7 +327,7 @@
val ord_info = K []
val lines = lines_of_atp_problem format ord ord_info atp_problem
in
- File.write_list (exploded_absolute_path out_file_name) lines
+ List.app Output.physical_stdout lines
end
end;
--- a/src/HOL/TPTP/lib/Tools/tptp_translate Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/TPTP/lib/Tools/tptp_translate Wed Nov 20 16:43:09 2013 +0100
@@ -9,20 +9,21 @@
function usage() {
echo
- echo "Usage: isabelle $PRG FORMAT IN_FILE OUT_FILE"
+ echo "Usage: isabelle $PRG FORMAT FILE"
echo
echo " Translates TPTP input file to the specified format (\"FOF\", \"TFF0\", \"THF0\", or \"DFG\")."
+ echo " Emits the result to standard output."
echo
exit 1
}
-[ "$#" -ne 3 -o "$1" = "-?" ] && usage
+[ "$#" -ne 2 -o "$1" = "-?" ] && usage
SCRATCH="Scratch_${PRG}_$$_${RANDOM}"
args=("$@")
echo "theory $SCRATCH imports \"$TPTP_HOME/ATP_Problem_Import\" begin \
-ML {* ATP_Problem_Import.translate_tptp_file @{theory} \"${args[0]}\" \"${args[1]}\" \"${args[2]}\" *} end;" \
+ML {* ATP_Problem_Import.translate_tptp_file @{theory} \"${args[0]}\" \"${args[1]}\" *} end;" \
> /tmp/$SCRATCH.thy
"$ISABELLE_PROCESS" -q -e "use_thy \"/tmp/$SCRATCH\"; exit 1;" HOL-TPTP | grep --line-buffered -v "^###\|^PROOF FAILED for depth\|^Failure node\|inferences so far. Searching to depth\|^val \|^Loading theory"
--- a/src/HOL/Tools/ATP/atp_proof_reconstruct.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/ATP/atp_proof_reconstruct.ML Wed Nov 20 16:43:09 2013 +0100
@@ -10,6 +10,9 @@
sig
type ('a, 'b) atp_term = ('a, 'b) ATP_Problem.atp_term
type ('a, 'b, 'c, 'd) atp_formula = ('a, 'b, 'c, 'd) ATP_Problem.atp_formula
+ type stature = ATP_Problem_Generate.stature
+ type ('a, 'b) atp_step = ('a, 'b) ATP_Proof.atp_step
+ type 'a atp_proof = 'a ATP_Proof.atp_proof
val metisN : string
val full_typesN : string
@@ -21,17 +24,26 @@
val no_type_enc : string
val full_type_encs : string list
val partial_type_encs : string list
- val metis_default_lam_trans : string
+ val default_metis_lam_trans : string
val metis_call : string -> string -> string
val forall_of : term -> term -> term
val exists_of : term -> term -> term
val unalias_type_enc : string -> string list
- val term_of_atp :
- Proof.context -> bool -> int Symtab.table -> typ option ->
+ val term_of_atp : Proof.context -> bool -> int Symtab.table -> typ option ->
(string, string) atp_term -> term
- val prop_of_atp :
- Proof.context -> bool -> int Symtab.table ->
+ val prop_of_atp : Proof.context -> bool -> int Symtab.table ->
(string, string, (string, string) atp_term, string) atp_formula -> term
+
+ val used_facts_in_atp_proof :
+ Proof.context -> (string * stature) list vector -> string atp_proof -> (string * stature) list
+ val used_facts_in_unsound_atp_proof : Proof.context -> (string * stature) list vector ->
+ 'a atp_proof -> string list option
+ val lam_trans_of_atp_proof : string atp_proof -> string -> string
+ val is_typed_helper_used_in_atp_proof : string atp_proof -> bool
+ val termify_atp_proof : Proof.context -> string Symtab.table -> (string * term) list ->
+ int Symtab.table -> string atp_proof -> (term, string) atp_step list
+ val factify_atp_proof : (string * 'a) list vector -> term list -> term ->
+ (term, string) atp_step list -> (term, string) atp_step list
end;
structure ATP_Proof_Reconstruct : ATP_PROOF_RECONSTRUCT =
@@ -64,7 +76,7 @@
fun unalias_type_enc s =
AList.lookup (op =) type_enc_aliases s |> the_default [s]
-val metis_default_lam_trans = combsN
+val default_metis_lam_trans = combsN
fun metis_call type_enc lam_trans =
let
@@ -74,7 +86,7 @@
[alias] => alias
| _ => type_enc
val opts = [] |> type_enc <> partial_typesN ? cons type_enc
- |> lam_trans <> metis_default_lam_trans ? cons lam_trans
+ |> lam_trans <> default_metis_lam_trans ? cons lam_trans
in metisN ^ (if null opts then "" else " (" ^ commas opts ^ ")") end
fun term_name' (Var ((s, _), _)) = perhaps (try Name.dest_skolem) s
@@ -339,4 +351,205 @@
| _ => raise ATP_FORMULA [phi]
in repair_tvar_sorts (do_formula true phi Vartab.empty) end
+fun find_first_in_list_vector vec key =
+ Vector.foldl (fn (ps, NONE) => AList.lookup (op =) ps key
+ | (_, value) => value) NONE vec
+
+val unprefix_fact_number = space_implode "_" o tl o space_explode "_"
+
+fun resolve_one_named_fact fact_names s =
+ case try (unprefix fact_prefix) s of
+ SOME s' =>
+ let val s' = s' |> unprefix_fact_number |> unascii_of in
+ s' |> find_first_in_list_vector fact_names |> Option.map (pair s')
+ end
+ | NONE => NONE
+
+fun resolve_fact fact_names = map_filter (resolve_one_named_fact fact_names)
+
+fun resolve_one_named_conjecture s =
+ case try (unprefix conjecture_prefix) s of
+ SOME s' => Int.fromString s'
+ | NONE => NONE
+
+val resolve_conjecture = map_filter resolve_one_named_conjecture
+
+fun is_axiom_used_in_proof pred =
+ exists (fn ((_, ss), _, _, _, []) => exists pred ss | _ => false)
+
+fun add_non_rec_defs fact_names accum =
+ Vector.foldl (fn (facts, facts') =>
+ union (op =) (filter (fn (_, (_, status)) => status = Non_Rec_Def) facts)
+ facts')
+ accum fact_names
+
+val isa_ext = Thm.get_name_hint @{thm ext}
+val isa_short_ext = Long_Name.base_name isa_ext
+
+fun ext_name ctxt =
+ if Thm.eq_thm_prop (@{thm ext},
+ singleton (Attrib.eval_thms ctxt) (Facts.named isa_short_ext, [])) then
+ isa_short_ext
+ else
+ isa_ext
+
+val leo2_extcnf_equal_neg_rule = "extcnf_equal_neg"
+val leo2_unfold_def_rule = "unfold_def"
+
+fun add_fact ctxt fact_names ((_, ss), _, _, rule, deps) =
+ (if rule = leo2_extcnf_equal_neg_rule then
+ insert (op =) (ext_name ctxt, (Global, General))
+ else if rule = leo2_unfold_def_rule then
+ (* LEO 1.3.3 does not record definitions properly, leading to missing
+ dependencies in the TSTP proof. Remove the next line once this is
+ fixed. *)
+ add_non_rec_defs fact_names
+ else if rule = agsyhol_coreN orelse rule = satallax_coreN then
+ (fn [] =>
+ (* agsyHOL and Satallax don't include definitions in their
+ unsatisfiable cores, so we assume the worst and include them all
+ here. *)
+ [(ext_name ctxt, (Global, General))] |> add_non_rec_defs fact_names
+ | facts => facts)
+ else
+ I)
+ #> (if null deps then union (op =) (resolve_fact fact_names ss) else I)
+
+fun used_facts_in_atp_proof ctxt fact_names atp_proof =
+ if null atp_proof then Vector.foldl (uncurry (union (op =))) [] fact_names
+ else fold (add_fact ctxt fact_names) atp_proof []
+
+fun used_facts_in_unsound_atp_proof _ _ [] = NONE
+ | used_facts_in_unsound_atp_proof ctxt fact_names atp_proof =
+ let val used_facts = used_facts_in_atp_proof ctxt fact_names atp_proof in
+ if forall (fn (_, (sc, _)) => sc = Global) used_facts andalso
+ not (is_axiom_used_in_proof (not o null o resolve_conjecture o single) atp_proof) then
+ SOME (map fst used_facts)
+ else
+ NONE
+ end
+
+val ascii_of_lam_fact_prefix = ascii_of lam_fact_prefix
+
+(* overapproximation (good enough) *)
+fun is_lam_lifted s =
+ String.isPrefix fact_prefix s andalso
+ String.isSubstring ascii_of_lam_fact_prefix s
+
+val is_combinator_def = String.isPrefix (helper_prefix ^ combinator_prefix)
+
+fun lam_trans_of_atp_proof atp_proof default =
+ case (is_axiom_used_in_proof is_combinator_def atp_proof,
+ is_axiom_used_in_proof is_lam_lifted atp_proof) of
+ (false, false) => default
+ | (false, true) => liftingN
+(* | (true, true) => combs_and_liftingN -- not supported by "metis" *)
+ | (true, _) => combsN
+
+val is_typed_helper_name =
+ String.isPrefix helper_prefix andf String.isSuffix typed_helper_suffix
+
+fun is_typed_helper_used_in_atp_proof atp_proof =
+ is_axiom_used_in_proof is_typed_helper_name atp_proof
+
+fun repair_name "$true" = "c_True"
+ | repair_name "$false" = "c_False"
+ | repair_name "$$e" = tptp_equal (* seen in Vampire proofs *)
+ | repair_name s =
+ if is_tptp_equal s orelse
+ (* seen in Vampire proofs *)
+ (String.isPrefix "sQ" s andalso String.isSuffix "_eqProxy" s) then
+ tptp_equal
+ else
+ s
+
+fun infer_formula_types ctxt =
+ Type.constraint HOLogic.boolT
+ #> Syntax.check_term
+ (Proof_Context.set_mode Proof_Context.mode_schematic ctxt)
+
+val combinator_table =
+ [(@{const_name Meson.COMBI}, @{thm Meson.COMBI_def [abs_def]}),
+ (@{const_name Meson.COMBK}, @{thm Meson.COMBK_def [abs_def]}),
+ (@{const_name Meson.COMBB}, @{thm Meson.COMBB_def [abs_def]}),
+ (@{const_name Meson.COMBC}, @{thm Meson.COMBC_def [abs_def]}),
+ (@{const_name Meson.COMBS}, @{thm Meson.COMBS_def [abs_def]})]
+
+fun uncombine_term thy =
+ let
+ fun aux (t1 $ t2) = betapply (pairself aux (t1, t2))
+ | aux (Abs (s, T, t')) = Abs (s, T, aux t')
+ | aux (t as Const (x as (s, _))) =
+ (case AList.lookup (op =) combinator_table s of
+ SOME thm => thm |> prop_of |> specialize_type thy x
+ |> Logic.dest_equals |> snd
+ | NONE => t)
+ | aux t = t
+ in aux end
+
+fun unlift_term lifted =
+ map_aterms (fn t as Const (s, _) =>
+ if String.isPrefix lam_lifted_prefix s then
+ case AList.lookup (op =) lifted s of
+ SOME t =>
+ (* FIXME: do something about the types *)
+ unlift_term lifted t
+ | NONE => t
+ else
+ t
+ | t => t)
+
+fun decode_line ctxt lifted sym_tab (name, role, u, rule, deps) =
+ let
+ val thy = Proof_Context.theory_of ctxt
+ val t =
+ u |> prop_of_atp ctxt true sym_tab
+ |> uncombine_term thy
+ |> unlift_term lifted
+ |> infer_formula_types ctxt
+ in (name, role, t, rule, deps) end
+
+val waldmeister_conjecture_num = "1.0.0.0"
+
+fun repair_waldmeister_endgame arg =
+ let
+ fun do_tail (name, _, t, rule, deps) =
+ (name, Negated_Conjecture, s_not t, rule, deps)
+ fun do_body [] = []
+ | do_body ((line as ((num, _), _, _, _, _)) :: lines) =
+ if num = waldmeister_conjecture_num then map do_tail (line :: lines)
+ else line :: do_body lines
+ in do_body arg end
+
+fun termify_atp_proof ctxt pool lifted sym_tab =
+ clean_up_atp_proof_dependencies
+ #> nasty_atp_proof pool
+ #> map_term_names_in_atp_proof repair_name
+ #> map (decode_line ctxt lifted sym_tab)
+ #> repair_waldmeister_endgame
+
+fun factify_atp_proof fact_names hyp_ts concl_t atp_proof =
+ let
+ fun factify_step ((num, ss), role, t, rule, deps) =
+ let
+ val (ss', role', t') =
+ (case resolve_conjecture ss of
+ [j] =>
+ if j = length hyp_ts then ([], Conjecture, concl_t) else ([], Hypothesis, nth hyp_ts j)
+ | _ =>
+ (case resolve_fact fact_names ss of
+ [] => (ss, Plain, t)
+ | facts => (map fst facts, Axiom, t)))
+ in
+ ((num, ss'), role', t', rule, deps)
+ end
+
+ val atp_proof = map factify_step atp_proof
+ val names = map #1 atp_proof
+
+ fun repair_dep (num, ss) = (num, the_default ss (AList.lookup (op =) names num))
+ fun repair_deps (name, role, t, rule, deps) = (name, role, t, rule, map repair_dep deps)
+
+ in map repair_deps atp_proof end
+
end;
--- a/src/HOL/Tools/Datatype/datatype_data.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/Datatype/datatype_data.ML Wed Nov 20 16:43:09 2013 +0100
@@ -118,7 +118,7 @@
expands = [],
sel_splits = [],
sel_split_asms = [],
- case_ifs = []};
+ case_eq_ifs = []};
fun register dt_infos =
Data.map (fn {types, constrs, cases} =>
--- a/src/HOL/Tools/Metis/metis_reconstruct.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/Metis/metis_reconstruct.ML Wed Nov 20 16:43:09 2013 +0100
@@ -376,7 +376,7 @@
val _ = trace_msg ctxt (fn () => "abstraction: " ^ Syntax.string_of_term ctxt tm_abs)
val _ = trace_msg ctxt (fn () => "i_tm: " ^ Syntax.string_of_term ctxt i_tm)
val _ = trace_msg ctxt (fn () => "located term: " ^ Syntax.string_of_term ctxt tm_subst)
- val imax = maxidx_of_term (i_tm $ tm_abs $ tm_subst) (*ill typed but gives right max*)
+ val imax = maxidx_of_term (i_tm $ tm_abs $ tm_subst) (*ill-typed but gives right max*)
val subst' = Thm.incr_indexes (imax+1) (if pos then subst_em else ssubst_em)
val _ = trace_msg ctxt (fn () => "subst' " ^ Display.string_of_thm ctxt subst')
val eq_terms = map (pairself (cterm_of thy))
--- a/src/HOL/Tools/Metis/metis_tactic.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/Metis/metis_tactic.ML Wed Nov 20 16:43:09 2013 +0100
@@ -294,7 +294,7 @@
()
val (schem_facts, nonschem_facts) = List.partition has_tvar facts
val type_encs = override_type_encs |> the_default default_type_encs
- val lam_trans = lam_trans |> the_default metis_default_lam_trans
+ val lam_trans = lam_trans |> the_default default_metis_lam_trans
in
HEADGOAL (Method.insert_tac nonschem_facts THEN'
CHANGED_PROP o generic_metis_tac type_encs lam_trans ctxt
--- a/src/HOL/Tools/Nitpick/nitpick_hol.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML Wed Nov 20 16:43:09 2013 +0100
@@ -1645,10 +1645,10 @@
(hol_ctxt as {thy, ctxt, stds, whacks, total_consts, case_names,
def_tables, ground_thm_table, ersatz_table, ...}) =
let
- fun do_numeral depth Ts mult T t0 t1 =
+ fun do_numeral depth Ts mult T some_t0 t1 t2 =
(if is_number_type ctxt T then
let
- val j = mult * HOLogic.dest_num t1
+ val j = mult * HOLogic.dest_num t2
in
if j = 1 then
raise SAME ()
@@ -1667,15 +1667,16 @@
handle TERM _ => raise SAME ()
else
raise SAME ())
- handle SAME () => s_betapply [] (do_term depth Ts t0, do_term depth Ts t1)
+ handle SAME () => (case some_t0 of NONE => s_betapply [] (do_term depth Ts t1, do_term depth Ts t2)
+ | SOME t0 => s_betapply [] (do_term depth Ts t0, s_betapply [] (do_term depth Ts t1, do_term depth Ts t2)))
and do_term depth Ts t =
case t of
- (t0 as Const (@{const_name Num.neg_numeral_class.neg_numeral},
- Type (@{type_name fun}, [_, ran_T]))) $ t1 =>
- do_numeral depth Ts ~1 ran_T t0 t1
- | (t0 as Const (@{const_name Num.numeral_class.numeral},
- Type (@{type_name fun}, [_, ran_T]))) $ t1 =>
- do_numeral depth Ts 1 ran_T t0 t1
+ (t0 as Const (@{const_name uminus}, _) $ ((t1 as Const (@{const_name numeral},
+ Type (@{type_name fun}, [_, ran_T]))) $ t2)) =>
+ do_numeral depth Ts ~1 ran_T (SOME t0) t1 t2
+ | (t1 as Const (@{const_name numeral},
+ Type (@{type_name fun}, [_, ran_T]))) $ t2 =>
+ do_numeral depth Ts 1 ran_T NONE t1 t2
| Const (@{const_name refl_on}, T) $ Const (@{const_name top}, _) $ t2 =>
do_const depth Ts t (@{const_name refl'}, range_type T) [t2]
| (t0 as Const (@{const_name Sigma}, Type (_, [T1, Type (_, [T2, T3])])))
--- a/src/HOL/Tools/Qelim/cooper.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/Qelim/cooper.ML Wed Nov 20 16:43:09 2013 +0100
@@ -42,7 +42,6 @@
@{term "nat"}, @{term "int"},
@{term "Num.One"}, @{term "Num.Bit0"}, @{term "Num.Bit1"},
@{term "Num.numeral :: num => int"}, @{term "Num.numeral :: num => nat"},
- @{term "Num.neg_numeral :: num => int"},
@{term "0::int"}, @{term "1::int"}, @{term "0::nat"}, @{term "1::nat"},
@{term "True"}, @{term "False"}];
@@ -610,8 +609,6 @@
| num_of_term vs @{term "1::int"} = Proc.C (Proc.Int_of_integer 1)
| num_of_term vs (t as Const (@{const_name numeral}, _) $ _) =
Proc.C (Proc.Int_of_integer (dest_number t))
- | num_of_term vs (t as Const (@{const_name neg_numeral}, _) $ _) =
- Proc.Neg (Proc.C (Proc.Int_of_integer (dest_number t)))
| num_of_term vs (Const (@{const_name Groups.uminus}, _) $ t') =
Proc.Neg (num_of_term vs t')
| num_of_term vs (Const (@{const_name Groups.plus}, _) $ t1 $ t2) =
--- a/src/HOL/Tools/SMT/smt_builtin.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/SMT/smt_builtin.ML Wed Nov 20 16:43:09 2013 +0100
@@ -144,9 +144,10 @@
(case try HOLogic.dest_number t of
NONE => NONE
| SOME (T, i) =>
- (case lookup_builtin_typ ctxt T of
- SOME (_, Int (_, g)) => g T i |> Option.map (rpair T)
- | _ => NONE))
+ if i < 0 then NONE else
+ (case lookup_builtin_typ ctxt T of
+ SOME (_, Int (_, g)) => g T i |> Option.map (rpair T)
+ | _ => NONE))
val is_builtin_num = is_some oo dest_builtin_num
--- a/src/HOL/Tools/SMT/smt_normalize.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/SMT/smt_normalize.ML Wed Nov 20 16:43:09 2013 +0100
@@ -526,23 +526,26 @@
local
(*
- rewrite negative numerals into positive numerals,
- rewrite Numeral0 into 0
rewrite Numeral1 into 1
+ rewrite - 0 into 0
*)
- fun is_strange_number ctxt (t as Const (@{const_name neg_numeral}, _) $ _) =
- SMT_Builtin.is_builtin_num ctxt t
- | is_strange_number _ _ = false
+ fun is_irregular_number (Const (@{const_name numeral}, _) $ Const (@{const_name num.One}, _)) =
+ true
+ | is_irregular_number (Const (@{const_name uminus}, _) $ Const (@{const_name Groups.zero}, _)) =
+ true
+ | is_irregular_number _ =
+ false;
- val pos_num_ss =
+ fun is_strange_number ctxt t = is_irregular_number t andalso SMT_Builtin.is_builtin_num ctxt t;
+
+ val proper_num_ss =
simpset_of (put_simpset HOL_ss @{context}
- addsimps [@{thm Num.numeral_One}]
- addsimps [@{thm Num.neg_numeral_def}])
+ addsimps @{thms Num.numeral_One minus_zero})
fun norm_num_conv ctxt =
SMT_Utils.if_conv (is_strange_number ctxt)
- (Simplifier.rewrite (put_simpset pos_num_ss ctxt)) Conv.no_conv
+ (Simplifier.rewrite (put_simpset proper_num_ss ctxt)) Conv.no_conv
in
fun normalize_numerals_conv ctxt =
--- a/src/HOL/Tools/SMT/smt_utils.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/SMT/smt_utils.ML Wed Nov 20 16:43:09 2013 +0100
@@ -140,7 +140,6 @@
is_num env t andalso is_num env u
| is_num env (Const (@{const_name Let}, _) $ t $ Abs (_, _, u)) =
is_num (t :: env) u
- | is_num env (Const (@{const_name uminus}, _) $ t) = is_num env t
| is_num env (Bound i) = i < length env andalso is_num env (nth env i)
| is_num _ t = can HOLogic.dest_number t
in is_num [] end
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_annotate.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_annotate.ML Wed Nov 20 16:43:09 2013 +0100
@@ -17,7 +17,7 @@
signature SLEDGEHAMMER_ANNOTATE =
sig
val annotate_types : Proof.context -> term -> term
-end
+end;
structure Sledgehammer_Annotate : SLEDGEHAMMER_ANNOTATE =
struct
@@ -215,4 +215,4 @@
|> sort int_ord
in introduce_annotations subst typing_spots t t' end
-end
+end;
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_compress.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_compress.ML Wed Nov 20 16:43:09 2013 +0100
@@ -15,8 +15,7 @@
type preplay_interface = Sledgehammer_Preplay.preplay_interface
val compress_proof : real -> preplay_interface -> isar_proof -> isar_proof
-end
-
+end;
structure Sledgehammer_Compress : SLEDGEHAMMER_COMPRESS =
struct
@@ -148,10 +147,9 @@
val (get_successors : label -> label list,
replace_successor: label -> label list -> label -> unit) =
let
-
fun add_refs (Let _) tab = tab
- | add_refs (Prove (_, _, l as v, _, _, By ((lfs, _), _))) tab =
- fold (fn lf as key => Canonical_Lbl_Tab.cons_list (key, v)) lfs tab
+ | add_refs (Prove (_, _, v, _, _, By ((lfs, _), _))) tab =
+ fold (fn key => Canonical_Lbl_Tab.cons_list (key, v)) lfs tab
val tab =
Canonical_Lbl_Tab.empty
@@ -180,7 +178,7 @@
if null subs orelse not (compress_further ()) then
(set_preplay_time l (false, time);
Prove (qs, fix, l, t, List.revAppend (nontriv_subs, subs),
- By_Metis (lfs, gfs)) )
+ By ((lfs, gfs), MetisM)))
else
case subs of
((sub as Proof(_, Assume assms, sub_steps)) :: subs) =>
@@ -199,7 +197,7 @@
subtract (op =) (map fst assms) lfs'
|> union (op =) lfs
val gfs'' = union (op =) gfs' gfs
- val by = By_Metis (lfs'', gfs'')
+ val by = By ((lfs'', gfs''), MetisM)
val step'' = Prove (qs, fix, l, t, subs'', by)
(* check if the modified step can be preplayed fast enough *)
@@ -251,7 +249,7 @@
val candidates =
let
- fun add_cand (i, Let _) = I
+ fun add_cand (_, Let _) = I
| add_cand (i, Prove (_, _, l, t, _, _)) =
cons (i, l, size_of_term t)
in
@@ -260,7 +258,7 @@
|> fold_index add_cand) []
end
- fun try_eliminate (cand as (i, l, _)) succ_lbls steps =
+ fun try_eliminate (i, l, _) succ_lbls steps =
let
(* only touch steps that can be preplayed successfully *)
val (false, time) = get_preplay_time l
@@ -371,5 +369,4 @@
do_proof proof
end
-
-end
+end;
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_fact.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_fact.ML Wed Nov 20 16:43:09 2013 +0100
@@ -527,7 +527,6 @@
[]
else
let
- val thy = Proof_Context.theory_of ctxt
val chained =
chained
|> maps (fn th => insert Thm.eq_thm_prop (zero_var_indexes th) [th])
@@ -536,10 +535,6 @@
maps (map (fn ((name, stature), th) => ((K name, stature), th))
o fact_of_ref ctxt reserved chained css) add
else
- (* The "fact_distinct" call would have cleaner semantics if it called "Pattern.equiv"
- instead of "Pattern.matches", but it would also be slower and less precise.
- "Pattern.matches" throws out theorems that are strict instances of other theorems, but
- only if the instance is met after the general version. *)
let
val (add, del) = pairself (Attrib.eval_thms ctxt) (add, del)
val facts =
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML Wed Nov 20 16:43:09 2013 +0100
@@ -368,7 +368,7 @@
val goal = prop_of (#goal (Proof.goal state))
val facts = nearly_all_facts ctxt false fact_override Symtab.empty
Termtab.empty [] [] goal
- fun learn prover = mash_learn_proof ctxt params prover goal facts
+ val learn = mash_learn_proof ctxt params goal facts
in run_minimize params learn i (#add fact_override) state end
else if subcommand = messagesN then
messages opt_i
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_mash.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_mash.ML Wed Nov 20 16:43:09 2013 +0100
@@ -82,11 +82,8 @@
-> ('b * thm) list -> ('b * thm) list * ('b * thm) list
val add_const_counts : term -> int Symtab.table -> int Symtab.table
val mash_suggested_facts :
- Proof.context -> params -> string -> int -> term list -> term
- -> raw_fact list -> fact list * fact list
- val mash_learn_proof :
- Proof.context -> params -> string -> term -> ('a * thm) list -> thm list
- -> unit
+ Proof.context -> params -> int -> term list -> term -> raw_fact list -> fact list * fact list
+ val mash_learn_proof : Proof.context -> params -> term -> ('a * thm) list -> thm list -> unit
val mash_learn :
Proof.context -> params -> fact_override -> thm list -> bool -> unit
@@ -568,8 +565,7 @@
{comment = "Goal: " ^ goal_name, state = Proof.init ctxt, goal = goal, subgoal = 1,
subgoal_count = 1, factss = [("", facts)]}
in
- get_minimizing_prover ctxt MaSh (K (K ())) prover params (K (K (K "")))
- problem
+ get_minimizing_prover ctxt MaSh (K ()) prover params (K (K (K ""))) problem
end
val bad_types = [@{type_name prop}, @{type_name bool}, @{type_name fun}]
@@ -930,8 +926,7 @@
fold (fn s => Symtab.map_default (s, 0) (Integer.add 1))
(Term.add_const_names t [])
-fun mash_suggested_facts ctxt ({debug, overlord, ...} : params) prover max_facts
- hyp_ts concl_t facts =
+fun mash_suggested_facts ctxt ({debug, overlord, ...} : params) max_facts hyp_ts concl_t facts =
let
val thy = Proof_Context.theory_of ctxt
val thy_name = Context.theory_name thy
@@ -1017,8 +1012,7 @@
val desc = ("Machine learner for Sledgehammer", "")
in Async_Manager.thread MaShN birth_time death_time desc task end
-fun mash_learn_proof ctxt ({overlord, timeout, ...} : params) prover t facts
- used_ths =
+fun mash_learn_proof ctxt ({overlord, timeout, ...} : params) t facts used_ths =
if is_mash_enabled () then
launch_thread (timeout |> the_default one_day) (fn () =>
let
@@ -1328,7 +1322,7 @@
(mepo_suggested_facts ctxt params max_facts NONE hyp_ts concl_t unique_facts
|> weight_facts_steeply, [])
fun mash () =
- mash_suggested_facts ctxt params prover (generous_max_facts max_facts) hyp_ts concl_t facts
+ mash_suggested_facts ctxt params (generous_max_facts max_facts) hyp_ts concl_t facts
|>> weight_facts_steeply
val mess =
(* the order is important for the "case" expression below *)
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML Wed Nov 20 16:43:09 2013 +0100
@@ -17,16 +17,14 @@
val auto_minimize_min_facts : int Config.T
val auto_minimize_max_time : real Config.T
val minimize_facts :
- (string -> thm list -> unit) -> string -> params -> bool -> int -> int
+ (thm list -> unit) -> string -> params -> bool -> int -> int
-> Proof.state -> thm -> play Lazy.lazy option
-> ((string * stature) * thm list) list
-> ((string * stature) * thm list) list option
* (play Lazy.lazy * (play -> string) * string)
- val get_minimizing_prover :
- Proof.context -> mode -> (string -> thm list -> unit) -> string -> prover
- val run_minimize :
- params -> (string -> thm list -> unit) -> int
- -> (Facts.ref * Attrib.src list) list -> Proof.state -> unit
+ val get_minimizing_prover : Proof.context -> mode -> (thm list -> unit) -> string -> prover
+ val run_minimize : params -> (thm list -> unit) -> int -> (Facts.ref * Attrib.src list) list ->
+ Proof.state -> unit
end;
structure Sledgehammer_Minimize : SLEDGEHAMMER_MINIMIZE =
@@ -221,7 +219,7 @@
(case min_facts |> filter is_fact_chained |> length of
0 => ""
| n => "\n(including " ^ string_of_int n ^ " chained)") ^ ".");
- (if learn then do_learn prover_name (maps snd min_facts) else ());
+ (if learn then do_learn (maps snd min_facts) else ());
(SOME min_facts,
(if is_some preplay0 andalso length min_facts = length facts then
the preplay0
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_minimize_isar.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_minimize_isar.ML Wed Nov 20 16:43:09 2013 +0100
@@ -11,7 +11,7 @@
type isar_proof = Sledgehammer_Proof.isar_proof
val minimize_dependencies_and_remove_unrefed_steps :
bool -> preplay_interface -> isar_proof -> isar_proof
-end
+end;
structure Sledgehammer_Minimize_Isar : SLEDGEHAMMER_MINIMIZE_ISAR =
struct
@@ -105,4 +105,4 @@
snd (do_proof proof)
end
-end
+end;
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_preplay.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_preplay.ML Wed Nov 20 16:43:09 2013 +0100
@@ -35,7 +35,7 @@
val proof_preplay_interface :
bool -> Proof.context -> string -> string -> bool -> Time.time
-> isar_proof -> preplay_interface
-end
+end;
structure Sledgehammer_Preplay : SLEDGEHAMMER_PREPLAY =
struct
@@ -312,4 +312,4 @@
overall_preplay_stats = overall_preplay_stats}
end
-end
+end;
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_print.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_print.ML Wed Nov 20 16:43:09 2013 +0100
@@ -293,4 +293,4 @@
of_indent 0 ^ (if n <> 1 then "next" else "qed")
end
-end
+end;
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_proof.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_proof.ML Wed Nov 20 16:43:09 2013 +0100
@@ -33,9 +33,6 @@
ArithM |
BlastM
- (* legacy *)
- val By_Metis : facts -> byline
-
val no_label : label
val no_facts : facts
@@ -70,7 +67,7 @@
structure Canonical_Lbl_Tab : TABLE
-end
+end;
structure Sledgehammer_Proof : SLEDGEHAMMER_PROOF =
struct
@@ -100,9 +97,6 @@
ArithM |
BlastM
-(* legacy *)
-fun By_Metis facts = By (facts, MetisM)
-
val no_label = ("", ~1)
val no_facts = ([],[])
@@ -202,4 +196,4 @@
fst (do_proof proof (0, []))
end
-end
+end;
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_provers.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_provers.ML Wed Nov 20 16:43:09 2013 +0100
@@ -367,7 +367,7 @@
[]
else
[("type_enc", [hd (unalias_type_enc type_enc')])]) @
- (if is_none lam_trans andalso lam_trans' = metis_default_lam_trans then
+ (if is_none lam_trans andalso lam_trans' = default_metis_lam_trans then
[]
else
[("lam_trans", [lam_trans'])])
@@ -820,7 +820,7 @@
bunch_of_reconstructors needs_full_types
(lam_trans_of_atp_proof atp_proof
o (fn desperate => if desperate then hide_lamsN
- else metis_default_lam_trans))
+ else default_metis_lam_trans))
in
(used_facts,
Lazy.lazy (fn () =>
@@ -838,18 +838,27 @@
Output.urgent_message "Generating proof text..."
else
()
- val isar_params =
- (debug, verbose, preplay_timeout, isar_compress, isar_try0,
- pool, fact_names, lifted, sym_tab, atp_proof, goal)
+ fun isar_params () =
+ let
+ val metis_type_enc =
+ if is_typed_helper_used_in_atp_proof atp_proof then full_typesN
+ else partial_typesN
+ val metis_lam_trans = lam_trans_of_atp_proof atp_proof default_metis_lam_trans
+ val atp_proof =
+ atp_proof
+ |> termify_atp_proof ctxt pool lifted sym_tab
+ |> factify_atp_proof fact_names hyp_ts concl_t
+ in
+ (debug, verbose, metis_type_enc, metis_lam_trans, preplay_timeout,
+ isar_compress, isar_try0, atp_proof, goal)
+ end
val one_line_params =
(preplay, proof_banner mode name, used_facts,
- choose_minimize_command ctxt params minimize_command name
- preplay,
+ choose_minimize_command ctxt params minimize_command name preplay,
subgoal, subgoal_count)
val num_chained = length (#facts (Proof.goal state))
in
- proof_text ctxt isar_proofs isar_params
- num_chained one_line_params
+ proof_text ctxt isar_proofs isar_params num_chained one_line_params
end,
(if verbose then
"\nATP real CPU time: " ^ string_of_time run_time ^ "."
@@ -1047,7 +1056,7 @@
play_one_line_proof mode debug verbose preplay_timeout used_pairs
state subgoal SMT
(bunch_of_reconstructors false (fn desperate =>
- if desperate then liftingN else metis_default_lam_trans))),
+ if desperate then liftingN else default_metis_lam_trans))),
fn preplay =>
let
val one_line_params =
@@ -1081,7 +1090,7 @@
val reconstr =
if name = metisN then
Metis (type_enc |> the_default (hd partial_type_encs),
- lam_trans |> the_default metis_default_lam_trans)
+ lam_trans |> the_default default_metis_lam_trans)
else if name = smtN then
SMT
else
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML Wed Nov 20 16:43:09 2013 +0100
@@ -7,28 +7,19 @@
signature SLEDGEHAMMER_RECONSTRUCT =
sig
+ type ('a, 'b) atp_step = ('a, 'b) ATP_Proof.atp_step
type 'a atp_proof = 'a ATP_Proof.atp_proof
type stature = ATP_Problem_Generate.stature
- type one_line_params = Sledgehammer_Print.one_line_params
+ type one_line_params = Sledgehammer_Reconstructor.one_line_params
type isar_params =
- bool * bool * Time.time option * real * bool * string Symtab.table
- * (string * stature) list vector * (string * term) list * int Symtab.table
- * string atp_proof * thm
+ bool * bool * string * string * Time.time option * real * bool * (term, string) atp_step list *
+ thm
- val lam_trans_of_atp_proof : string atp_proof -> string -> string
- val is_typed_helper_used_in_atp_proof : string atp_proof -> bool
- val used_facts_in_atp_proof :
- Proof.context -> (string * stature) list vector -> string atp_proof ->
- (string * stature) list
- val used_facts_in_unsound_atp_proof :
- Proof.context -> (string * stature) list vector -> 'a atp_proof ->
- string list option
val isar_proof_text :
Proof.context -> bool option -> isar_params -> one_line_params -> string
val proof_text :
- Proof.context -> bool option -> isar_params -> int -> one_line_params
- -> string
+ Proof.context -> bool option -> (unit -> isar_params) -> int -> one_line_params -> string
end;
structure Sledgehammer_Reconstruct : SLEDGEHAMMER_RECONSTRUCT =
@@ -56,190 +47,13 @@
open String_Redirect
-(** fact extraction from ATP proofs **)
-
-fun find_first_in_list_vector vec key =
- Vector.foldl (fn (ps, NONE) => AList.lookup (op =) ps key
- | (_, value) => value) NONE vec
-
-val unprefix_fact_number = space_implode "_" o tl o space_explode "_"
-
-fun resolve_one_named_fact fact_names s =
- case try (unprefix fact_prefix) s of
- SOME s' =>
- let val s' = s' |> unprefix_fact_number |> unascii_of in
- s' |> find_first_in_list_vector fact_names |> Option.map (pair s')
- end
- | NONE => NONE
-fun resolve_fact fact_names = map_filter (resolve_one_named_fact fact_names)
-fun is_fact fact_names = not o null o resolve_fact fact_names
-
-fun resolve_one_named_conjecture s =
- case try (unprefix conjecture_prefix) s of
- SOME s' => Int.fromString s'
- | NONE => NONE
-
-val resolve_conjecture = map_filter resolve_one_named_conjecture
-val is_conjecture = not o null o resolve_conjecture
-
-val ascii_of_lam_fact_prefix = ascii_of lam_fact_prefix
-
-(* overapproximation (good enough) *)
-fun is_lam_lifted s =
- String.isPrefix fact_prefix s andalso
- String.isSubstring ascii_of_lam_fact_prefix s
-
-val is_combinator_def = String.isPrefix (helper_prefix ^ combinator_prefix)
-
-fun is_axiom_used_in_proof pred =
- exists (fn ((_, ss), _, _, _, []) => exists pred ss | _ => false)
-
-fun lam_trans_of_atp_proof atp_proof default =
- case (is_axiom_used_in_proof is_combinator_def atp_proof,
- is_axiom_used_in_proof is_lam_lifted atp_proof) of
- (false, false) => default
- | (false, true) => liftingN
-(* | (true, true) => combs_and_liftingN -- not supported by "metis" *)
- | (true, _) => combsN
-
-val is_typed_helper_name =
- String.isPrefix helper_prefix andf String.isSuffix typed_helper_suffix
-
-fun is_typed_helper_used_in_atp_proof atp_proof =
- is_axiom_used_in_proof is_typed_helper_name atp_proof
-
-fun add_non_rec_defs fact_names accum =
- Vector.foldl (fn (facts, facts') =>
- union (op =) (filter (fn (_, (_, status)) => status = Non_Rec_Def) facts)
- facts')
- accum fact_names
-
-val isa_ext = Thm.get_name_hint @{thm ext}
-val isa_short_ext = Long_Name.base_name isa_ext
-
-fun ext_name ctxt =
- if Thm.eq_thm_prop (@{thm ext},
- singleton (Attrib.eval_thms ctxt) (Facts.named isa_short_ext, [])) then
- isa_short_ext
- else
- isa_ext
-
-val leo2_extcnf_equal_neg_rule = "extcnf_equal_neg"
-val leo2_unfold_def_rule = "unfold_def"
-
-fun add_fact ctxt fact_names ((_, ss), _, _, rule, deps) =
- (if rule = leo2_extcnf_equal_neg_rule then
- insert (op =) (ext_name ctxt, (Global, General))
- else if rule = leo2_unfold_def_rule then
- (* LEO 1.3.3 does not record definitions properly, leading to missing
- dependencies in the TSTP proof. Remove the next line once this is
- fixed. *)
- add_non_rec_defs fact_names
- else if rule = agsyhol_coreN orelse rule = satallax_coreN then
- (fn [] =>
- (* agsyHOL and Satallax don't include definitions in their
- unsatisfiable cores, so we assume the worst and include them all
- here. *)
- [(ext_name ctxt, (Global, General))] |> add_non_rec_defs fact_names
- | facts => facts)
- else
- I)
- #> (if null deps then union (op =) (resolve_fact fact_names ss) else I)
+fun raw_label_of_num num = (num, 0)
-fun used_facts_in_atp_proof ctxt fact_names atp_proof =
- if null atp_proof then Vector.foldl (uncurry (union (op =))) [] fact_names
- else fold (add_fact ctxt fact_names) atp_proof []
-
-fun used_facts_in_unsound_atp_proof _ _ [] = NONE
- | used_facts_in_unsound_atp_proof ctxt fact_names atp_proof =
- let val used_facts = used_facts_in_atp_proof ctxt fact_names atp_proof in
- if forall (fn (_, (sc, _)) => sc = Global) used_facts andalso
- not (is_axiom_used_in_proof (is_conjecture o single) atp_proof) then
- SOME (map fst used_facts)
- else
- NONE
- end
-
-
-(** Isar proof construction and manipulation **)
-
-val assume_prefix = "a"
-val have_prefix = "f"
-val raw_prefix = "x"
-
-fun raw_label_of_name (num, ss) =
- case resolve_conjecture ss of
- [j] => (conjecture_prefix, j)
- | _ => (raw_prefix ^ ascii_of num, 0)
-
-fun label_of_clause [name] = raw_label_of_name name
- | label_of_clause c =
- (space_implode "___" (map (fst o raw_label_of_name) c), 0)
-
-fun add_fact_of_dependencies fact_names (names as [(_, ss)]) =
- if is_fact fact_names ss then
- apsnd (union (op =) (map fst (resolve_fact fact_names ss)))
- else
- apfst (insert (op =) (label_of_clause names))
- | add_fact_of_dependencies _ names =
- apfst (insert (op =) (label_of_clause names))
+fun label_of_clause [(num, _)] = raw_label_of_num num
+ | label_of_clause c = (space_implode "___" (map (fst o raw_label_of_num o fst) c), 0)
-fun repair_name "$true" = "c_True"
- | repair_name "$false" = "c_False"
- | repair_name "$$e" = tptp_equal (* seen in Vampire proofs *)
- | repair_name s =
- if is_tptp_equal s orelse
- (* seen in Vampire proofs *)
- (String.isPrefix "sQ" s andalso String.isSuffix "_eqProxy" s) then
- tptp_equal
- else
- s
-
-fun infer_formula_types ctxt =
- Type.constraint HOLogic.boolT
- #> Syntax.check_term
- (Proof_Context.set_mode Proof_Context.mode_schematic ctxt)
-
-val combinator_table =
- [(@{const_name Meson.COMBI}, @{thm Meson.COMBI_def [abs_def]}),
- (@{const_name Meson.COMBK}, @{thm Meson.COMBK_def [abs_def]}),
- (@{const_name Meson.COMBB}, @{thm Meson.COMBB_def [abs_def]}),
- (@{const_name Meson.COMBC}, @{thm Meson.COMBC_def [abs_def]}),
- (@{const_name Meson.COMBS}, @{thm Meson.COMBS_def [abs_def]})]
-
-fun uncombine_term thy =
- let
- fun aux (t1 $ t2) = betapply (pairself aux (t1, t2))
- | aux (Abs (s, T, t')) = Abs (s, T, aux t')
- | aux (t as Const (x as (s, _))) =
- (case AList.lookup (op =) combinator_table s of
- SOME thm => thm |> prop_of |> specialize_type thy x
- |> Logic.dest_equals |> snd
- | NONE => t)
- | aux t = t
- in aux end
-
-fun unlift_term lifted =
- map_aterms (fn t as Const (s, _) =>
- if String.isPrefix lam_lifted_prefix s then
- case AList.lookup (op =) lifted s of
- SOME t =>
- (* FIXME: do something about the types *)
- unlift_term lifted t
- | NONE => t
- else
- t
- | t => t)
-
-fun decode_line ctxt lifted sym_tab (name, role, u, rule, deps) =
- let
- val thy = Proof_Context.theory_of ctxt
- val t =
- u |> prop_of_atp ctxt true sym_tab
- |> uncombine_term thy
- |> unlift_term lifted
- |> infer_formula_types ctxt
- in (name, role, t, rule, deps) end
+fun add_fact_of_dependencies [(_, ss as _ :: _)] = apsnd (union (op =) ss)
+ | add_fact_of_dependencies names = apfst (insert (op =) (label_of_clause names))
fun replace_one_dependency (old, new) dep =
if is_same_atp_step dep old then new else [dep]
@@ -257,20 +71,17 @@
(* Discard facts; consolidate adjacent lines that prove the same formula, since
they differ only in type information.*)
-fun add_line fact_names (name as (_, ss), role, t, rule, []) lines =
+fun add_line (name as (_, ss), role, t, rule, []) lines =
(* No dependencies: fact, conjecture, or (for Vampire) internal facts or
definitions. *)
- if is_conjecture ss then
+ if role = Conjecture orelse role = Hypothesis then
(name, role, t, rule, []) :: lines
- else if is_fact fact_names ss then
+ else if role = Axiom then
(* Facts are not proof lines. *)
- if is_only_type_information t then
- map (replace_dependencies_in_line (name, [])) lines
- else
- lines
+ lines |> is_only_type_information t ? map (replace_dependencies_in_line (name, []))
else
map (replace_dependencies_in_line (name, [])) lines
- | add_line _ (line as (name, role, t, _, _)) lines =
+ | add_line (line as (name, role, t, _, _)) lines =
(* Type information will be deleted later; skip repetition test. *)
if is_only_type_information t then
line :: lines
@@ -280,22 +91,9 @@
| (pre, (name', _, _, _, _) :: post) =>
line :: pre @ map (replace_dependencies_in_line (name', [name])) post
-val waldmeister_conjecture_num = "1.0.0.0"
-
-fun repair_waldmeister_endgame arg =
- let
- fun do_tail (name, _, t, rule, deps) =
- (name, Negated_Conjecture, s_not t, rule, deps)
- fun do_body [] = []
- | do_body ((line as ((num, _), _, _, _, _)) :: lines) =
- if num = waldmeister_conjecture_num then map do_tail (line :: lines)
- else line :: do_body lines
- in do_body arg end
-
(* Recursively delete empty lines (type information) from the proof. *)
fun add_nontrivial_line (line as (name, _, t, _, [])) lines =
- if is_only_type_information t then delete_dependency name lines
- else line :: lines
+ if is_only_type_information t then delete_dependency name lines else line :: lines
| add_nontrivial_line line lines = line :: lines
and delete_dependency name lines =
fold_rev add_nontrivial_line
@@ -307,12 +105,9 @@
val is_skolemize_rule =
member (op =) [e_skolemize_rule, vampire_skolemisation_rule]
-fun add_desired_line fact_names (name as (_, ss), role, t, rule, deps)
- (j, lines) =
+fun add_desired_line (name as (_, ss), role, t, rule, deps) (j, lines) =
(j + 1,
- if is_fact fact_names ss orelse
- is_conjecture ss orelse
- is_skolemize_rule rule orelse
+ if role <> Plain orelse is_skolemize_rule rule orelse
(* the last line must be kept *)
j = 0 orelse
(not (is_only_type_information t) andalso
@@ -335,14 +130,17 @@
fun do_label l = if member (op =) used_ls l then l else no_label
fun do_assms (Assume assms) = Assume (map (apfst do_label) assms)
fun do_step (Prove (qs, xs, l, t, subproofs, by)) =
- Prove (qs, xs, do_label l, t, map do_proof subproofs, by)
+ Prove (qs, xs, do_label l, t, map do_proof subproofs, by)
| do_step step = step
and do_proof (Proof (fix, assms, steps)) =
- Proof (fix, do_assms assms, map do_step steps)
+ Proof (fix, do_assms assms, map do_step steps)
in do_proof proof end
fun prefix_of_depth n = replicate_string (n + 1)
+val assume_prefix = "a"
+val have_prefix = "f"
+
val relabel_proof =
let
fun fresh_label depth prefix (old as (l, subst, next)) =
@@ -352,24 +150,17 @@
let val l' = (prefix_of_depth depth prefix, next) in
(l', (l, l') :: subst, next + 1)
end
- fun do_facts subst =
- apfst (maps (the_list o AList.lookup (op =) subst))
+ fun do_facts subst = apfst (maps (the_list o AList.lookup (op =) subst))
fun do_assm depth (l, t) (subst, next) =
- let
- val (l, subst, next) =
- (l, subst, next) |> fresh_label depth assume_prefix
- in
+ let val (l, subst, next) = (l, subst, next) |> fresh_label depth assume_prefix in
((l, t), (subst, next))
end
fun do_assms subst depth (Assume assms) =
- fold_map (do_assm depth) assms (subst, 1)
- |> apfst Assume
- |> apsnd fst
+ fold_map (do_assm depth) assms (subst, 1) |>> Assume ||> fst
fun do_steps _ _ _ [] = []
| do_steps subst depth next (Prove (qs, xs, l, t, sub, by) :: steps) =
let
- val (l, subst, next) =
- (l, subst, next) |> fresh_label depth have_prefix
+ val (l, subst, next) = (l, subst, next) |> fresh_label depth have_prefix
val sub = do_proofs subst depth sub
val by = by |> do_byline subst
in Prove (qs, xs, l, t, sub, by) :: do_steps subst depth next steps end
@@ -407,59 +198,38 @@
and chain_proofs proofs = map (chain_proof) proofs
in chain_proof end
+fun maybe_mk_Trueprop t = t |> fastype_of t = HOLogic.boolT ? HOLogic.mk_Trueprop
+
type isar_params =
- bool * bool * Time.time option * real * bool * string Symtab.table
- * (string * stature) list vector * (string * term) list * int Symtab.table
- * string atp_proof * thm
+ bool * bool * string * string * Time.time option * real * bool * (term, string) atp_step list *
+ thm
fun isar_proof_text ctxt isar_proofs
- (debug, verbose, preplay_timeout, isar_compress, isar_try0, pool,
- fact_names, lifted, sym_tab, atp_proof, goal)
+ (debug, verbose, metis_type_enc, metis_lam_trans, preplay_timeout, isar_compress,
+ isar_try0, atp_proof, goal)
(one_line_params as (_, _, _, _, subgoal, subgoal_count)) =
let
val (params, hyp_ts, concl_t) = strip_subgoal goal subgoal ctxt
val (_, ctxt) =
params
|> map (fn (s, T) => (Binding.name s, SOME T, NoSyn))
- |> (fn fixes =>
- ctxt |> Variable.set_body false
- |> Proof_Context.add_fixes fixes)
+ |> (fn fixes => ctxt |> Variable.set_body false |> Proof_Context.add_fixes fixes)
val one_line_proof = one_line_proof_text 0 one_line_params
- val type_enc =
- if is_typed_helper_used_in_atp_proof atp_proof then full_typesN
- else partial_typesN
- val lam_trans = lam_trans_of_atp_proof atp_proof metis_default_lam_trans
val do_preplay = preplay_timeout <> SOME Time.zeroTime
fun isar_proof_of () =
let
val atp_proof =
atp_proof
- |> clean_up_atp_proof_dependencies
- |> nasty_atp_proof pool
- |> map_term_names_in_atp_proof repair_name
- |> map (decode_line ctxt lifted sym_tab)
- |> repair_waldmeister_endgame
- |> rpair [] |-> fold_rev (add_line fact_names)
+ |> rpair [] |-> fold_rev add_line
|> rpair [] |-> fold_rev add_nontrivial_line
|> rpair (0, [])
- |-> fold_rev (add_desired_line fact_names)
+ |-> fold_rev add_desired_line
|> snd
- val conj_name = conjecture_prefix ^ string_of_int (length hyp_ts)
- val conjs =
- atp_proof |> map_filter
- (fn (name as (_, ss), _, _, _, []) =>
- if member (op =) ss conj_name then SOME name else NONE
- | _ => NONE)
+ val conjs = atp_proof |> map_filter (try (fn (name, Conjecture, _, _, []) => name))
val assms =
- atp_proof |> map_filter
- (fn (name as (_, ss), _, _, _, []) =>
- (case resolve_conjecture ss of
- [j] =>
- if j = length hyp_ts then NONE
- else SOME (raw_label_of_name name, nth hyp_ts j)
- | _ => NONE)
- | _ => NONE)
+ atp_proof
+ |> map_filter (try (fn ((num, _), Hypothesis, t, _, []) => (raw_label_of_num num, t)))
val bot = atp_proof |> List.last |> #1
val refute_graph =
atp_proof
@@ -480,41 +250,32 @@
I))))
atp_proof
fun is_clause_skolemize_rule [(s, _)] =
- Option.map (is_skolemize_rule o fst) (Symtab.lookup steps s) =
- SOME true
+ Option.map (is_skolemize_rule o fst) (Symtab.lookup steps s) = SOME true
| is_clause_skolemize_rule _ = false
- (* The assumptions and conjecture are "prop"s; the other formulas are
- "bool"s. *)
- fun prop_of_clause [(s, ss)] =
- (case resolve_conjecture ss of
- [j] => if j = length hyp_ts then concl_t else nth hyp_ts j
- | _ => the_default ("", @{term False}) (Symtab.lookup steps s)
- |> snd |> HOLogic.mk_Trueprop |> close_form)
+ (* The assumptions and conjecture are "prop"s; the other formulas are "bool"s. *)
+ fun prop_of_clause [(num, _)] =
+ Symtab.lookup steps num |> the |> snd |> maybe_mk_Trueprop |> close_form
| prop_of_clause names =
let
val lits = map snd (map_filter (Symtab.lookup steps o fst) names)
in
case List.partition (can HOLogic.dest_not) lits of
(negs as _ :: _, pos as _ :: _) =>
- s_imp (Library.foldr1 s_conj (map HOLogic.dest_not negs),
- Library.foldr1 s_disj pos)
+ s_imp (Library.foldr1 s_conj (map HOLogic.dest_not negs), Library.foldr1 s_disj pos)
| _ => fold (curry s_disj) lits @{term False}
end
|> HOLogic.mk_Trueprop |> close_form
fun isar_proof_of_direct_proof infs =
let
fun maybe_show outer c =
- (outer andalso length c = 1 andalso subset (op =) (c, conjs))
- ? cons Show
+ (outer andalso length c = 1 andalso subset (op =) (c, conjs)) ? cons Show
val is_fixed = Variable.is_declared ctxt orf can Name.dest_skolem
- fun skolems_of t =
- Term.add_frees t [] |> filter_out (is_fixed o fst) |> rev
+ fun skolems_of t = Term.add_frees t [] |> filter_out (is_fixed o fst) |> rev
fun do_steps outer predecessor accum [] =
accum
|> (if tainted = [] then
- cons (Prove (if outer then [Show] else [],
- Fix [], no_label, concl_t, [],
- By_Metis ([the predecessor], [])))
+ cons (Prove (if outer then [Show] else [], Fix [], no_label, concl_t, [],
+ By (([the predecessor], []), MetisM)))
else
I)
|> rev
@@ -522,26 +283,19 @@
let
val l = label_of_clause c
val t = prop_of_clause c
- val by =
- By_Metis
- (fold (add_fact_of_dependencies fact_names) gamma no_facts)
- fun prove sub by =
- Prove (maybe_show outer c [], Fix [], l, t, sub, by)
- fun do_rest l step =
- do_steps outer (SOME l) (step :: accum) infs
+ val by = By (fold add_fact_of_dependencies gamma no_facts, MetisM)
+ fun prove sub by = Prove (maybe_show outer c [], Fix [], l, t, sub, by)
+ fun do_rest l step = do_steps outer (SOME l) (step :: accum) infs
in
if is_clause_tainted c then
case gamma of
[g] =>
- if is_clause_skolemize_rule g andalso
- is_clause_tainted g then
+ if is_clause_skolemize_rule g andalso is_clause_tainted g then
let
val subproof =
- Proof (Fix (skolems_of (prop_of_clause g)),
- Assume [], rev accum)
+ Proof (Fix (skolems_of (prop_of_clause g)), Assume [], rev accum)
in
- do_steps outer (SOME l)
- [prove [subproof] (By_Metis no_facts)] []
+ do_steps outer (SOME l) [prove [subproof] (By (no_facts, MetisM))] []
end
else
do_rest l (prove [] by)
@@ -555,14 +309,13 @@
| do_steps outer predecessor accum (Cases cases :: infs) =
let
fun do_case (c, infs) =
- do_proof false [] [(label_of_clause c, prop_of_clause c)]
- infs
+ do_proof false [] [(label_of_clause c, prop_of_clause c)] infs
val c = succedent_of_cases cases
val l = label_of_clause c
val t = prop_of_clause c
val step =
- Prove (maybe_show outer c [], Fix [], l, t,
- map do_case cases, By_Metis (the_list predecessor, []))
+ Prove (maybe_show outer c [], Fix [], l, t, map do_case cases,
+ By ((the_list predecessor, []), MetisM))
in
do_steps outer (SOME l) (step :: accum) infs
end
@@ -584,20 +337,18 @@
|> redirect_graph axioms tainted bot
|> isar_proof_of_direct_proof
|> relabel_proof_canonically
- |> `(proof_preplay_interface debug ctxt type_enc lam_trans do_preplay
+ |> `(proof_preplay_interface debug ctxt metis_type_enc metis_lam_trans do_preplay
preplay_timeout)
val ((preplay_time, preplay_fail), isar_proof) =
isar_proof
- |> compress_proof
- (if isar_proofs = SOME true then isar_compress else 1000.0)
+ |> compress_proof (if isar_proofs = SOME true then isar_compress else 1000.0)
preplay_interface
|> isar_try0 ? try0 preplay_timeout preplay_interface
- |> minimize_dependencies_and_remove_unrefed_steps isar_try0
- preplay_interface
+ |> minimize_dependencies_and_remove_unrefed_steps isar_try0 preplay_interface
|> `overall_preplay_stats
||> clean_up_labels_in_proof
- val isar_text = string_of_proof ctxt type_enc lam_trans subgoal
- subgoal_count isar_proof
+ val isar_text =
+ string_of_proof ctxt metis_type_enc metis_lam_trans subgoal subgoal_count isar_proof
in
case isar_text of
"" =>
@@ -620,10 +371,8 @@
else
[])
in
- "\n\nStructured proof"
- ^ (commas msg |> not (null msg) ? enclose " (" ")")
- ^ ":\n" ^
- Active.sendback_markup [Markup.padding_command] isar_text
+ "\n\nStructured proof" ^ (commas msg |> not (null msg) ? enclose " (" ")") ^ ":\n" ^
+ Active.sendback_markup [Markup.padding_command] isar_text
end
end
val isar_proof =
@@ -647,7 +396,7 @@
(one_line_params as (preplay, _, _, _, _, _)) =
(if isar_proofs = SOME true orelse
(isar_proofs = NONE andalso isar_proof_would_be_a_good_idea preplay) then
- isar_proof_text ctxt isar_proofs isar_params
+ isar_proof_text ctxt isar_proofs (isar_params ())
else
one_line_proof_text num_chained) one_line_params
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_reconstructor.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_reconstructor.ML Wed Nov 20 16:43:09 2013 +0100
@@ -7,7 +7,6 @@
signature SLEDGEHAMMER_RECONSTRUCTOR =
sig
-
type stature = ATP_Problem_Generate.stature
datatype reconstructor =
@@ -25,8 +24,7 @@
play * string * (string * stature) list * minimize_command * int * int
val smtN : string
-
-end
+end;
structure Sledgehammer_Reconstructor : SLEDGEHAMMER_RECONSTRUCTOR =
struct
@@ -49,4 +47,4 @@
val smtN = "smt"
-end
+end;
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_run.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_run.ML Wed Nov 20 16:43:09 2013 +0100
@@ -284,10 +284,8 @@
val problem =
{comment = "", state = state, goal = goal, subgoal = i, subgoal_count = n,
factss = factss}
- fun learn prover =
- mash_learn_proof ctxt params prover (prop_of goal) all_facts
- val launch =
- launch_prover params mode output_result minimize_command only learn
+ val learn = mash_learn_proof ctxt params (prop_of goal) all_facts
+ val launch = launch_prover params mode output_result minimize_command only learn
in
if mode = Auto_Try then
(unknownN, state)
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_try0.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_try0.ML Wed Nov 20 16:43:09 2013 +0100
@@ -12,7 +12,7 @@
type preplay_interface = Sledgehammer_Preplay.preplay_interface
val try0 : Time.time -> preplay_interface -> isar_proof -> isar_proof
-end
+end;
structure Sledgehammer_Try0 : SLEDGEHAMMER_TRY0 =
struct
@@ -59,4 +59,4 @@
fun try0 preplay_timeout preplay_interface proof =
map_isar_steps (try0_step preplay_timeout preplay_interface) proof
-end
+end;
--- a/src/HOL/Tools/ctr_sugar.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/ctr_sugar.ML Wed Nov 20 16:43:09 2013 +0100
@@ -30,7 +30,7 @@
expands: thm list,
sel_splits: thm list,
sel_split_asms: thm list,
- case_ifs: thm list};
+ case_eq_ifs: thm list};
val morph_ctr_sugar: morphism -> ctr_sugar -> ctr_sugar
val transfer_ctr_sugar: Proof.context -> ctr_sugar -> ctr_sugar
@@ -90,7 +90,7 @@
expands: thm list,
sel_splits: thm list,
sel_split_asms: thm list,
- case_ifs: thm list};
+ case_eq_ifs: thm list};
fun eq_ctr_sugar ({ctrs = ctrs1, casex = case1, discs = discs1, selss = selss1, ...} : ctr_sugar,
{ctrs = ctrs2, casex = case2, discs = discs2, selss = selss2, ...} : ctr_sugar) =
@@ -98,7 +98,7 @@
fun morph_ctr_sugar phi {ctrs, casex, discs, selss, exhaust, nchotomy, injects, distincts,
case_thms, case_cong, weak_case_cong, split, split_asm, disc_thmss, discIs, sel_thmss,
- disc_exhausts, sel_exhausts, collapses, expands, sel_splits, sel_split_asms, case_ifs} =
+ disc_exhausts, sel_exhausts, collapses, expands, sel_splits, sel_split_asms, case_eq_ifs} =
{ctrs = map (Morphism.term phi) ctrs,
casex = Morphism.term phi casex,
discs = map (Morphism.term phi) discs,
@@ -121,7 +121,7 @@
expands = map (Morphism.thm phi) expands,
sel_splits = map (Morphism.thm phi) sel_splits,
sel_split_asms = map (Morphism.thm phi) sel_split_asms,
- case_ifs = map (Morphism.thm phi) case_ifs};
+ case_eq_ifs = map (Morphism.thm phi) case_eq_ifs};
val transfer_ctr_sugar =
morph_ctr_sugar o Morphism.thm_morphism o Thm.transfer o Proof_Context.theory_of;
@@ -160,7 +160,7 @@
val caseN = "case";
val case_congN = "case_cong";
-val case_ifN = "case_if";
+val case_eq_ifN = "case_eq_if";
val collapseN = "collapse";
val disc_excludeN = "disc_exclude";
val disc_exhaustN = "disc_exhaust";
@@ -390,7 +390,7 @@
qualify false
(if Binding.is_empty raw_case_binding orelse
Binding.eq_name (raw_case_binding, standard_binding) then
- Binding.suffix_name ("_" ^ caseN) fc_b
+ Binding.prefix_name (caseN ^ "_") fc_b
else
raw_case_binding);
@@ -657,7 +657,7 @@
val (all_sel_thms, sel_thmss, disc_thmss, nontriv_disc_thms, discI_thms, nontriv_discI_thms,
disc_exclude_thms, disc_exhaust_thms, sel_exhaust_thms, all_collapse_thms,
- safe_collapse_thms, expand_thms, sel_split_thms, sel_split_asm_thms, case_if_thms) =
+ safe_collapse_thms, expand_thms, sel_split_thms, sel_split_asm_thms, case_eq_if_thms) =
if no_discs_sels then
([], [], [], [], [], [], [], [], [], [], [], [], [], [], [])
else
@@ -861,12 +861,12 @@
(thm, asm_thm)
end;
- val case_if_thm =
+ val case_eq_if_thm =
let
val goal = mk_Trueprop_eq (ufcase, mk_IfN B udiscs usel_fs);
in
Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} =>
- mk_case_if_tac ctxt n uexhaust_thm case_thms disc_thmss' sel_thmss)
+ mk_case_eq_if_tac ctxt n uexhaust_thm case_thms disc_thmss' sel_thmss)
|> Thm.close_derivation
|> singleton (Proof_Context.export names_lthy lthy)
end;
@@ -874,7 +874,7 @@
(all_sel_thms, sel_thmss, disc_thmss, nontriv_disc_thms, discI_thms,
nontriv_discI_thms, disc_exclude_thms, [disc_exhaust_thm], [sel_exhaust_thm],
all_collapse_thms, safe_collapse_thms, [expand_thm], [sel_split_thm],
- [sel_split_asm_thm], [case_if_thm])
+ [sel_split_asm_thm], [case_eq_if_thm])
end;
val exhaust_case_names_attr = Attrib.internal (K (Rule_Cases.case_names exhaust_cases));
@@ -890,7 +890,7 @@
val notes =
[(caseN, case_thms, code_nitpicksimp_simp_attrs),
(case_congN, [case_cong_thm], []),
- (case_ifN, case_if_thms, []),
+ (case_eq_ifN, case_eq_if_thms, []),
(collapseN, safe_collapse_thms, simp_attrs),
(discN, nontriv_disc_thms, simp_attrs),
(discIN, nontriv_discI_thms, []),
@@ -921,7 +921,7 @@
discIs = discI_thms, sel_thmss = sel_thmss, disc_exhausts = disc_exhaust_thms,
sel_exhausts = sel_exhaust_thms, collapses = all_collapse_thms, expands = expand_thms,
sel_splits = sel_split_thms, sel_split_asms = sel_split_asm_thms,
- case_ifs = case_if_thms};
+ case_eq_ifs = case_eq_if_thms};
in
(ctr_sugar,
lthy
--- a/src/HOL/Tools/ctr_sugar_tactics.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/ctr_sugar_tactics.ML Wed Nov 20 16:43:09 2013 +0100
@@ -18,8 +18,8 @@
val mk_alternate_disc_def_tac: Proof.context -> int -> thm -> thm -> thm -> tactic
val mk_case_tac: Proof.context -> int -> int -> thm -> thm list -> thm list list -> tactic
val mk_case_cong_tac: Proof.context -> thm -> thm list -> tactic
- val mk_case_if_tac: Proof.context -> int -> thm -> thm list -> thm list list -> thm list list ->
- tactic
+ val mk_case_eq_if_tac: Proof.context -> int -> thm -> thm list -> thm list list ->
+ thm list list -> tactic
val mk_collapse_tac: Proof.context -> int -> thm -> thm list -> tactic
val mk_disc_exhaust_tac: int -> thm -> thm list -> tactic
val mk_expand_tac: Proof.context -> int -> int list -> thm -> thm -> thm list ->
@@ -28,8 +28,8 @@
val mk_nchotomy_tac: int -> thm -> tactic
val mk_other_half_disc_exclude_tac: thm -> tactic
val mk_sel_exhaust_tac: int -> thm -> thm list -> tactic
- val mk_split_tac: Proof.context -> thm -> thm list -> thm list list -> thm list list -> thm list
- list list -> tactic
+ val mk_split_tac: Proof.context -> thm -> thm list -> thm list list -> thm list list ->
+ thm list list list -> tactic
val mk_split_asm_tac: Proof.context -> thm -> tactic
val mk_unique_disc_def_tac: int -> thm -> tactic
end;
@@ -143,17 +143,17 @@
else mk_case_distinct_ctrs_tac ctxt distincts)) ks distinctss))
end;
-fun mk_case_if_tac ctxt n uexhaust cases discss' selss =
+fun mk_case_cong_tac ctxt uexhaust cases =
+ HEADGOAL (rtac uexhaust THEN'
+ EVERY' (maps (fn casex => [dtac sym, asm_simp_tac (ss_only [casex] ctxt)]) cases));
+
+fun mk_case_eq_if_tac ctxt n uexhaust cases discss' selss =
HEADGOAL (rtac uexhaust THEN'
EVERY' (map3 (fn casex => fn if_discs => fn sels =>
EVERY' [hyp_subst_tac ctxt, SELECT_GOAL (unfold_thms_tac ctxt (if_discs @ sels)),
rtac casex])
cases (map2 (seq_conds if_P_or_not_P_OF n) (1 upto n) discss') selss));
-fun mk_case_cong_tac ctxt uexhaust cases =
- HEADGOAL (rtac uexhaust THEN'
- EVERY' (maps (fn casex => [dtac sym, asm_simp_tac (ss_only [casex] ctxt)]) cases));
-
fun mk_split_tac ctxt uexhaust cases selss injectss distinctsss =
HEADGOAL (rtac uexhaust) THEN
ALLGOALS (fn k => (hyp_subst_tac ctxt THEN'
@@ -169,4 +169,4 @@
end;
-structure Ctr_Sugar_General_Tactics: CTR_SUGAR_GENERAL_TACTICS = Ctr_Sugar_Tactics;
+structure Ctr_Sugar_General_Tactics : CTR_SUGAR_GENERAL_TACTICS = Ctr_Sugar_Tactics;
--- a/src/HOL/Tools/ctr_sugar_util.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/ctr_sugar_util.ML Wed Nov 20 16:43:09 2013 +0100
@@ -176,9 +176,7 @@
fun rapp u t = betapply (t, u);
fun list_quant_free quant_const =
- fold_rev (fn free => fn P =>
- let val (x, T) = Term.dest_Free free;
- in quant_const T $ Term.absfree (x, T) P end);
+ fold_rev (fn Free (xT as (_, T)) => fn P => quant_const T $ Term.absfree xT P);
val list_all_free = list_quant_free HOLogic.all_const;
val list_exists_free = list_quant_free HOLogic.exists_const;
--- a/src/HOL/Tools/hologic.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/hologic.ML Wed Nov 20 16:43:09 2013 +0100
@@ -104,7 +104,6 @@
val mk_numeral: int -> term
val dest_num: term -> int
val numeral_const: typ -> term
- val neg_numeral_const: typ -> term
val add_numerals: term -> (term * typ) list -> (term * typ) list
val mk_number: typ -> int -> term
val dest_number: term -> typ * int
@@ -548,7 +547,6 @@
| dest_num t = raise TERM ("dest_num", [t]);
fun numeral_const T = Const ("Num.numeral_class.numeral", numT --> T);
-fun neg_numeral_const T = Const ("Num.neg_numeral_class.neg_numeral", numT --> T);
fun add_numerals (Const ("Num.numeral_class.numeral", Type (_, [_, T])) $ t) = cons (t, T)
| add_numerals (t $ u) = add_numerals t #> add_numerals u
@@ -559,14 +557,14 @@
| mk_number T 1 = Const ("Groups.one_class.one", T)
| mk_number T i =
if i > 0 then numeral_const T $ mk_numeral i
- else neg_numeral_const T $ mk_numeral (~ i);
+ else Const ("Groups.uminus_class.uminus", T --> T) $ mk_number T (~ i);
fun dest_number (Const ("Groups.zero_class.zero", T)) = (T, 0)
| dest_number (Const ("Groups.one_class.one", T)) = (T, 1)
| dest_number (Const ("Num.numeral_class.numeral", Type ("fun", [_, T])) $ t) =
(T, dest_num t)
- | dest_number (Const ("Num.neg_numeral_class.neg_numeral", Type ("fun", [_, T])) $ t) =
- (T, ~ (dest_num t))
+ | dest_number (Const ("Groups.uminus_class.uminus", Type ("fun", [_, T])) $ t) =
+ apsnd (op ~) (dest_number t)
| dest_number t = raise TERM ("dest_number", [t]);
--- a/src/HOL/Tools/lin_arith.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/lin_arith.ML Wed Nov 20 16:43:09 2013 +0100
@@ -183,9 +183,6 @@
| demult (t as Const ("Num.numeral_class.numeral", _) $ n, m) =
((NONE, Rat.mult m (Rat.rat_of_int (HOLogic.dest_num n)))
handle TERM _ => (SOME t, m))
- | demult (t as Const ("Num.neg_numeral_class.neg_numeral", _) $ n, m) =
- ((NONE, Rat.mult m (Rat.rat_of_int (~ (HOLogic.dest_num n))))
- handle TERM _ => (SOME t, m))
| demult (t as Const (@{const_name Suc}, _) $ _, m) =
((NONE, Rat.mult m (Rat.rat_of_int (HOLogic.dest_nat t)))
handle TERM _ => (SOME t, m))
@@ -212,6 +209,10 @@
pi
| poly (Const (@{const_name Groups.one}, _), m, (p, i)) =
(p, Rat.add i m)
+ | poly (all as Const ("Num.numeral_class.numeral", Type(_,[_,_])) $ t, m, pi as (p, i)) =
+ (let val k = HOLogic.dest_num t
+ in (p, Rat.add i (Rat.mult m (Rat.rat_of_int k))) end
+ handle TERM _ => add_atom all m pi)
| poly (Const (@{const_name Suc}, _) $ t, m, (p, i)) =
poly (t, m, (p, Rat.add i m))
| poly (all as Const (@{const_name Groups.times}, _) $ _ $ _, m, pi as (p, i)) =
@@ -222,14 +223,6 @@
(case demult inj_consts (all, m) of
(NONE, m') => (p, Rat.add i m')
| (SOME u, m') => add_atom u m' pi)
- | poly (all as Const ("Num.numeral_class.numeral", Type(_,[_,_])) $ t, m, pi as (p, i)) =
- (let val k = HOLogic.dest_num t
- in (p, Rat.add i (Rat.mult m (Rat.rat_of_int k))) end
- handle TERM _ => add_atom all m pi)
- | poly (all as Const ("Num.neg_numeral_class.neg_numeral", Type(_,[_,_])) $ t, m, pi as (p, i)) =
- (let val k = HOLogic.dest_num t
- in (p, Rat.add i (Rat.mult m (Rat.rat_of_int (~ k)))) end
- handle TERM _ => add_atom all m pi)
| poly (all as Const f $ x, m, pi) =
if member (op =) inj_consts f then poly (x, m, pi) else add_atom all m pi
| poly (all, m, pi) =
--- a/src/HOL/Tools/numeral.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/numeral.ML Wed Nov 20 16:43:09 2013 +0100
@@ -45,8 +45,8 @@
val numeral = @{cpat "numeral"};
val numeralT = Thm.ctyp_of @{theory} (Term.range_type (Thm.typ_of (Thm.ctyp_of_term numeral)));
-val neg_numeral = @{cpat "neg_numeral"};
-val neg_numeralT = Thm.ctyp_of @{theory} (Term.range_type (Thm.typ_of (Thm.ctyp_of_term neg_numeral)));
+val uminus = @{cpat "uminus"};
+val uminusT = Thm.ctyp_of @{theory} (Term.range_type (Thm.typ_of (Thm.ctyp_of_term uminus)));
fun instT T V = Thm.instantiate_cterm ([(V, T)], []);
@@ -56,7 +56,7 @@
| mk_cnumber T 1 = instT T oneT one
| mk_cnumber T i =
if i > 0 then Thm.apply (instT T numeralT numeral) (mk_cnumeral i)
- else Thm.apply (instT T neg_numeralT neg_numeral) (mk_cnumeral (~i));
+ else Thm.apply (instT T uminusT uminus) (Thm.apply (instT T numeralT numeral) (mk_cnumeral (~i)));
end;
--- a/src/HOL/Tools/numeral_simprocs.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Tools/numeral_simprocs.ML Wed Nov 20 16:43:09 2013 +0100
@@ -56,9 +56,6 @@
val long_mk_sum = Arith_Data.long_mk_sum;
val dest_sum = Arith_Data.dest_sum;
-val mk_diff = HOLogic.mk_binop @{const_name Groups.minus};
-val dest_diff = HOLogic.dest_bin @{const_name Groups.minus} dummyT;
-
val mk_times = HOLogic.mk_binop @{const_name Groups.times};
fun one_of T = Const(@{const_name Groups.one}, T);
@@ -181,7 +178,7 @@
(*Simplify 0+n, n+0, Numeral1*n, n*Numeral1, 1*x, x*1, x/1 *)
val add_0s = @{thms add_0s};
-val mult_1s = @{thms mult_1s mult_1_left mult_1_right divide_1};
+val mult_1s = @{thms mult_1s divide_numeral_1 mult_1_left mult_1_right mult_minus1 mult_minus1_right divide_1};
(* For post-simplification of the rhs of simproc-generated rules *)
val post_simps =
@@ -194,9 +191,8 @@
val field_post_simps =
post_simps @ [@{thm divide_zero_left}, @{thm divide_1}]
-(*Simplify inverse Numeral1, a/Numeral1*)
+(*Simplify inverse Numeral1*)
val inverse_1s = [@{thm inverse_numeral_1}];
-val divide_1s = [@{thm divide_numeral_1}];
(*To perform binary arithmetic. The "left" rewriting handles patterns
created by the Numeral_Simprocs, such as 3 * (5 * x). *)
@@ -217,7 +213,7 @@
@{thms add_neg_numeral_simps}) simps;
(*To evaluate binary negations of coefficients*)
-val minus_simps = [@{thm minus_zero}, @{thm minus_one}, @{thm minus_numeral}, @{thm minus_neg_numeral}];
+val minus_simps = [@{thm minus_zero}, @{thm minus_minus}];
(*To let us treat subtraction as addition*)
val diff_simps = [@{thm diff_conv_add_uminus}, @{thm minus_add_distrib}, @{thm minus_minus}];
@@ -225,16 +221,13 @@
(*To let us treat division as multiplication*)
val divide_simps = [@{thm divide_inverse}, @{thm inverse_mult_distrib}, @{thm inverse_inverse_eq}];
-(*push the unary minus down*)
-val minus_mult_eq_1_to_2 = @{lemma "- (a::'a::ring) * b = a * - b" by simp};
-
(*to extract again any uncancelled minuses*)
val minus_from_mult_simps =
[@{thm minus_minus}, @{thm mult_minus_left}, @{thm mult_minus_right}];
(*combine unary minus with numeric literals, however nested within a product*)
val mult_minus_simps =
- [@{thm mult_assoc}, @{thm minus_mult_left}, minus_mult_eq_1_to_2];
+ [@{thm mult_assoc}, @{thm minus_mult_right}, @{thm minus_mult_commute}];
val norm_ss1 =
simpset_of (put_simpset num_ss @{context}
@@ -247,7 +240,7 @@
val norm_ss3 =
simpset_of (put_simpset num_ss @{context}
- addsimps minus_from_mult_simps @ @{thms add_ac} @ @{thms mult_ac})
+ addsimps minus_from_mult_simps @ @{thms add_ac} @ @{thms mult_ac minus_mult_commute})
structure CancelNumeralsCommon =
struct
@@ -330,7 +323,7 @@
structure FieldCombineNumeralsData =
struct
type coeff = int * int
- val iszero = (fn (p, q) => p = 0)
+ val iszero = (fn (p, _) => p = 0)
val add = add_frac
val mk_sum = long_mk_sum
val dest_sum = dest_sum
@@ -368,7 +361,7 @@
structure Semiring_Times_Assoc_Data : ASSOC_FOLD_DATA =
struct
- val assoc_ss = simpset_of (put_simpset HOL_basic_ss @{context} addsimps @{thms mult_ac})
+ val assoc_ss = simpset_of (put_simpset HOL_basic_ss @{context} addsimps @{thms mult_ac minus_mult_commute})
val eq_reflection = eq_reflection
val is_numeral = can HOLogic.dest_number
end;
@@ -388,7 +381,7 @@
val norm_ss2 =
simpset_of (put_simpset HOL_basic_ss @{context} addsimps simps @ mult_minus_simps)
val norm_ss3 =
- simpset_of (put_simpset HOL_basic_ss @{context} addsimps @{thms mult_ac})
+ simpset_of (put_simpset HOL_basic_ss @{context} addsimps @{thms mult_ac minus_mult_commute})
fun norm_tac ctxt =
ALLGOALS (simp_tac (put_simpset norm_ss1 ctxt))
THEN ALLGOALS (simp_tac (put_simpset norm_ss2 ctxt))
@@ -463,9 +456,9 @@
["((l::'a::field_inverse_zero) * m) / n",
"(l::'a::field_inverse_zero) / (m * n)",
"((numeral v)::'a::field_inverse_zero) / (numeral w)",
- "((numeral v)::'a::field_inverse_zero) / (neg_numeral w)",
- "((neg_numeral v)::'a::field_inverse_zero) / (numeral w)",
- "((neg_numeral v)::'a::field_inverse_zero) / (neg_numeral w)"],
+ "((numeral v)::'a::field_inverse_zero) / (- numeral w)",
+ "((- numeral v)::'a::field_inverse_zero) / (numeral w)",
+ "((- numeral v)::'a::field_inverse_zero) / (- numeral w)"],
DivideCancelNumeralFactor.proc)]
@@ -516,7 +509,7 @@
val find_first = find_first_t []
val trans_tac = trans_tac
val norm_ss =
- simpset_of (put_simpset HOL_basic_ss @{context} addsimps mult_1s @ @{thms mult_ac})
+ simpset_of (put_simpset HOL_basic_ss @{context} addsimps mult_1s @ @{thms mult_ac minus_mult_commute})
fun norm_tac ctxt =
ALLGOALS (simp_tac (put_simpset norm_ss ctxt))
val simplify_meta_eq = cancel_simplify_meta_eq
--- a/src/HOL/Transcendental.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Transcendental.thy Wed Nov 20 16:43:09 2013 +0100
@@ -2000,8 +2000,8 @@
apply (subst powr_add, simp, simp)
done
-lemma powr_realpow_numeral: "0 < x \<Longrightarrow> x powr (numeral n :: real) = x^(numeral n)"
- unfolding real_of_nat_numeral[symmetric] by (rule powr_realpow)
+lemma powr_realpow_numeral: "0 < x \<Longrightarrow> x powr (numeral n :: real) = x ^ (numeral n)"
+ unfolding real_of_nat_numeral [symmetric] by (rule powr_realpow)
lemma powr_realpow2: "0 <= x ==> 0 < n ==> x^n = (if (x = 0) then 0 else x powr (real n))"
apply (case_tac "x = 0", simp, simp)
@@ -2020,11 +2020,17 @@
then show ?thesis by (simp add: assms powr_realpow[symmetric])
qed
-lemma powr_numeral: "0 < x \<Longrightarrow> x powr numeral n = x^numeral n"
- using powr_realpow[of x "numeral n"] by simp
-
-lemma powr_neg_numeral: "0 < x \<Longrightarrow> x powr neg_numeral n = 1 / x^numeral n"
- using powr_int[of x "neg_numeral n"] by simp
+lemma powr_one: "0 < x \<Longrightarrow> x powr 1 = x"
+ using powr_realpow [of x 1] by simp
+
+lemma powr_numeral: "0 < x \<Longrightarrow> x powr numeral n = x ^ numeral n"
+ by (fact powr_realpow_numeral)
+
+lemma powr_neg_one: "0 < x \<Longrightarrow> x powr - 1 = 1 / x"
+ using powr_int [of x "- 1"] by simp
+
+lemma powr_neg_numeral: "0 < x \<Longrightarrow> x powr - numeral n = 1 / x ^ numeral n"
+ using powr_int [of x "- numeral n"] by simp
lemma root_powr_inverse: "0 < n \<Longrightarrow> 0 < x \<Longrightarrow> root n x = x powr (1/n)"
by (rule real_root_pos_unique) (auto simp: powr_realpow[symmetric] powr_powr)
@@ -4047,7 +4053,7 @@
show "sgn x * pi / 2 - arctan x < pi / 2"
using arctan_bounded [of "- x"] assms
unfolding sgn_real_def arctan_minus
- by auto
+ by (auto simp add: algebra_simps)
show "tan (sgn x * pi / 2 - arctan x) = 1 / x"
unfolding tan_inverse [of "arctan x", unfolded tan_arctan]
unfolding sgn_real_def
--- a/src/HOL/Word/Bit_Int.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Word/Bit_Int.thy Wed Nov 20 16:43:09 2013 +0100
@@ -52,10 +52,10 @@
lemma int_not_simps [simp]:
"NOT (0::int) = -1"
"NOT (1::int) = -2"
- "NOT (-1::int) = 0"
- "NOT (numeral w::int) = neg_numeral (w + Num.One)"
- "NOT (neg_numeral (Num.Bit0 w)::int) = numeral (Num.BitM w)"
- "NOT (neg_numeral (Num.Bit1 w)::int) = numeral (Num.Bit0 w)"
+ "NOT (- 1::int) = 0"
+ "NOT (numeral w::int) = - numeral (w + Num.One)"
+ "NOT (- numeral (Num.Bit0 w)::int) = numeral (Num.BitM w)"
+ "NOT (- numeral (Num.Bit1 w)::int) = numeral (Num.Bit0 w)"
unfolding int_not_def by simp_all
lemma int_not_not [simp]: "NOT (NOT (x::int)) = x"
@@ -228,11 +228,11 @@
by (metis bin_rl_simp)
lemma bin_rest_neg_numeral_BitM [simp]:
- "bin_rest (neg_numeral (Num.BitM w)) = neg_numeral w"
+ "bin_rest (- numeral (Num.BitM w)) = - numeral w"
by (simp only: BIT_bin_simps [symmetric] bin_rest_BIT)
lemma bin_last_neg_numeral_BitM [simp]:
- "bin_last (neg_numeral (Num.BitM w)) = 1"
+ "bin_last (- numeral (Num.BitM w)) = 1"
by (simp only: BIT_bin_simps [symmetric] bin_last_BIT)
text {* FIXME: The rule sets below are very large (24 rules for each
@@ -243,26 +243,26 @@
"numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (numeral x AND numeral y) BIT 0"
"numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (numeral x AND numeral y) BIT 0"
"numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = (numeral x AND numeral y) BIT 1"
- "numeral (Num.Bit0 x) AND neg_numeral (Num.Bit0 y) = (numeral x AND neg_numeral y) BIT 0"
- "numeral (Num.Bit0 x) AND neg_numeral (Num.Bit1 y) = (numeral x AND neg_numeral (y + Num.One)) BIT 0"
- "numeral (Num.Bit1 x) AND neg_numeral (Num.Bit0 y) = (numeral x AND neg_numeral y) BIT 0"
- "numeral (Num.Bit1 x) AND neg_numeral (Num.Bit1 y) = (numeral x AND neg_numeral (y + Num.One)) BIT 1"
- "neg_numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (neg_numeral x AND numeral y) BIT 0"
- "neg_numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (neg_numeral x AND numeral y) BIT 0"
- "neg_numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (neg_numeral (x + Num.One) AND numeral y) BIT 0"
- "neg_numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = (neg_numeral (x + Num.One) AND numeral y) BIT 1"
- "neg_numeral (Num.Bit0 x) AND neg_numeral (Num.Bit0 y) = (neg_numeral x AND neg_numeral y) BIT 0"
- "neg_numeral (Num.Bit0 x) AND neg_numeral (Num.Bit1 y) = (neg_numeral x AND neg_numeral (y + Num.One)) BIT 0"
- "neg_numeral (Num.Bit1 x) AND neg_numeral (Num.Bit0 y) = (neg_numeral (x + Num.One) AND neg_numeral y) BIT 0"
- "neg_numeral (Num.Bit1 x) AND neg_numeral (Num.Bit1 y) = (neg_numeral (x + Num.One) AND neg_numeral (y + Num.One)) BIT 1"
+ "numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (numeral x AND - numeral y) BIT 0"
+ "numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (numeral x AND - numeral (y + Num.One)) BIT 0"
+ "numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (numeral x AND - numeral y) BIT 0"
+ "numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = (numeral x AND - numeral (y + Num.One)) BIT 1"
+ "- numeral (Num.Bit0 x) AND numeral (Num.Bit0 y) = (- numeral x AND numeral y) BIT 0"
+ "- numeral (Num.Bit0 x) AND numeral (Num.Bit1 y) = (- numeral x AND numeral y) BIT 0"
+ "- numeral (Num.Bit1 x) AND numeral (Num.Bit0 y) = (- numeral (x + Num.One) AND numeral y) BIT 0"
+ "- numeral (Num.Bit1 x) AND numeral (Num.Bit1 y) = (- numeral (x + Num.One) AND numeral y) BIT 1"
+ "- numeral (Num.Bit0 x) AND - numeral (Num.Bit0 y) = (- numeral x AND - numeral y) BIT 0"
+ "- numeral (Num.Bit0 x) AND - numeral (Num.Bit1 y) = (- numeral x AND - numeral (y + Num.One)) BIT 0"
+ "- numeral (Num.Bit1 x) AND - numeral (Num.Bit0 y) = (- numeral (x + Num.One) AND - numeral y) BIT 0"
+ "- numeral (Num.Bit1 x) AND - numeral (Num.Bit1 y) = (- numeral (x + Num.One) AND - numeral (y + Num.One)) BIT 1"
"(1::int) AND numeral (Num.Bit0 y) = 0"
"(1::int) AND numeral (Num.Bit1 y) = 1"
- "(1::int) AND neg_numeral (Num.Bit0 y) = 0"
- "(1::int) AND neg_numeral (Num.Bit1 y) = 1"
+ "(1::int) AND - numeral (Num.Bit0 y) = 0"
+ "(1::int) AND - numeral (Num.Bit1 y) = 1"
"numeral (Num.Bit0 x) AND (1::int) = 0"
"numeral (Num.Bit1 x) AND (1::int) = 1"
- "neg_numeral (Num.Bit0 x) AND (1::int) = 0"
- "neg_numeral (Num.Bit1 x) AND (1::int) = 1"
+ "- numeral (Num.Bit0 x) AND (1::int) = 0"
+ "- numeral (Num.Bit1 x) AND (1::int) = 1"
by (rule bin_rl_eqI, simp, simp)+
lemma int_or_numerals [simp]:
@@ -270,26 +270,26 @@
"numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = (numeral x OR numeral y) BIT 1"
"numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = (numeral x OR numeral y) BIT 1"
"numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = (numeral x OR numeral y) BIT 1"
- "numeral (Num.Bit0 x) OR neg_numeral (Num.Bit0 y) = (numeral x OR neg_numeral y) BIT 0"
- "numeral (Num.Bit0 x) OR neg_numeral (Num.Bit1 y) = (numeral x OR neg_numeral (y + Num.One)) BIT 1"
- "numeral (Num.Bit1 x) OR neg_numeral (Num.Bit0 y) = (numeral x OR neg_numeral y) BIT 1"
- "numeral (Num.Bit1 x) OR neg_numeral (Num.Bit1 y) = (numeral x OR neg_numeral (y + Num.One)) BIT 1"
- "neg_numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (neg_numeral x OR numeral y) BIT 0"
- "neg_numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = (neg_numeral x OR numeral y) BIT 1"
- "neg_numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = (neg_numeral (x + Num.One) OR numeral y) BIT 1"
- "neg_numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = (neg_numeral (x + Num.One) OR numeral y) BIT 1"
- "neg_numeral (Num.Bit0 x) OR neg_numeral (Num.Bit0 y) = (neg_numeral x OR neg_numeral y) BIT 0"
- "neg_numeral (Num.Bit0 x) OR neg_numeral (Num.Bit1 y) = (neg_numeral x OR neg_numeral (y + Num.One)) BIT 1"
- "neg_numeral (Num.Bit1 x) OR neg_numeral (Num.Bit0 y) = (neg_numeral (x + Num.One) OR neg_numeral y) BIT 1"
- "neg_numeral (Num.Bit1 x) OR neg_numeral (Num.Bit1 y) = (neg_numeral (x + Num.One) OR neg_numeral (y + Num.One)) BIT 1"
+ "numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (numeral x OR - numeral y) BIT 0"
+ "numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = (numeral x OR - numeral (y + Num.One)) BIT 1"
+ "numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = (numeral x OR - numeral y) BIT 1"
+ "numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = (numeral x OR - numeral (y + Num.One)) BIT 1"
+ "- numeral (Num.Bit0 x) OR numeral (Num.Bit0 y) = (- numeral x OR numeral y) BIT 0"
+ "- numeral (Num.Bit0 x) OR numeral (Num.Bit1 y) = (- numeral x OR numeral y) BIT 1"
+ "- numeral (Num.Bit1 x) OR numeral (Num.Bit0 y) = (- numeral (x + Num.One) OR numeral y) BIT 1"
+ "- numeral (Num.Bit1 x) OR numeral (Num.Bit1 y) = (- numeral (x + Num.One) OR numeral y) BIT 1"
+ "- numeral (Num.Bit0 x) OR - numeral (Num.Bit0 y) = (- numeral x OR - numeral y) BIT 0"
+ "- numeral (Num.Bit0 x) OR - numeral (Num.Bit1 y) = (- numeral x OR - numeral (y + Num.One)) BIT 1"
+ "- numeral (Num.Bit1 x) OR - numeral (Num.Bit0 y) = (- numeral (x + Num.One) OR - numeral y) BIT 1"
+ "- numeral (Num.Bit1 x) OR - numeral (Num.Bit1 y) = (- numeral (x + Num.One) OR - numeral (y + Num.One)) BIT 1"
"(1::int) OR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)"
"(1::int) OR numeral (Num.Bit1 y) = numeral (Num.Bit1 y)"
- "(1::int) OR neg_numeral (Num.Bit0 y) = neg_numeral (Num.BitM y)"
- "(1::int) OR neg_numeral (Num.Bit1 y) = neg_numeral (Num.Bit1 y)"
+ "(1::int) OR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)"
+ "(1::int) OR - numeral (Num.Bit1 y) = - numeral (Num.Bit1 y)"
"numeral (Num.Bit0 x) OR (1::int) = numeral (Num.Bit1 x)"
"numeral (Num.Bit1 x) OR (1::int) = numeral (Num.Bit1 x)"
- "neg_numeral (Num.Bit0 x) OR (1::int) = neg_numeral (Num.BitM x)"
- "neg_numeral (Num.Bit1 x) OR (1::int) = neg_numeral (Num.Bit1 x)"
+ "- numeral (Num.Bit0 x) OR (1::int) = - numeral (Num.BitM x)"
+ "- numeral (Num.Bit1 x) OR (1::int) = - numeral (Num.Bit1 x)"
by (rule bin_rl_eqI, simp, simp)+
lemma int_xor_numerals [simp]:
@@ -297,26 +297,26 @@
"numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = (numeral x XOR numeral y) BIT 1"
"numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = (numeral x XOR numeral y) BIT 1"
"numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (numeral x XOR numeral y) BIT 0"
- "numeral (Num.Bit0 x) XOR neg_numeral (Num.Bit0 y) = (numeral x XOR neg_numeral y) BIT 0"
- "numeral (Num.Bit0 x) XOR neg_numeral (Num.Bit1 y) = (numeral x XOR neg_numeral (y + Num.One)) BIT 1"
- "numeral (Num.Bit1 x) XOR neg_numeral (Num.Bit0 y) = (numeral x XOR neg_numeral y) BIT 1"
- "numeral (Num.Bit1 x) XOR neg_numeral (Num.Bit1 y) = (numeral x XOR neg_numeral (y + Num.One)) BIT 0"
- "neg_numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (neg_numeral x XOR numeral y) BIT 0"
- "neg_numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = (neg_numeral x XOR numeral y) BIT 1"
- "neg_numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = (neg_numeral (x + Num.One) XOR numeral y) BIT 1"
- "neg_numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (neg_numeral (x + Num.One) XOR numeral y) BIT 0"
- "neg_numeral (Num.Bit0 x) XOR neg_numeral (Num.Bit0 y) = (neg_numeral x XOR neg_numeral y) BIT 0"
- "neg_numeral (Num.Bit0 x) XOR neg_numeral (Num.Bit1 y) = (neg_numeral x XOR neg_numeral (y + Num.One)) BIT 1"
- "neg_numeral (Num.Bit1 x) XOR neg_numeral (Num.Bit0 y) = (neg_numeral (x + Num.One) XOR neg_numeral y) BIT 1"
- "neg_numeral (Num.Bit1 x) XOR neg_numeral (Num.Bit1 y) = (neg_numeral (x + Num.One) XOR neg_numeral (y + Num.One)) BIT 0"
+ "numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (numeral x XOR - numeral y) BIT 0"
+ "numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = (numeral x XOR - numeral (y + Num.One)) BIT 1"
+ "numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = (numeral x XOR - numeral y) BIT 1"
+ "numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (numeral x XOR - numeral (y + Num.One)) BIT 0"
+ "- numeral (Num.Bit0 x) XOR numeral (Num.Bit0 y) = (- numeral x XOR numeral y) BIT 0"
+ "- numeral (Num.Bit0 x) XOR numeral (Num.Bit1 y) = (- numeral x XOR numeral y) BIT 1"
+ "- numeral (Num.Bit1 x) XOR numeral (Num.Bit0 y) = (- numeral (x + Num.One) XOR numeral y) BIT 1"
+ "- numeral (Num.Bit1 x) XOR numeral (Num.Bit1 y) = (- numeral (x + Num.One) XOR numeral y) BIT 0"
+ "- numeral (Num.Bit0 x) XOR - numeral (Num.Bit0 y) = (- numeral x XOR - numeral y) BIT 0"
+ "- numeral (Num.Bit0 x) XOR - numeral (Num.Bit1 y) = (- numeral x XOR - numeral (y + Num.One)) BIT 1"
+ "- numeral (Num.Bit1 x) XOR - numeral (Num.Bit0 y) = (- numeral (x + Num.One) XOR - numeral y) BIT 1"
+ "- numeral (Num.Bit1 x) XOR - numeral (Num.Bit1 y) = (- numeral (x + Num.One) XOR - numeral (y + Num.One)) BIT 0"
"(1::int) XOR numeral (Num.Bit0 y) = numeral (Num.Bit1 y)"
"(1::int) XOR numeral (Num.Bit1 y) = numeral (Num.Bit0 y)"
- "(1::int) XOR neg_numeral (Num.Bit0 y) = neg_numeral (Num.BitM y)"
- "(1::int) XOR neg_numeral (Num.Bit1 y) = neg_numeral (Num.Bit0 (y + Num.One))"
+ "(1::int) XOR - numeral (Num.Bit0 y) = - numeral (Num.BitM y)"
+ "(1::int) XOR - numeral (Num.Bit1 y) = - numeral (Num.Bit0 (y + Num.One))"
"numeral (Num.Bit0 x) XOR (1::int) = numeral (Num.Bit1 x)"
"numeral (Num.Bit1 x) XOR (1::int) = numeral (Num.Bit0 x)"
- "neg_numeral (Num.Bit0 x) XOR (1::int) = neg_numeral (Num.BitM x)"
- "neg_numeral (Num.Bit1 x) XOR (1::int) = neg_numeral (Num.Bit0 (x + Num.One))"
+ "- numeral (Num.Bit0 x) XOR (1::int) = - numeral (Num.BitM x)"
+ "- numeral (Num.Bit1 x) XOR (1::int) = - numeral (Num.Bit0 (x + Num.One))"
by (rule bin_rl_eqI, simp, simp)+
subsubsection {* Interactions with arithmetic *}
--- a/src/HOL/Word/Bit_Representation.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Word/Bit_Representation.thy Wed Nov 20 16:43:09 2013 +0100
@@ -61,21 +61,23 @@
lemma BIT_bin_simps [simp]:
"numeral k BIT 0 = numeral (Num.Bit0 k)"
"numeral k BIT 1 = numeral (Num.Bit1 k)"
- "neg_numeral k BIT 0 = neg_numeral (Num.Bit0 k)"
- "neg_numeral k BIT 1 = neg_numeral (Num.BitM k)"
- unfolding neg_numeral_def numeral.simps numeral_BitM
+ "(- numeral k) BIT 0 = - numeral (Num.Bit0 k)"
+ "(- numeral k) BIT 1 = - numeral (Num.BitM k)"
+ unfolding numeral.simps numeral_BitM
unfolding Bit_def bitval_simps
by (simp_all del: arith_simps add_numeral_special diff_numeral_special)
lemma BIT_special_simps [simp]:
- shows "0 BIT 0 = 0" and "0 BIT 1 = 1" and "1 BIT 0 = 2" and "1 BIT 1 = 3"
+ shows "0 BIT 0 = 0" and "0 BIT 1 = 1"
+ and "1 BIT 0 = 2" and "1 BIT 1 = 3"
+ and "(- 1) BIT 0 = - 2" and "(- 1) BIT 1 = - 1"
unfolding Bit_def by simp_all
lemma Bit_eq_0_iff: "w BIT b = 0 \<longleftrightarrow> w = 0 \<and> b = 0"
by (subst BIT_eq_iff [symmetric], simp)
-lemma Bit_eq_m1_iff: "w BIT b = -1 \<longleftrightarrow> w = -1 \<and> b = 1"
- by (subst BIT_eq_iff [symmetric], simp)
+lemma Bit_eq_m1_iff: "w BIT b = - 1 \<longleftrightarrow> w = - 1 \<and> b = 1"
+ by (cases b) (auto simp add: Bit_def, arith)
lemma BitM_inc: "Num.BitM (Num.inc w) = Num.Bit1 w"
by (induct w, simp_all)
@@ -83,8 +85,8 @@
lemma expand_BIT:
"numeral (Num.Bit0 w) = numeral w BIT 0"
"numeral (Num.Bit1 w) = numeral w BIT 1"
- "neg_numeral (Num.Bit0 w) = neg_numeral w BIT 0"
- "neg_numeral (Num.Bit1 w) = neg_numeral (w + Num.One) BIT 1"
+ "- numeral (Num.Bit0 w) = - numeral w BIT 0"
+ "- numeral (Num.Bit1 w) = (- numeral (w + Num.One)) BIT 1"
unfolding add_One by (simp_all add: BitM_inc)
lemma bin_last_numeral_simps [simp]:
@@ -94,9 +96,9 @@
"bin_last Numeral1 = 1"
"bin_last (numeral (Num.Bit0 w)) = 0"
"bin_last (numeral (Num.Bit1 w)) = 1"
- "bin_last (neg_numeral (Num.Bit0 w)) = 0"
- "bin_last (neg_numeral (Num.Bit1 w)) = 1"
- unfolding expand_BIT bin_last_BIT by (simp_all add: bin_last_def)
+ "bin_last (- numeral (Num.Bit0 w)) = 0"
+ "bin_last (- numeral (Num.Bit1 w)) = 1"
+ unfolding expand_BIT bin_last_BIT by (simp_all add: bin_last_def zmod_zminus1_eq_if)
lemma bin_rest_numeral_simps [simp]:
"bin_rest 0 = 0"
@@ -105,9 +107,9 @@
"bin_rest Numeral1 = 0"
"bin_rest (numeral (Num.Bit0 w)) = numeral w"
"bin_rest (numeral (Num.Bit1 w)) = numeral w"
- "bin_rest (neg_numeral (Num.Bit0 w)) = neg_numeral w"
- "bin_rest (neg_numeral (Num.Bit1 w)) = neg_numeral (w + Num.One)"
- unfolding expand_BIT bin_rest_BIT by (simp_all add: bin_rest_def)
+ "bin_rest (- numeral (Num.Bit0 w)) = - numeral w"
+ "bin_rest (- numeral (Num.Bit1 w)) = - numeral (w + Num.One)"
+ unfolding expand_BIT bin_rest_BIT by (simp_all add: bin_rest_def zdiv_zminus1_eq_if)
lemma less_Bits:
"(v BIT b < w BIT c) = (v < w | v <= w & b = (0::bit) & c = (1::bit))"
@@ -197,42 +199,45 @@
lemma Bit_div2 [simp]: "(w BIT b) div 2 = w"
unfolding bin_rest_def [symmetric] by (rule bin_rest_BIT)
-lemma bin_nth_lem [rule_format]:
- "ALL y. bin_nth x = bin_nth y --> x = y"
- apply (induct x rule: bin_induct)
- apply safe
- apply (erule rev_mp)
- apply (induct_tac y rule: bin_induct)
+lemma bin_nth_eq_iff:
+ "bin_nth x = bin_nth y \<longleftrightarrow> x = y"
+proof -
+ have bin_nth_lem [rule_format]: "ALL y. bin_nth x = bin_nth y --> x = y"
+ apply (induct x rule: bin_induct)
apply safe
- apply (drule_tac x=0 in fun_cong, force)
- apply (erule notE, rule ext,
+ apply (erule rev_mp)
+ apply (induct_tac y rule: bin_induct)
+ apply safe
+ apply (drule_tac x=0 in fun_cong, force)
+ apply (erule notE, rule ext,
drule_tac x="Suc x" in fun_cong, force)
- apply (drule_tac x=0 in fun_cong, force)
- apply (erule rev_mp)
- apply (induct_tac y rule: bin_induct)
- apply safe
+ apply (drule_tac x=0 in fun_cong, force)
+ apply (erule rev_mp)
+ apply (induct_tac y rule: bin_induct)
+ apply safe
+ apply (drule_tac x=0 in fun_cong, force)
+ apply (erule notE, rule ext,
+ drule_tac x="Suc x" in fun_cong, force)
+ apply (metis Bit_eq_m1_iff Z bin_last_BIT)
+ apply (case_tac y rule: bin_exhaust)
+ apply clarify
+ apply (erule allE)
+ apply (erule impE)
+ prefer 2
+ apply (erule conjI)
apply (drule_tac x=0 in fun_cong, force)
- apply (erule notE, rule ext,
- drule_tac x="Suc x" in fun_cong, force)
- apply (drule_tac x=0 in fun_cong, force)
- apply (case_tac y rule: bin_exhaust)
- apply clarify
- apply (erule allE)
- apply (erule impE)
- prefer 2
- apply (erule conjI)
- apply (drule_tac x=0 in fun_cong, force)
- apply (rule ext)
- apply (drule_tac x="Suc ?x" in fun_cong, force)
- done
-
-lemma bin_nth_eq_iff: "(bin_nth x = bin_nth y) = (x = y)"
+ apply (rule ext)
+ apply (drule_tac x="Suc ?x" in fun_cong, force)
+ done
+ show ?thesis
by (auto elim: bin_nth_lem)
+qed
lemmas bin_eqI = ext [THEN bin_nth_eq_iff [THEN iffD1]]
-lemma bin_eq_iff: "x = y \<longleftrightarrow> (\<forall>n. bin_nth x n = bin_nth y n)"
- by (auto intro!: bin_nth_lem del: equalityI)
+lemma bin_eq_iff:
+ "x = y \<longleftrightarrow> (\<forall>n. bin_nth x n = bin_nth y n)"
+ using bin_nth_eq_iff by auto
lemma bin_nth_zero [simp]: "\<not> bin_nth 0 n"
by (induct n) auto
@@ -276,8 +281,9 @@
lemma bin_sign_simps [simp]:
"bin_sign 0 = 0"
"bin_sign 1 = 0"
+ "bin_sign (- 1) = - 1"
"bin_sign (numeral k) = 0"
- "bin_sign (neg_numeral k) = -1"
+ "bin_sign (- numeral k) = -1"
"bin_sign (w BIT b) = bin_sign w"
unfolding bin_sign_def Bit_def bitval_def
by (simp_all split: bit.split)
@@ -331,18 +337,18 @@
"bintrunc (Suc n) -1 = bintrunc n -1 BIT 1"
"bintrunc (Suc n) (numeral (Num.Bit0 w)) = bintrunc n (numeral w) BIT 0"
"bintrunc (Suc n) (numeral (Num.Bit1 w)) = bintrunc n (numeral w) BIT 1"
- "bintrunc (Suc n) (neg_numeral (Num.Bit0 w)) =
- bintrunc n (neg_numeral w) BIT 0"
- "bintrunc (Suc n) (neg_numeral (Num.Bit1 w)) =
- bintrunc n (neg_numeral (w + Num.One)) BIT 1"
+ "bintrunc (Suc n) (- numeral (Num.Bit0 w)) =
+ bintrunc n (- numeral w) BIT 0"
+ "bintrunc (Suc n) (- numeral (Num.Bit1 w)) =
+ bintrunc n (- numeral (w + Num.One)) BIT 1"
by simp_all
lemma sbintrunc_0_numeral [simp]:
"sbintrunc 0 1 = -1"
"sbintrunc 0 (numeral (Num.Bit0 w)) = 0"
"sbintrunc 0 (numeral (Num.Bit1 w)) = -1"
- "sbintrunc 0 (neg_numeral (Num.Bit0 w)) = 0"
- "sbintrunc 0 (neg_numeral (Num.Bit1 w)) = -1"
+ "sbintrunc 0 (- numeral (Num.Bit0 w)) = 0"
+ "sbintrunc 0 (- numeral (Num.Bit1 w)) = -1"
by simp_all
lemma sbintrunc_Suc_numeral:
@@ -351,10 +357,10 @@
sbintrunc n (numeral w) BIT 0"
"sbintrunc (Suc n) (numeral (Num.Bit1 w)) =
sbintrunc n (numeral w) BIT 1"
- "sbintrunc (Suc n) (neg_numeral (Num.Bit0 w)) =
- sbintrunc n (neg_numeral w) BIT 0"
- "sbintrunc (Suc n) (neg_numeral (Num.Bit1 w)) =
- sbintrunc n (neg_numeral (w + Num.One)) BIT 1"
+ "sbintrunc (Suc n) (- numeral (Num.Bit0 w)) =
+ sbintrunc n (- numeral w) BIT 0"
+ "sbintrunc (Suc n) (- numeral (Num.Bit1 w)) =
+ sbintrunc n (- numeral (w + Num.One)) BIT 1"
by simp_all
lemma bit_bool:
@@ -580,10 +586,10 @@
bintrunc (pred_numeral k) (numeral w) BIT 0"
"bintrunc (numeral k) (numeral (Num.Bit1 w)) =
bintrunc (pred_numeral k) (numeral w) BIT 1"
- "bintrunc (numeral k) (neg_numeral (Num.Bit0 w)) =
- bintrunc (pred_numeral k) (neg_numeral w) BIT 0"
- "bintrunc (numeral k) (neg_numeral (Num.Bit1 w)) =
- bintrunc (pred_numeral k) (neg_numeral (w + Num.One)) BIT 1"
+ "bintrunc (numeral k) (- numeral (Num.Bit0 w)) =
+ bintrunc (pred_numeral k) (- numeral w) BIT 0"
+ "bintrunc (numeral k) (- numeral (Num.Bit1 w)) =
+ bintrunc (pred_numeral k) (- numeral (w + Num.One)) BIT 1"
"bintrunc (numeral k) 1 = 1"
by (simp_all add: bintrunc_numeral)
@@ -592,10 +598,10 @@
sbintrunc (pred_numeral k) (numeral w) BIT 0"
"sbintrunc (numeral k) (numeral (Num.Bit1 w)) =
sbintrunc (pred_numeral k) (numeral w) BIT 1"
- "sbintrunc (numeral k) (neg_numeral (Num.Bit0 w)) =
- sbintrunc (pred_numeral k) (neg_numeral w) BIT 0"
- "sbintrunc (numeral k) (neg_numeral (Num.Bit1 w)) =
- sbintrunc (pred_numeral k) (neg_numeral (w + Num.One)) BIT 1"
+ "sbintrunc (numeral k) (- numeral (Num.Bit0 w)) =
+ sbintrunc (pred_numeral k) (- numeral w) BIT 0"
+ "sbintrunc (numeral k) (- numeral (Num.Bit1 w)) =
+ sbintrunc (pred_numeral k) (- numeral (w + Num.One)) BIT 1"
"sbintrunc (numeral k) 1 = 1"
by (simp_all add: sbintrunc_numeral)
--- a/src/HOL/Word/Word.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Word/Word.thy Wed Nov 20 16:43:09 2013 +0100
@@ -591,24 +591,24 @@
declare word_numeral_alt [symmetric, code_abbrev]
lemma word_neg_numeral_alt:
- "neg_numeral b = word_of_int (neg_numeral b)"
- by (simp only: neg_numeral_def word_numeral_alt wi_hom_neg)
+ "- numeral b = word_of_int (- numeral b)"
+ by (simp only: word_numeral_alt wi_hom_neg)
declare word_neg_numeral_alt [symmetric, code_abbrev]
lemma word_numeral_transfer [transfer_rule]:
"(fun_rel op = pcr_word) numeral numeral"
- "(fun_rel op = pcr_word) neg_numeral neg_numeral"
- unfolding fun_rel_def word.pcr_cr_eq cr_word_def word_numeral_alt word_neg_numeral_alt
- by simp_all
+ "(fun_rel op = pcr_word) (- numeral) (- numeral)"
+ apply (simp_all add: fun_rel_def word.pcr_cr_eq cr_word_def)
+ using word_numeral_alt [symmetric] word_neg_numeral_alt [symmetric] by blast+
lemma uint_bintrunc [simp]:
"uint (numeral bin :: 'a word) =
bintrunc (len_of TYPE ('a :: len0)) (numeral bin)"
unfolding word_numeral_alt by (rule word_ubin.eq_norm)
-lemma uint_bintrunc_neg [simp]: "uint (neg_numeral bin :: 'a word) =
- bintrunc (len_of TYPE ('a :: len0)) (neg_numeral bin)"
+lemma uint_bintrunc_neg [simp]: "uint (- numeral bin :: 'a word) =
+ bintrunc (len_of TYPE ('a :: len0)) (- numeral bin)"
by (simp only: word_neg_numeral_alt word_ubin.eq_norm)
lemma sint_sbintrunc [simp]:
@@ -616,8 +616,8 @@
sbintrunc (len_of TYPE ('a :: len) - 1) (numeral bin)"
by (simp only: word_numeral_alt word_sbin.eq_norm)
-lemma sint_sbintrunc_neg [simp]: "sint (neg_numeral bin :: 'a word) =
- sbintrunc (len_of TYPE ('a :: len) - 1) (neg_numeral bin)"
+lemma sint_sbintrunc_neg [simp]: "sint (- numeral bin :: 'a word) =
+ sbintrunc (len_of TYPE ('a :: len) - 1) (- numeral bin)"
by (simp only: word_neg_numeral_alt word_sbin.eq_norm)
lemma unat_bintrunc [simp]:
@@ -626,8 +626,8 @@
by (simp only: unat_def uint_bintrunc)
lemma unat_bintrunc_neg [simp]:
- "unat (neg_numeral bin :: 'a :: len0 word) =
- nat (bintrunc (len_of TYPE('a)) (neg_numeral bin))"
+ "unat (- numeral bin :: 'a :: len0 word) =
+ nat (bintrunc (len_of TYPE('a)) (- numeral bin))"
by (simp only: unat_def uint_bintrunc_neg)
lemma size_0_eq: "size (w :: 'a :: len0 word) = 0 \<Longrightarrow> v = w"
@@ -678,7 +678,7 @@
by (simp only: int_word_uint)
lemma uint_neg_numeral:
- "uint (neg_numeral b :: 'a :: len0 word) = neg_numeral b mod 2 ^ len_of TYPE('a)"
+ "uint (- numeral b :: 'a :: len0 word) = - numeral b mod 2 ^ len_of TYPE('a)"
unfolding word_neg_numeral_alt
by (simp only: int_word_uint)
@@ -702,13 +702,16 @@
lemma word_of_int_1 [simp, code_post]: "word_of_int 1 = 1"
unfolding word_1_wi ..
+lemma word_of_int_neg_1 [simp]: "word_of_int (- 1) = - 1"
+ by (simp add: wi_hom_syms)
+
lemma word_of_int_numeral [simp] :
"(word_of_int (numeral bin) :: 'a :: len0 word) = (numeral bin)"
unfolding word_numeral_alt ..
lemma word_of_int_neg_numeral [simp]:
- "(word_of_int (neg_numeral bin) :: 'a :: len0 word) = (neg_numeral bin)"
- unfolding neg_numeral_def word_numeral_alt wi_hom_syms ..
+ "(word_of_int (- numeral bin) :: 'a :: len0 word) = (- numeral bin)"
+ unfolding word_numeral_alt wi_hom_syms ..
lemma word_int_case_wi:
"word_int_case f (word_of_int i :: 'b word) =
@@ -880,8 +883,8 @@
unfolding word_numeral_alt by (rule to_bl_of_bin)
lemma to_bl_neg_numeral [simp]:
- "to_bl (neg_numeral bin::'a::len0 word) =
- bin_to_bl (len_of TYPE('a)) (neg_numeral bin)"
+ "to_bl (- numeral bin::'a::len0 word) =
+ bin_to_bl (len_of TYPE('a)) (- numeral bin)"
unfolding word_neg_numeral_alt by (rule to_bl_of_bin)
lemma to_bl_to_bin [simp] : "bl_to_bin (to_bl w) = uint w"
@@ -1156,11 +1159,8 @@
lemmas word_sle_no [simp] = word_sle_def [of "numeral a" "numeral b"] for a b
-lemma word_1_no: "(1::'a::len0 word) = Numeral1"
- by (simp add: word_numeral_alt)
-
-lemma word_m1_wi: "-1 = word_of_int -1"
- by (rule word_neg_numeral_alt)
+lemma word_m1_wi: "- 1 = word_of_int (- 1)"
+ using word_neg_numeral_alt [of Num.One] by simp
lemma word_0_bl [simp]: "of_bl [] = 0"
unfolding of_bl_def by simp
@@ -1215,9 +1215,9 @@
unfolding scast_def by simp
lemma sint_n1 [simp] : "sint -1 = -1"
- unfolding word_m1_wi by (simp add: word_sbin.eq_norm)
-
-lemma scast_n1 [simp]: "scast -1 = -1"
+ unfolding word_m1_wi word_sbin.eq_norm by simp
+
+lemma scast_n1 [simp]: "scast (- 1) = - 1"
unfolding scast_def by simp
lemma uint_1 [simp]: "uint (1::'a::len word) = 1"
@@ -1270,8 +1270,8 @@
lemma succ_pred_no [simp]:
"word_succ (numeral w) = numeral w + 1"
"word_pred (numeral w) = numeral w - 1"
- "word_succ (neg_numeral w) = neg_numeral w + 1"
- "word_pred (neg_numeral w) = neg_numeral w - 1"
+ "word_succ (- numeral w) = - numeral w + 1"
+ "word_pred (- numeral w) = - numeral w - 1"
unfolding word_succ_p1 word_pred_m1 by simp_all
lemma word_sp_01 [simp] :
@@ -2151,19 +2151,19 @@
lemma word_no_log_defs [simp]:
"NOT (numeral a) = word_of_int (NOT (numeral a))"
- "NOT (neg_numeral a) = word_of_int (NOT (neg_numeral a))"
+ "NOT (- numeral a) = word_of_int (NOT (- numeral a))"
"numeral a AND numeral b = word_of_int (numeral a AND numeral b)"
- "numeral a AND neg_numeral b = word_of_int (numeral a AND neg_numeral b)"
- "neg_numeral a AND numeral b = word_of_int (neg_numeral a AND numeral b)"
- "neg_numeral a AND neg_numeral b = word_of_int (neg_numeral a AND neg_numeral b)"
+ "numeral a AND - numeral b = word_of_int (numeral a AND - numeral b)"
+ "- numeral a AND numeral b = word_of_int (- numeral a AND numeral b)"
+ "- numeral a AND - numeral b = word_of_int (- numeral a AND - numeral b)"
"numeral a OR numeral b = word_of_int (numeral a OR numeral b)"
- "numeral a OR neg_numeral b = word_of_int (numeral a OR neg_numeral b)"
- "neg_numeral a OR numeral b = word_of_int (neg_numeral a OR numeral b)"
- "neg_numeral a OR neg_numeral b = word_of_int (neg_numeral a OR neg_numeral b)"
+ "numeral a OR - numeral b = word_of_int (numeral a OR - numeral b)"
+ "- numeral a OR numeral b = word_of_int (- numeral a OR numeral b)"
+ "- numeral a OR - numeral b = word_of_int (- numeral a OR - numeral b)"
"numeral a XOR numeral b = word_of_int (numeral a XOR numeral b)"
- "numeral a XOR neg_numeral b = word_of_int (numeral a XOR neg_numeral b)"
- "neg_numeral a XOR numeral b = word_of_int (neg_numeral a XOR numeral b)"
- "neg_numeral a XOR neg_numeral b = word_of_int (neg_numeral a XOR neg_numeral b)"
+ "numeral a XOR - numeral b = word_of_int (numeral a XOR - numeral b)"
+ "- numeral a XOR numeral b = word_of_int (- numeral a XOR numeral b)"
+ "- numeral a XOR - numeral b = word_of_int (- numeral a XOR - numeral b)"
by (transfer, rule refl)+
text {* Special cases for when one of the arguments equals 1. *}
@@ -2171,17 +2171,17 @@
lemma word_bitwise_1_simps [simp]:
"NOT (1::'a::len0 word) = -2"
"1 AND numeral b = word_of_int (1 AND numeral b)"
- "1 AND neg_numeral b = word_of_int (1 AND neg_numeral b)"
+ "1 AND - numeral b = word_of_int (1 AND - numeral b)"
"numeral a AND 1 = word_of_int (numeral a AND 1)"
- "neg_numeral a AND 1 = word_of_int (neg_numeral a AND 1)"
+ "- numeral a AND 1 = word_of_int (- numeral a AND 1)"
"1 OR numeral b = word_of_int (1 OR numeral b)"
- "1 OR neg_numeral b = word_of_int (1 OR neg_numeral b)"
+ "1 OR - numeral b = word_of_int (1 OR - numeral b)"
"numeral a OR 1 = word_of_int (numeral a OR 1)"
- "neg_numeral a OR 1 = word_of_int (neg_numeral a OR 1)"
+ "- numeral a OR 1 = word_of_int (- numeral a OR 1)"
"1 XOR numeral b = word_of_int (1 XOR numeral b)"
- "1 XOR neg_numeral b = word_of_int (1 XOR neg_numeral b)"
+ "1 XOR - numeral b = word_of_int (1 XOR - numeral b)"
"numeral a XOR 1 = word_of_int (numeral a XOR 1)"
- "neg_numeral a XOR 1 = word_of_int (neg_numeral a XOR 1)"
+ "- numeral a XOR 1 = word_of_int (- numeral a XOR 1)"
by (transfer, simp)+
lemma uint_or: "uint (x OR y) = (uint x) OR (uint y)"
@@ -2220,8 +2220,8 @@
by transfer (rule refl)
lemma test_bit_neg_numeral [simp]:
- "(neg_numeral w :: 'a::len0 word) !! n \<longleftrightarrow>
- n < len_of TYPE('a) \<and> bin_nth (neg_numeral w) n"
+ "(- numeral w :: 'a::len0 word) !! n \<longleftrightarrow>
+ n < len_of TYPE('a) \<and> bin_nth (- numeral w) n"
by transfer (rule refl)
lemma test_bit_1 [simp]: "(1::'a::len word) !! n \<longleftrightarrow> n = 0"
@@ -2398,7 +2398,7 @@
unfolding word_numeral_alt by (rule msb_word_of_int)
lemma word_msb_neg_numeral [simp]:
- "msb (neg_numeral w::'a::len word) = bin_nth (neg_numeral w) (len_of TYPE('a) - 1)"
+ "msb (- numeral w::'a::len word) = bin_nth (- numeral w) (len_of TYPE('a) - 1)"
unfolding word_neg_numeral_alt by (rule msb_word_of_int)
lemma word_msb_0 [simp]: "\<not> msb (0::'a::len word)"
@@ -2528,7 +2528,7 @@
unfolding word_lsb_alt test_bit_numeral by simp
lemma word_lsb_neg_numeral [simp]:
- "lsb (neg_numeral bin :: 'a :: len word) = (bin_last (neg_numeral bin) = 1)"
+ "lsb (- numeral bin :: 'a :: len word) = (bin_last (- numeral bin) = 1)"
unfolding word_lsb_alt test_bit_neg_numeral by simp
lemma set_bit_word_of_int:
@@ -2544,8 +2544,8 @@
unfolding word_numeral_alt by (rule set_bit_word_of_int)
lemma word_set_neg_numeral [simp]:
- "set_bit (neg_numeral bin::'a::len0 word) n b =
- word_of_int (bin_sc n (if b then 1 else 0) (neg_numeral bin))"
+ "set_bit (- numeral bin::'a::len0 word) n b =
+ word_of_int (bin_sc n (if b then 1 else 0) (- numeral bin))"
unfolding word_neg_numeral_alt by (rule set_bit_word_of_int)
lemma word_set_bit_0 [simp]:
@@ -2612,8 +2612,14 @@
apply clarsimp
apply clarsimp
apply (drule word_gt_0 [THEN iffD1])
- apply (safe intro!: word_eqI bin_nth_lem)
- apply (auto simp add: test_bit_2p nth_2p_bin word_test_bit_def [symmetric])
+ apply (safe intro!: word_eqI)
+ apply (auto simp add: nth_2p_bin)
+ apply (erule notE)
+ apply (simp (no_asm_use) add: uint_word_of_int word_size)
+ apply (subst mod_pos_pos_trivial)
+ apply simp
+ apply (rule power_strict_increasing)
+ apply simp_all
done
lemma word_of_int_2p: "(word_of_int (2 ^ n) :: 'a :: len word) = 2 ^ n"
@@ -2670,7 +2676,7 @@
unfolding word_numeral_alt shiftl1_wi by simp
lemma shiftl1_neg_numeral [simp]:
- "shiftl1 (neg_numeral w) = neg_numeral (Num.Bit0 w)"
+ "shiftl1 (- numeral w) = - numeral (Num.Bit0 w)"
unfolding word_neg_numeral_alt shiftl1_wi by simp
lemma shiftl1_0 [simp] : "shiftl1 0 = 0"
@@ -4638,9 +4644,6 @@
"1 + n \<noteq> (0::'a::len word) \<Longrightarrow> unat (1 + n) = Suc (unat n)"
by unat_arith
-lemma word_no_1 [simp]: "(Numeral1::'a::len0 word) = 1"
- by (fact word_1_no [symmetric])
-
declare bin_to_bl_def [simp]
ML_file "Tools/word_lib.ML"
--- a/src/HOL/Word/WordBitwise.thy Wed Nov 20 16:15:54 2013 +0100
+++ b/src/HOL/Word/WordBitwise.thy Wed Nov 20 16:43:09 2013 +0100
@@ -461,18 +461,18 @@
= True # rev (bin_to_bl n (numeral nm))"
"rev (bin_to_bl (Suc n) (numeral (num.One)))
= True # replicate n False"
- "rev (bin_to_bl (Suc n) (neg_numeral (num.Bit0 nm)))
- = False # rev (bin_to_bl n (neg_numeral nm))"
- "rev (bin_to_bl (Suc n) (neg_numeral (num.Bit1 nm)))
- = True # rev (bin_to_bl n (neg_numeral (nm + num.One)))"
- "rev (bin_to_bl (Suc n) (neg_numeral (num.One)))
+ "rev (bin_to_bl (Suc n) (- numeral (num.Bit0 nm)))
+ = False # rev (bin_to_bl n (- numeral nm))"
+ "rev (bin_to_bl (Suc n) (- numeral (num.Bit1 nm)))
+ = True # rev (bin_to_bl n (- numeral (nm + num.One)))"
+ "rev (bin_to_bl (Suc n) (- numeral (num.One)))
= True # replicate n True"
- "rev (bin_to_bl (Suc n) (neg_numeral (num.Bit0 nm + num.One)))
- = True # rev (bin_to_bl n (neg_numeral (nm + num.One)))"
- "rev (bin_to_bl (Suc n) (neg_numeral (num.Bit1 nm + num.One)))
- = False # rev (bin_to_bl n (neg_numeral (nm + num.One)))"
- "rev (bin_to_bl (Suc n) (neg_numeral (num.One + num.One)))
- = False # rev (bin_to_bl n (neg_numeral num.One))"
+ "rev (bin_to_bl (Suc n) (- numeral (num.Bit0 nm + num.One)))
+ = True # rev (bin_to_bl n (- numeral (nm + num.One)))"
+ "rev (bin_to_bl (Suc n) (- numeral (num.Bit1 nm + num.One)))
+ = False # rev (bin_to_bl n (- numeral (nm + num.One)))"
+ "rev (bin_to_bl (Suc n) (- numeral (num.One + num.One)))
+ = False # rev (bin_to_bl n (- numeral num.One))"
apply (simp_all add: bin_to_bl_def)
apply (simp_all only: bin_to_bl_aux_alt)
apply (simp_all)
--- a/src/Tools/adhoc_overloading.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/Tools/adhoc_overloading.ML Wed Nov 20 16:43:09 2013 +0100
@@ -178,9 +178,9 @@
fun check ctxt =
map (fn t => Term.map_aterms (insert_variants ctxt t) t);
-fun uncheck ctxt =
- if Config.get ctxt show_variants then I
- else map (insert_overloaded ctxt);
+fun uncheck ctxt ts =
+ if Config.get ctxt show_variants orelse exists (is_none o try Term.type_of) ts then ts
+ else map (insert_overloaded ctxt) ts;
fun reject_unresolved ctxt =
let
--- a/src/Tools/subtyping.ML Wed Nov 20 16:15:54 2013 +0100
+++ b/src/Tools/subtyping.ML Wed Nov 20 16:43:09 2013 +0100
@@ -597,8 +597,8 @@
val assignment =
if null bound orelse null not_params then NONE
else SOME (tightest lower S styps_and_sorts (map nameT not_params)
- handle BOUND_ERROR msg =>
- err_bound ctxt (gen_msg err msg) tye (find_error_pack lower key))
+ handle BOUND_ERROR msg => err_bound ctxt (gen_msg err msg) tye
+ (maps (find_error_pack (not lower)) raw_bound))
in
(case assignment of
NONE => tye_idx
@@ -614,7 +614,8 @@
else err_bound ctxt (gen_msg err ("assigned base type " ^
quote (Syntax.string_of_typ ctxt T) ^
" clashes with the upper bound of variable " ^
- Syntax.string_of_typ ctxt (TVar(xi, S)))) tye (find_error_pack (not lower) key)
+ Syntax.string_of_typ ctxt (TVar(xi, S)))) tye
+ (maps (find_error_pack lower) other_bound)
end
else apfst (Vartab.update (xi, T)) tye_idx)
end