--- a/src/HOL/HOL.thy Tue May 11 08:29:42 2010 +0200
+++ b/src/HOL/HOL.thy Tue May 11 08:30:02 2010 +0200
@@ -1869,7 +1869,7 @@
proof
assume "PROP ?ofclass"
show "PROP ?eq"
- by (tactic {* ALLGOALS (rtac (Thm.unconstrain_allTs @{thm equals_eq})) *})
+ by (tactic {* ALLGOALS (rtac (Thm.unconstrainT @{thm equals_eq})) *})
(fact `PROP ?ofclass`)
next
assume "PROP ?eq"
--- a/src/HOL/IsaMakefile Tue May 11 08:29:42 2010 +0200
+++ b/src/HOL/IsaMakefile Tue May 11 08:30:02 2010 +0200
@@ -352,7 +352,6 @@
Lubs.thy \
MacLaurin.thy \
NthRoot.thy \
- PReal.thy \
Parity.thy \
RComplete.thy \
Rat.thy \
@@ -960,7 +959,7 @@
ex/Binary.thy ex/CTL.thy ex/Chinese.thy ex/Classical.thy \
ex/CodegenSML_Test.thy ex/Codegenerator_Candidates.thy \
ex/Codegenerator_Pretty.thy ex/Codegenerator_Pretty_Test.thy \
- ex/Codegenerator_Test.thy ex/Coherent.thy \
+ ex/Codegenerator_Test.thy ex/Coherent.thy ex/Dedekind_Real.thy \
ex/Efficient_Nat_examples.thy ex/Eval_Examples.thy ex/Fundefs.thy \
ex/Gauge_Integration.thy ex/Groebner_Examples.thy ex/Guess.thy \
ex/HarmonicSeries.thy ex/Hebrew.thy ex/Hex_Bin_Examples.thy \
--- a/src/HOL/NSA/NSA.thy Tue May 11 08:29:42 2010 +0200
+++ b/src/HOL/NSA/NSA.thy Tue May 11 08:30:02 2010 +0200
@@ -2188,7 +2188,7 @@
apply (subst pos_less_divide_eq, assumption)
apply (subst pos_less_divide_eq)
apply (simp add: real_of_nat_Suc_gt_zero)
-apply (simp add: real_mult_commute)
+apply (simp add: mult_commute)
done
lemma finite_inverse_real_of_posnat_gt_real:
--- a/src/HOL/PReal.thy Tue May 11 08:29:42 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1163 +0,0 @@
-(* Title: HOL/PReal.thy
- Author: Jacques D. Fleuriot, University of Cambridge
-
-The positive reals as Dedekind sections of positive
-rationals. Fundamentals of Abstract Analysis [Gleason- p. 121]
-provides some of the definitions.
-*)
-
-header {* Positive real numbers *}
-
-theory PReal
-imports Rat
-begin
-
-text{*Could be generalized and moved to @{text Groups}*}
-lemma add_eq_exists: "\<exists>x. a+x = (b::rat)"
-by (rule_tac x="b-a" in exI, simp)
-
-definition
- cut :: "rat set => bool" where
- [code del]: "cut A = ({} \<subset> A &
- A < {r. 0 < r} &
- (\<forall>y \<in> A. ((\<forall>z. 0<z & z < y --> z \<in> A) & (\<exists>u \<in> A. y < u))))"
-
-lemma interval_empty_iff:
- "{y. (x::'a::dense_linorder) < y \<and> y < z} = {} \<longleftrightarrow> \<not> x < z"
- by (auto dest: dense)
-
-
-lemma cut_of_rat:
- assumes q: "0 < q" shows "cut {r::rat. 0 < r & r < q}" (is "cut ?A")
-proof -
- from q have pos: "?A < {r. 0 < r}" by force
- have nonempty: "{} \<subset> ?A"
- proof
- show "{} \<subseteq> ?A" by simp
- show "{} \<noteq> ?A"
- by (force simp only: q eq_commute [of "{}"] interval_empty_iff)
- qed
- show ?thesis
- by (simp add: cut_def pos nonempty,
- blast dest: dense intro: order_less_trans)
-qed
-
-
-typedef preal = "{A. cut A}"
- by (blast intro: cut_of_rat [OF zero_less_one])
-
-definition
- psup :: "preal set => preal" where
- [code del]: "psup P = Abs_preal (\<Union>X \<in> P. Rep_preal X)"
-
-definition
- add_set :: "[rat set,rat set] => rat set" where
- "add_set A B = {w. \<exists>x \<in> A. \<exists>y \<in> B. w = x + y}"
-
-definition
- diff_set :: "[rat set,rat set] => rat set" where
- [code del]: "diff_set A B = {w. \<exists>x. 0 < w & 0 < x & x \<notin> B & x + w \<in> A}"
-
-definition
- mult_set :: "[rat set,rat set] => rat set" where
- "mult_set A B = {w. \<exists>x \<in> A. \<exists>y \<in> B. w = x * y}"
-
-definition
- inverse_set :: "rat set => rat set" where
- [code del]: "inverse_set A = {x. \<exists>y. 0 < x & x < y & inverse y \<notin> A}"
-
-instantiation preal :: "{ord, plus, minus, times, inverse, one}"
-begin
-
-definition
- preal_less_def [code del]:
- "R < S == Rep_preal R < Rep_preal S"
-
-definition
- preal_le_def [code del]:
- "R \<le> S == Rep_preal R \<subseteq> Rep_preal S"
-
-definition
- preal_add_def:
- "R + S == Abs_preal (add_set (Rep_preal R) (Rep_preal S))"
-
-definition
- preal_diff_def:
- "R - S == Abs_preal (diff_set (Rep_preal R) (Rep_preal S))"
-
-definition
- preal_mult_def:
- "R * S == Abs_preal (mult_set (Rep_preal R) (Rep_preal S))"
-
-definition
- preal_inverse_def:
- "inverse R == Abs_preal (inverse_set (Rep_preal R))"
-
-definition "R / S = R * inverse (S\<Colon>preal)"
-
-definition
- preal_one_def:
- "1 == Abs_preal {x. 0 < x & x < 1}"
-
-instance ..
-
-end
-
-
-text{*Reduces equality on abstractions to equality on representatives*}
-declare Abs_preal_inject [simp]
-declare Abs_preal_inverse [simp]
-
-lemma rat_mem_preal: "0 < q ==> {r::rat. 0 < r & r < q} \<in> preal"
-by (simp add: preal_def cut_of_rat)
-
-lemma preal_nonempty: "A \<in> preal ==> \<exists>x\<in>A. 0 < x"
-by (unfold preal_def cut_def, blast)
-
-lemma preal_Ex_mem: "A \<in> preal \<Longrightarrow> \<exists>x. x \<in> A"
-by (drule preal_nonempty, fast)
-
-lemma preal_imp_psubset_positives: "A \<in> preal ==> A < {r. 0 < r}"
-by (force simp add: preal_def cut_def)
-
-lemma preal_exists_bound: "A \<in> preal ==> \<exists>x. 0 < x & x \<notin> A"
-by (drule preal_imp_psubset_positives, auto)
-
-lemma preal_exists_greater: "[| A \<in> preal; y \<in> A |] ==> \<exists>u \<in> A. y < u"
-by (unfold preal_def cut_def, blast)
-
-lemma preal_downwards_closed: "[| A \<in> preal; y \<in> A; 0 < z; z < y |] ==> z \<in> A"
-by (unfold preal_def cut_def, blast)
-
-text{*Relaxing the final premise*}
-lemma preal_downwards_closed':
- "[| A \<in> preal; y \<in> A; 0 < z; z \<le> y |] ==> z \<in> A"
-apply (simp add: order_le_less)
-apply (blast intro: preal_downwards_closed)
-done
-
-text{*A positive fraction not in a positive real is an upper bound.
- Gleason p. 122 - Remark (1)*}
-
-lemma not_in_preal_ub:
- assumes A: "A \<in> preal"
- and notx: "x \<notin> A"
- and y: "y \<in> A"
- and pos: "0 < x"
- shows "y < x"
-proof (cases rule: linorder_cases)
- assume "x<y"
- with notx show ?thesis
- by (simp add: preal_downwards_closed [OF A y] pos)
-next
- assume "x=y"
- with notx and y show ?thesis by simp
-next
- assume "y<x"
- thus ?thesis .
-qed
-
-text {* preal lemmas instantiated to @{term "Rep_preal X"} *}
-
-lemma mem_Rep_preal_Ex: "\<exists>x. x \<in> Rep_preal X"
-by (rule preal_Ex_mem [OF Rep_preal])
-
-lemma Rep_preal_exists_bound: "\<exists>x>0. x \<notin> Rep_preal X"
-by (rule preal_exists_bound [OF Rep_preal])
-
-lemmas not_in_Rep_preal_ub = not_in_preal_ub [OF Rep_preal]
-
-
-subsection{*Properties of Ordering*}
-
-instance preal :: order
-proof
- fix w :: preal
- show "w \<le> w" by (simp add: preal_le_def)
-next
- fix i j k :: preal
- assume "i \<le> j" and "j \<le> k"
- then show "i \<le> k" by (simp add: preal_le_def)
-next
- fix z w :: preal
- assume "z \<le> w" and "w \<le> z"
- then show "z = w" by (simp add: preal_le_def Rep_preal_inject)
-next
- fix z w :: preal
- show "z < w \<longleftrightarrow> z \<le> w \<and> \<not> w \<le> z"
- by (auto simp add: preal_le_def preal_less_def Rep_preal_inject)
-qed
-
-lemma preal_imp_pos: "[|A \<in> preal; r \<in> A|] ==> 0 < r"
-by (insert preal_imp_psubset_positives, blast)
-
-instance preal :: linorder
-proof
- fix x y :: preal
- show "x <= y | y <= x"
- apply (auto simp add: preal_le_def)
- apply (rule ccontr)
- apply (blast dest: not_in_Rep_preal_ub intro: preal_imp_pos [OF Rep_preal]
- elim: order_less_asym)
- done
-qed
-
-instantiation preal :: distrib_lattice
-begin
-
-definition
- "(inf \<Colon> preal \<Rightarrow> preal \<Rightarrow> preal) = min"
-
-definition
- "(sup \<Colon> preal \<Rightarrow> preal \<Rightarrow> preal) = max"
-
-instance
- by intro_classes
- (auto simp add: inf_preal_def sup_preal_def min_max.sup_inf_distrib1)
-
-end
-
-subsection{*Properties of Addition*}
-
-lemma preal_add_commute: "(x::preal) + y = y + x"
-apply (unfold preal_add_def add_set_def)
-apply (rule_tac f = Abs_preal in arg_cong)
-apply (force simp add: add_commute)
-done
-
-text{*Lemmas for proving that addition of two positive reals gives
- a positive real*}
-
-lemma empty_psubset_nonempty: "a \<in> A ==> {} \<subset> A"
-by blast
-
-text{*Part 1 of Dedekind sections definition*}
-lemma add_set_not_empty:
- "[|A \<in> preal; B \<in> preal|] ==> {} \<subset> add_set A B"
-apply (drule preal_nonempty)+
-apply (auto simp add: add_set_def)
-done
-
-text{*Part 2 of Dedekind sections definition. A structured version of
-this proof is @{text preal_not_mem_mult_set_Ex} below.*}
-lemma preal_not_mem_add_set_Ex:
- "[|A \<in> preal; B \<in> preal|] ==> \<exists>q>0. q \<notin> add_set A B"
-apply (insert preal_exists_bound [of A] preal_exists_bound [of B], auto)
-apply (rule_tac x = "x+xa" in exI)
-apply (simp add: add_set_def, clarify)
-apply (drule (3) not_in_preal_ub)+
-apply (force dest: add_strict_mono)
-done
-
-lemma add_set_not_rat_set:
- assumes A: "A \<in> preal"
- and B: "B \<in> preal"
- shows "add_set A B < {r. 0 < r}"
-proof
- from preal_imp_pos [OF A] preal_imp_pos [OF B]
- show "add_set A B \<subseteq> {r. 0 < r}" by (force simp add: add_set_def)
-next
- show "add_set A B \<noteq> {r. 0 < r}"
- by (insert preal_not_mem_add_set_Ex [OF A B], blast)
-qed
-
-text{*Part 3 of Dedekind sections definition*}
-lemma add_set_lemma3:
- "[|A \<in> preal; B \<in> preal; u \<in> add_set A B; 0 < z; z < u|]
- ==> z \<in> add_set A B"
-proof (unfold add_set_def, clarify)
- fix x::rat and y::rat
- assume A: "A \<in> preal"
- and B: "B \<in> preal"
- and [simp]: "0 < z"
- and zless: "z < x + y"
- and x: "x \<in> A"
- and y: "y \<in> B"
- have xpos [simp]: "0<x" by (rule preal_imp_pos [OF A x])
- have ypos [simp]: "0<y" by (rule preal_imp_pos [OF B y])
- have xypos [simp]: "0 < x+y" by (simp add: pos_add_strict)
- let ?f = "z/(x+y)"
- have fless: "?f < 1" by (simp add: zless pos_divide_less_eq)
- show "\<exists>x' \<in> A. \<exists>y'\<in>B. z = x' + y'"
- proof (intro bexI)
- show "z = x*?f + y*?f"
- by (simp add: left_distrib [symmetric] divide_inverse mult_ac
- order_less_imp_not_eq2)
- next
- show "y * ?f \<in> B"
- proof (rule preal_downwards_closed [OF B y])
- show "0 < y * ?f"
- by (simp add: divide_inverse zero_less_mult_iff)
- next
- show "y * ?f < y"
- by (insert mult_strict_left_mono [OF fless ypos], simp)
- qed
- next
- show "x * ?f \<in> A"
- proof (rule preal_downwards_closed [OF A x])
- show "0 < x * ?f"
- by (simp add: divide_inverse zero_less_mult_iff)
- next
- show "x * ?f < x"
- by (insert mult_strict_left_mono [OF fless xpos], simp)
- qed
- qed
-qed
-
-text{*Part 4 of Dedekind sections definition*}
-lemma add_set_lemma4:
- "[|A \<in> preal; B \<in> preal; y \<in> add_set A B|] ==> \<exists>u \<in> add_set A B. y < u"
-apply (auto simp add: add_set_def)
-apply (frule preal_exists_greater [of A], auto)
-apply (rule_tac x="u + y" in exI)
-apply (auto intro: add_strict_left_mono)
-done
-
-lemma mem_add_set:
- "[|A \<in> preal; B \<in> preal|] ==> add_set A B \<in> preal"
-apply (simp (no_asm_simp) add: preal_def cut_def)
-apply (blast intro!: add_set_not_empty add_set_not_rat_set
- add_set_lemma3 add_set_lemma4)
-done
-
-lemma preal_add_assoc: "((x::preal) + y) + z = x + (y + z)"
-apply (simp add: preal_add_def mem_add_set Rep_preal)
-apply (force simp add: add_set_def add_ac)
-done
-
-instance preal :: ab_semigroup_add
-proof
- fix a b c :: preal
- show "(a + b) + c = a + (b + c)" by (rule preal_add_assoc)
- show "a + b = b + a" by (rule preal_add_commute)
-qed
-
-
-subsection{*Properties of Multiplication*}
-
-text{*Proofs essentially same as for addition*}
-
-lemma preal_mult_commute: "(x::preal) * y = y * x"
-apply (unfold preal_mult_def mult_set_def)
-apply (rule_tac f = Abs_preal in arg_cong)
-apply (force simp add: mult_commute)
-done
-
-text{*Multiplication of two positive reals gives a positive real.*}
-
-text{*Lemmas for proving positive reals multiplication set in @{typ preal}*}
-
-text{*Part 1 of Dedekind sections definition*}
-lemma mult_set_not_empty:
- "[|A \<in> preal; B \<in> preal|] ==> {} \<subset> mult_set A B"
-apply (insert preal_nonempty [of A] preal_nonempty [of B])
-apply (auto simp add: mult_set_def)
-done
-
-text{*Part 2 of Dedekind sections definition*}
-lemma preal_not_mem_mult_set_Ex:
- assumes A: "A \<in> preal"
- and B: "B \<in> preal"
- shows "\<exists>q. 0 < q & q \<notin> mult_set A B"
-proof -
- from preal_exists_bound [OF A]
- obtain x where [simp]: "0 < x" "x \<notin> A" by blast
- from preal_exists_bound [OF B]
- obtain y where [simp]: "0 < y" "y \<notin> B" by blast
- show ?thesis
- proof (intro exI conjI)
- show "0 < x*y" by (simp add: mult_pos_pos)
- show "x * y \<notin> mult_set A B"
- proof -
- { fix u::rat and v::rat
- assume "u \<in> A" and "v \<in> B" and "x*y = u*v"
- moreover
- with prems have "u<x" and "v<y" by (blast dest: not_in_preal_ub)+
- moreover
- with prems have "0\<le>v"
- by (blast intro: preal_imp_pos [OF B] order_less_imp_le prems)
- moreover
- from calculation
- have "u*v < x*y" by (blast intro: mult_strict_mono prems)
- ultimately have False by force }
- thus ?thesis by (auto simp add: mult_set_def)
- qed
- qed
-qed
-
-lemma mult_set_not_rat_set:
- assumes A: "A \<in> preal"
- and B: "B \<in> preal"
- shows "mult_set A B < {r. 0 < r}"
-proof
- show "mult_set A B \<subseteq> {r. 0 < r}"
- by (force simp add: mult_set_def
- intro: preal_imp_pos [OF A] preal_imp_pos [OF B] mult_pos_pos)
- show "mult_set A B \<noteq> {r. 0 < r}"
- using preal_not_mem_mult_set_Ex [OF A B] by blast
-qed
-
-
-
-text{*Part 3 of Dedekind sections definition*}
-lemma mult_set_lemma3:
- "[|A \<in> preal; B \<in> preal; u \<in> mult_set A B; 0 < z; z < u|]
- ==> z \<in> mult_set A B"
-proof (unfold mult_set_def, clarify)
- fix x::rat and y::rat
- assume A: "A \<in> preal"
- and B: "B \<in> preal"
- and [simp]: "0 < z"
- and zless: "z < x * y"
- and x: "x \<in> A"
- and y: "y \<in> B"
- have [simp]: "0<y" by (rule preal_imp_pos [OF B y])
- show "\<exists>x' \<in> A. \<exists>y' \<in> B. z = x' * y'"
- proof
- show "\<exists>y'\<in>B. z = (z/y) * y'"
- proof
- show "z = (z/y)*y"
- by (simp add: divide_inverse mult_commute [of y] mult_assoc
- order_less_imp_not_eq2)
- show "y \<in> B" by fact
- qed
- next
- show "z/y \<in> A"
- proof (rule preal_downwards_closed [OF A x])
- show "0 < z/y"
- by (simp add: zero_less_divide_iff)
- show "z/y < x" by (simp add: pos_divide_less_eq zless)
- qed
- qed
-qed
-
-text{*Part 4 of Dedekind sections definition*}
-lemma mult_set_lemma4:
- "[|A \<in> preal; B \<in> preal; y \<in> mult_set A B|] ==> \<exists>u \<in> mult_set A B. y < u"
-apply (auto simp add: mult_set_def)
-apply (frule preal_exists_greater [of A], auto)
-apply (rule_tac x="u * y" in exI)
-apply (auto intro: preal_imp_pos [of A] preal_imp_pos [of B]
- mult_strict_right_mono)
-done
-
-
-lemma mem_mult_set:
- "[|A \<in> preal; B \<in> preal|] ==> mult_set A B \<in> preal"
-apply (simp (no_asm_simp) add: preal_def cut_def)
-apply (blast intro!: mult_set_not_empty mult_set_not_rat_set
- mult_set_lemma3 mult_set_lemma4)
-done
-
-lemma preal_mult_assoc: "((x::preal) * y) * z = x * (y * z)"
-apply (simp add: preal_mult_def mem_mult_set Rep_preal)
-apply (force simp add: mult_set_def mult_ac)
-done
-
-instance preal :: ab_semigroup_mult
-proof
- fix a b c :: preal
- show "(a * b) * c = a * (b * c)" by (rule preal_mult_assoc)
- show "a * b = b * a" by (rule preal_mult_commute)
-qed
-
-
-text{* Positive real 1 is the multiplicative identity element *}
-
-lemma preal_mult_1: "(1::preal) * z = z"
-proof (induct z)
- fix A :: "rat set"
- assume A: "A \<in> preal"
- have "{w. \<exists>u. 0 < u \<and> u < 1 & (\<exists>v \<in> A. w = u * v)} = A" (is "?lhs = A")
- proof
- show "?lhs \<subseteq> A"
- proof clarify
- fix x::rat and u::rat and v::rat
- assume upos: "0<u" and "u<1" and v: "v \<in> A"
- have vpos: "0<v" by (rule preal_imp_pos [OF A v])
- hence "u*v < 1*v" by (simp only: mult_strict_right_mono prems)
- thus "u * v \<in> A"
- by (force intro: preal_downwards_closed [OF A v] mult_pos_pos
- upos vpos)
- qed
- next
- show "A \<subseteq> ?lhs"
- proof clarify
- fix x::rat
- assume x: "x \<in> A"
- have xpos: "0<x" by (rule preal_imp_pos [OF A x])
- from preal_exists_greater [OF A x]
- obtain v where v: "v \<in> A" and xlessv: "x < v" ..
- have vpos: "0<v" by (rule preal_imp_pos [OF A v])
- show "\<exists>u. 0 < u \<and> u < 1 \<and> (\<exists>v\<in>A. x = u * v)"
- proof (intro exI conjI)
- show "0 < x/v"
- by (simp add: zero_less_divide_iff xpos vpos)
- show "x / v < 1"
- by (simp add: pos_divide_less_eq vpos xlessv)
- show "\<exists>v'\<in>A. x = (x / v) * v'"
- proof
- show "x = (x/v)*v"
- by (simp add: divide_inverse mult_assoc vpos
- order_less_imp_not_eq2)
- show "v \<in> A" by fact
- qed
- qed
- qed
- qed
- thus "1 * Abs_preal A = Abs_preal A"
- by (simp add: preal_one_def preal_mult_def mult_set_def
- rat_mem_preal A)
-qed
-
-instance preal :: comm_monoid_mult
-by intro_classes (rule preal_mult_1)
-
-
-subsection{*Distribution of Multiplication across Addition*}
-
-lemma mem_Rep_preal_add_iff:
- "(z \<in> Rep_preal(R+S)) = (\<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. z = x + y)"
-apply (simp add: preal_add_def mem_add_set Rep_preal)
-apply (simp add: add_set_def)
-done
-
-lemma mem_Rep_preal_mult_iff:
- "(z \<in> Rep_preal(R*S)) = (\<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. z = x * y)"
-apply (simp add: preal_mult_def mem_mult_set Rep_preal)
-apply (simp add: mult_set_def)
-done
-
-lemma distrib_subset1:
- "Rep_preal (w * (x + y)) \<subseteq> Rep_preal (w * x + w * y)"
-apply (auto simp add: Bex_def mem_Rep_preal_add_iff mem_Rep_preal_mult_iff)
-apply (force simp add: right_distrib)
-done
-
-lemma preal_add_mult_distrib_mean:
- assumes a: "a \<in> Rep_preal w"
- and b: "b \<in> Rep_preal w"
- and d: "d \<in> Rep_preal x"
- and e: "e \<in> Rep_preal y"
- shows "\<exists>c \<in> Rep_preal w. a * d + b * e = c * (d + e)"
-proof
- let ?c = "(a*d + b*e)/(d+e)"
- have [simp]: "0<a" "0<b" "0<d" "0<e" "0<d+e"
- by (blast intro: preal_imp_pos [OF Rep_preal] a b d e pos_add_strict)+
- have cpos: "0 < ?c"
- by (simp add: zero_less_divide_iff zero_less_mult_iff pos_add_strict)
- show "a * d + b * e = ?c * (d + e)"
- by (simp add: divide_inverse mult_assoc order_less_imp_not_eq2)
- show "?c \<in> Rep_preal w"
- proof (cases rule: linorder_le_cases)
- assume "a \<le> b"
- hence "?c \<le> b"
- by (simp add: pos_divide_le_eq right_distrib mult_right_mono
- order_less_imp_le)
- thus ?thesis by (rule preal_downwards_closed' [OF Rep_preal b cpos])
- next
- assume "b \<le> a"
- hence "?c \<le> a"
- by (simp add: pos_divide_le_eq right_distrib mult_right_mono
- order_less_imp_le)
- thus ?thesis by (rule preal_downwards_closed' [OF Rep_preal a cpos])
- qed
-qed
-
-lemma distrib_subset2:
- "Rep_preal (w * x + w * y) \<subseteq> Rep_preal (w * (x + y))"
-apply (auto simp add: Bex_def mem_Rep_preal_add_iff mem_Rep_preal_mult_iff)
-apply (drule_tac w=w and x=x and y=y in preal_add_mult_distrib_mean, auto)
-done
-
-lemma preal_add_mult_distrib2: "(w * ((x::preal) + y)) = (w * x) + (w * y)"
-apply (rule Rep_preal_inject [THEN iffD1])
-apply (rule equalityI [OF distrib_subset1 distrib_subset2])
-done
-
-lemma preal_add_mult_distrib: "(((x::preal) + y) * w) = (x * w) + (y * w)"
-by (simp add: preal_mult_commute preal_add_mult_distrib2)
-
-instance preal :: comm_semiring
-by intro_classes (rule preal_add_mult_distrib)
-
-
-subsection{*Existence of Inverse, a Positive Real*}
-
-lemma mem_inv_set_ex:
- assumes A: "A \<in> preal" shows "\<exists>x y. 0 < x & x < y & inverse y \<notin> A"
-proof -
- from preal_exists_bound [OF A]
- obtain x where [simp]: "0<x" "x \<notin> A" by blast
- show ?thesis
- proof (intro exI conjI)
- show "0 < inverse (x+1)"
- by (simp add: order_less_trans [OF _ less_add_one])
- show "inverse(x+1) < inverse x"
- by (simp add: less_imp_inverse_less less_add_one)
- show "inverse (inverse x) \<notin> A"
- by (simp add: order_less_imp_not_eq2)
- qed
-qed
-
-text{*Part 1 of Dedekind sections definition*}
-lemma inverse_set_not_empty:
- "A \<in> preal ==> {} \<subset> inverse_set A"
-apply (insert mem_inv_set_ex [of A])
-apply (auto simp add: inverse_set_def)
-done
-
-text{*Part 2 of Dedekind sections definition*}
-
-lemma preal_not_mem_inverse_set_Ex:
- assumes A: "A \<in> preal" shows "\<exists>q. 0 < q & q \<notin> inverse_set A"
-proof -
- from preal_nonempty [OF A]
- obtain x where x: "x \<in> A" and xpos [simp]: "0<x" ..
- show ?thesis
- proof (intro exI conjI)
- show "0 < inverse x" by simp
- show "inverse x \<notin> inverse_set A"
- proof -
- { fix y::rat
- assume ygt: "inverse x < y"
- have [simp]: "0 < y" by (simp add: order_less_trans [OF _ ygt])
- have iyless: "inverse y < x"
- by (simp add: inverse_less_imp_less [of x] ygt)
- have "inverse y \<in> A"
- by (simp add: preal_downwards_closed [OF A x] iyless)}
- thus ?thesis by (auto simp add: inverse_set_def)
- qed
- qed
-qed
-
-lemma inverse_set_not_rat_set:
- assumes A: "A \<in> preal" shows "inverse_set A < {r. 0 < r}"
-proof
- show "inverse_set A \<subseteq> {r. 0 < r}" by (force simp add: inverse_set_def)
-next
- show "inverse_set A \<noteq> {r. 0 < r}"
- by (insert preal_not_mem_inverse_set_Ex [OF A], blast)
-qed
-
-text{*Part 3 of Dedekind sections definition*}
-lemma inverse_set_lemma3:
- "[|A \<in> preal; u \<in> inverse_set A; 0 < z; z < u|]
- ==> z \<in> inverse_set A"
-apply (auto simp add: inverse_set_def)
-apply (auto intro: order_less_trans)
-done
-
-text{*Part 4 of Dedekind sections definition*}
-lemma inverse_set_lemma4:
- "[|A \<in> preal; y \<in> inverse_set A|] ==> \<exists>u \<in> inverse_set A. y < u"
-apply (auto simp add: inverse_set_def)
-apply (drule dense [of y])
-apply (blast intro: order_less_trans)
-done
-
-
-lemma mem_inverse_set:
- "A \<in> preal ==> inverse_set A \<in> preal"
-apply (simp (no_asm_simp) add: preal_def cut_def)
-apply (blast intro!: inverse_set_not_empty inverse_set_not_rat_set
- inverse_set_lemma3 inverse_set_lemma4)
-done
-
-
-subsection{*Gleason's Lemma 9-3.4, page 122*}
-
-lemma Gleason9_34_exists:
- assumes A: "A \<in> preal"
- and "\<forall>x\<in>A. x + u \<in> A"
- and "0 \<le> z"
- shows "\<exists>b\<in>A. b + (of_int z) * u \<in> A"
-proof (cases z rule: int_cases)
- case (nonneg n)
- show ?thesis
- proof (simp add: prems, induct n)
- case 0
- from preal_nonempty [OF A]
- show ?case by force
- case (Suc k)
- from this obtain b where "b \<in> A" "b + of_nat k * u \<in> A" ..
- hence "b + of_int (int k)*u + u \<in> A" by (simp add: prems)
- thus ?case by (force simp add: algebra_simps prems)
- qed
-next
- case (neg n)
- with prems show ?thesis by simp
-qed
-
-lemma Gleason9_34_contra:
- assumes A: "A \<in> preal"
- shows "[|\<forall>x\<in>A. x + u \<in> A; 0 < u; 0 < y; y \<notin> A|] ==> False"
-proof (induct u, induct y)
- fix a::int and b::int
- fix c::int and d::int
- assume bpos [simp]: "0 < b"
- and dpos [simp]: "0 < d"
- and closed: "\<forall>x\<in>A. x + (Fract c d) \<in> A"
- and upos: "0 < Fract c d"
- and ypos: "0 < Fract a b"
- and notin: "Fract a b \<notin> A"
- have cpos [simp]: "0 < c"
- by (simp add: zero_less_Fract_iff [OF dpos, symmetric] upos)
- have apos [simp]: "0 < a"
- by (simp add: zero_less_Fract_iff [OF bpos, symmetric] ypos)
- let ?k = "a*d"
- have frle: "Fract a b \<le> Fract ?k 1 * (Fract c d)"
- proof -
- have "?thesis = ((a * d * b * d) \<le> c * b * (a * d * b * d))"
- by (simp add: order_less_imp_not_eq2 mult_ac)
- moreover
- have "(1 * (a * d * b * d)) \<le> c * b * (a * d * b * d)"
- by (rule mult_mono,
- simp_all add: int_one_le_iff_zero_less zero_less_mult_iff
- order_less_imp_le)
- ultimately
- show ?thesis by simp
- qed
- have k: "0 \<le> ?k" by (simp add: order_less_imp_le zero_less_mult_iff)
- from Gleason9_34_exists [OF A closed k]
- obtain z where z: "z \<in> A"
- and mem: "z + of_int ?k * Fract c d \<in> A" ..
- have less: "z + of_int ?k * Fract c d < Fract a b"
- by (rule not_in_preal_ub [OF A notin mem ypos])
- have "0<z" by (rule preal_imp_pos [OF A z])
- with frle and less show False by (simp add: Fract_of_int_eq)
-qed
-
-
-lemma Gleason9_34:
- assumes A: "A \<in> preal"
- and upos: "0 < u"
- shows "\<exists>r \<in> A. r + u \<notin> A"
-proof (rule ccontr, simp)
- assume closed: "\<forall>r\<in>A. r + u \<in> A"
- from preal_exists_bound [OF A]
- obtain y where y: "y \<notin> A" and ypos: "0 < y" by blast
- show False
- by (rule Gleason9_34_contra [OF A closed upos ypos y])
-qed
-
-
-
-subsection{*Gleason's Lemma 9-3.6*}
-
-lemma lemma_gleason9_36:
- assumes A: "A \<in> preal"
- and x: "1 < x"
- shows "\<exists>r \<in> A. r*x \<notin> A"
-proof -
- from preal_nonempty [OF A]
- obtain y where y: "y \<in> A" and ypos: "0<y" ..
- show ?thesis
- proof (rule classical)
- assume "~(\<exists>r\<in>A. r * x \<notin> A)"
- with y have ymem: "y * x \<in> A" by blast
- from ypos mult_strict_left_mono [OF x]
- have yless: "y < y*x" by simp
- let ?d = "y*x - y"
- from yless have dpos: "0 < ?d" and eq: "y + ?d = y*x" by auto
- from Gleason9_34 [OF A dpos]
- obtain r where r: "r\<in>A" and notin: "r + ?d \<notin> A" ..
- have rpos: "0<r" by (rule preal_imp_pos [OF A r])
- with dpos have rdpos: "0 < r + ?d" by arith
- have "~ (r + ?d \<le> y + ?d)"
- proof
- assume le: "r + ?d \<le> y + ?d"
- from ymem have yd: "y + ?d \<in> A" by (simp add: eq)
- have "r + ?d \<in> A" by (rule preal_downwards_closed' [OF A yd rdpos le])
- with notin show False by simp
- qed
- hence "y < r" by simp
- with ypos have dless: "?d < (r * ?d)/y"
- by (simp add: pos_less_divide_eq mult_commute [of ?d]
- mult_strict_right_mono dpos)
- have "r + ?d < r*x"
- proof -
- have "r + ?d < r + (r * ?d)/y" by (simp add: dless)
- also with ypos have "... = (r/y) * (y + ?d)"
- by (simp only: algebra_simps divide_inverse, simp)
- also have "... = r*x" using ypos
- by simp
- finally show "r + ?d < r*x" .
- qed
- with r notin rdpos
- show "\<exists>r\<in>A. r * x \<notin> A" by (blast dest: preal_downwards_closed [OF A])
- qed
-qed
-
-subsection{*Existence of Inverse: Part 2*}
-
-lemma mem_Rep_preal_inverse_iff:
- "(z \<in> Rep_preal(inverse R)) =
- (0 < z \<and> (\<exists>y. z < y \<and> inverse y \<notin> Rep_preal R))"
-apply (simp add: preal_inverse_def mem_inverse_set Rep_preal)
-apply (simp add: inverse_set_def)
-done
-
-lemma Rep_preal_one:
- "Rep_preal 1 = {x. 0 < x \<and> x < 1}"
-by (simp add: preal_one_def rat_mem_preal)
-
-lemma subset_inverse_mult_lemma:
- assumes xpos: "0 < x" and xless: "x < 1"
- shows "\<exists>r u y. 0 < r & r < y & inverse y \<notin> Rep_preal R &
- u \<in> Rep_preal R & x = r * u"
-proof -
- from xpos and xless have "1 < inverse x" by (simp add: one_less_inverse_iff)
- from lemma_gleason9_36 [OF Rep_preal this]
- obtain r where r: "r \<in> Rep_preal R"
- and notin: "r * (inverse x) \<notin> Rep_preal R" ..
- have rpos: "0<r" by (rule preal_imp_pos [OF Rep_preal r])
- from preal_exists_greater [OF Rep_preal r]
- obtain u where u: "u \<in> Rep_preal R" and rless: "r < u" ..
- have upos: "0<u" by (rule preal_imp_pos [OF Rep_preal u])
- show ?thesis
- proof (intro exI conjI)
- show "0 < x/u" using xpos upos
- by (simp add: zero_less_divide_iff)
- show "x/u < x/r" using xpos upos rpos
- by (simp add: divide_inverse mult_less_cancel_left rless)
- show "inverse (x / r) \<notin> Rep_preal R" using notin
- by (simp add: divide_inverse mult_commute)
- show "u \<in> Rep_preal R" by (rule u)
- show "x = x / u * u" using upos
- by (simp add: divide_inverse mult_commute)
- qed
-qed
-
-lemma subset_inverse_mult:
- "Rep_preal 1 \<subseteq> Rep_preal(inverse R * R)"
-apply (auto simp add: Bex_def Rep_preal_one mem_Rep_preal_inverse_iff
- mem_Rep_preal_mult_iff)
-apply (blast dest: subset_inverse_mult_lemma)
-done
-
-lemma inverse_mult_subset_lemma:
- assumes rpos: "0 < r"
- and rless: "r < y"
- and notin: "inverse y \<notin> Rep_preal R"
- and q: "q \<in> Rep_preal R"
- shows "r*q < 1"
-proof -
- have "q < inverse y" using rpos rless
- by (simp add: not_in_preal_ub [OF Rep_preal notin] q)
- hence "r * q < r/y" using rpos
- by (simp add: divide_inverse mult_less_cancel_left)
- also have "... \<le> 1" using rpos rless
- by (simp add: pos_divide_le_eq)
- finally show ?thesis .
-qed
-
-lemma inverse_mult_subset:
- "Rep_preal(inverse R * R) \<subseteq> Rep_preal 1"
-apply (auto simp add: Bex_def Rep_preal_one mem_Rep_preal_inverse_iff
- mem_Rep_preal_mult_iff)
-apply (simp add: zero_less_mult_iff preal_imp_pos [OF Rep_preal])
-apply (blast intro: inverse_mult_subset_lemma)
-done
-
-lemma preal_mult_inverse: "inverse R * R = (1::preal)"
-apply (rule Rep_preal_inject [THEN iffD1])
-apply (rule equalityI [OF inverse_mult_subset subset_inverse_mult])
-done
-
-lemma preal_mult_inverse_right: "R * inverse R = (1::preal)"
-apply (rule preal_mult_commute [THEN subst])
-apply (rule preal_mult_inverse)
-done
-
-
-text{*Theorems needing @{text Gleason9_34}*}
-
-lemma Rep_preal_self_subset: "Rep_preal (R) \<subseteq> Rep_preal(R + S)"
-proof
- fix r
- assume r: "r \<in> Rep_preal R"
- have rpos: "0<r" by (rule preal_imp_pos [OF Rep_preal r])
- from mem_Rep_preal_Ex
- obtain y where y: "y \<in> Rep_preal S" ..
- have ypos: "0<y" by (rule preal_imp_pos [OF Rep_preal y])
- have ry: "r+y \<in> Rep_preal(R + S)" using r y
- by (auto simp add: mem_Rep_preal_add_iff)
- show "r \<in> Rep_preal(R + S)" using r ypos rpos
- by (simp add: preal_downwards_closed [OF Rep_preal ry])
-qed
-
-lemma Rep_preal_sum_not_subset: "~ Rep_preal (R + S) \<subseteq> Rep_preal(R)"
-proof -
- from mem_Rep_preal_Ex
- obtain y where y: "y \<in> Rep_preal S" ..
- have ypos: "0<y" by (rule preal_imp_pos [OF Rep_preal y])
- from Gleason9_34 [OF Rep_preal ypos]
- obtain r where r: "r \<in> Rep_preal R" and notin: "r + y \<notin> Rep_preal R" ..
- have "r + y \<in> Rep_preal (R + S)" using r y
- by (auto simp add: mem_Rep_preal_add_iff)
- thus ?thesis using notin by blast
-qed
-
-lemma Rep_preal_sum_not_eq: "Rep_preal (R + S) \<noteq> Rep_preal(R)"
-by (insert Rep_preal_sum_not_subset, blast)
-
-text{*at last, Gleason prop. 9-3.5(iii) page 123*}
-lemma preal_self_less_add_left: "(R::preal) < R + S"
-apply (unfold preal_less_def less_le)
-apply (simp add: Rep_preal_self_subset Rep_preal_sum_not_eq [THEN not_sym])
-done
-
-
-subsection{*Subtraction for Positive Reals*}
-
-text{*Gleason prop. 9-3.5(iv), page 123: proving @{prop "A < B ==> \<exists>D. A + D =
-B"}. We define the claimed @{term D} and show that it is a positive real*}
-
-text{*Part 1 of Dedekind sections definition*}
-lemma diff_set_not_empty:
- "R < S ==> {} \<subset> diff_set (Rep_preal S) (Rep_preal R)"
-apply (auto simp add: preal_less_def diff_set_def elim!: equalityE)
-apply (frule_tac x1 = S in Rep_preal [THEN preal_exists_greater])
-apply (drule preal_imp_pos [OF Rep_preal], clarify)
-apply (cut_tac a=x and b=u in add_eq_exists, force)
-done
-
-text{*Part 2 of Dedekind sections definition*}
-lemma diff_set_nonempty:
- "\<exists>q. 0 < q & q \<notin> diff_set (Rep_preal S) (Rep_preal R)"
-apply (cut_tac X = S in Rep_preal_exists_bound)
-apply (erule exE)
-apply (rule_tac x = x in exI, auto)
-apply (simp add: diff_set_def)
-apply (auto dest: Rep_preal [THEN preal_downwards_closed])
-done
-
-lemma diff_set_not_rat_set:
- "diff_set (Rep_preal S) (Rep_preal R) < {r. 0 < r}" (is "?lhs < ?rhs")
-proof
- show "?lhs \<subseteq> ?rhs" by (auto simp add: diff_set_def)
- show "?lhs \<noteq> ?rhs" using diff_set_nonempty by blast
-qed
-
-text{*Part 3 of Dedekind sections definition*}
-lemma diff_set_lemma3:
- "[|R < S; u \<in> diff_set (Rep_preal S) (Rep_preal R); 0 < z; z < u|]
- ==> z \<in> diff_set (Rep_preal S) (Rep_preal R)"
-apply (auto simp add: diff_set_def)
-apply (rule_tac x=x in exI)
-apply (drule Rep_preal [THEN preal_downwards_closed], auto)
-done
-
-text{*Part 4 of Dedekind sections definition*}
-lemma diff_set_lemma4:
- "[|R < S; y \<in> diff_set (Rep_preal S) (Rep_preal R)|]
- ==> \<exists>u \<in> diff_set (Rep_preal S) (Rep_preal R). y < u"
-apply (auto simp add: diff_set_def)
-apply (drule Rep_preal [THEN preal_exists_greater], clarify)
-apply (cut_tac a="x+y" and b=u in add_eq_exists, clarify)
-apply (rule_tac x="y+xa" in exI)
-apply (auto simp add: add_ac)
-done
-
-lemma mem_diff_set:
- "R < S ==> diff_set (Rep_preal S) (Rep_preal R) \<in> preal"
-apply (unfold preal_def cut_def)
-apply (blast intro!: diff_set_not_empty diff_set_not_rat_set
- diff_set_lemma3 diff_set_lemma4)
-done
-
-lemma mem_Rep_preal_diff_iff:
- "R < S ==>
- (z \<in> Rep_preal(S-R)) =
- (\<exists>x. 0 < x & 0 < z & x \<notin> Rep_preal R & x + z \<in> Rep_preal S)"
-apply (simp add: preal_diff_def mem_diff_set Rep_preal)
-apply (force simp add: diff_set_def)
-done
-
-
-text{*proving that @{term "R + D \<le> S"}*}
-
-lemma less_add_left_lemma:
- assumes Rless: "R < S"
- and a: "a \<in> Rep_preal R"
- and cb: "c + b \<in> Rep_preal S"
- and "c \<notin> Rep_preal R"
- and "0 < b"
- and "0 < c"
- shows "a + b \<in> Rep_preal S"
-proof -
- have "0<a" by (rule preal_imp_pos [OF Rep_preal a])
- moreover
- have "a < c" using prems
- by (blast intro: not_in_Rep_preal_ub )
- ultimately show ?thesis using prems
- by (simp add: preal_downwards_closed [OF Rep_preal cb])
-qed
-
-lemma less_add_left_le1:
- "R < (S::preal) ==> R + (S-R) \<le> S"
-apply (auto simp add: Bex_def preal_le_def mem_Rep_preal_add_iff
- mem_Rep_preal_diff_iff)
-apply (blast intro: less_add_left_lemma)
-done
-
-subsection{*proving that @{term "S \<le> R + D"} --- trickier*}
-
-lemma lemma_sum_mem_Rep_preal_ex:
- "x \<in> Rep_preal S ==> \<exists>e. 0 < e & x + e \<in> Rep_preal S"
-apply (drule Rep_preal [THEN preal_exists_greater], clarify)
-apply (cut_tac a=x and b=u in add_eq_exists, auto)
-done
-
-lemma less_add_left_lemma2:
- assumes Rless: "R < S"
- and x: "x \<in> Rep_preal S"
- and xnot: "x \<notin> Rep_preal R"
- shows "\<exists>u v z. 0 < v & 0 < z & u \<in> Rep_preal R & z \<notin> Rep_preal R &
- z + v \<in> Rep_preal S & x = u + v"
-proof -
- have xpos: "0<x" by (rule preal_imp_pos [OF Rep_preal x])
- from lemma_sum_mem_Rep_preal_ex [OF x]
- obtain e where epos: "0 < e" and xe: "x + e \<in> Rep_preal S" by blast
- from Gleason9_34 [OF Rep_preal epos]
- obtain r where r: "r \<in> Rep_preal R" and notin: "r + e \<notin> Rep_preal R" ..
- with x xnot xpos have rless: "r < x" by (blast intro: not_in_Rep_preal_ub)
- from add_eq_exists [of r x]
- obtain y where eq: "x = r+y" by auto
- show ?thesis
- proof (intro exI conjI)
- show "r \<in> Rep_preal R" by (rule r)
- show "r + e \<notin> Rep_preal R" by (rule notin)
- show "r + e + y \<in> Rep_preal S" using xe eq by (simp add: add_ac)
- show "x = r + y" by (simp add: eq)
- show "0 < r + e" using epos preal_imp_pos [OF Rep_preal r]
- by simp
- show "0 < y" using rless eq by arith
- qed
-qed
-
-lemma less_add_left_le2: "R < (S::preal) ==> S \<le> R + (S-R)"
-apply (auto simp add: preal_le_def)
-apply (case_tac "x \<in> Rep_preal R")
-apply (cut_tac Rep_preal_self_subset [of R], force)
-apply (auto simp add: Bex_def mem_Rep_preal_add_iff mem_Rep_preal_diff_iff)
-apply (blast dest: less_add_left_lemma2)
-done
-
-lemma less_add_left: "R < (S::preal) ==> R + (S-R) = S"
-by (blast intro: antisym [OF less_add_left_le1 less_add_left_le2])
-
-lemma less_add_left_Ex: "R < (S::preal) ==> \<exists>D. R + D = S"
-by (fast dest: less_add_left)
-
-lemma preal_add_less2_mono1: "R < (S::preal) ==> R + T < S + T"
-apply (auto dest!: less_add_left_Ex simp add: preal_add_assoc)
-apply (rule_tac y1 = D in preal_add_commute [THEN subst])
-apply (auto intro: preal_self_less_add_left simp add: preal_add_assoc [symmetric])
-done
-
-lemma preal_add_less2_mono2: "R < (S::preal) ==> T + R < T + S"
-by (auto intro: preal_add_less2_mono1 simp add: preal_add_commute [of T])
-
-lemma preal_add_right_less_cancel: "R + T < S + T ==> R < (S::preal)"
-apply (insert linorder_less_linear [of R S], auto)
-apply (drule_tac R = S and T = T in preal_add_less2_mono1)
-apply (blast dest: order_less_trans)
-done
-
-lemma preal_add_left_less_cancel: "T + R < T + S ==> R < (S::preal)"
-by (auto elim: preal_add_right_less_cancel simp add: preal_add_commute [of T])
-
-lemma preal_add_less_cancel_left: "(T + (R::preal) < T + S) = (R < S)"
-by (blast intro: preal_add_less2_mono2 preal_add_left_less_cancel)
-
-lemma preal_add_le_cancel_left: "(T + (R::preal) \<le> T + S) = (R \<le> S)"
-by (simp add: linorder_not_less [symmetric] preal_add_less_cancel_left)
-
-lemma preal_add_right_cancel: "(R::preal) + T = S + T ==> R = S"
-apply (insert linorder_less_linear [of R S], safe)
-apply (drule_tac [!] T = T in preal_add_less2_mono1, auto)
-done
-
-lemma preal_add_left_cancel: "C + A = C + B ==> A = (B::preal)"
-by (auto intro: preal_add_right_cancel simp add: preal_add_commute)
-
-instance preal :: linordered_cancel_ab_semigroup_add
-proof
- fix a b c :: preal
- show "a + b = a + c \<Longrightarrow> b = c" by (rule preal_add_left_cancel)
- show "a \<le> b \<Longrightarrow> c + a \<le> c + b" by (simp only: preal_add_le_cancel_left)
-qed
-
-
-subsection{*Completeness of type @{typ preal}*}
-
-text{*Prove that supremum is a cut*}
-
-text{*Part 1 of Dedekind sections definition*}
-
-lemma preal_sup_set_not_empty:
- "P \<noteq> {} ==> {} \<subset> (\<Union>X \<in> P. Rep_preal(X))"
-apply auto
-apply (cut_tac X = x in mem_Rep_preal_Ex, auto)
-done
-
-
-text{*Part 2 of Dedekind sections definition*}
-
-lemma preal_sup_not_exists:
- "\<forall>X \<in> P. X \<le> Y ==> \<exists>q. 0 < q & q \<notin> (\<Union>X \<in> P. Rep_preal(X))"
-apply (cut_tac X = Y in Rep_preal_exists_bound)
-apply (auto simp add: preal_le_def)
-done
-
-lemma preal_sup_set_not_rat_set:
- "\<forall>X \<in> P. X \<le> Y ==> (\<Union>X \<in> P. Rep_preal(X)) < {r. 0 < r}"
-apply (drule preal_sup_not_exists)
-apply (blast intro: preal_imp_pos [OF Rep_preal])
-done
-
-text{*Part 3 of Dedekind sections definition*}
-lemma preal_sup_set_lemma3:
- "[|P \<noteq> {}; \<forall>X \<in> P. X \<le> Y; u \<in> (\<Union>X \<in> P. Rep_preal(X)); 0 < z; z < u|]
- ==> z \<in> (\<Union>X \<in> P. Rep_preal(X))"
-by (auto elim: Rep_preal [THEN preal_downwards_closed])
-
-text{*Part 4 of Dedekind sections definition*}
-lemma preal_sup_set_lemma4:
- "[|P \<noteq> {}; \<forall>X \<in> P. X \<le> Y; y \<in> (\<Union>X \<in> P. Rep_preal(X)) |]
- ==> \<exists>u \<in> (\<Union>X \<in> P. Rep_preal(X)). y < u"
-by (blast dest: Rep_preal [THEN preal_exists_greater])
-
-lemma preal_sup:
- "[|P \<noteq> {}; \<forall>X \<in> P. X \<le> Y|] ==> (\<Union>X \<in> P. Rep_preal(X)) \<in> preal"
-apply (unfold preal_def cut_def)
-apply (blast intro!: preal_sup_set_not_empty preal_sup_set_not_rat_set
- preal_sup_set_lemma3 preal_sup_set_lemma4)
-done
-
-lemma preal_psup_le:
- "[| \<forall>X \<in> P. X \<le> Y; x \<in> P |] ==> x \<le> psup P"
-apply (simp (no_asm_simp) add: preal_le_def)
-apply (subgoal_tac "P \<noteq> {}")
-apply (auto simp add: psup_def preal_sup)
-done
-
-lemma psup_le_ub: "[| P \<noteq> {}; \<forall>X \<in> P. X \<le> Y |] ==> psup P \<le> Y"
-apply (simp (no_asm_simp) add: preal_le_def)
-apply (simp add: psup_def preal_sup)
-apply (auto simp add: preal_le_def)
-done
-
-text{*Supremum property*}
-lemma preal_complete:
- "[| P \<noteq> {}; \<forall>X \<in> P. X \<le> Y |] ==> (\<exists>X \<in> P. Z < X) = (Z < psup P)"
-apply (simp add: preal_less_def psup_def preal_sup)
-apply (auto simp add: preal_le_def)
-apply (rename_tac U)
-apply (cut_tac x = U and y = Z in linorder_less_linear)
-apply (auto simp add: preal_less_def)
-done
-
-end
--- a/src/HOL/RComplete.thy Tue May 11 08:29:42 2010 +0200
+++ b/src/HOL/RComplete.thy Tue May 11 08:30:02 2010 +0200
@@ -30,92 +30,27 @@
FIXME: Can the premise be weakened to @{text "\<forall>x \<in> P. x\<le> y"}?
*}
+text {* Only used in HOL/Import/HOL4Compat.thy; delete? *}
+
lemma posreal_complete:
assumes positive_P: "\<forall>x \<in> P. (0::real) < x"
and not_empty_P: "\<exists>x. x \<in> P"
and upper_bound_Ex: "\<exists>y. \<forall>x \<in> P. x<y"
shows "\<exists>S. \<forall>y. (\<exists>x \<in> P. y < x) = (y < S)"
-proof (rule exI, rule allI)
- fix y
- let ?pP = "{w. real_of_preal w \<in> P}"
-
- show "(\<exists>x\<in>P. y < x) = (y < real_of_preal (psup ?pP))"
- proof (cases "0 < y")
- assume neg_y: "\<not> 0 < y"
- show ?thesis
- proof
- assume "\<exists>x\<in>P. y < x"
- have "\<forall>x. y < real_of_preal x"
- using neg_y by (rule real_less_all_real2)
- thus "y < real_of_preal (psup ?pP)" ..
- next
- assume "y < real_of_preal (psup ?pP)"
- obtain "x" where x_in_P: "x \<in> P" using not_empty_P ..
- hence "0 < x" using positive_P by simp
- hence "y < x" using neg_y by simp
- thus "\<exists>x \<in> P. y < x" using x_in_P ..
- qed
- next
- assume pos_y: "0 < y"
-
- then obtain py where y_is_py: "y = real_of_preal py"
- by (auto simp add: real_gt_zero_preal_Ex)
-
- obtain a where "a \<in> P" using not_empty_P ..
- with positive_P have a_pos: "0 < a" ..
- then obtain pa where "a = real_of_preal pa"
- by (auto simp add: real_gt_zero_preal_Ex)
- hence "pa \<in> ?pP" using `a \<in> P` by auto
- hence pP_not_empty: "?pP \<noteq> {}" by auto
-
- obtain sup where sup: "\<forall>x \<in> P. x < sup"
- using upper_bound_Ex ..
- from this and `a \<in> P` have "a < sup" ..
- hence "0 < sup" using a_pos by arith
- then obtain possup where "sup = real_of_preal possup"
- by (auto simp add: real_gt_zero_preal_Ex)
- hence "\<forall>X \<in> ?pP. X \<le> possup"
- using sup by (auto simp add: real_of_preal_lessI)
- with pP_not_empty have psup: "\<And>Z. (\<exists>X \<in> ?pP. Z < X) = (Z < psup ?pP)"
- by (rule preal_complete)
-
- show ?thesis
- proof
- assume "\<exists>x \<in> P. y < x"
- then obtain x where x_in_P: "x \<in> P" and y_less_x: "y < x" ..
- hence "0 < x" using pos_y by arith
- then obtain px where x_is_px: "x = real_of_preal px"
- by (auto simp add: real_gt_zero_preal_Ex)
-
- have py_less_X: "\<exists>X \<in> ?pP. py < X"
- proof
- show "py < px" using y_is_py and x_is_px and y_less_x
- by (simp add: real_of_preal_lessI)
- show "px \<in> ?pP" using x_in_P and x_is_px by simp
- qed
-
- have "(\<exists>X \<in> ?pP. py < X) ==> (py < psup ?pP)"
- using psup by simp
- hence "py < psup ?pP" using py_less_X by simp
- thus "y < real_of_preal (psup {w. real_of_preal w \<in> P})"
- using y_is_py and pos_y by (simp add: real_of_preal_lessI)
- next
- assume y_less_psup: "y < real_of_preal (psup ?pP)"
-
- hence "py < psup ?pP" using y_is_py
- by (simp add: real_of_preal_lessI)
- then obtain "X" where py_less_X: "py < X" and X_in_pP: "X \<in> ?pP"
- using psup by auto
- then obtain x where x_is_X: "x = real_of_preal X"
- by (simp add: real_gt_zero_preal_Ex)
- hence "y < x" using py_less_X and y_is_py
- by (simp add: real_of_preal_lessI)
-
- moreover have "x \<in> P" using x_is_X and X_in_pP by simp
-
- ultimately show "\<exists> x \<in> P. y < x" ..
- qed
+proof -
+ from upper_bound_Ex have "\<exists>z. \<forall>x\<in>P. x \<le> z"
+ by (auto intro: less_imp_le)
+ from complete_real [OF not_empty_P this] obtain S
+ where S1: "\<And>x. x \<in> P \<Longrightarrow> x \<le> S" and S2: "\<And>z. \<forall>x\<in>P. x \<le> z \<Longrightarrow> S \<le> z" by fast
+ have "\<forall>y. (\<exists>x \<in> P. y < x) = (y < S)"
+ proof
+ fix y show "(\<exists>x\<in>P. y < x) = (y < S)"
+ apply (cases "\<exists>x\<in>P. y < x", simp_all)
+ apply (clarify, drule S1, simp)
+ apply (simp add: not_less S2)
+ done
qed
+ thus ?thesis ..
qed
text {*
@@ -130,89 +65,6 @@
text {*
- \medskip Completeness theorem for the positive reals (again).
-*}
-
-lemma posreals_complete:
- assumes positive_S: "\<forall>x \<in> S. 0 < x"
- and not_empty_S: "\<exists>x. x \<in> S"
- and upper_bound_Ex: "\<exists>u. isUb (UNIV::real set) S u"
- shows "\<exists>t. isLub (UNIV::real set) S t"
-proof
- let ?pS = "{w. real_of_preal w \<in> S}"
-
- obtain u where "isUb UNIV S u" using upper_bound_Ex ..
- hence sup: "\<forall>x \<in> S. x \<le> u" by (simp add: isUb_def setle_def)
-
- obtain x where x_in_S: "x \<in> S" using not_empty_S ..
- hence x_gt_zero: "0 < x" using positive_S by simp
- have "x \<le> u" using sup and x_in_S ..
- hence "0 < u" using x_gt_zero by arith
-
- then obtain pu where u_is_pu: "u = real_of_preal pu"
- by (auto simp add: real_gt_zero_preal_Ex)
-
- have pS_less_pu: "\<forall>pa \<in> ?pS. pa \<le> pu"
- proof
- fix pa
- assume "pa \<in> ?pS"
- then obtain a where "a \<in> S" and "a = real_of_preal pa"
- by simp
- moreover hence "a \<le> u" using sup by simp
- ultimately show "pa \<le> pu"
- using sup and u_is_pu by (simp add: real_of_preal_le_iff)
- qed
-
- have "\<forall>y \<in> S. y \<le> real_of_preal (psup ?pS)"
- proof
- fix y
- assume y_in_S: "y \<in> S"
- hence "0 < y" using positive_S by simp
- then obtain py where y_is_py: "y = real_of_preal py"
- by (auto simp add: real_gt_zero_preal_Ex)
- hence py_in_pS: "py \<in> ?pS" using y_in_S by simp
- with pS_less_pu have "py \<le> psup ?pS"
- by (rule preal_psup_le)
- thus "y \<le> real_of_preal (psup ?pS)"
- using y_is_py by (simp add: real_of_preal_le_iff)
- qed
-
- moreover {
- fix x
- assume x_ub_S: "\<forall>y\<in>S. y \<le> x"
- have "real_of_preal (psup ?pS) \<le> x"
- proof -
- obtain "s" where s_in_S: "s \<in> S" using not_empty_S ..
- hence s_pos: "0 < s" using positive_S by simp
-
- hence "\<exists> ps. s = real_of_preal ps" by (simp add: real_gt_zero_preal_Ex)
- then obtain "ps" where s_is_ps: "s = real_of_preal ps" ..
- hence ps_in_pS: "ps \<in> {w. real_of_preal w \<in> S}" using s_in_S by simp
-
- from x_ub_S have "s \<le> x" using s_in_S ..
- hence "0 < x" using s_pos by simp
- hence "\<exists> px. x = real_of_preal px" by (simp add: real_gt_zero_preal_Ex)
- then obtain "px" where x_is_px: "x = real_of_preal px" ..
-
- have "\<forall>pe \<in> ?pS. pe \<le> px"
- proof
- fix pe
- assume "pe \<in> ?pS"
- hence "real_of_preal pe \<in> S" by simp
- hence "real_of_preal pe \<le> x" using x_ub_S by simp
- thus "pe \<le> px" using x_is_px by (simp add: real_of_preal_le_iff)
- qed
-
- moreover have "?pS \<noteq> {}" using ps_in_pS by auto
- ultimately have "(psup ?pS) \<le> px" by (simp add: psup_le_ub)
- thus "real_of_preal (psup ?pS) \<le> x" using x_is_px by (simp add: real_of_preal_le_iff)
- qed
- }
- ultimately show "isLub UNIV S (real_of_preal (psup ?pS))"
- by (simp add: isLub_def leastP_def isUb_def setle_def setge_def)
-qed
-
-text {*
\medskip reals Completeness (again!)
*}
@@ -221,87 +73,11 @@
and exists_Ub: "\<exists>Y. isUb (UNIV::real set) S Y"
shows "\<exists>t. isLub (UNIV :: real set) S t"
proof -
- obtain X where X_in_S: "X \<in> S" using notempty_S ..
- obtain Y where Y_isUb: "isUb (UNIV::real set) S Y"
- using exists_Ub ..
- let ?SHIFT = "{z. \<exists>x \<in>S. z = x + (-X) + 1} \<inter> {x. 0 < x}"
-
- {
- fix x
- assume "isUb (UNIV::real set) S x"
- hence S_le_x: "\<forall> y \<in> S. y <= x"
- by (simp add: isUb_def setle_def)
- {
- fix s
- assume "s \<in> {z. \<exists>x\<in>S. z = x + - X + 1}"
- hence "\<exists> x \<in> S. s = x + -X + 1" ..
- then obtain x1 where "x1 \<in> S" and "s = x1 + (-X) + 1" ..
- moreover hence "x1 \<le> x" using S_le_x by simp
- ultimately have "s \<le> x + - X + 1" by arith
- }
- then have "isUb (UNIV::real set) ?SHIFT (x + (-X) + 1)"
- by (auto simp add: isUb_def setle_def)
- } note S_Ub_is_SHIFT_Ub = this
-
- hence "isUb UNIV ?SHIFT (Y + (-X) + 1)" using Y_isUb by simp
- hence "\<exists>Z. isUb UNIV ?SHIFT Z" ..
- moreover have "\<forall>y \<in> ?SHIFT. 0 < y" by auto
- moreover have shifted_not_empty: "\<exists>u. u \<in> ?SHIFT"
- using X_in_S and Y_isUb by auto
- ultimately obtain t where t_is_Lub: "isLub UNIV ?SHIFT t"
- using posreals_complete [of ?SHIFT] by blast
-
- show ?thesis
- proof
- show "isLub UNIV S (t + X + (-1))"
- proof (rule isLubI2)
- {
- fix x
- assume "isUb (UNIV::real set) S x"
- hence "isUb (UNIV::real set) (?SHIFT) (x + (-X) + 1)"
- using S_Ub_is_SHIFT_Ub by simp
- hence "t \<le> (x + (-X) + 1)"
- using t_is_Lub by (simp add: isLub_le_isUb)
- hence "t + X + -1 \<le> x" by arith
- }
- then show "(t + X + -1) <=* Collect (isUb UNIV S)"
- by (simp add: setgeI)
- next
- show "isUb UNIV S (t + X + -1)"
- proof -
- {
- fix y
- assume y_in_S: "y \<in> S"
- have "y \<le> t + X + -1"
- proof -
- obtain "u" where u_in_shift: "u \<in> ?SHIFT" using shifted_not_empty ..
- hence "\<exists> x \<in> S. u = x + - X + 1" by simp
- then obtain "x" where x_and_u: "u = x + - X + 1" ..
- have u_le_t: "u \<le> t" using u_in_shift and t_is_Lub by (simp add: isLubD2)
-
- show ?thesis
- proof cases
- assume "y \<le> x"
- moreover have "x = u + X + - 1" using x_and_u by arith
- moreover have "u + X + - 1 \<le> t + X + -1" using u_le_t by arith
- ultimately show "y \<le> t + X + -1" by arith
- next
- assume "~(y \<le> x)"
- hence x_less_y: "x < y" by arith
-
- have "x + (-X) + 1 \<in> ?SHIFT" using x_and_u and u_in_shift by simp
- hence "0 < x + (-X) + 1" by simp
- hence "0 < y + (-X) + 1" using x_less_y by arith
- hence "y + (-X) + 1 \<in> ?SHIFT" using y_in_S by simp
- hence "y + (-X) + 1 \<le> t" using t_is_Lub by (simp add: isLubD2)
- thus ?thesis by simp
- qed
- qed
- }
- then show ?thesis by (simp add: isUb_def setle_def)
- qed
- qed
- qed
+ from assms have "\<exists>X. X \<in> S" and "\<exists>Y. \<forall>x\<in>S. x \<le> Y"
+ unfolding isUb_def setle_def by simp_all
+ from complete_real [OF this] show ?thesis
+ unfolding isLub_def leastP_def setle_def setge_def Ball_def
+ Collect_def mem_def isUb_def UNIV_def by simp
qed
text{*A version of the same theorem without all those predicates!*}
@@ -310,13 +86,7 @@
assumes "\<exists>y. y\<in>S" and "\<exists>(x::real). \<forall>y\<in>S. y \<le> x"
shows "\<exists>x. (\<forall>y\<in>S. y \<le> x) &
(\<forall>z. ((\<forall>y\<in>S. y \<le> z) --> x \<le> z))"
-proof -
- have "\<exists>x. isLub UNIV S x"
- by (rule reals_complete)
- (auto simp add: isLub_def isUb_def leastP_def setle_def setge_def prems)
- thus ?thesis
- by (metis UNIV_I isLub_isUb isLub_le_isUb isUbD isUb_def setleI)
-qed
+using assms by (rule complete_real)
subsection {* The Archimedean Property of the Reals *}
@@ -324,88 +94,11 @@
theorem reals_Archimedean:
assumes x_pos: "0 < x"
shows "\<exists>n. inverse (real (Suc n)) < x"
-proof (rule ccontr)
- assume contr: "\<not> ?thesis"
- have "\<forall>n. x * real (Suc n) <= 1"
- proof
- fix n
- from contr have "x \<le> inverse (real (Suc n))"
- by (simp add: linorder_not_less)
- hence "x \<le> (1 / (real (Suc n)))"
- by (simp add: inverse_eq_divide)
- moreover have "0 \<le> real (Suc n)"
- by (rule real_of_nat_ge_zero)
- ultimately have "x * real (Suc n) \<le> (1 / real (Suc n)) * real (Suc n)"
- by (rule mult_right_mono)
- thus "x * real (Suc n) \<le> 1" by simp
- qed
- hence "{z. \<exists>n. z = x * (real (Suc n))} *<= 1"
- by (simp add: setle_def, safe, rule spec)
- hence "isUb (UNIV::real set) {z. \<exists>n. z = x * (real (Suc n))} 1"
- by (simp add: isUbI)
- hence "\<exists>Y. isUb (UNIV::real set) {z. \<exists>n. z = x* (real (Suc n))} Y" ..
- moreover have "\<exists>X. X \<in> {z. \<exists>n. z = x* (real (Suc n))}" by auto
- ultimately have "\<exists>t. isLub UNIV {z. \<exists>n. z = x * real (Suc n)} t"
- by (simp add: reals_complete)
- then obtain "t" where
- t_is_Lub: "isLub UNIV {z. \<exists>n. z = x * real (Suc n)} t" ..
-
- have "\<forall>n::nat. x * real n \<le> t + - x"
- proof
- fix n
- from t_is_Lub have "x * real (Suc n) \<le> t"
- by (simp add: isLubD2)
- hence "x * (real n) + x \<le> t"
- by (simp add: right_distrib real_of_nat_Suc)
- thus "x * (real n) \<le> t + - x" by arith
- qed
-
- hence "\<forall>m. x * real (Suc m) \<le> t + - x" by simp
- hence "{z. \<exists>n. z = x * (real (Suc n))} *<= (t + - x)"
- by (auto simp add: setle_def)
- hence "isUb (UNIV::real set) {z. \<exists>n. z = x * (real (Suc n))} (t + (-x))"
- by (simp add: isUbI)
- hence "t \<le> t + - x"
- using t_is_Lub by (simp add: isLub_le_isUb)
- thus False using x_pos by arith
-qed
-
-text {*
- There must be other proofs, e.g. @{text "Suc"} of the largest
- integer in the cut representing @{text "x"}.
-*}
+ unfolding real_of_nat_def using x_pos
+ by (rule ex_inverse_of_nat_Suc_less)
lemma reals_Archimedean2: "\<exists>n. (x::real) < real (n::nat)"
-proof cases
- assume "x \<le> 0"
- hence "x < real (1::nat)" by simp
- thus ?thesis ..
-next
- assume "\<not> x \<le> 0"
- hence x_greater_zero: "0 < x" by simp
- hence "0 < inverse x" by simp
- then obtain n where "inverse (real (Suc n)) < inverse x"
- using reals_Archimedean by blast
- hence "inverse (real (Suc n)) * x < inverse x * x"
- using x_greater_zero by (rule mult_strict_right_mono)
- hence "inverse (real (Suc n)) * x < 1"
- using x_greater_zero by simp
- hence "real (Suc n) * (inverse (real (Suc n)) * x) < real (Suc n) * 1"
- by (rule mult_strict_left_mono) simp
- hence "x < real (Suc n)"
- by (simp add: algebra_simps)
- thus "\<exists>(n::nat). x < real n" ..
-qed
-
-instance real :: archimedean_field
-proof
- fix r :: real
- obtain n :: nat where "r < real n"
- using reals_Archimedean2 ..
- then have "r \<le> of_int (int n)"
- unfolding real_eq_of_nat by simp
- then show "\<exists>z. r \<le> of_int z" ..
-qed
+ unfolding real_of_nat_def by (rule ex_less_of_nat)
lemma reals_Archimedean3:
assumes x_greater_zero: "0 < x"
@@ -458,7 +151,7 @@
have "x = y-(y-x)" by simp
also from suc q have "\<dots> < real (Suc p)/real q - inverse (real q)" by arith
also have "\<dots> = real p / real q"
- by (simp only: inverse_eq_divide real_diff_def real_of_nat_Suc
+ by (simp only: inverse_eq_divide diff_def real_of_nat_Suc
minus_divide_left add_divide_distrib[THEN sym]) simp
finally have "x<r" by (unfold r_def)
have "p<Suc p" .. also note main[THEN sym]
--- a/src/HOL/RealDef.thy Tue May 11 08:29:42 2010 +0200
+++ b/src/HOL/RealDef.thy Tue May 11 08:30:02 2010 +0200
@@ -1,508 +1,1065 @@
(* Title : HOL/RealDef.thy
- Author : Jacques D. Fleuriot
- Copyright : 1998 University of Cambridge
+ Author : Jacques D. Fleuriot, 1998
Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4
Additional contributions by Jeremy Avigad
+ Construction of Cauchy Reals by Brian Huffman, 2010
*)
-header{*Defining the Reals from the Positive Reals*}
+header {* Development of the Reals using Cauchy Sequences *}
theory RealDef
-imports PReal
+imports Rat
begin
+text {*
+ This theory contains a formalization of the real numbers as
+ equivalence classes of Cauchy sequences of rationals. See
+ \url{HOL/ex/Dedekind_Real.thy} for an alternative construction
+ using Dedekind cuts.
+*}
+
+subsection {* Preliminary lemmas *}
+
+lemma add_diff_add:
+ fixes a b c d :: "'a::ab_group_add"
+ shows "(a + c) - (b + d) = (a - b) + (c - d)"
+ by simp
+
+lemma minus_diff_minus:
+ fixes a b :: "'a::ab_group_add"
+ shows "- a - - b = - (a - b)"
+ by simp
+
+lemma mult_diff_mult:
+ fixes x y a b :: "'a::ring"
+ shows "(x * y - a * b) = x * (y - b) + (x - a) * b"
+ by (simp add: algebra_simps)
+
+lemma inverse_diff_inverse:
+ fixes a b :: "'a::division_ring"
+ assumes "a \<noteq> 0" and "b \<noteq> 0"
+ shows "inverse a - inverse b = - (inverse a * (a - b) * inverse b)"
+ using assms by (simp add: algebra_simps)
+
+lemma obtain_pos_sum:
+ fixes r :: rat assumes r: "0 < r"
+ obtains s t where "0 < s" and "0 < t" and "r = s + t"
+proof
+ from r show "0 < r/2" by simp
+ from r show "0 < r/2" by simp
+ show "r = r/2 + r/2" by simp
+qed
+
+subsection {* Sequences that converge to zero *}
+
definition
- realrel :: "((preal * preal) * (preal * preal)) set" where
- [code del]: "realrel = {p. \<exists>x1 y1 x2 y2. p = ((x1,y1),(x2,y2)) & x1+y2 = x2+y1}"
+ vanishes :: "(nat \<Rightarrow> rat) \<Rightarrow> bool"
+where
+ "vanishes X = (\<forall>r>0. \<exists>k. \<forall>n\<ge>k. \<bar>X n\<bar> < r)"
+
+lemma vanishesI: "(\<And>r. 0 < r \<Longrightarrow> \<exists>k. \<forall>n\<ge>k. \<bar>X n\<bar> < r) \<Longrightarrow> vanishes X"
+ unfolding vanishes_def by simp
+
+lemma vanishesD: "\<lbrakk>vanishes X; 0 < r\<rbrakk> \<Longrightarrow> \<exists>k. \<forall>n\<ge>k. \<bar>X n\<bar> < r"
+ unfolding vanishes_def by simp
+
+lemma vanishes_const [simp]: "vanishes (\<lambda>n. c) \<longleftrightarrow> c = 0"
+ unfolding vanishes_def
+ apply (cases "c = 0", auto)
+ apply (rule exI [where x="\<bar>c\<bar>"], auto)
+ done
+
+lemma vanishes_minus: "vanishes X \<Longrightarrow> vanishes (\<lambda>n. - X n)"
+ unfolding vanishes_def by simp
-typedef (Real) real = "UNIV//realrel"
- by (auto simp add: quotient_def)
+lemma vanishes_add:
+ assumes X: "vanishes X" and Y: "vanishes Y"
+ shows "vanishes (\<lambda>n. X n + Y n)"
+proof (rule vanishesI)
+ fix r :: rat assume "0 < r"
+ then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
+ by (rule obtain_pos_sum)
+ obtain i where i: "\<forall>n\<ge>i. \<bar>X n\<bar> < s"
+ using vanishesD [OF X s] ..
+ obtain j where j: "\<forall>n\<ge>j. \<bar>Y n\<bar> < t"
+ using vanishesD [OF Y t] ..
+ have "\<forall>n\<ge>max i j. \<bar>X n + Y n\<bar> < r"
+ proof (clarsimp)
+ fix n assume n: "i \<le> n" "j \<le> n"
+ have "\<bar>X n + Y n\<bar> \<le> \<bar>X n\<bar> + \<bar>Y n\<bar>" by (rule abs_triangle_ineq)
+ also have "\<dots> < s + t" by (simp add: add_strict_mono i j n)
+ finally show "\<bar>X n + Y n\<bar> < r" unfolding r .
+ qed
+ thus "\<exists>k. \<forall>n\<ge>k. \<bar>X n + Y n\<bar> < r" ..
+qed
+
+lemma vanishes_diff:
+ assumes X: "vanishes X" and Y: "vanishes Y"
+ shows "vanishes (\<lambda>n. X n - Y n)"
+unfolding diff_minus by (intro vanishes_add vanishes_minus X Y)
+
+lemma vanishes_mult_bounded:
+ assumes X: "\<exists>a>0. \<forall>n. \<bar>X n\<bar> < a"
+ assumes Y: "vanishes (\<lambda>n. Y n)"
+ shows "vanishes (\<lambda>n. X n * Y n)"
+proof (rule vanishesI)
+ fix r :: rat assume r: "0 < r"
+ obtain a where a: "0 < a" "\<forall>n. \<bar>X n\<bar> < a"
+ using X by fast
+ obtain b where b: "0 < b" "r = a * b"
+ proof
+ show "0 < r / a" using r a by (simp add: divide_pos_pos)
+ show "r = a * (r / a)" using a by simp
+ qed
+ obtain k where k: "\<forall>n\<ge>k. \<bar>Y n\<bar> < b"
+ using vanishesD [OF Y b(1)] ..
+ have "\<forall>n\<ge>k. \<bar>X n * Y n\<bar> < r"
+ by (simp add: b(2) abs_mult mult_strict_mono' a k)
+ thus "\<exists>k. \<forall>n\<ge>k. \<bar>X n * Y n\<bar> < r" ..
+qed
+
+subsection {* Cauchy sequences *}
definition
- (** these don't use the overloaded "real" function: users don't see them **)
- real_of_preal :: "preal => real" where
- [code del]: "real_of_preal m = Abs_Real (realrel `` {(m + 1, 1)})"
+ cauchy :: "(nat \<Rightarrow> rat) set"
+where
+ "cauchy X \<longleftrightarrow> (\<forall>r>0. \<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m - X n\<bar> < r)"
+
+lemma cauchyI:
+ "(\<And>r. 0 < r \<Longrightarrow> \<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m - X n\<bar> < r) \<Longrightarrow> cauchy X"
+ unfolding cauchy_def by simp
+
+lemma cauchyD:
+ "\<lbrakk>cauchy X; 0 < r\<rbrakk> \<Longrightarrow> \<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m - X n\<bar> < r"
+ unfolding cauchy_def by simp
+
+lemma cauchy_const [simp]: "cauchy (\<lambda>n. x)"
+ unfolding cauchy_def by simp
+
+lemma cauchy_add [simp]:
+ assumes X: "cauchy X" and Y: "cauchy Y"
+ shows "cauchy (\<lambda>n. X n + Y n)"
+proof (rule cauchyI)
+ fix r :: rat assume "0 < r"
+ then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
+ by (rule obtain_pos_sum)
+ obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>X m - X n\<bar> < s"
+ using cauchyD [OF X s] ..
+ obtain j where j: "\<forall>m\<ge>j. \<forall>n\<ge>j. \<bar>Y m - Y n\<bar> < t"
+ using cauchyD [OF Y t] ..
+ have "\<forall>m\<ge>max i j. \<forall>n\<ge>max i j. \<bar>(X m + Y m) - (X n + Y n)\<bar> < r"
+ proof (clarsimp)
+ fix m n assume *: "i \<le> m" "j \<le> m" "i \<le> n" "j \<le> n"
+ have "\<bar>(X m + Y m) - (X n + Y n)\<bar> \<le> \<bar>X m - X n\<bar> + \<bar>Y m - Y n\<bar>"
+ unfolding add_diff_add by (rule abs_triangle_ineq)
+ also have "\<dots> < s + t"
+ by (rule add_strict_mono, simp_all add: i j *)
+ finally show "\<bar>(X m + Y m) - (X n + Y n)\<bar> < r" unfolding r .
+ qed
+ thus "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>(X m + Y m) - (X n + Y n)\<bar> < r" ..
+qed
+
+lemma cauchy_minus [simp]:
+ assumes X: "cauchy X"
+ shows "cauchy (\<lambda>n. - X n)"
+using assms unfolding cauchy_def
+unfolding minus_diff_minus abs_minus_cancel .
+
+lemma cauchy_diff [simp]:
+ assumes X: "cauchy X" and Y: "cauchy Y"
+ shows "cauchy (\<lambda>n. X n - Y n)"
+using assms unfolding diff_minus by simp
+
+lemma cauchy_imp_bounded:
+ assumes "cauchy X" shows "\<exists>b>0. \<forall>n. \<bar>X n\<bar> < b"
+proof -
+ obtain k where k: "\<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m - X n\<bar> < 1"
+ using cauchyD [OF assms zero_less_one] ..
+ show "\<exists>b>0. \<forall>n. \<bar>X n\<bar> < b"
+ proof (intro exI conjI allI)
+ have "0 \<le> \<bar>X 0\<bar>" by simp
+ also have "\<bar>X 0\<bar> \<le> Max (abs ` X ` {..k})" by simp
+ finally have "0 \<le> Max (abs ` X ` {..k})" .
+ thus "0 < Max (abs ` X ` {..k}) + 1" by simp
+ next
+ fix n :: nat
+ show "\<bar>X n\<bar> < Max (abs ` X ` {..k}) + 1"
+ proof (rule linorder_le_cases)
+ assume "n \<le> k"
+ hence "\<bar>X n\<bar> \<le> Max (abs ` X ` {..k})" by simp
+ thus "\<bar>X n\<bar> < Max (abs ` X ` {..k}) + 1" by simp
+ next
+ assume "k \<le> n"
+ have "\<bar>X n\<bar> = \<bar>X k + (X n - X k)\<bar>" by simp
+ also have "\<bar>X k + (X n - X k)\<bar> \<le> \<bar>X k\<bar> + \<bar>X n - X k\<bar>"
+ by (rule abs_triangle_ineq)
+ also have "\<dots> < Max (abs ` X ` {..k}) + 1"
+ by (rule add_le_less_mono, simp, simp add: k `k \<le> n`)
+ finally show "\<bar>X n\<bar> < Max (abs ` X ` {..k}) + 1" .
+ qed
+ qed
+qed
+
+lemma cauchy_mult [simp]:
+ assumes X: "cauchy X" and Y: "cauchy Y"
+ shows "cauchy (\<lambda>n. X n * Y n)"
+proof (rule cauchyI)
+ fix r :: rat assume "0 < r"
+ then obtain u v where u: "0 < u" and v: "0 < v" and "r = u + v"
+ by (rule obtain_pos_sum)
+ obtain a where a: "0 < a" "\<forall>n. \<bar>X n\<bar> < a"
+ using cauchy_imp_bounded [OF X] by fast
+ obtain b where b: "0 < b" "\<forall>n. \<bar>Y n\<bar> < b"
+ using cauchy_imp_bounded [OF Y] by fast
+ obtain s t where s: "0 < s" and t: "0 < t" and r: "r = a * t + s * b"
+ proof
+ show "0 < v/b" using v b(1) by (rule divide_pos_pos)
+ show "0 < u/a" using u a(1) by (rule divide_pos_pos)
+ show "r = a * (u/a) + (v/b) * b"
+ using a(1) b(1) `r = u + v` by simp
+ qed
+ obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>X m - X n\<bar> < s"
+ using cauchyD [OF X s] ..
+ obtain j where j: "\<forall>m\<ge>j. \<forall>n\<ge>j. \<bar>Y m - Y n\<bar> < t"
+ using cauchyD [OF Y t] ..
+ have "\<forall>m\<ge>max i j. \<forall>n\<ge>max i j. \<bar>X m * Y m - X n * Y n\<bar> < r"
+ proof (clarsimp)
+ fix m n assume *: "i \<le> m" "j \<le> m" "i \<le> n" "j \<le> n"
+ have "\<bar>X m * Y m - X n * Y n\<bar> = \<bar>X m * (Y m - Y n) + (X m - X n) * Y n\<bar>"
+ unfolding mult_diff_mult ..
+ also have "\<dots> \<le> \<bar>X m * (Y m - Y n)\<bar> + \<bar>(X m - X n) * Y n\<bar>"
+ by (rule abs_triangle_ineq)
+ also have "\<dots> = \<bar>X m\<bar> * \<bar>Y m - Y n\<bar> + \<bar>X m - X n\<bar> * \<bar>Y n\<bar>"
+ unfolding abs_mult ..
+ also have "\<dots> < a * t + s * b"
+ by (simp_all add: add_strict_mono mult_strict_mono' a b i j *)
+ finally show "\<bar>X m * Y m - X n * Y n\<bar> < r" unfolding r .
+ qed
+ thus "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>X m * Y m - X n * Y n\<bar> < r" ..
+qed
+
+lemma cauchy_not_vanishes_cases:
+ assumes X: "cauchy X"
+ assumes nz: "\<not> vanishes X"
+ shows "\<exists>b>0. \<exists>k. (\<forall>n\<ge>k. b < - X n) \<or> (\<forall>n\<ge>k. b < X n)"
+proof -
+ obtain r where "0 < r" and r: "\<forall>k. \<exists>n\<ge>k. r \<le> \<bar>X n\<bar>"
+ using nz unfolding vanishes_def by (auto simp add: not_less)
+ obtain s t where s: "0 < s" and t: "0 < t" and "r = s + t"
+ using `0 < r` by (rule obtain_pos_sum)
+ obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>X m - X n\<bar> < s"
+ using cauchyD [OF X s] ..
+ obtain k where "i \<le> k" and "r \<le> \<bar>X k\<bar>"
+ using r by fast
+ have k: "\<forall>n\<ge>k. \<bar>X n - X k\<bar> < s"
+ using i `i \<le> k` by auto
+ have "X k \<le> - r \<or> r \<le> X k"
+ using `r \<le> \<bar>X k\<bar>` by auto
+ hence "(\<forall>n\<ge>k. t < - X n) \<or> (\<forall>n\<ge>k. t < X n)"
+ unfolding `r = s + t` using k by auto
+ hence "\<exists>k. (\<forall>n\<ge>k. t < - X n) \<or> (\<forall>n\<ge>k. t < X n)" ..
+ thus "\<exists>t>0. \<exists>k. (\<forall>n\<ge>k. t < - X n) \<or> (\<forall>n\<ge>k. t < X n)"
+ using t by auto
+qed
+
+lemma cauchy_not_vanishes:
+ assumes X: "cauchy X"
+ assumes nz: "\<not> vanishes X"
+ shows "\<exists>b>0. \<exists>k. \<forall>n\<ge>k. b < \<bar>X n\<bar>"
+using cauchy_not_vanishes_cases [OF assms]
+by clarify (rule exI, erule conjI, rule_tac x=k in exI, auto)
+
+lemma cauchy_inverse [simp]:
+ assumes X: "cauchy X"
+ assumes nz: "\<not> vanishes X"
+ shows "cauchy (\<lambda>n. inverse (X n))"
+proof (rule cauchyI)
+ fix r :: rat assume "0 < r"
+ obtain b i where b: "0 < b" and i: "\<forall>n\<ge>i. b < \<bar>X n\<bar>"
+ using cauchy_not_vanishes [OF X nz] by fast
+ from b i have nz: "\<forall>n\<ge>i. X n \<noteq> 0" by auto
+ obtain s where s: "0 < s" and r: "r = inverse b * s * inverse b"
+ proof
+ show "0 < b * r * b"
+ by (simp add: `0 < r` b mult_pos_pos)
+ show "r = inverse b * (b * r * b) * inverse b"
+ using b by simp
+ qed
+ obtain j where j: "\<forall>m\<ge>j. \<forall>n\<ge>j. \<bar>X m - X n\<bar> < s"
+ using cauchyD [OF X s] ..
+ have "\<forall>m\<ge>max i j. \<forall>n\<ge>max i j. \<bar>inverse (X m) - inverse (X n)\<bar> < r"
+ proof (clarsimp)
+ fix m n assume *: "i \<le> m" "j \<le> m" "i \<le> n" "j \<le> n"
+ have "\<bar>inverse (X m) - inverse (X n)\<bar> =
+ inverse \<bar>X m\<bar> * \<bar>X m - X n\<bar> * inverse \<bar>X n\<bar>"
+ by (simp add: inverse_diff_inverse nz * abs_mult)
+ also have "\<dots> < inverse b * s * inverse b"
+ by (simp add: mult_strict_mono less_imp_inverse_less
+ mult_pos_pos i j b * s)
+ finally show "\<bar>inverse (X m) - inverse (X n)\<bar> < r" unfolding r .
+ qed
+ thus "\<exists>k. \<forall>m\<ge>k. \<forall>n\<ge>k. \<bar>inverse (X m) - inverse (X n)\<bar> < r" ..
+qed
+
+subsection {* Equivalence relation on Cauchy sequences *}
+
+definition
+ realrel :: "((nat \<Rightarrow> rat) \<times> (nat \<Rightarrow> rat)) set"
+where
+ "realrel = {(X, Y). cauchy X \<and> cauchy Y \<and> vanishes (\<lambda>n. X n - Y n)}"
+
+lemma refl_realrel: "refl_on {X. cauchy X} realrel"
+ unfolding realrel_def by (rule refl_onI, clarsimp, simp)
+
+lemma sym_realrel: "sym realrel"
+ unfolding realrel_def
+ by (rule symI, clarify, drule vanishes_minus, simp)
-instantiation real :: "{zero, one, plus, minus, uminus, times, inverse, ord, abs, sgn}"
+lemma trans_realrel: "trans realrel"
+ unfolding realrel_def
+ apply (rule transI, clarify)
+ apply (drule (1) vanishes_add)
+ apply (simp add: algebra_simps)
+ done
+
+lemma equiv_realrel: "equiv {X. cauchy X} realrel"
+ using refl_realrel sym_realrel trans_realrel
+ by (rule equiv.intro)
+
+subsection {* The field of real numbers *}
+
+typedef (open) real = "{X. cauchy X} // realrel"
+ by (fast intro: quotientI cauchy_const)
+
+definition
+ Real :: "(nat \<Rightarrow> rat) \<Rightarrow> real"
+where
+ "Real X = Abs_real (realrel `` {X})"
+
+definition
+ real_case :: "((nat \<Rightarrow> rat) \<Rightarrow> 'a) \<Rightarrow> real \<Rightarrow> 'a"
+where
+ "real_case f x = (THE y. \<forall>X\<in>Rep_real x. y = f X)"
+
+lemma Real_induct [induct type: real]:
+ "(\<And>X. cauchy X \<Longrightarrow> P (Real X)) \<Longrightarrow> P x"
+ unfolding Real_def
+ apply (induct x)
+ apply (erule quotientE)
+ apply (simp)
+ done
+
+lemma real_case_1:
+ assumes f: "congruent realrel f"
+ assumes X: "cauchy X"
+ shows "real_case f (Real X) = f X"
+ unfolding real_case_def Real_def
+ apply (subst Abs_real_inverse)
+ apply (simp add: quotientI X)
+ apply (rule the_equality)
+ apply clarsimp
+ apply (erule congruent.congruent [OF f])
+ apply (erule bspec)
+ apply simp
+ apply (rule refl_onD [OF refl_realrel])
+ apply (simp add: X)
+ done
+
+lemma real_case_2:
+ assumes f: "congruent2 realrel realrel f"
+ assumes X: "cauchy X" and Y: "cauchy Y"
+ shows "real_case (\<lambda>X. real_case (\<lambda>Y. f X Y) (Real Y)) (Real X) = f X Y"
+ apply (subst real_case_1 [OF _ X])
+ apply (rule congruent.intro)
+ apply (subst real_case_1 [OF _ Y])
+ apply (rule congruent2_implies_congruent [OF equiv_realrel f])
+ apply (simp add: realrel_def)
+ apply (subst real_case_1 [OF _ Y])
+ apply (rule congruent2_implies_congruent [OF equiv_realrel f])
+ apply (simp add: realrel_def)
+ apply (erule congruent2.congruent2 [OF f])
+ apply (rule refl_onD [OF refl_realrel])
+ apply (simp add: Y)
+ apply (rule real_case_1 [OF _ Y])
+ apply (rule congruent2_implies_congruent [OF equiv_realrel f])
+ apply (simp add: X)
+ done
+
+lemma eq_Real:
+ "cauchy X \<Longrightarrow> cauchy Y \<Longrightarrow> Real X = Real Y \<longleftrightarrow> vanishes (\<lambda>n. X n - Y n)"
+ unfolding Real_def
+ apply (subst Abs_real_inject)
+ apply (simp add: quotientI)
+ apply (simp add: quotientI)
+ apply (simp add: eq_equiv_class_iff [OF equiv_realrel])
+ apply (simp add: realrel_def)
+ done
+
+lemma add_respects2_realrel:
+ "(\<lambda>X Y. Real (\<lambda>n. X n + Y n)) respects2 realrel"
+proof (rule congruent2_commuteI [OF equiv_realrel, unfolded mem_Collect_eq])
+ fix X Y show "Real (\<lambda>n. X n + Y n) = Real (\<lambda>n. Y n + X n)"
+ by (simp add: add_commute)
+next
+ fix X assume X: "cauchy X"
+ fix Y Z assume "(Y, Z) \<in> realrel"
+ hence Y: "cauchy Y" and Z: "cauchy Z" and YZ: "vanishes (\<lambda>n. Y n - Z n)"
+ unfolding realrel_def by simp_all
+ show "Real (\<lambda>n. X n + Y n) = Real (\<lambda>n. X n + Z n)"
+ proof (rule eq_Real [THEN iffD2])
+ show "cauchy (\<lambda>n. X n + Y n)" using X Y by (rule cauchy_add)
+ show "cauchy (\<lambda>n. X n + Z n)" using X Z by (rule cauchy_add)
+ show "vanishes (\<lambda>n. (X n + Y n) - (X n + Z n))"
+ unfolding add_diff_add using YZ by simp
+ qed
+qed
+
+lemma minus_respects_realrel:
+ "(\<lambda>X. Real (\<lambda>n. - X n)) respects realrel"
+proof (rule congruent.intro)
+ fix X Y assume "(X, Y) \<in> realrel"
+ hence X: "cauchy X" and Y: "cauchy Y" and XY: "vanishes (\<lambda>n. X n - Y n)"
+ unfolding realrel_def by simp_all
+ show "Real (\<lambda>n. - X n) = Real (\<lambda>n. - Y n)"
+ proof (rule eq_Real [THEN iffD2])
+ show "cauchy (\<lambda>n. - X n)" using X by (rule cauchy_minus)
+ show "cauchy (\<lambda>n. - Y n)" using Y by (rule cauchy_minus)
+ show "vanishes (\<lambda>n. (- X n) - (- Y n))"
+ unfolding minus_diff_minus using XY by (rule vanishes_minus)
+ qed
+qed
+
+lemma mult_respects2_realrel:
+ "(\<lambda>X Y. Real (\<lambda>n. X n * Y n)) respects2 realrel"
+proof (rule congruent2_commuteI [OF equiv_realrel, unfolded mem_Collect_eq])
+ fix X Y
+ show "Real (\<lambda>n. X n * Y n) = Real (\<lambda>n. Y n * X n)"
+ by (simp add: mult_commute)
+next
+ fix X assume X: "cauchy X"
+ fix Y Z assume "(Y, Z) \<in> realrel"
+ hence Y: "cauchy Y" and Z: "cauchy Z" and YZ: "vanishes (\<lambda>n. Y n - Z n)"
+ unfolding realrel_def by simp_all
+ show "Real (\<lambda>n. X n * Y n) = Real (\<lambda>n. X n * Z n)"
+ proof (rule eq_Real [THEN iffD2])
+ show "cauchy (\<lambda>n. X n * Y n)" using X Y by (rule cauchy_mult)
+ show "cauchy (\<lambda>n. X n * Z n)" using X Z by (rule cauchy_mult)
+ have "vanishes (\<lambda>n. X n * (Y n - Z n))"
+ by (intro vanishes_mult_bounded cauchy_imp_bounded X YZ)
+ thus "vanishes (\<lambda>n. X n * Y n - X n * Z n)"
+ by (simp add: right_diff_distrib)
+ qed
+qed
+
+lemma vanishes_diff_inverse:
+ assumes X: "cauchy X" "\<not> vanishes X"
+ assumes Y: "cauchy Y" "\<not> vanishes Y"
+ assumes XY: "vanishes (\<lambda>n. X n - Y n)"
+ shows "vanishes (\<lambda>n. inverse (X n) - inverse (Y n))"
+proof (rule vanishesI)
+ fix r :: rat assume r: "0 < r"
+ obtain a i where a: "0 < a" and i: "\<forall>n\<ge>i. a < \<bar>X n\<bar>"
+ using cauchy_not_vanishes [OF X] by fast
+ obtain b j where b: "0 < b" and j: "\<forall>n\<ge>j. b < \<bar>Y n\<bar>"
+ using cauchy_not_vanishes [OF Y] by fast
+ obtain s where s: "0 < s" and "inverse a * s * inverse b = r"
+ proof
+ show "0 < a * r * b"
+ using a r b by (simp add: mult_pos_pos)
+ show "inverse a * (a * r * b) * inverse b = r"
+ using a r b by simp
+ qed
+ obtain k where k: "\<forall>n\<ge>k. \<bar>X n - Y n\<bar> < s"
+ using vanishesD [OF XY s] ..
+ have "\<forall>n\<ge>max (max i j) k. \<bar>inverse (X n) - inverse (Y n)\<bar> < r"
+ proof (clarsimp)
+ fix n assume n: "i \<le> n" "j \<le> n" "k \<le> n"
+ have "X n \<noteq> 0" and "Y n \<noteq> 0"
+ using i j a b n by auto
+ hence "\<bar>inverse (X n) - inverse (Y n)\<bar> =
+ inverse \<bar>X n\<bar> * \<bar>X n - Y n\<bar> * inverse \<bar>Y n\<bar>"
+ by (simp add: inverse_diff_inverse abs_mult)
+ also have "\<dots> < inverse a * s * inverse b"
+ apply (intro mult_strict_mono' less_imp_inverse_less)
+ apply (simp_all add: a b i j k n mult_nonneg_nonneg)
+ done
+ also note `inverse a * s * inverse b = r`
+ finally show "\<bar>inverse (X n) - inverse (Y n)\<bar> < r" .
+ qed
+ thus "\<exists>k. \<forall>n\<ge>k. \<bar>inverse (X n) - inverse (Y n)\<bar> < r" ..
+qed
+
+lemma inverse_respects_realrel:
+ "(\<lambda>X. if vanishes X then c else Real (\<lambda>n. inverse (X n))) respects realrel"
+ (is "?inv respects realrel")
+proof (rule congruent.intro)
+ fix X Y assume "(X, Y) \<in> realrel"
+ hence X: "cauchy X" and Y: "cauchy Y" and XY: "vanishes (\<lambda>n. X n - Y n)"
+ unfolding realrel_def by simp_all
+ have "vanishes X \<longleftrightarrow> vanishes Y"
+ proof
+ assume "vanishes X"
+ from vanishes_diff [OF this XY] show "vanishes Y" by simp
+ next
+ assume "vanishes Y"
+ from vanishes_add [OF this XY] show "vanishes X" by simp
+ qed
+ thus "?inv X = ?inv Y"
+ by (simp add: vanishes_diff_inverse eq_Real X Y XY)
+qed
+
+instantiation real :: field_inverse_zero
begin
definition
- real_zero_def [code del]: "0 = Abs_Real(realrel``{(1, 1)})"
+ "0 = Real (\<lambda>n. 0)"
definition
- real_one_def [code del]: "1 = Abs_Real(realrel``{(1 + 1, 1)})"
+ "1 = Real (\<lambda>n. 1)"
+
+definition
+ "x + y = real_case (\<lambda>X. real_case (\<lambda>Y. Real (\<lambda>n. X n + Y n)) y) x"
definition
- real_add_def [code del]: "z + w =
- contents (\<Union>(x,y) \<in> Rep_Real(z). \<Union>(u,v) \<in> Rep_Real(w).
- { Abs_Real(realrel``{(x+u, y+v)}) })"
+ "- x = real_case (\<lambda>X. Real (\<lambda>n. - X n)) x"
+
+definition
+ "x - y = (x::real) + - y"
definition
- real_minus_def [code del]: "- r = contents (\<Union>(x,y) \<in> Rep_Real(r). { Abs_Real(realrel``{(y,x)}) })"
+ "x * y = real_case (\<lambda>X. real_case (\<lambda>Y. Real (\<lambda>n. X n * Y n)) y) x"
definition
- real_diff_def [code del]: "r - (s::real) = r + - s"
+ "inverse =
+ real_case (\<lambda>X. if vanishes X then 0 else Real (\<lambda>n. inverse (X n)))"
definition
- real_mult_def [code del]:
- "z * w =
- contents (\<Union>(x,y) \<in> Rep_Real(z). \<Union>(u,v) \<in> Rep_Real(w).
- { Abs_Real(realrel``{(x*u + y*v, x*v + y*u)}) })"
+ "x / y = (x::real) * inverse y"
+
+lemma add_Real:
+ assumes X: "cauchy X" and Y: "cauchy Y"
+ shows "Real X + Real Y = Real (\<lambda>n. X n + Y n)"
+ unfolding plus_real_def
+ by (rule real_case_2 [OF add_respects2_realrel X Y])
+
+lemma minus_Real:
+ assumes X: "cauchy X"
+ shows "- Real X = Real (\<lambda>n. - X n)"
+ unfolding uminus_real_def
+ by (rule real_case_1 [OF minus_respects_realrel X])
-definition
- real_inverse_def [code del]: "inverse (R::real) = (THE S. (R = 0 & S = 0) | S * R = 1)"
+lemma diff_Real:
+ assumes X: "cauchy X" and Y: "cauchy Y"
+ shows "Real X - Real Y = Real (\<lambda>n. X n - Y n)"
+ unfolding minus_real_def diff_minus
+ by (simp add: minus_Real add_Real X Y)
-definition
- real_divide_def [code del]: "R / (S::real) = R * inverse S"
+lemma mult_Real:
+ assumes X: "cauchy X" and Y: "cauchy Y"
+ shows "Real X * Real Y = Real (\<lambda>n. X n * Y n)"
+ unfolding times_real_def
+ by (rule real_case_2 [OF mult_respects2_realrel X Y])
+
+lemma inverse_Real:
+ assumes X: "cauchy X"
+ shows "inverse (Real X) =
+ (if vanishes X then 0 else Real (\<lambda>n. inverse (X n)))"
+ unfolding inverse_real_def
+ by (rule real_case_1 [OF inverse_respects_realrel X])
-definition
- real_le_def [code del]: "z \<le> (w::real) \<longleftrightarrow>
- (\<exists>x y u v. x+v \<le> u+y & (x,y) \<in> Rep_Real z & (u,v) \<in> Rep_Real w)"
-
-definition
- real_less_def [code del]: "x < (y\<Colon>real) \<longleftrightarrow> x \<le> y \<and> x \<noteq> y"
-
-definition
- real_abs_def: "abs (r::real) = (if r < 0 then - r else r)"
-
-definition
- real_sgn_def: "sgn (x::real) = (if x=0 then 0 else if 0<x then 1 else - 1)"
-
-instance ..
+instance proof
+ fix a b c :: real
+ show "a + b = b + a"
+ by (induct a, induct b) (simp add: add_Real add_ac)
+ show "(a + b) + c = a + (b + c)"
+ by (induct a, induct b, induct c) (simp add: add_Real add_ac)
+ show "0 + a = a"
+ unfolding zero_real_def
+ by (induct a) (simp add: add_Real)
+ show "- a + a = 0"
+ unfolding zero_real_def
+ by (induct a) (simp add: minus_Real add_Real)
+ show "a - b = a + - b"
+ by (rule minus_real_def)
+ show "(a * b) * c = a * (b * c)"
+ by (induct a, induct b, induct c) (simp add: mult_Real mult_ac)
+ show "a * b = b * a"
+ by (induct a, induct b) (simp add: mult_Real mult_ac)
+ show "1 * a = a"
+ unfolding one_real_def
+ by (induct a) (simp add: mult_Real)
+ show "(a + b) * c = a * c + b * c"
+ by (induct a, induct b, induct c)
+ (simp add: mult_Real add_Real algebra_simps)
+ show "(0\<Colon>real) \<noteq> (1\<Colon>real)"
+ unfolding zero_real_def one_real_def
+ by (simp add: eq_Real)
+ show "a \<noteq> 0 \<Longrightarrow> inverse a * a = 1"
+ unfolding zero_real_def one_real_def
+ apply (induct a)
+ apply (simp add: eq_Real inverse_Real mult_Real)
+ apply (rule vanishesI)
+ apply (frule (1) cauchy_not_vanishes, clarify)
+ apply (rule_tac x=k in exI, clarify)
+ apply (drule_tac x=n in spec, simp)
+ done
+ show "a / b = a * inverse b"
+ by (rule divide_real_def)
+ show "inverse (0::real) = 0"
+ by (simp add: zero_real_def inverse_Real)
+qed
end
-subsection {* Equivalence relation over positive reals *}
+subsection {* Positive reals *}
-lemma preal_trans_lemma:
- assumes "x + y1 = x1 + y"
- and "x + y2 = x2 + y"
- shows "x1 + y2 = x2 + (y1::preal)"
-proof -
- have "(x1 + y2) + x = (x + y2) + x1" by (simp add: add_ac)
- also have "... = (x2 + y) + x1" by (simp add: prems)
- also have "... = x2 + (x1 + y)" by (simp add: add_ac)
- also have "... = x2 + (x + y1)" by (simp add: prems)
- also have "... = (x2 + y1) + x" by (simp add: add_ac)
- finally have "(x1 + y2) + x = (x2 + y1) + x" .
- thus ?thesis by (rule add_right_imp_eq)
-qed
-
-
-lemma realrel_iff [simp]: "(((x1,y1),(x2,y2)) \<in> realrel) = (x1 + y2 = x2 + y1)"
-by (simp add: realrel_def)
-
-lemma equiv_realrel: "equiv UNIV realrel"
-apply (auto simp add: equiv_def refl_on_def sym_def trans_def realrel_def)
-apply (blast dest: preal_trans_lemma)
-done
-
-text{*Reduces equality of equivalence classes to the @{term realrel} relation:
- @{term "(realrel `` {x} = realrel `` {y}) = ((x,y) \<in> realrel)"} *}
-lemmas equiv_realrel_iff =
- eq_equiv_class_iff [OF equiv_realrel UNIV_I UNIV_I]
+definition
+ positive :: "real \<Rightarrow> bool"
+where
+ "positive = real_case (\<lambda>X. \<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < X n)"
-declare equiv_realrel_iff [simp]
-
-
-lemma realrel_in_real [simp]: "realrel``{(x,y)}: Real"
-by (simp add: Real_def realrel_def quotient_def, blast)
-
-declare Abs_Real_inject [simp]
-declare Abs_Real_inverse [simp]
-
-
-text{*Case analysis on the representation of a real number as an equivalence
- class of pairs of positive reals.*}
-lemma eq_Abs_Real [case_names Abs_Real, cases type: real]:
- "(!!x y. z = Abs_Real(realrel``{(x,y)}) ==> P) ==> P"
-apply (rule Rep_Real [of z, unfolded Real_def, THEN quotientE])
-apply (drule arg_cong [where f=Abs_Real])
-apply (auto simp add: Rep_Real_inverse)
-done
-
-
-subsection {* Addition and Subtraction *}
-
-lemma real_add_congruent2_lemma:
- "[|a + ba = aa + b; ab + bc = ac + bb|]
- ==> a + ab + (ba + bc) = aa + ac + (b + (bb::preal))"
-apply (simp add: add_assoc)
-apply (rule add_left_commute [of ab, THEN ssubst])
-apply (simp add: add_assoc [symmetric])
-apply (simp add: add_ac)
+lemma bool_congruentI:
+ assumes sym: "sym r"
+ assumes P: "\<And>x y. (x, y) \<in> r \<Longrightarrow> P x \<Longrightarrow> P y"
+ shows "P respects r"
+apply (rule congruent.intro)
+apply (rule iffI)
+apply (erule (1) P)
+apply (erule (1) P [OF symD [OF sym]])
done
-lemma real_add:
- "Abs_Real (realrel``{(x,y)}) + Abs_Real (realrel``{(u,v)}) =
- Abs_Real (realrel``{(x+u, y+v)})"
-proof -
- have "(\<lambda>z w. (\<lambda>(x,y). (\<lambda>(u,v). {Abs_Real (realrel `` {(x+u, y+v)})}) w) z)
- respects2 realrel"
- by (simp add: congruent2_def, blast intro: real_add_congruent2_lemma)
- thus ?thesis
- by (simp add: real_add_def UN_UN_split_split_eq
- UN_equiv_class2 [OF equiv_realrel equiv_realrel])
-qed
-
-lemma real_minus: "- Abs_Real(realrel``{(x,y)}) = Abs_Real(realrel `` {(y,x)})"
-proof -
- have "(\<lambda>(x,y). {Abs_Real (realrel``{(y,x)})}) respects realrel"
- by (simp add: congruent_def add_commute)
- thus ?thesis
- by (simp add: real_minus_def UN_equiv_class [OF equiv_realrel])
-qed
-
-instance real :: ab_group_add
-proof
- fix x y z :: real
- show "(x + y) + z = x + (y + z)"
- by (cases x, cases y, cases z, simp add: real_add add_assoc)
- show "x + y = y + x"
- by (cases x, cases y, simp add: real_add add_commute)
- show "0 + x = x"
- by (cases x, simp add: real_add real_zero_def add_ac)
- show "- x + x = 0"
- by (cases x, simp add: real_minus real_add real_zero_def add_commute)
- show "x - y = x + - y"
- by (simp add: real_diff_def)
-qed
-
-
-subsection {* Multiplication *}
-
-lemma real_mult_congruent2_lemma:
- "!!(x1::preal). [| x1 + y2 = x2 + y1 |] ==>
- x * x1 + y * y1 + (x * y2 + y * x2) =
- x * x2 + y * y2 + (x * y1 + y * x1)"
-apply (simp add: add_left_commute add_assoc [symmetric])
-apply (simp add: add_assoc right_distrib [symmetric])
-apply (simp add: add_commute)
-done
-
-lemma real_mult_congruent2:
- "(%p1 p2.
- (%(x1,y1). (%(x2,y2).
- { Abs_Real (realrel``{(x1*x2 + y1*y2, x1*y2+y1*x2)}) }) p2) p1)
- respects2 realrel"
-apply (rule congruent2_commuteI [OF equiv_realrel], clarify)
-apply (simp add: mult_commute add_commute)
-apply (auto simp add: real_mult_congruent2_lemma)
-done
-
-lemma real_mult:
- "Abs_Real((realrel``{(x1,y1)})) * Abs_Real((realrel``{(x2,y2)})) =
- Abs_Real(realrel `` {(x1*x2+y1*y2,x1*y2+y1*x2)})"
-by (simp add: real_mult_def UN_UN_split_split_eq
- UN_equiv_class2 [OF equiv_realrel equiv_realrel real_mult_congruent2])
-
-lemma real_mult_commute: "(z::real) * w = w * z"
-by (cases z, cases w, simp add: real_mult add_ac mult_ac)
-
-lemma real_mult_assoc: "((z1::real) * z2) * z3 = z1 * (z2 * z3)"
-apply (cases z1, cases z2, cases z3)
-apply (simp add: real_mult algebra_simps)
-done
-
-lemma real_mult_1: "(1::real) * z = z"
-apply (cases z)
-apply (simp add: real_mult real_one_def algebra_simps)
-done
-
-lemma real_add_mult_distrib: "((z1::real) + z2) * w = (z1 * w) + (z2 * w)"
-apply (cases z1, cases z2, cases w)
-apply (simp add: real_add real_mult algebra_simps)
-done
-
-text{*one and zero are distinct*}
-lemma real_zero_not_eq_one: "0 \<noteq> (1::real)"
-proof -
- have "(1::preal) < 1 + 1"
- by (simp add: preal_self_less_add_left)
- thus ?thesis
- by (simp add: real_zero_def real_one_def)
+lemma positive_respects_realrel:
+ "(\<lambda>X. \<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < X n) respects realrel"
+proof (rule bool_congruentI)
+ show "sym realrel" by (rule sym_realrel)
+next
+ fix X Y assume "(X, Y) \<in> realrel"
+ hence XY: "vanishes (\<lambda>n. X n - Y n)"
+ unfolding realrel_def by simp_all
+ assume "\<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < X n"
+ then obtain r i where "0 < r" and i: "\<forall>n\<ge>i. r < X n"
+ by fast
+ obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
+ using `0 < r` by (rule obtain_pos_sum)
+ obtain j where j: "\<forall>n\<ge>j. \<bar>X n - Y n\<bar> < s"
+ using vanishesD [OF XY s] ..
+ have "\<forall>n\<ge>max i j. t < Y n"
+ proof (clarsimp)
+ fix n assume n: "i \<le> n" "j \<le> n"
+ have "\<bar>X n - Y n\<bar> < s" and "r < X n"
+ using i j n by simp_all
+ thus "t < Y n" unfolding r by simp
+ qed
+ thus "\<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < Y n" using t by fast
qed
-instance real :: comm_ring_1
-proof
- fix x y z :: real
- show "(x * y) * z = x * (y * z)" by (rule real_mult_assoc)
- show "x * y = y * x" by (rule real_mult_commute)
- show "1 * x = x" by (rule real_mult_1)
- show "(x + y) * z = x * z + y * z" by (rule real_add_mult_distrib)
- show "0 \<noteq> (1::real)" by (rule real_zero_not_eq_one)
-qed
+lemma positive_Real:
+ assumes X: "cauchy X"
+ shows "positive (Real X) \<longleftrightarrow> (\<exists>r>0. \<exists>k. \<forall>n\<ge>k. r < X n)"
+unfolding positive_def
+by (rule real_case_1 [OF positive_respects_realrel X])
-subsection {* Inverse and Division *}
-
-lemma real_zero_iff: "Abs_Real (realrel `` {(x, x)}) = 0"
-by (simp add: real_zero_def add_commute)
+lemma positive_zero: "\<not> positive 0"
+unfolding zero_real_def by (auto simp add: positive_Real)
-text{*Instead of using an existential quantifier and constructing the inverse
-within the proof, we could define the inverse explicitly.*}
-
-lemma real_mult_inverse_left_ex: "x \<noteq> 0 ==> \<exists>y. y*x = (1::real)"
-apply (simp add: real_zero_def real_one_def, cases x)
-apply (cut_tac x = xa and y = y in linorder_less_linear)
-apply (auto dest!: less_add_left_Ex simp add: real_zero_iff)
-apply (rule_tac
- x = "Abs_Real (realrel``{(1, inverse (D) + 1)})"
- in exI)
-apply (rule_tac [2]
- x = "Abs_Real (realrel``{(inverse (D) + 1, 1)})"
- in exI)
-apply (auto simp add: real_mult preal_mult_inverse_right algebra_simps)
+lemma positive_add:
+ "positive x \<Longrightarrow> positive y \<Longrightarrow> positive (x + y)"
+apply (induct x, induct y, rename_tac Y X)
+apply (simp add: add_Real positive_Real)
+apply (clarify, rename_tac a b i j)
+apply (rule_tac x="a + b" in exI, simp)
+apply (rule_tac x="max i j" in exI, clarsimp)
+apply (simp add: add_strict_mono)
done
-lemma real_mult_inverse_left: "x \<noteq> 0 ==> inverse(x)*x = (1::real)"
-apply (simp add: real_inverse_def)
-apply (drule real_mult_inverse_left_ex, safe)
-apply (rule theI, assumption, rename_tac z)
-apply (subgoal_tac "(z * x) * y = z * (x * y)")
-apply (simp add: mult_commute)
-apply (rule mult_assoc)
+lemma positive_mult:
+ "positive x \<Longrightarrow> positive y \<Longrightarrow> positive (x * y)"
+apply (induct x, induct y, rename_tac Y X)
+apply (simp add: mult_Real positive_Real)
+apply (clarify, rename_tac a b i j)
+apply (rule_tac x="a * b" in exI, simp add: mult_pos_pos)
+apply (rule_tac x="max i j" in exI, clarsimp)
+apply (rule mult_strict_mono, auto)
+done
+
+lemma positive_minus:
+ "\<not> positive x \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> positive (- x)"
+apply (induct x, rename_tac X)
+apply (simp add: zero_real_def eq_Real minus_Real positive_Real)
+apply (drule (1) cauchy_not_vanishes_cases, safe, fast, fast)
done
+instantiation real :: linordered_field_inverse_zero
+begin
-subsection{*The Real Numbers form a Field*}
+definition
+ "x < y \<longleftrightarrow> positive (y - x)"
-instance real :: field_inverse_zero
-proof
- fix x y z :: real
- show "x \<noteq> 0 ==> inverse x * x = 1" by (rule real_mult_inverse_left)
- show "x / y = x * inverse y" by (simp add: real_divide_def)
- show "inverse 0 = (0::real)" by (simp add: real_inverse_def)
-qed
+definition
+ "x \<le> (y::real) \<longleftrightarrow> x < y \<or> x = y"
-lemma INVERSE_ZERO: "inverse 0 = (0::real)"
- by (fact inverse_zero)
-
+definition
+ "abs (a::real) = (if a < 0 then - a else a)"
-subsection{*The @{text "\<le>"} Ordering*}
-
-lemma real_le_refl: "w \<le> (w::real)"
-by (cases w, force simp add: real_le_def)
+definition
+ "sgn (a::real) = (if a = 0 then 0 else if 0 < a then 1 else - 1)"
-text{*The arithmetic decision procedure is not set up for type preal.
- This lemma is currently unused, but it could simplify the proofs of the
- following two lemmas.*}
-lemma preal_eq_le_imp_le:
- assumes eq: "a+b = c+d" and le: "c \<le> a"
- shows "b \<le> (d::preal)"
-proof -
- have "c+d \<le> a+d" by (simp add: prems)
- hence "a+b \<le> a+d" by (simp add: prems)
- thus "b \<le> d" by simp
-qed
-
-lemma real_le_lemma:
- assumes l: "u1 + v2 \<le> u2 + v1"
- and "x1 + v1 = u1 + y1"
- and "x2 + v2 = u2 + y2"
- shows "x1 + y2 \<le> x2 + (y1::preal)"
-proof -
- have "(x1+v1) + (u2+y2) = (u1+y1) + (x2+v2)" by (simp add: prems)
- hence "(x1+y2) + (u2+v1) = (x2+y1) + (u1+v2)" by (simp add: add_ac)
- also have "... \<le> (x2+y1) + (u2+v1)" by (simp add: prems)
- finally show ?thesis by simp
+instance proof
+ fix a b c :: real
+ show "\<bar>a\<bar> = (if a < 0 then - a else a)"
+ by (rule abs_real_def)
+ show "a < b \<longleftrightarrow> a \<le> b \<and> \<not> b \<le> a"
+ unfolding less_eq_real_def less_real_def
+ by (auto, drule (1) positive_add, simp_all add: positive_zero)
+ show "a \<le> a"
+ unfolding less_eq_real_def by simp
+ show "a \<le> b \<Longrightarrow> b \<le> c \<Longrightarrow> a \<le> c"
+ unfolding less_eq_real_def less_real_def
+ by (auto, drule (1) positive_add, simp add: algebra_simps)
+ show "a \<le> b \<Longrightarrow> b \<le> a \<Longrightarrow> a = b"
+ unfolding less_eq_real_def less_real_def
+ by (auto, drule (1) positive_add, simp add: positive_zero)
+ show "a \<le> b \<Longrightarrow> c + a \<le> c + b"
+ unfolding less_eq_real_def less_real_def by auto
+ show "sgn a = (if a = 0 then 0 else if 0 < a then 1 else - 1)"
+ by (rule sgn_real_def)
+ show "a \<le> b \<or> b \<le> a"
+ unfolding less_eq_real_def less_real_def
+ by (auto dest!: positive_minus)
+ show "a < b \<Longrightarrow> 0 < c \<Longrightarrow> c * a < c * b"
+ unfolding less_real_def
+ by (drule (1) positive_mult, simp add: algebra_simps)
qed
-lemma real_le:
- "(Abs_Real(realrel``{(x1,y1)}) \<le> Abs_Real(realrel``{(x2,y2)})) =
- (x1 + y2 \<le> x2 + y1)"
-apply (simp add: real_le_def)
-apply (auto intro: real_le_lemma)
-done
-
-lemma real_le_antisym: "[| z \<le> w; w \<le> z |] ==> z = (w::real)"
-by (cases z, cases w, simp add: real_le)
-
-lemma real_trans_lemma:
- assumes "x + v \<le> u + y"
- and "u + v' \<le> u' + v"
- and "x2 + v2 = u2 + y2"
- shows "x + v' \<le> u' + (y::preal)"
-proof -
- have "(x+v') + (u+v) = (x+v) + (u+v')" by (simp add: add_ac)
- also have "... \<le> (u+y) + (u+v')" by (simp add: prems)
- also have "... \<le> (u+y) + (u'+v)" by (simp add: prems)
- also have "... = (u'+y) + (u+v)" by (simp add: add_ac)
- finally show ?thesis by simp
-qed
-
-lemma real_le_trans: "[| i \<le> j; j \<le> k |] ==> i \<le> (k::real)"
-apply (cases i, cases j, cases k)
-apply (simp add: real_le)
-apply (blast intro: real_trans_lemma)
-done
-
-instance real :: order
-proof
- fix u v :: real
- show "u < v \<longleftrightarrow> u \<le> v \<and> \<not> v \<le> u"
- by (auto simp add: real_less_def intro: real_le_antisym)
-qed (assumption | rule real_le_refl real_le_trans real_le_antisym)+
-
-(* Axiom 'linorder_linear' of class 'linorder': *)
-lemma real_le_linear: "(z::real) \<le> w | w \<le> z"
-apply (cases z, cases w)
-apply (auto simp add: real_le real_zero_def add_ac)
-done
-
-instance real :: linorder
- by (intro_classes, rule real_le_linear)
-
-
-lemma real_le_eq_diff: "(x \<le> y) = (x-y \<le> (0::real))"
-apply (cases x, cases y)
-apply (auto simp add: real_le real_zero_def real_diff_def real_add real_minus
- add_ac)
-apply (simp_all add: add_assoc [symmetric])
-done
-
-lemma real_add_left_mono:
- assumes le: "x \<le> y" shows "z + x \<le> z + (y::real)"
-proof -
- have "z + x - (z + y) = (z + -z) + (x - y)"
- by (simp add: algebra_simps)
- with le show ?thesis
- by (simp add: real_le_eq_diff[of x] real_le_eq_diff[of "z+x"] diff_minus)
-qed
-
-lemma real_sum_gt_zero_less: "(0 < S + (-W::real)) ==> (W < S)"
-by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S] diff_minus)
-
-lemma real_less_sum_gt_zero: "(W < S) ==> (0 < S + (-W::real))"
-by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S] diff_minus)
-
-lemma real_mult_order: "[| 0 < x; 0 < y |] ==> (0::real) < x * y"
-apply (cases x, cases y)
-apply (simp add: linorder_not_le [where 'a = real, symmetric]
- linorder_not_le [where 'a = preal]
- real_zero_def real_le real_mult)
- --{*Reduce to the (simpler) @{text "\<le>"} relation *}
-apply (auto dest!: less_add_left_Ex
- simp add: algebra_simps preal_self_less_add_left)
-done
-
-lemma real_mult_less_mono2: "[| (0::real) < z; x < y |] ==> z * x < z * y"
-apply (rule real_sum_gt_zero_less)
-apply (drule real_less_sum_gt_zero [of x y])
-apply (drule real_mult_order, assumption)
-apply (simp add: right_distrib)
-done
+end
instantiation real :: distrib_lattice
begin
definition
- "(inf \<Colon> real \<Rightarrow> real \<Rightarrow> real) = min"
+ "(inf :: real \<Rightarrow> real \<Rightarrow> real) = min"
definition
- "(sup \<Colon> real \<Rightarrow> real \<Rightarrow> real) = max"
+ "(sup :: real \<Rightarrow> real \<Rightarrow> real) = max"
+
+instance proof
+qed (auto simp add: inf_real_def sup_real_def min_max.sup_inf_distrib1)
+
+end
-instance
- by default (auto simp add: inf_real_def sup_real_def min_max.sup_inf_distrib1)
+instantiation real :: number_ring
+begin
+
+definition
+ "(number_of x :: real) = of_int x"
+
+instance proof
+qed (rule number_of_real_def)
end
+lemma of_nat_Real: "of_nat x = Real (\<lambda>n. of_nat x)"
+apply (induct x)
+apply (simp add: zero_real_def)
+apply (simp add: one_real_def add_Real)
+done
-subsection{*The Reals Form an Ordered Field*}
+lemma of_int_Real: "of_int x = Real (\<lambda>n. of_int x)"
+apply (cases x rule: int_diff_cases)
+apply (simp add: of_nat_Real diff_Real)
+done
-instance real :: linordered_field_inverse_zero
+lemma of_rat_Real: "of_rat x = Real (\<lambda>n. x)"
+apply (induct x)
+apply (simp add: Fract_of_int_quotient of_rat_divide)
+apply (simp add: of_int_Real divide_inverse)
+apply (simp add: inverse_Real mult_Real)
+done
+
+instance real :: archimedean_field
proof
- fix x y z :: real
- show "x \<le> y ==> z + x \<le> z + y" by (rule real_add_left_mono)
- show "x < y ==> 0 < z ==> z * x < z * y" by (rule real_mult_less_mono2)
- show "\<bar>x\<bar> = (if x < 0 then -x else x)" by (simp only: real_abs_def)
- show "sgn x = (if x=0 then 0 else if 0<x then 1 else - 1)"
- by (simp only: real_sgn_def)
+ fix x :: real
+ show "\<exists>z. x \<le> of_int z"
+ apply (induct x)
+ apply (frule cauchy_imp_bounded, clarify)
+ apply (rule_tac x="ceiling b + 1" in exI)
+ apply (rule less_imp_le)
+ apply (simp add: of_int_Real less_real_def diff_Real positive_Real)
+ apply (rule_tac x=1 in exI, simp add: algebra_simps)
+ apply (rule_tac x=0 in exI, clarsimp)
+ apply (rule le_less_trans [OF abs_ge_self])
+ apply (rule less_le_trans [OF _ le_of_int_ceiling])
+ apply simp
+ done
qed
-text{*The function @{term real_of_preal} requires many proofs, but it seems
-to be essential for proving completeness of the reals from that of the
-positive reals.*}
+subsection {* Completeness *}
-lemma real_of_preal_add:
- "real_of_preal ((x::preal) + y) = real_of_preal x + real_of_preal y"
-by (simp add: real_of_preal_def real_add algebra_simps)
+lemma not_positive_Real:
+ assumes X: "cauchy X"
+ shows "\<not> positive (Real X) \<longleftrightarrow> (\<forall>r>0. \<exists>k. \<forall>n\<ge>k. X n \<le> r)"
+unfolding positive_Real [OF X]
+apply (auto, unfold not_less)
+apply (erule obtain_pos_sum)
+apply (drule_tac x=s in spec, simp)
+apply (drule_tac r=t in cauchyD [OF X], clarify)
+apply (drule_tac x=k in spec, clarsimp)
+apply (rule_tac x=n in exI, clarify, rename_tac m)
+apply (drule_tac x=m in spec, simp)
+apply (drule_tac x=n in spec, simp)
+apply (drule spec, drule (1) mp, clarify, rename_tac i)
+apply (rule_tac x="max i k" in exI, simp)
+done
+
+lemma le_Real:
+ assumes X: "cauchy X" and Y: "cauchy Y"
+ shows "Real X \<le> Real Y = (\<forall>r>0. \<exists>k. \<forall>n\<ge>k. X n \<le> Y n + r)"
+unfolding not_less [symmetric, where 'a=real] less_real_def
+apply (simp add: diff_Real not_positive_Real X Y)
+apply (simp add: diff_le_eq add_ac)
+done
-lemma real_of_preal_mult:
- "real_of_preal ((x::preal) * y) = real_of_preal x* real_of_preal y"
-by (simp add: real_of_preal_def real_mult algebra_simps)
-
+lemma le_RealI:
+ assumes Y: "cauchy Y"
+ shows "\<forall>n. x \<le> of_rat (Y n) \<Longrightarrow> x \<le> Real Y"
+proof (induct x)
+ fix X assume X: "cauchy X" and "\<forall>n. Real X \<le> of_rat (Y n)"
+ hence le: "\<And>m r. 0 < r \<Longrightarrow> \<exists>k. \<forall>n\<ge>k. X n \<le> Y m + r"
+ by (simp add: of_rat_Real le_Real)
+ {
+ fix r :: rat assume "0 < r"
+ then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t"
+ by (rule obtain_pos_sum)
+ obtain i where i: "\<forall>m\<ge>i. \<forall>n\<ge>i. \<bar>Y m - Y n\<bar> < s"
+ using cauchyD [OF Y s] ..
+ obtain j where j: "\<forall>n\<ge>j. X n \<le> Y i + t"
+ using le [OF t] ..
+ have "\<forall>n\<ge>max i j. X n \<le> Y n + r"
+ proof (clarsimp)
+ fix n assume n: "i \<le> n" "j \<le> n"
+ have "X n \<le> Y i + t" using n j by simp
+ moreover have "\<bar>Y i - Y n\<bar> < s" using n i by simp
+ ultimately show "X n \<le> Y n + r" unfolding r by simp
+ qed
+ hence "\<exists>k. \<forall>n\<ge>k. X n \<le> Y n + r" ..
+ }
+ thus "Real X \<le> Real Y"
+ by (simp add: of_rat_Real le_Real X Y)
+qed
-text{*Gleason prop 9-4.4 p 127*}
-lemma real_of_preal_trichotomy:
- "\<exists>m. (x::real) = real_of_preal m | x = 0 | x = -(real_of_preal m)"
-apply (simp add: real_of_preal_def real_zero_def, cases x)
-apply (auto simp add: real_minus add_ac)
-apply (cut_tac x = x and y = y in linorder_less_linear)
-apply (auto dest!: less_add_left_Ex simp add: add_assoc [symmetric])
+lemma Real_leI:
+ assumes X: "cauchy X"
+ assumes le: "\<forall>n. of_rat (X n) \<le> y"
+ shows "Real X \<le> y"
+proof -
+ have "- y \<le> - Real X"
+ by (simp add: minus_Real X le_RealI of_rat_minus le)
+ thus ?thesis by simp
+qed
+
+lemma less_RealD:
+ assumes Y: "cauchy Y"
+ shows "x < Real Y \<Longrightarrow> \<exists>n. x < of_rat (Y n)"
+by (erule contrapos_pp, simp add: not_less, erule Real_leI [OF Y])
+
+lemma of_nat_less_two_power:
+ "of_nat n < (2::'a::{linordered_idom,number_ring}) ^ n"
+apply (induct n)
+apply simp
+apply (subgoal_tac "(1::'a) \<le> 2 ^ n")
+apply (drule (1) add_le_less_mono, simp)
+apply simp
done
-lemma real_of_preal_leD:
- "real_of_preal m1 \<le> real_of_preal m2 ==> m1 \<le> m2"
-by (simp add: real_of_preal_def real_le)
+lemma complete_real:
+ fixes S :: "real set"
+ assumes "\<exists>x. x \<in> S" and "\<exists>z. \<forall>x\<in>S. x \<le> z"
+ shows "\<exists>y. (\<forall>x\<in>S. x \<le> y) \<and> (\<forall>z. (\<forall>x\<in>S. x \<le> z) \<longrightarrow> y \<le> z)"
+proof -
+ obtain x where x: "x \<in> S" using assms(1) ..
+ obtain z where z: "\<forall>x\<in>S. x \<le> z" using assms(2) ..
-lemma real_of_preal_lessI: "m1 < m2 ==> real_of_preal m1 < real_of_preal m2"
-by (auto simp add: real_of_preal_leD linorder_not_le [symmetric])
+ def P \<equiv> "\<lambda>x. \<forall>y\<in>S. y \<le> of_rat x"
+ obtain a where a: "\<not> P a"
+ proof
+ have "of_int (floor (x - 1)) \<le> x - 1" by (rule of_int_floor_le)
+ also have "x - 1 < x" by simp
+ finally have "of_int (floor (x - 1)) < x" .
+ hence "\<not> x \<le> of_int (floor (x - 1))" by (simp only: not_le)
+ then show "\<not> P (of_int (floor (x - 1)))"
+ unfolding P_def of_rat_of_int_eq using x by fast
+ qed
+ obtain b where b: "P b"
+ proof
+ show "P (of_int (ceiling z))"
+ unfolding P_def of_rat_of_int_eq
+ proof
+ fix y assume "y \<in> S"
+ hence "y \<le> z" using z by simp
+ also have "z \<le> of_int (ceiling z)" by (rule le_of_int_ceiling)
+ finally show "y \<le> of_int (ceiling z)" .
+ qed
+ qed
-lemma real_of_preal_lessD:
- "real_of_preal m1 < real_of_preal m2 ==> m1 < m2"
-by (simp add: real_of_preal_def real_le linorder_not_le [symmetric])
+ def avg \<equiv> "\<lambda>x y :: rat. x/2 + y/2"
+ def bisect \<equiv> "\<lambda>(x, y). if P (avg x y) then (x, avg x y) else (avg x y, y)"
+ def A \<equiv> "\<lambda>n. fst ((bisect ^^ n) (a, b))"
+ def B \<equiv> "\<lambda>n. snd ((bisect ^^ n) (a, b))"
+ def C \<equiv> "\<lambda>n. avg (A n) (B n)"
+ have A_0 [simp]: "A 0 = a" unfolding A_def by simp
+ have B_0 [simp]: "B 0 = b" unfolding B_def by simp
+ have A_Suc [simp]: "\<And>n. A (Suc n) = (if P (C n) then A n else C n)"
+ unfolding A_def B_def C_def bisect_def split_def by simp
+ have B_Suc [simp]: "\<And>n. B (Suc n) = (if P (C n) then C n else B n)"
+ unfolding A_def B_def C_def bisect_def split_def by simp
-lemma real_of_preal_less_iff [simp]:
- "(real_of_preal m1 < real_of_preal m2) = (m1 < m2)"
-by (blast intro: real_of_preal_lessI real_of_preal_lessD)
+ have width: "\<And>n. B n - A n = (b - a) / 2^n"
+ apply (simp add: eq_divide_eq)
+ apply (induct_tac n, simp)
+ apply (simp add: C_def avg_def algebra_simps)
+ done
+
+ have twos: "\<And>y r :: rat. 0 < r \<Longrightarrow> \<exists>n. y / 2 ^ n < r"
+ apply (simp add: divide_less_eq)
+ apply (subst mult_commute)
+ apply (frule_tac y=y in ex_less_of_nat_mult)
+ apply clarify
+ apply (rule_tac x=n in exI)
+ apply (erule less_trans)
+ apply (rule mult_strict_right_mono)
+ apply (rule le_less_trans [OF _ of_nat_less_two_power])
+ apply simp
+ apply assumption
+ done
-lemma real_of_preal_le_iff:
- "(real_of_preal m1 \<le> real_of_preal m2) = (m1 \<le> m2)"
-by (simp add: linorder_not_less [symmetric])
-
-lemma real_of_preal_zero_less: "0 < real_of_preal m"
-apply (insert preal_self_less_add_left [of 1 m])
-apply (auto simp add: real_zero_def real_of_preal_def
- real_less_def real_le_def add_ac)
-apply (rule_tac x="m + 1" in exI, rule_tac x="1" in exI)
-apply (simp add: add_ac)
-done
-
-lemma real_of_preal_minus_less_zero: "- real_of_preal m < 0"
-by (simp add: real_of_preal_zero_less)
-
-lemma real_of_preal_not_minus_gt_zero: "~ 0 < - real_of_preal m"
-proof -
- from real_of_preal_minus_less_zero
- show ?thesis by (blast dest: order_less_trans)
+ have PA: "\<And>n. \<not> P (A n)"
+ by (induct_tac n, simp_all add: a)
+ have PB: "\<And>n. P (B n)"
+ by (induct_tac n, simp_all add: b)
+ have ab: "a < b"
+ using a b unfolding P_def
+ apply (clarsimp simp add: not_le)
+ apply (drule (1) bspec)
+ apply (drule (1) less_le_trans)
+ apply (simp add: of_rat_less)
+ done
+ have AB: "\<And>n. A n < B n"
+ by (induct_tac n, simp add: ab, simp add: C_def avg_def)
+ have A_mono: "\<And>i j. i \<le> j \<Longrightarrow> A i \<le> A j"
+ apply (auto simp add: le_less [where 'a=nat])
+ apply (erule less_Suc_induct)
+ apply (clarsimp simp add: C_def avg_def)
+ apply (simp add: add_divide_distrib [symmetric])
+ apply (rule AB [THEN less_imp_le])
+ apply simp
+ done
+ have B_mono: "\<And>i j. i \<le> j \<Longrightarrow> B j \<le> B i"
+ apply (auto simp add: le_less [where 'a=nat])
+ apply (erule less_Suc_induct)
+ apply (clarsimp simp add: C_def avg_def)
+ apply (simp add: add_divide_distrib [symmetric])
+ apply (rule AB [THEN less_imp_le])
+ apply simp
+ done
+ have cauchy_lemma:
+ "\<And>X. \<forall>n. \<forall>i\<ge>n. A n \<le> X i \<and> X i \<le> B n \<Longrightarrow> cauchy X"
+ apply (rule cauchyI)
+ apply (drule twos [where y="b - a"])
+ apply (erule exE)
+ apply (rule_tac x=n in exI, clarify, rename_tac i j)
+ apply (rule_tac y="B n - A n" in le_less_trans) defer
+ apply (simp add: width)
+ apply (drule_tac x=n in spec)
+ apply (frule_tac x=i in spec, drule (1) mp)
+ apply (frule_tac x=j in spec, drule (1) mp)
+ apply (frule A_mono, drule B_mono)
+ apply (frule A_mono, drule B_mono)
+ apply arith
+ done
+ have "cauchy A"
+ apply (rule cauchy_lemma [rule_format])
+ apply (simp add: A_mono)
+ apply (erule order_trans [OF less_imp_le [OF AB] B_mono])
+ done
+ have "cauchy B"
+ apply (rule cauchy_lemma [rule_format])
+ apply (simp add: B_mono)
+ apply (erule order_trans [OF A_mono less_imp_le [OF AB]])
+ done
+ have 1: "\<forall>x\<in>S. x \<le> Real B"
+ proof
+ fix x assume "x \<in> S"
+ then show "x \<le> Real B"
+ using PB [unfolded P_def] `cauchy B`
+ by (simp add: le_RealI)
+ qed
+ have 2: "\<forall>z. (\<forall>x\<in>S. x \<le> z) \<longrightarrow> Real A \<le> z"
+ apply clarify
+ apply (erule contrapos_pp)
+ apply (simp add: not_le)
+ apply (drule less_RealD [OF `cauchy A`], clarify)
+ apply (subgoal_tac "\<not> P (A n)")
+ apply (simp add: P_def not_le, clarify)
+ apply (erule rev_bexI)
+ apply (erule (1) less_trans)
+ apply (simp add: PA)
+ done
+ have "vanishes (\<lambda>n. (b - a) / 2 ^ n)"
+ proof (rule vanishesI)
+ fix r :: rat assume "0 < r"
+ then obtain k where k: "\<bar>b - a\<bar> / 2 ^ k < r"
+ using twos by fast
+ have "\<forall>n\<ge>k. \<bar>(b - a) / 2 ^ n\<bar> < r"
+ proof (clarify)
+ fix n assume n: "k \<le> n"
+ have "\<bar>(b - a) / 2 ^ n\<bar> = \<bar>b - a\<bar> / 2 ^ n"
+ by simp
+ also have "\<dots> \<le> \<bar>b - a\<bar> / 2 ^ k"
+ using n by (simp add: divide_left_mono mult_pos_pos)
+ also note k
+ finally show "\<bar>(b - a) / 2 ^ n\<bar> < r" .
+ qed
+ thus "\<exists>k. \<forall>n\<ge>k. \<bar>(b - a) / 2 ^ n\<bar> < r" ..
+ qed
+ hence 3: "Real B = Real A"
+ by (simp add: eq_Real `cauchy A` `cauchy B` width)
+ show "\<exists>y. (\<forall>x\<in>S. x \<le> y) \<and> (\<forall>z. (\<forall>x\<in>S. x \<le> z) \<longrightarrow> y \<le> z)"
+ using 1 2 3 by (rule_tac x="Real B" in exI, simp)
qed
+subsection {* Hiding implementation details *}
-subsection{*Theorems About the Ordering*}
+hide_const (open) vanishes cauchy positive Real real_case
-lemma real_gt_zero_preal_Ex: "(0 < x) = (\<exists>y. x = real_of_preal y)"
-apply (auto simp add: real_of_preal_zero_less)
-apply (cut_tac x = x in real_of_preal_trichotomy)
-apply (blast elim!: real_of_preal_not_minus_gt_zero [THEN notE])
-done
+declare Real_induct [induct del]
+declare Abs_real_induct [induct del]
+declare Abs_real_cases [cases del]
+
+subsection {* Legacy theorem names *}
+
+text {* TODO: Could we have a way to mark theorem names as deprecated,
+and have Isabelle issue a warning and display the preferred name? *}
-lemma real_gt_preal_preal_Ex:
- "real_of_preal z < x ==> \<exists>y. x = real_of_preal y"
-by (blast dest!: real_of_preal_zero_less [THEN order_less_trans]
- intro: real_gt_zero_preal_Ex [THEN iffD1])
-
-lemma real_ge_preal_preal_Ex:
- "real_of_preal z \<le> x ==> \<exists>y. x = real_of_preal y"
-by (blast dest: order_le_imp_less_or_eq real_gt_preal_preal_Ex)
-
-lemma real_less_all_preal: "y \<le> 0 ==> \<forall>x. y < real_of_preal x"
-by (auto elim: order_le_imp_less_or_eq [THEN disjE]
- intro: real_of_preal_zero_less [THEN [2] order_less_trans]
- simp add: real_of_preal_zero_less)
-
-lemma real_less_all_real2: "~ 0 < y ==> \<forall>x. y < real_of_preal x"
-by (blast intro!: real_less_all_preal linorder_not_less [THEN iffD1])
-
+lemmas real_diff_def = minus_real_def [of "r" "s", standard]
+lemmas real_divide_def = divide_real_def [of "R" "S", standard]
+lemmas real_less_def = less_le [of "x::real" "y", standard]
+lemmas real_abs_def = abs_real_def [of "r", standard]
+lemmas real_sgn_def = sgn_real_def [of "x", standard]
+lemmas real_mult_commute = mult_commute [of "z::real" "w", standard]
+lemmas real_mult_assoc = mult_assoc [of "z1::real" "z2" "z3", standard]
+lemmas real_mult_1 = mult_1_left [of "z::real", standard]
+lemmas real_add_mult_distrib = left_distrib [of "z1::real" "z2" "w", standard]
+lemmas real_zero_not_eq_one = zero_neq_one [where 'a=real]
+lemmas real_mult_inverse_left = left_inverse [of "x::real", standard]
+lemmas INVERSE_ZERO = inverse_zero [where 'a=real]
+lemmas real_le_refl = order_refl [of "w::real", standard]
+lemmas real_le_antisym = order_antisym [of "z::real" "w", standard]
+lemmas real_le_trans = order_trans [of "i::real" "j" "k", standard]
+lemmas real_le_eq_diff = le_iff_diff_le_0 [of "x::real" "y", standard]
+lemmas real_add_left_mono = add_left_mono [of "x::real" "y" "z", standard]
+lemmas real_mult_order = mult_pos_pos [of "x::real" "y", standard]
+lemmas real_mult_less_mono2 =
+ mult_strict_left_mono [of "x::real" "y" "z", COMP swap_prems_rl, standard]
subsection{*More Lemmas*}
@@ -940,20 +1497,11 @@
subsection{*Numerals and Arithmetic*}
-instantiation real :: number_ring
-begin
-
-definition
- real_number_of_def [code del]: "number_of w = real_of_int w"
-
-instance
- by intro_classes (simp add: real_number_of_def)
-
-end
+declare number_of_real_def [code del]
lemma [code_unfold_post]:
"number_of k = real_of_int (number_of k)"
- unfolding number_of_is_id real_number_of_def ..
+ unfolding number_of_is_id number_of_real_def ..
text{*Collapse applications of @{term real} to @{term number_of}*}
@@ -976,8 +1524,8 @@
@{thm real_of_int_add}, @{thm real_of_int_minus}, @{thm real_of_int_diff},
@{thm real_of_int_mult}, @{thm real_of_int_of_nat_eq},
@{thm real_of_nat_number_of}, @{thm real_number_of}]
- #> Lin_Arith.add_inj_const (@{const_name real}, HOLogic.natT --> HOLogic.realT)
- #> Lin_Arith.add_inj_const (@{const_name real}, HOLogic.intT --> HOLogic.realT))
+ #> Lin_Arith.add_inj_const (@{const_name real}, @{typ "nat \<Rightarrow> real"})
+ #> Lin_Arith.add_inj_const (@{const_name real}, @{typ "int \<Rightarrow> real"}))
*}
@@ -986,7 +1534,7 @@
text{*Needed in this non-standard form by Hyperreal/Transcendental*}
lemma real_0_le_divide_iff:
"((0::real) \<le> x/y) = ((x \<le> 0 | 0 \<le> y) & (0 \<le> x | y \<le> 0))"
-by (simp add: real_divide_def zero_le_mult_iff, auto)
+by (auto simp add: zero_le_divide_iff)
lemma real_add_minus_iff [simp]: "(x + - a = (0::real)) = (x=a)"
by arith
--- a/src/HOL/RealVector.thy Tue May 11 08:29:42 2010 +0200
+++ b/src/HOL/RealVector.thy Tue May 11 08:30:02 2010 +0200
@@ -793,7 +793,7 @@
apply (intro_classes, unfold real_norm_def real_scaleR_def)
apply (rule dist_real_def)
apply (rule open_real_def)
-apply (simp add: real_sgn_def)
+apply (simp add: sgn_real_def)
apply (rule abs_ge_zero)
apply (rule abs_eq_0)
apply (rule abs_triangle_ineq)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/Dedekind_Real.thy Tue May 11 08:30:02 2010 +0200
@@ -0,0 +1,2045 @@
+(* Title: HOL/ex/Dedekind_Real.thy
+ Author: Jacques D. Fleuriot, University of Cambridge
+ Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4
+
+The positive reals as Dedekind sections of positive
+rationals. Fundamentals of Abstract Analysis [Gleason- p. 121]
+provides some of the definitions.
+*)
+
+theory Dedekind_Real
+imports Rat Lubs
+begin
+
+section {* Positive real numbers *}
+
+text{*Could be generalized and moved to @{text Groups}*}
+lemma add_eq_exists: "\<exists>x. a+x = (b::rat)"
+by (rule_tac x="b-a" in exI, simp)
+
+definition
+ cut :: "rat set => bool" where
+ [code del]: "cut A = ({} \<subset> A &
+ A < {r. 0 < r} &
+ (\<forall>y \<in> A. ((\<forall>z. 0<z & z < y --> z \<in> A) & (\<exists>u \<in> A. y < u))))"
+
+lemma interval_empty_iff:
+ "{y. (x::'a::dense_linorder) < y \<and> y < z} = {} \<longleftrightarrow> \<not> x < z"
+ by (auto dest: dense)
+
+
+lemma cut_of_rat:
+ assumes q: "0 < q" shows "cut {r::rat. 0 < r & r < q}" (is "cut ?A")
+proof -
+ from q have pos: "?A < {r. 0 < r}" by force
+ have nonempty: "{} \<subset> ?A"
+ proof
+ show "{} \<subseteq> ?A" by simp
+ show "{} \<noteq> ?A"
+ by (force simp only: q eq_commute [of "{}"] interval_empty_iff)
+ qed
+ show ?thesis
+ by (simp add: cut_def pos nonempty,
+ blast dest: dense intro: order_less_trans)
+qed
+
+
+typedef preal = "{A. cut A}"
+ by (blast intro: cut_of_rat [OF zero_less_one])
+
+definition
+ psup :: "preal set => preal" where
+ [code del]: "psup P = Abs_preal (\<Union>X \<in> P. Rep_preal X)"
+
+definition
+ add_set :: "[rat set,rat set] => rat set" where
+ "add_set A B = {w. \<exists>x \<in> A. \<exists>y \<in> B. w = x + y}"
+
+definition
+ diff_set :: "[rat set,rat set] => rat set" where
+ [code del]: "diff_set A B = {w. \<exists>x. 0 < w & 0 < x & x \<notin> B & x + w \<in> A}"
+
+definition
+ mult_set :: "[rat set,rat set] => rat set" where
+ "mult_set A B = {w. \<exists>x \<in> A. \<exists>y \<in> B. w = x * y}"
+
+definition
+ inverse_set :: "rat set => rat set" where
+ [code del]: "inverse_set A = {x. \<exists>y. 0 < x & x < y & inverse y \<notin> A}"
+
+instantiation preal :: "{ord, plus, minus, times, inverse, one}"
+begin
+
+definition
+ preal_less_def [code del]:
+ "R < S == Rep_preal R < Rep_preal S"
+
+definition
+ preal_le_def [code del]:
+ "R \<le> S == Rep_preal R \<subseteq> Rep_preal S"
+
+definition
+ preal_add_def:
+ "R + S == Abs_preal (add_set (Rep_preal R) (Rep_preal S))"
+
+definition
+ preal_diff_def:
+ "R - S == Abs_preal (diff_set (Rep_preal R) (Rep_preal S))"
+
+definition
+ preal_mult_def:
+ "R * S == Abs_preal (mult_set (Rep_preal R) (Rep_preal S))"
+
+definition
+ preal_inverse_def:
+ "inverse R == Abs_preal (inverse_set (Rep_preal R))"
+
+definition "R / S = R * inverse (S\<Colon>preal)"
+
+definition
+ preal_one_def:
+ "1 == Abs_preal {x. 0 < x & x < 1}"
+
+instance ..
+
+end
+
+
+text{*Reduces equality on abstractions to equality on representatives*}
+declare Abs_preal_inject [simp]
+declare Abs_preal_inverse [simp]
+
+lemma rat_mem_preal: "0 < q ==> {r::rat. 0 < r & r < q} \<in> preal"
+by (simp add: preal_def cut_of_rat)
+
+lemma preal_nonempty: "A \<in> preal ==> \<exists>x\<in>A. 0 < x"
+by (unfold preal_def cut_def, blast)
+
+lemma preal_Ex_mem: "A \<in> preal \<Longrightarrow> \<exists>x. x \<in> A"
+by (drule preal_nonempty, fast)
+
+lemma preal_imp_psubset_positives: "A \<in> preal ==> A < {r. 0 < r}"
+by (force simp add: preal_def cut_def)
+
+lemma preal_exists_bound: "A \<in> preal ==> \<exists>x. 0 < x & x \<notin> A"
+by (drule preal_imp_psubset_positives, auto)
+
+lemma preal_exists_greater: "[| A \<in> preal; y \<in> A |] ==> \<exists>u \<in> A. y < u"
+by (unfold preal_def cut_def, blast)
+
+lemma preal_downwards_closed: "[| A \<in> preal; y \<in> A; 0 < z; z < y |] ==> z \<in> A"
+by (unfold preal_def cut_def, blast)
+
+text{*Relaxing the final premise*}
+lemma preal_downwards_closed':
+ "[| A \<in> preal; y \<in> A; 0 < z; z \<le> y |] ==> z \<in> A"
+apply (simp add: order_le_less)
+apply (blast intro: preal_downwards_closed)
+done
+
+text{*A positive fraction not in a positive real is an upper bound.
+ Gleason p. 122 - Remark (1)*}
+
+lemma not_in_preal_ub:
+ assumes A: "A \<in> preal"
+ and notx: "x \<notin> A"
+ and y: "y \<in> A"
+ and pos: "0 < x"
+ shows "y < x"
+proof (cases rule: linorder_cases)
+ assume "x<y"
+ with notx show ?thesis
+ by (simp add: preal_downwards_closed [OF A y] pos)
+next
+ assume "x=y"
+ with notx and y show ?thesis by simp
+next
+ assume "y<x"
+ thus ?thesis .
+qed
+
+text {* preal lemmas instantiated to @{term "Rep_preal X"} *}
+
+lemma mem_Rep_preal_Ex: "\<exists>x. x \<in> Rep_preal X"
+by (rule preal_Ex_mem [OF Rep_preal])
+
+lemma Rep_preal_exists_bound: "\<exists>x>0. x \<notin> Rep_preal X"
+by (rule preal_exists_bound [OF Rep_preal])
+
+lemmas not_in_Rep_preal_ub = not_in_preal_ub [OF Rep_preal]
+
+
+subsection{*Properties of Ordering*}
+
+instance preal :: order
+proof
+ fix w :: preal
+ show "w \<le> w" by (simp add: preal_le_def)
+next
+ fix i j k :: preal
+ assume "i \<le> j" and "j \<le> k"
+ then show "i \<le> k" by (simp add: preal_le_def)
+next
+ fix z w :: preal
+ assume "z \<le> w" and "w \<le> z"
+ then show "z = w" by (simp add: preal_le_def Rep_preal_inject)
+next
+ fix z w :: preal
+ show "z < w \<longleftrightarrow> z \<le> w \<and> \<not> w \<le> z"
+ by (auto simp add: preal_le_def preal_less_def Rep_preal_inject)
+qed
+
+lemma preal_imp_pos: "[|A \<in> preal; r \<in> A|] ==> 0 < r"
+by (insert preal_imp_psubset_positives, blast)
+
+instance preal :: linorder
+proof
+ fix x y :: preal
+ show "x <= y | y <= x"
+ apply (auto simp add: preal_le_def)
+ apply (rule ccontr)
+ apply (blast dest: not_in_Rep_preal_ub intro: preal_imp_pos [OF Rep_preal]
+ elim: order_less_asym)
+ done
+qed
+
+instantiation preal :: distrib_lattice
+begin
+
+definition
+ "(inf \<Colon> preal \<Rightarrow> preal \<Rightarrow> preal) = min"
+
+definition
+ "(sup \<Colon> preal \<Rightarrow> preal \<Rightarrow> preal) = max"
+
+instance
+ by intro_classes
+ (auto simp add: inf_preal_def sup_preal_def min_max.sup_inf_distrib1)
+
+end
+
+subsection{*Properties of Addition*}
+
+lemma preal_add_commute: "(x::preal) + y = y + x"
+apply (unfold preal_add_def add_set_def)
+apply (rule_tac f = Abs_preal in arg_cong)
+apply (force simp add: add_commute)
+done
+
+text{*Lemmas for proving that addition of two positive reals gives
+ a positive real*}
+
+text{*Part 1 of Dedekind sections definition*}
+lemma add_set_not_empty:
+ "[|A \<in> preal; B \<in> preal|] ==> {} \<subset> add_set A B"
+apply (drule preal_nonempty)+
+apply (auto simp add: add_set_def)
+done
+
+text{*Part 2 of Dedekind sections definition. A structured version of
+this proof is @{text preal_not_mem_mult_set_Ex} below.*}
+lemma preal_not_mem_add_set_Ex:
+ "[|A \<in> preal; B \<in> preal|] ==> \<exists>q>0. q \<notin> add_set A B"
+apply (insert preal_exists_bound [of A] preal_exists_bound [of B], auto)
+apply (rule_tac x = "x+xa" in exI)
+apply (simp add: add_set_def, clarify)
+apply (drule (3) not_in_preal_ub)+
+apply (force dest: add_strict_mono)
+done
+
+lemma add_set_not_rat_set:
+ assumes A: "A \<in> preal"
+ and B: "B \<in> preal"
+ shows "add_set A B < {r. 0 < r}"
+proof
+ from preal_imp_pos [OF A] preal_imp_pos [OF B]
+ show "add_set A B \<subseteq> {r. 0 < r}" by (force simp add: add_set_def)
+next
+ show "add_set A B \<noteq> {r. 0 < r}"
+ by (insert preal_not_mem_add_set_Ex [OF A B], blast)
+qed
+
+text{*Part 3 of Dedekind sections definition*}
+lemma add_set_lemma3:
+ "[|A \<in> preal; B \<in> preal; u \<in> add_set A B; 0 < z; z < u|]
+ ==> z \<in> add_set A B"
+proof (unfold add_set_def, clarify)
+ fix x::rat and y::rat
+ assume A: "A \<in> preal"
+ and B: "B \<in> preal"
+ and [simp]: "0 < z"
+ and zless: "z < x + y"
+ and x: "x \<in> A"
+ and y: "y \<in> B"
+ have xpos [simp]: "0<x" by (rule preal_imp_pos [OF A x])
+ have ypos [simp]: "0<y" by (rule preal_imp_pos [OF B y])
+ have xypos [simp]: "0 < x+y" by (simp add: pos_add_strict)
+ let ?f = "z/(x+y)"
+ have fless: "?f < 1" by (simp add: zless pos_divide_less_eq)
+ show "\<exists>x' \<in> A. \<exists>y'\<in>B. z = x' + y'"
+ proof (intro bexI)
+ show "z = x*?f + y*?f"
+ by (simp add: left_distrib [symmetric] divide_inverse mult_ac
+ order_less_imp_not_eq2)
+ next
+ show "y * ?f \<in> B"
+ proof (rule preal_downwards_closed [OF B y])
+ show "0 < y * ?f"
+ by (simp add: divide_inverse zero_less_mult_iff)
+ next
+ show "y * ?f < y"
+ by (insert mult_strict_left_mono [OF fless ypos], simp)
+ qed
+ next
+ show "x * ?f \<in> A"
+ proof (rule preal_downwards_closed [OF A x])
+ show "0 < x * ?f"
+ by (simp add: divide_inverse zero_less_mult_iff)
+ next
+ show "x * ?f < x"
+ by (insert mult_strict_left_mono [OF fless xpos], simp)
+ qed
+ qed
+qed
+
+text{*Part 4 of Dedekind sections definition*}
+lemma add_set_lemma4:
+ "[|A \<in> preal; B \<in> preal; y \<in> add_set A B|] ==> \<exists>u \<in> add_set A B. y < u"
+apply (auto simp add: add_set_def)
+apply (frule preal_exists_greater [of A], auto)
+apply (rule_tac x="u + y" in exI)
+apply (auto intro: add_strict_left_mono)
+done
+
+lemma mem_add_set:
+ "[|A \<in> preal; B \<in> preal|] ==> add_set A B \<in> preal"
+apply (simp (no_asm_simp) add: preal_def cut_def)
+apply (blast intro!: add_set_not_empty add_set_not_rat_set
+ add_set_lemma3 add_set_lemma4)
+done
+
+lemma preal_add_assoc: "((x::preal) + y) + z = x + (y + z)"
+apply (simp add: preal_add_def mem_add_set Rep_preal)
+apply (force simp add: add_set_def add_ac)
+done
+
+instance preal :: ab_semigroup_add
+proof
+ fix a b c :: preal
+ show "(a + b) + c = a + (b + c)" by (rule preal_add_assoc)
+ show "a + b = b + a" by (rule preal_add_commute)
+qed
+
+
+subsection{*Properties of Multiplication*}
+
+text{*Proofs essentially same as for addition*}
+
+lemma preal_mult_commute: "(x::preal) * y = y * x"
+apply (unfold preal_mult_def mult_set_def)
+apply (rule_tac f = Abs_preal in arg_cong)
+apply (force simp add: mult_commute)
+done
+
+text{*Multiplication of two positive reals gives a positive real.*}
+
+text{*Lemmas for proving positive reals multiplication set in @{typ preal}*}
+
+text{*Part 1 of Dedekind sections definition*}
+lemma mult_set_not_empty:
+ "[|A \<in> preal; B \<in> preal|] ==> {} \<subset> mult_set A B"
+apply (insert preal_nonempty [of A] preal_nonempty [of B])
+apply (auto simp add: mult_set_def)
+done
+
+text{*Part 2 of Dedekind sections definition*}
+lemma preal_not_mem_mult_set_Ex:
+ assumes A: "A \<in> preal"
+ and B: "B \<in> preal"
+ shows "\<exists>q. 0 < q & q \<notin> mult_set A B"
+proof -
+ from preal_exists_bound [OF A]
+ obtain x where [simp]: "0 < x" "x \<notin> A" by blast
+ from preal_exists_bound [OF B]
+ obtain y where [simp]: "0 < y" "y \<notin> B" by blast
+ show ?thesis
+ proof (intro exI conjI)
+ show "0 < x*y" by (simp add: mult_pos_pos)
+ show "x * y \<notin> mult_set A B"
+ proof -
+ { fix u::rat and v::rat
+ assume "u \<in> A" and "v \<in> B" and "x*y = u*v"
+ moreover
+ with prems have "u<x" and "v<y" by (blast dest: not_in_preal_ub)+
+ moreover
+ with prems have "0\<le>v"
+ by (blast intro: preal_imp_pos [OF B] order_less_imp_le prems)
+ moreover
+ from calculation
+ have "u*v < x*y" by (blast intro: mult_strict_mono prems)
+ ultimately have False by force }
+ thus ?thesis by (auto simp add: mult_set_def)
+ qed
+ qed
+qed
+
+lemma mult_set_not_rat_set:
+ assumes A: "A \<in> preal"
+ and B: "B \<in> preal"
+ shows "mult_set A B < {r. 0 < r}"
+proof
+ show "mult_set A B \<subseteq> {r. 0 < r}"
+ by (force simp add: mult_set_def
+ intro: preal_imp_pos [OF A] preal_imp_pos [OF B] mult_pos_pos)
+ show "mult_set A B \<noteq> {r. 0 < r}"
+ using preal_not_mem_mult_set_Ex [OF A B] by blast
+qed
+
+
+
+text{*Part 3 of Dedekind sections definition*}
+lemma mult_set_lemma3:
+ "[|A \<in> preal; B \<in> preal; u \<in> mult_set A B; 0 < z; z < u|]
+ ==> z \<in> mult_set A B"
+proof (unfold mult_set_def, clarify)
+ fix x::rat and y::rat
+ assume A: "A \<in> preal"
+ and B: "B \<in> preal"
+ and [simp]: "0 < z"
+ and zless: "z < x * y"
+ and x: "x \<in> A"
+ and y: "y \<in> B"
+ have [simp]: "0<y" by (rule preal_imp_pos [OF B y])
+ show "\<exists>x' \<in> A. \<exists>y' \<in> B. z = x' * y'"
+ proof
+ show "\<exists>y'\<in>B. z = (z/y) * y'"
+ proof
+ show "z = (z/y)*y"
+ by (simp add: divide_inverse mult_commute [of y] mult_assoc
+ order_less_imp_not_eq2)
+ show "y \<in> B" by fact
+ qed
+ next
+ show "z/y \<in> A"
+ proof (rule preal_downwards_closed [OF A x])
+ show "0 < z/y"
+ by (simp add: zero_less_divide_iff)
+ show "z/y < x" by (simp add: pos_divide_less_eq zless)
+ qed
+ qed
+qed
+
+text{*Part 4 of Dedekind sections definition*}
+lemma mult_set_lemma4:
+ "[|A \<in> preal; B \<in> preal; y \<in> mult_set A B|] ==> \<exists>u \<in> mult_set A B. y < u"
+apply (auto simp add: mult_set_def)
+apply (frule preal_exists_greater [of A], auto)
+apply (rule_tac x="u * y" in exI)
+apply (auto intro: preal_imp_pos [of A] preal_imp_pos [of B]
+ mult_strict_right_mono)
+done
+
+
+lemma mem_mult_set:
+ "[|A \<in> preal; B \<in> preal|] ==> mult_set A B \<in> preal"
+apply (simp (no_asm_simp) add: preal_def cut_def)
+apply (blast intro!: mult_set_not_empty mult_set_not_rat_set
+ mult_set_lemma3 mult_set_lemma4)
+done
+
+lemma preal_mult_assoc: "((x::preal) * y) * z = x * (y * z)"
+apply (simp add: preal_mult_def mem_mult_set Rep_preal)
+apply (force simp add: mult_set_def mult_ac)
+done
+
+instance preal :: ab_semigroup_mult
+proof
+ fix a b c :: preal
+ show "(a * b) * c = a * (b * c)" by (rule preal_mult_assoc)
+ show "a * b = b * a" by (rule preal_mult_commute)
+qed
+
+
+text{* Positive real 1 is the multiplicative identity element *}
+
+lemma preal_mult_1: "(1::preal) * z = z"
+proof (induct z)
+ fix A :: "rat set"
+ assume A: "A \<in> preal"
+ have "{w. \<exists>u. 0 < u \<and> u < 1 & (\<exists>v \<in> A. w = u * v)} = A" (is "?lhs = A")
+ proof
+ show "?lhs \<subseteq> A"
+ proof clarify
+ fix x::rat and u::rat and v::rat
+ assume upos: "0<u" and "u<1" and v: "v \<in> A"
+ have vpos: "0<v" by (rule preal_imp_pos [OF A v])
+ hence "u*v < 1*v" by (simp only: mult_strict_right_mono prems)
+ thus "u * v \<in> A"
+ by (force intro: preal_downwards_closed [OF A v] mult_pos_pos
+ upos vpos)
+ qed
+ next
+ show "A \<subseteq> ?lhs"
+ proof clarify
+ fix x::rat
+ assume x: "x \<in> A"
+ have xpos: "0<x" by (rule preal_imp_pos [OF A x])
+ from preal_exists_greater [OF A x]
+ obtain v where v: "v \<in> A" and xlessv: "x < v" ..
+ have vpos: "0<v" by (rule preal_imp_pos [OF A v])
+ show "\<exists>u. 0 < u \<and> u < 1 \<and> (\<exists>v\<in>A. x = u * v)"
+ proof (intro exI conjI)
+ show "0 < x/v"
+ by (simp add: zero_less_divide_iff xpos vpos)
+ show "x / v < 1"
+ by (simp add: pos_divide_less_eq vpos xlessv)
+ show "\<exists>v'\<in>A. x = (x / v) * v'"
+ proof
+ show "x = (x/v)*v"
+ by (simp add: divide_inverse mult_assoc vpos
+ order_less_imp_not_eq2)
+ show "v \<in> A" by fact
+ qed
+ qed
+ qed
+ qed
+ thus "1 * Abs_preal A = Abs_preal A"
+ by (simp add: preal_one_def preal_mult_def mult_set_def
+ rat_mem_preal A)
+qed
+
+instance preal :: comm_monoid_mult
+by intro_classes (rule preal_mult_1)
+
+
+subsection{*Distribution of Multiplication across Addition*}
+
+lemma mem_Rep_preal_add_iff:
+ "(z \<in> Rep_preal(R+S)) = (\<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. z = x + y)"
+apply (simp add: preal_add_def mem_add_set Rep_preal)
+apply (simp add: add_set_def)
+done
+
+lemma mem_Rep_preal_mult_iff:
+ "(z \<in> Rep_preal(R*S)) = (\<exists>x \<in> Rep_preal R. \<exists>y \<in> Rep_preal S. z = x * y)"
+apply (simp add: preal_mult_def mem_mult_set Rep_preal)
+apply (simp add: mult_set_def)
+done
+
+lemma distrib_subset1:
+ "Rep_preal (w * (x + y)) \<subseteq> Rep_preal (w * x + w * y)"
+apply (auto simp add: Bex_def mem_Rep_preal_add_iff mem_Rep_preal_mult_iff)
+apply (force simp add: right_distrib)
+done
+
+lemma preal_add_mult_distrib_mean:
+ assumes a: "a \<in> Rep_preal w"
+ and b: "b \<in> Rep_preal w"
+ and d: "d \<in> Rep_preal x"
+ and e: "e \<in> Rep_preal y"
+ shows "\<exists>c \<in> Rep_preal w. a * d + b * e = c * (d + e)"
+proof
+ let ?c = "(a*d + b*e)/(d+e)"
+ have [simp]: "0<a" "0<b" "0<d" "0<e" "0<d+e"
+ by (blast intro: preal_imp_pos [OF Rep_preal] a b d e pos_add_strict)+
+ have cpos: "0 < ?c"
+ by (simp add: zero_less_divide_iff zero_less_mult_iff pos_add_strict)
+ show "a * d + b * e = ?c * (d + e)"
+ by (simp add: divide_inverse mult_assoc order_less_imp_not_eq2)
+ show "?c \<in> Rep_preal w"
+ proof (cases rule: linorder_le_cases)
+ assume "a \<le> b"
+ hence "?c \<le> b"
+ by (simp add: pos_divide_le_eq right_distrib mult_right_mono
+ order_less_imp_le)
+ thus ?thesis by (rule preal_downwards_closed' [OF Rep_preal b cpos])
+ next
+ assume "b \<le> a"
+ hence "?c \<le> a"
+ by (simp add: pos_divide_le_eq right_distrib mult_right_mono
+ order_less_imp_le)
+ thus ?thesis by (rule preal_downwards_closed' [OF Rep_preal a cpos])
+ qed
+qed
+
+lemma distrib_subset2:
+ "Rep_preal (w * x + w * y) \<subseteq> Rep_preal (w * (x + y))"
+apply (auto simp add: Bex_def mem_Rep_preal_add_iff mem_Rep_preal_mult_iff)
+apply (drule_tac w=w and x=x and y=y in preal_add_mult_distrib_mean, auto)
+done
+
+lemma preal_add_mult_distrib2: "(w * ((x::preal) + y)) = (w * x) + (w * y)"
+apply (rule Rep_preal_inject [THEN iffD1])
+apply (rule equalityI [OF distrib_subset1 distrib_subset2])
+done
+
+lemma preal_add_mult_distrib: "(((x::preal) + y) * w) = (x * w) + (y * w)"
+by (simp add: preal_mult_commute preal_add_mult_distrib2)
+
+instance preal :: comm_semiring
+by intro_classes (rule preal_add_mult_distrib)
+
+
+subsection{*Existence of Inverse, a Positive Real*}
+
+lemma mem_inv_set_ex:
+ assumes A: "A \<in> preal" shows "\<exists>x y. 0 < x & x < y & inverse y \<notin> A"
+proof -
+ from preal_exists_bound [OF A]
+ obtain x where [simp]: "0<x" "x \<notin> A" by blast
+ show ?thesis
+ proof (intro exI conjI)
+ show "0 < inverse (x+1)"
+ by (simp add: order_less_trans [OF _ less_add_one])
+ show "inverse(x+1) < inverse x"
+ by (simp add: less_imp_inverse_less less_add_one)
+ show "inverse (inverse x) \<notin> A"
+ by (simp add: order_less_imp_not_eq2)
+ qed
+qed
+
+text{*Part 1 of Dedekind sections definition*}
+lemma inverse_set_not_empty:
+ "A \<in> preal ==> {} \<subset> inverse_set A"
+apply (insert mem_inv_set_ex [of A])
+apply (auto simp add: inverse_set_def)
+done
+
+text{*Part 2 of Dedekind sections definition*}
+
+lemma preal_not_mem_inverse_set_Ex:
+ assumes A: "A \<in> preal" shows "\<exists>q. 0 < q & q \<notin> inverse_set A"
+proof -
+ from preal_nonempty [OF A]
+ obtain x where x: "x \<in> A" and xpos [simp]: "0<x" ..
+ show ?thesis
+ proof (intro exI conjI)
+ show "0 < inverse x" by simp
+ show "inverse x \<notin> inverse_set A"
+ proof -
+ { fix y::rat
+ assume ygt: "inverse x < y"
+ have [simp]: "0 < y" by (simp add: order_less_trans [OF _ ygt])
+ have iyless: "inverse y < x"
+ by (simp add: inverse_less_imp_less [of x] ygt)
+ have "inverse y \<in> A"
+ by (simp add: preal_downwards_closed [OF A x] iyless)}
+ thus ?thesis by (auto simp add: inverse_set_def)
+ qed
+ qed
+qed
+
+lemma inverse_set_not_rat_set:
+ assumes A: "A \<in> preal" shows "inverse_set A < {r. 0 < r}"
+proof
+ show "inverse_set A \<subseteq> {r. 0 < r}" by (force simp add: inverse_set_def)
+next
+ show "inverse_set A \<noteq> {r. 0 < r}"
+ by (insert preal_not_mem_inverse_set_Ex [OF A], blast)
+qed
+
+text{*Part 3 of Dedekind sections definition*}
+lemma inverse_set_lemma3:
+ "[|A \<in> preal; u \<in> inverse_set A; 0 < z; z < u|]
+ ==> z \<in> inverse_set A"
+apply (auto simp add: inverse_set_def)
+apply (auto intro: order_less_trans)
+done
+
+text{*Part 4 of Dedekind sections definition*}
+lemma inverse_set_lemma4:
+ "[|A \<in> preal; y \<in> inverse_set A|] ==> \<exists>u \<in> inverse_set A. y < u"
+apply (auto simp add: inverse_set_def)
+apply (drule dense [of y])
+apply (blast intro: order_less_trans)
+done
+
+
+lemma mem_inverse_set:
+ "A \<in> preal ==> inverse_set A \<in> preal"
+apply (simp (no_asm_simp) add: preal_def cut_def)
+apply (blast intro!: inverse_set_not_empty inverse_set_not_rat_set
+ inverse_set_lemma3 inverse_set_lemma4)
+done
+
+
+subsection{*Gleason's Lemma 9-3.4, page 122*}
+
+lemma Gleason9_34_exists:
+ assumes A: "A \<in> preal"
+ and "\<forall>x\<in>A. x + u \<in> A"
+ and "0 \<le> z"
+ shows "\<exists>b\<in>A. b + (of_int z) * u \<in> A"
+proof (cases z rule: int_cases)
+ case (nonneg n)
+ show ?thesis
+ proof (simp add: prems, induct n)
+ case 0
+ from preal_nonempty [OF A]
+ show ?case by force
+ case (Suc k)
+ from this obtain b where "b \<in> A" "b + of_nat k * u \<in> A" ..
+ hence "b + of_int (int k)*u + u \<in> A" by (simp add: prems)
+ thus ?case by (force simp add: algebra_simps prems)
+ qed
+next
+ case (neg n)
+ with prems show ?thesis by simp
+qed
+
+lemma Gleason9_34_contra:
+ assumes A: "A \<in> preal"
+ shows "[|\<forall>x\<in>A. x + u \<in> A; 0 < u; 0 < y; y \<notin> A|] ==> False"
+proof (induct u, induct y)
+ fix a::int and b::int
+ fix c::int and d::int
+ assume bpos [simp]: "0 < b"
+ and dpos [simp]: "0 < d"
+ and closed: "\<forall>x\<in>A. x + (Fract c d) \<in> A"
+ and upos: "0 < Fract c d"
+ and ypos: "0 < Fract a b"
+ and notin: "Fract a b \<notin> A"
+ have cpos [simp]: "0 < c"
+ by (simp add: zero_less_Fract_iff [OF dpos, symmetric] upos)
+ have apos [simp]: "0 < a"
+ by (simp add: zero_less_Fract_iff [OF bpos, symmetric] ypos)
+ let ?k = "a*d"
+ have frle: "Fract a b \<le> Fract ?k 1 * (Fract c d)"
+ proof -
+ have "?thesis = ((a * d * b * d) \<le> c * b * (a * d * b * d))"
+ by (simp add: order_less_imp_not_eq2 mult_ac)
+ moreover
+ have "(1 * (a * d * b * d)) \<le> c * b * (a * d * b * d)"
+ by (rule mult_mono,
+ simp_all add: int_one_le_iff_zero_less zero_less_mult_iff
+ order_less_imp_le)
+ ultimately
+ show ?thesis by simp
+ qed
+ have k: "0 \<le> ?k" by (simp add: order_less_imp_le zero_less_mult_iff)
+ from Gleason9_34_exists [OF A closed k]
+ obtain z where z: "z \<in> A"
+ and mem: "z + of_int ?k * Fract c d \<in> A" ..
+ have less: "z + of_int ?k * Fract c d < Fract a b"
+ by (rule not_in_preal_ub [OF A notin mem ypos])
+ have "0<z" by (rule preal_imp_pos [OF A z])
+ with frle and less show False by (simp add: Fract_of_int_eq)
+qed
+
+
+lemma Gleason9_34:
+ assumes A: "A \<in> preal"
+ and upos: "0 < u"
+ shows "\<exists>r \<in> A. r + u \<notin> A"
+proof (rule ccontr, simp)
+ assume closed: "\<forall>r\<in>A. r + u \<in> A"
+ from preal_exists_bound [OF A]
+ obtain y where y: "y \<notin> A" and ypos: "0 < y" by blast
+ show False
+ by (rule Gleason9_34_contra [OF A closed upos ypos y])
+qed
+
+
+
+subsection{*Gleason's Lemma 9-3.6*}
+
+lemma lemma_gleason9_36:
+ assumes A: "A \<in> preal"
+ and x: "1 < x"
+ shows "\<exists>r \<in> A. r*x \<notin> A"
+proof -
+ from preal_nonempty [OF A]
+ obtain y where y: "y \<in> A" and ypos: "0<y" ..
+ show ?thesis
+ proof (rule classical)
+ assume "~(\<exists>r\<in>A. r * x \<notin> A)"
+ with y have ymem: "y * x \<in> A" by blast
+ from ypos mult_strict_left_mono [OF x]
+ have yless: "y < y*x" by simp
+ let ?d = "y*x - y"
+ from yless have dpos: "0 < ?d" and eq: "y + ?d = y*x" by auto
+ from Gleason9_34 [OF A dpos]
+ obtain r where r: "r\<in>A" and notin: "r + ?d \<notin> A" ..
+ have rpos: "0<r" by (rule preal_imp_pos [OF A r])
+ with dpos have rdpos: "0 < r + ?d" by arith
+ have "~ (r + ?d \<le> y + ?d)"
+ proof
+ assume le: "r + ?d \<le> y + ?d"
+ from ymem have yd: "y + ?d \<in> A" by (simp add: eq)
+ have "r + ?d \<in> A" by (rule preal_downwards_closed' [OF A yd rdpos le])
+ with notin show False by simp
+ qed
+ hence "y < r" by simp
+ with ypos have dless: "?d < (r * ?d)/y"
+ by (simp add: pos_less_divide_eq mult_commute [of ?d]
+ mult_strict_right_mono dpos)
+ have "r + ?d < r*x"
+ proof -
+ have "r + ?d < r + (r * ?d)/y" by (simp add: dless)
+ also with ypos have "... = (r/y) * (y + ?d)"
+ by (simp only: algebra_simps divide_inverse, simp)
+ also have "... = r*x" using ypos
+ by simp
+ finally show "r + ?d < r*x" .
+ qed
+ with r notin rdpos
+ show "\<exists>r\<in>A. r * x \<notin> A" by (blast dest: preal_downwards_closed [OF A])
+ qed
+qed
+
+subsection{*Existence of Inverse: Part 2*}
+
+lemma mem_Rep_preal_inverse_iff:
+ "(z \<in> Rep_preal(inverse R)) =
+ (0 < z \<and> (\<exists>y. z < y \<and> inverse y \<notin> Rep_preal R))"
+apply (simp add: preal_inverse_def mem_inverse_set Rep_preal)
+apply (simp add: inverse_set_def)
+done
+
+lemma Rep_preal_one:
+ "Rep_preal 1 = {x. 0 < x \<and> x < 1}"
+by (simp add: preal_one_def rat_mem_preal)
+
+lemma subset_inverse_mult_lemma:
+ assumes xpos: "0 < x" and xless: "x < 1"
+ shows "\<exists>r u y. 0 < r & r < y & inverse y \<notin> Rep_preal R &
+ u \<in> Rep_preal R & x = r * u"
+proof -
+ from xpos and xless have "1 < inverse x" by (simp add: one_less_inverse_iff)
+ from lemma_gleason9_36 [OF Rep_preal this]
+ obtain r where r: "r \<in> Rep_preal R"
+ and notin: "r * (inverse x) \<notin> Rep_preal R" ..
+ have rpos: "0<r" by (rule preal_imp_pos [OF Rep_preal r])
+ from preal_exists_greater [OF Rep_preal r]
+ obtain u where u: "u \<in> Rep_preal R" and rless: "r < u" ..
+ have upos: "0<u" by (rule preal_imp_pos [OF Rep_preal u])
+ show ?thesis
+ proof (intro exI conjI)
+ show "0 < x/u" using xpos upos
+ by (simp add: zero_less_divide_iff)
+ show "x/u < x/r" using xpos upos rpos
+ by (simp add: divide_inverse mult_less_cancel_left rless)
+ show "inverse (x / r) \<notin> Rep_preal R" using notin
+ by (simp add: divide_inverse mult_commute)
+ show "u \<in> Rep_preal R" by (rule u)
+ show "x = x / u * u" using upos
+ by (simp add: divide_inverse mult_commute)
+ qed
+qed
+
+lemma subset_inverse_mult:
+ "Rep_preal 1 \<subseteq> Rep_preal(inverse R * R)"
+apply (auto simp add: Bex_def Rep_preal_one mem_Rep_preal_inverse_iff
+ mem_Rep_preal_mult_iff)
+apply (blast dest: subset_inverse_mult_lemma)
+done
+
+lemma inverse_mult_subset_lemma:
+ assumes rpos: "0 < r"
+ and rless: "r < y"
+ and notin: "inverse y \<notin> Rep_preal R"
+ and q: "q \<in> Rep_preal R"
+ shows "r*q < 1"
+proof -
+ have "q < inverse y" using rpos rless
+ by (simp add: not_in_preal_ub [OF Rep_preal notin] q)
+ hence "r * q < r/y" using rpos
+ by (simp add: divide_inverse mult_less_cancel_left)
+ also have "... \<le> 1" using rpos rless
+ by (simp add: pos_divide_le_eq)
+ finally show ?thesis .
+qed
+
+lemma inverse_mult_subset:
+ "Rep_preal(inverse R * R) \<subseteq> Rep_preal 1"
+apply (auto simp add: Bex_def Rep_preal_one mem_Rep_preal_inverse_iff
+ mem_Rep_preal_mult_iff)
+apply (simp add: zero_less_mult_iff preal_imp_pos [OF Rep_preal])
+apply (blast intro: inverse_mult_subset_lemma)
+done
+
+lemma preal_mult_inverse: "inverse R * R = (1::preal)"
+apply (rule Rep_preal_inject [THEN iffD1])
+apply (rule equalityI [OF inverse_mult_subset subset_inverse_mult])
+done
+
+lemma preal_mult_inverse_right: "R * inverse R = (1::preal)"
+apply (rule preal_mult_commute [THEN subst])
+apply (rule preal_mult_inverse)
+done
+
+
+text{*Theorems needing @{text Gleason9_34}*}
+
+lemma Rep_preal_self_subset: "Rep_preal (R) \<subseteq> Rep_preal(R + S)"
+proof
+ fix r
+ assume r: "r \<in> Rep_preal R"
+ have rpos: "0<r" by (rule preal_imp_pos [OF Rep_preal r])
+ from mem_Rep_preal_Ex
+ obtain y where y: "y \<in> Rep_preal S" ..
+ have ypos: "0<y" by (rule preal_imp_pos [OF Rep_preal y])
+ have ry: "r+y \<in> Rep_preal(R + S)" using r y
+ by (auto simp add: mem_Rep_preal_add_iff)
+ show "r \<in> Rep_preal(R + S)" using r ypos rpos
+ by (simp add: preal_downwards_closed [OF Rep_preal ry])
+qed
+
+lemma Rep_preal_sum_not_subset: "~ Rep_preal (R + S) \<subseteq> Rep_preal(R)"
+proof -
+ from mem_Rep_preal_Ex
+ obtain y where y: "y \<in> Rep_preal S" ..
+ have ypos: "0<y" by (rule preal_imp_pos [OF Rep_preal y])
+ from Gleason9_34 [OF Rep_preal ypos]
+ obtain r where r: "r \<in> Rep_preal R" and notin: "r + y \<notin> Rep_preal R" ..
+ have "r + y \<in> Rep_preal (R + S)" using r y
+ by (auto simp add: mem_Rep_preal_add_iff)
+ thus ?thesis using notin by blast
+qed
+
+lemma Rep_preal_sum_not_eq: "Rep_preal (R + S) \<noteq> Rep_preal(R)"
+by (insert Rep_preal_sum_not_subset, blast)
+
+text{*at last, Gleason prop. 9-3.5(iii) page 123*}
+lemma preal_self_less_add_left: "(R::preal) < R + S"
+apply (unfold preal_less_def less_le)
+apply (simp add: Rep_preal_self_subset Rep_preal_sum_not_eq [THEN not_sym])
+done
+
+
+subsection{*Subtraction for Positive Reals*}
+
+text{*Gleason prop. 9-3.5(iv), page 123: proving @{prop "A < B ==> \<exists>D. A + D =
+B"}. We define the claimed @{term D} and show that it is a positive real*}
+
+text{*Part 1 of Dedekind sections definition*}
+lemma diff_set_not_empty:
+ "R < S ==> {} \<subset> diff_set (Rep_preal S) (Rep_preal R)"
+apply (auto simp add: preal_less_def diff_set_def elim!: equalityE)
+apply (frule_tac x1 = S in Rep_preal [THEN preal_exists_greater])
+apply (drule preal_imp_pos [OF Rep_preal], clarify)
+apply (cut_tac a=x and b=u in add_eq_exists, force)
+done
+
+text{*Part 2 of Dedekind sections definition*}
+lemma diff_set_nonempty:
+ "\<exists>q. 0 < q & q \<notin> diff_set (Rep_preal S) (Rep_preal R)"
+apply (cut_tac X = S in Rep_preal_exists_bound)
+apply (erule exE)
+apply (rule_tac x = x in exI, auto)
+apply (simp add: diff_set_def)
+apply (auto dest: Rep_preal [THEN preal_downwards_closed])
+done
+
+lemma diff_set_not_rat_set:
+ "diff_set (Rep_preal S) (Rep_preal R) < {r. 0 < r}" (is "?lhs < ?rhs")
+proof
+ show "?lhs \<subseteq> ?rhs" by (auto simp add: diff_set_def)
+ show "?lhs \<noteq> ?rhs" using diff_set_nonempty by blast
+qed
+
+text{*Part 3 of Dedekind sections definition*}
+lemma diff_set_lemma3:
+ "[|R < S; u \<in> diff_set (Rep_preal S) (Rep_preal R); 0 < z; z < u|]
+ ==> z \<in> diff_set (Rep_preal S) (Rep_preal R)"
+apply (auto simp add: diff_set_def)
+apply (rule_tac x=x in exI)
+apply (drule Rep_preal [THEN preal_downwards_closed], auto)
+done
+
+text{*Part 4 of Dedekind sections definition*}
+lemma diff_set_lemma4:
+ "[|R < S; y \<in> diff_set (Rep_preal S) (Rep_preal R)|]
+ ==> \<exists>u \<in> diff_set (Rep_preal S) (Rep_preal R). y < u"
+apply (auto simp add: diff_set_def)
+apply (drule Rep_preal [THEN preal_exists_greater], clarify)
+apply (cut_tac a="x+y" and b=u in add_eq_exists, clarify)
+apply (rule_tac x="y+xa" in exI)
+apply (auto simp add: add_ac)
+done
+
+lemma mem_diff_set:
+ "R < S ==> diff_set (Rep_preal S) (Rep_preal R) \<in> preal"
+apply (unfold preal_def cut_def)
+apply (blast intro!: diff_set_not_empty diff_set_not_rat_set
+ diff_set_lemma3 diff_set_lemma4)
+done
+
+lemma mem_Rep_preal_diff_iff:
+ "R < S ==>
+ (z \<in> Rep_preal(S-R)) =
+ (\<exists>x. 0 < x & 0 < z & x \<notin> Rep_preal R & x + z \<in> Rep_preal S)"
+apply (simp add: preal_diff_def mem_diff_set Rep_preal)
+apply (force simp add: diff_set_def)
+done
+
+
+text{*proving that @{term "R + D \<le> S"}*}
+
+lemma less_add_left_lemma:
+ assumes Rless: "R < S"
+ and a: "a \<in> Rep_preal R"
+ and cb: "c + b \<in> Rep_preal S"
+ and "c \<notin> Rep_preal R"
+ and "0 < b"
+ and "0 < c"
+ shows "a + b \<in> Rep_preal S"
+proof -
+ have "0<a" by (rule preal_imp_pos [OF Rep_preal a])
+ moreover
+ have "a < c" using prems
+ by (blast intro: not_in_Rep_preal_ub )
+ ultimately show ?thesis using prems
+ by (simp add: preal_downwards_closed [OF Rep_preal cb])
+qed
+
+lemma less_add_left_le1:
+ "R < (S::preal) ==> R + (S-R) \<le> S"
+apply (auto simp add: Bex_def preal_le_def mem_Rep_preal_add_iff
+ mem_Rep_preal_diff_iff)
+apply (blast intro: less_add_left_lemma)
+done
+
+subsection{*proving that @{term "S \<le> R + D"} --- trickier*}
+
+lemma lemma_sum_mem_Rep_preal_ex:
+ "x \<in> Rep_preal S ==> \<exists>e. 0 < e & x + e \<in> Rep_preal S"
+apply (drule Rep_preal [THEN preal_exists_greater], clarify)
+apply (cut_tac a=x and b=u in add_eq_exists, auto)
+done
+
+lemma less_add_left_lemma2:
+ assumes Rless: "R < S"
+ and x: "x \<in> Rep_preal S"
+ and xnot: "x \<notin> Rep_preal R"
+ shows "\<exists>u v z. 0 < v & 0 < z & u \<in> Rep_preal R & z \<notin> Rep_preal R &
+ z + v \<in> Rep_preal S & x = u + v"
+proof -
+ have xpos: "0<x" by (rule preal_imp_pos [OF Rep_preal x])
+ from lemma_sum_mem_Rep_preal_ex [OF x]
+ obtain e where epos: "0 < e" and xe: "x + e \<in> Rep_preal S" by blast
+ from Gleason9_34 [OF Rep_preal epos]
+ obtain r where r: "r \<in> Rep_preal R" and notin: "r + e \<notin> Rep_preal R" ..
+ with x xnot xpos have rless: "r < x" by (blast intro: not_in_Rep_preal_ub)
+ from add_eq_exists [of r x]
+ obtain y where eq: "x = r+y" by auto
+ show ?thesis
+ proof (intro exI conjI)
+ show "r \<in> Rep_preal R" by (rule r)
+ show "r + e \<notin> Rep_preal R" by (rule notin)
+ show "r + e + y \<in> Rep_preal S" using xe eq by (simp add: add_ac)
+ show "x = r + y" by (simp add: eq)
+ show "0 < r + e" using epos preal_imp_pos [OF Rep_preal r]
+ by simp
+ show "0 < y" using rless eq by arith
+ qed
+qed
+
+lemma less_add_left_le2: "R < (S::preal) ==> S \<le> R + (S-R)"
+apply (auto simp add: preal_le_def)
+apply (case_tac "x \<in> Rep_preal R")
+apply (cut_tac Rep_preal_self_subset [of R], force)
+apply (auto simp add: Bex_def mem_Rep_preal_add_iff mem_Rep_preal_diff_iff)
+apply (blast dest: less_add_left_lemma2)
+done
+
+lemma less_add_left: "R < (S::preal) ==> R + (S-R) = S"
+by (blast intro: antisym [OF less_add_left_le1 less_add_left_le2])
+
+lemma less_add_left_Ex: "R < (S::preal) ==> \<exists>D. R + D = S"
+by (fast dest: less_add_left)
+
+lemma preal_add_less2_mono1: "R < (S::preal) ==> R + T < S + T"
+apply (auto dest!: less_add_left_Ex simp add: preal_add_assoc)
+apply (rule_tac y1 = D in preal_add_commute [THEN subst])
+apply (auto intro: preal_self_less_add_left simp add: preal_add_assoc [symmetric])
+done
+
+lemma preal_add_less2_mono2: "R < (S::preal) ==> T + R < T + S"
+by (auto intro: preal_add_less2_mono1 simp add: preal_add_commute [of T])
+
+lemma preal_add_right_less_cancel: "R + T < S + T ==> R < (S::preal)"
+apply (insert linorder_less_linear [of R S], auto)
+apply (drule_tac R = S and T = T in preal_add_less2_mono1)
+apply (blast dest: order_less_trans)
+done
+
+lemma preal_add_left_less_cancel: "T + R < T + S ==> R < (S::preal)"
+by (auto elim: preal_add_right_less_cancel simp add: preal_add_commute [of T])
+
+lemma preal_add_less_cancel_left: "(T + (R::preal) < T + S) = (R < S)"
+by (blast intro: preal_add_less2_mono2 preal_add_left_less_cancel)
+
+lemma preal_add_le_cancel_left: "(T + (R::preal) \<le> T + S) = (R \<le> S)"
+by (simp add: linorder_not_less [symmetric] preal_add_less_cancel_left)
+
+lemma preal_add_right_cancel: "(R::preal) + T = S + T ==> R = S"
+apply (insert linorder_less_linear [of R S], safe)
+apply (drule_tac [!] T = T in preal_add_less2_mono1, auto)
+done
+
+lemma preal_add_left_cancel: "C + A = C + B ==> A = (B::preal)"
+by (auto intro: preal_add_right_cancel simp add: preal_add_commute)
+
+instance preal :: linordered_cancel_ab_semigroup_add
+proof
+ fix a b c :: preal
+ show "a + b = a + c \<Longrightarrow> b = c" by (rule preal_add_left_cancel)
+ show "a \<le> b \<Longrightarrow> c + a \<le> c + b" by (simp only: preal_add_le_cancel_left)
+qed
+
+
+subsection{*Completeness of type @{typ preal}*}
+
+text{*Prove that supremum is a cut*}
+
+text{*Part 1 of Dedekind sections definition*}
+
+lemma preal_sup_set_not_empty:
+ "P \<noteq> {} ==> {} \<subset> (\<Union>X \<in> P. Rep_preal(X))"
+apply auto
+apply (cut_tac X = x in mem_Rep_preal_Ex, auto)
+done
+
+
+text{*Part 2 of Dedekind sections definition*}
+
+lemma preal_sup_not_exists:
+ "\<forall>X \<in> P. X \<le> Y ==> \<exists>q. 0 < q & q \<notin> (\<Union>X \<in> P. Rep_preal(X))"
+apply (cut_tac X = Y in Rep_preal_exists_bound)
+apply (auto simp add: preal_le_def)
+done
+
+lemma preal_sup_set_not_rat_set:
+ "\<forall>X \<in> P. X \<le> Y ==> (\<Union>X \<in> P. Rep_preal(X)) < {r. 0 < r}"
+apply (drule preal_sup_not_exists)
+apply (blast intro: preal_imp_pos [OF Rep_preal])
+done
+
+text{*Part 3 of Dedekind sections definition*}
+lemma preal_sup_set_lemma3:
+ "[|P \<noteq> {}; \<forall>X \<in> P. X \<le> Y; u \<in> (\<Union>X \<in> P. Rep_preal(X)); 0 < z; z < u|]
+ ==> z \<in> (\<Union>X \<in> P. Rep_preal(X))"
+by (auto elim: Rep_preal [THEN preal_downwards_closed])
+
+text{*Part 4 of Dedekind sections definition*}
+lemma preal_sup_set_lemma4:
+ "[|P \<noteq> {}; \<forall>X \<in> P. X \<le> Y; y \<in> (\<Union>X \<in> P. Rep_preal(X)) |]
+ ==> \<exists>u \<in> (\<Union>X \<in> P. Rep_preal(X)). y < u"
+by (blast dest: Rep_preal [THEN preal_exists_greater])
+
+lemma preal_sup:
+ "[|P \<noteq> {}; \<forall>X \<in> P. X \<le> Y|] ==> (\<Union>X \<in> P. Rep_preal(X)) \<in> preal"
+apply (unfold preal_def cut_def)
+apply (blast intro!: preal_sup_set_not_empty preal_sup_set_not_rat_set
+ preal_sup_set_lemma3 preal_sup_set_lemma4)
+done
+
+lemma preal_psup_le:
+ "[| \<forall>X \<in> P. X \<le> Y; x \<in> P |] ==> x \<le> psup P"
+apply (simp (no_asm_simp) add: preal_le_def)
+apply (subgoal_tac "P \<noteq> {}")
+apply (auto simp add: psup_def preal_sup)
+done
+
+lemma psup_le_ub: "[| P \<noteq> {}; \<forall>X \<in> P. X \<le> Y |] ==> psup P \<le> Y"
+apply (simp (no_asm_simp) add: preal_le_def)
+apply (simp add: psup_def preal_sup)
+apply (auto simp add: preal_le_def)
+done
+
+text{*Supremum property*}
+lemma preal_complete:
+ "[| P \<noteq> {}; \<forall>X \<in> P. X \<le> Y |] ==> (\<exists>X \<in> P. Z < X) = (Z < psup P)"
+apply (simp add: preal_less_def psup_def preal_sup)
+apply (auto simp add: preal_le_def)
+apply (rename_tac U)
+apply (cut_tac x = U and y = Z in linorder_less_linear)
+apply (auto simp add: preal_less_def)
+done
+
+section {*Defining the Reals from the Positive Reals*}
+
+definition
+ realrel :: "((preal * preal) * (preal * preal)) set" where
+ [code del]: "realrel = {p. \<exists>x1 y1 x2 y2. p = ((x1,y1),(x2,y2)) & x1+y2 = x2+y1}"
+
+typedef (Real) real = "UNIV//realrel"
+ by (auto simp add: quotient_def)
+
+definition
+ (** these don't use the overloaded "real" function: users don't see them **)
+ real_of_preal :: "preal => real" where
+ [code del]: "real_of_preal m = Abs_Real (realrel `` {(m + 1, 1)})"
+
+instantiation real :: "{zero, one, plus, minus, uminus, times, inverse, ord, abs, sgn}"
+begin
+
+definition
+ real_zero_def [code del]: "0 = Abs_Real(realrel``{(1, 1)})"
+
+definition
+ real_one_def [code del]: "1 = Abs_Real(realrel``{(1 + 1, 1)})"
+
+definition
+ real_add_def [code del]: "z + w =
+ contents (\<Union>(x,y) \<in> Rep_Real(z). \<Union>(u,v) \<in> Rep_Real(w).
+ { Abs_Real(realrel``{(x+u, y+v)}) })"
+
+definition
+ real_minus_def [code del]: "- r = contents (\<Union>(x,y) \<in> Rep_Real(r). { Abs_Real(realrel``{(y,x)}) })"
+
+definition
+ real_diff_def [code del]: "r - (s::real) = r + - s"
+
+definition
+ real_mult_def [code del]:
+ "z * w =
+ contents (\<Union>(x,y) \<in> Rep_Real(z). \<Union>(u,v) \<in> Rep_Real(w).
+ { Abs_Real(realrel``{(x*u + y*v, x*v + y*u)}) })"
+
+definition
+ real_inverse_def [code del]: "inverse (R::real) = (THE S. (R = 0 & S = 0) | S * R = 1)"
+
+definition
+ real_divide_def [code del]: "R / (S::real) = R * inverse S"
+
+definition
+ real_le_def [code del]: "z \<le> (w::real) \<longleftrightarrow>
+ (\<exists>x y u v. x+v \<le> u+y & (x,y) \<in> Rep_Real z & (u,v) \<in> Rep_Real w)"
+
+definition
+ real_less_def [code del]: "x < (y\<Colon>real) \<longleftrightarrow> x \<le> y \<and> x \<noteq> y"
+
+definition
+ real_abs_def: "abs (r::real) = (if r < 0 then - r else r)"
+
+definition
+ real_sgn_def: "sgn (x::real) = (if x=0 then 0 else if 0<x then 1 else - 1)"
+
+instance ..
+
+end
+
+subsection {* Equivalence relation over positive reals *}
+
+lemma preal_trans_lemma:
+ assumes "x + y1 = x1 + y"
+ and "x + y2 = x2 + y"
+ shows "x1 + y2 = x2 + (y1::preal)"
+proof -
+ have "(x1 + y2) + x = (x + y2) + x1" by (simp add: add_ac)
+ also have "... = (x2 + y) + x1" by (simp add: prems)
+ also have "... = x2 + (x1 + y)" by (simp add: add_ac)
+ also have "... = x2 + (x + y1)" by (simp add: prems)
+ also have "... = (x2 + y1) + x" by (simp add: add_ac)
+ finally have "(x1 + y2) + x = (x2 + y1) + x" .
+ thus ?thesis by (rule add_right_imp_eq)
+qed
+
+
+lemma realrel_iff [simp]: "(((x1,y1),(x2,y2)) \<in> realrel) = (x1 + y2 = x2 + y1)"
+by (simp add: realrel_def)
+
+lemma equiv_realrel: "equiv UNIV realrel"
+apply (auto simp add: equiv_def refl_on_def sym_def trans_def realrel_def)
+apply (blast dest: preal_trans_lemma)
+done
+
+text{*Reduces equality of equivalence classes to the @{term realrel} relation:
+ @{term "(realrel `` {x} = realrel `` {y}) = ((x,y) \<in> realrel)"} *}
+lemmas equiv_realrel_iff =
+ eq_equiv_class_iff [OF equiv_realrel UNIV_I UNIV_I]
+
+declare equiv_realrel_iff [simp]
+
+
+lemma realrel_in_real [simp]: "realrel``{(x,y)}: Real"
+by (simp add: Real_def realrel_def quotient_def, blast)
+
+declare Abs_Real_inject [simp]
+declare Abs_Real_inverse [simp]
+
+
+text{*Case analysis on the representation of a real number as an equivalence
+ class of pairs of positive reals.*}
+lemma eq_Abs_Real [case_names Abs_Real, cases type: real]:
+ "(!!x y. z = Abs_Real(realrel``{(x,y)}) ==> P) ==> P"
+apply (rule Rep_Real [of z, unfolded Real_def, THEN quotientE])
+apply (drule arg_cong [where f=Abs_Real])
+apply (auto simp add: Rep_Real_inverse)
+done
+
+
+subsection {* Addition and Subtraction *}
+
+lemma real_add_congruent2_lemma:
+ "[|a + ba = aa + b; ab + bc = ac + bb|]
+ ==> a + ab + (ba + bc) = aa + ac + (b + (bb::preal))"
+apply (simp add: add_assoc)
+apply (rule add_left_commute [of ab, THEN ssubst])
+apply (simp add: add_assoc [symmetric])
+apply (simp add: add_ac)
+done
+
+lemma real_add:
+ "Abs_Real (realrel``{(x,y)}) + Abs_Real (realrel``{(u,v)}) =
+ Abs_Real (realrel``{(x+u, y+v)})"
+proof -
+ have "(\<lambda>z w. (\<lambda>(x,y). (\<lambda>(u,v). {Abs_Real (realrel `` {(x+u, y+v)})}) w) z)
+ respects2 realrel"
+ by (simp add: congruent2_def, blast intro: real_add_congruent2_lemma)
+ thus ?thesis
+ by (simp add: real_add_def UN_UN_split_split_eq
+ UN_equiv_class2 [OF equiv_realrel equiv_realrel])
+qed
+
+lemma real_minus: "- Abs_Real(realrel``{(x,y)}) = Abs_Real(realrel `` {(y,x)})"
+proof -
+ have "(\<lambda>(x,y). {Abs_Real (realrel``{(y,x)})}) respects realrel"
+ by (simp add: congruent_def add_commute)
+ thus ?thesis
+ by (simp add: real_minus_def UN_equiv_class [OF equiv_realrel])
+qed
+
+instance real :: ab_group_add
+proof
+ fix x y z :: real
+ show "(x + y) + z = x + (y + z)"
+ by (cases x, cases y, cases z, simp add: real_add add_assoc)
+ show "x + y = y + x"
+ by (cases x, cases y, simp add: real_add add_commute)
+ show "0 + x = x"
+ by (cases x, simp add: real_add real_zero_def add_ac)
+ show "- x + x = 0"
+ by (cases x, simp add: real_minus real_add real_zero_def add_commute)
+ show "x - y = x + - y"
+ by (simp add: real_diff_def)
+qed
+
+
+subsection {* Multiplication *}
+
+lemma real_mult_congruent2_lemma:
+ "!!(x1::preal). [| x1 + y2 = x2 + y1 |] ==>
+ x * x1 + y * y1 + (x * y2 + y * x2) =
+ x * x2 + y * y2 + (x * y1 + y * x1)"
+apply (simp add: add_left_commute add_assoc [symmetric])
+apply (simp add: add_assoc right_distrib [symmetric])
+apply (simp add: add_commute)
+done
+
+lemma real_mult_congruent2:
+ "(%p1 p2.
+ (%(x1,y1). (%(x2,y2).
+ { Abs_Real (realrel``{(x1*x2 + y1*y2, x1*y2+y1*x2)}) }) p2) p1)
+ respects2 realrel"
+apply (rule congruent2_commuteI [OF equiv_realrel], clarify)
+apply (simp add: mult_commute add_commute)
+apply (auto simp add: real_mult_congruent2_lemma)
+done
+
+lemma real_mult:
+ "Abs_Real((realrel``{(x1,y1)})) * Abs_Real((realrel``{(x2,y2)})) =
+ Abs_Real(realrel `` {(x1*x2+y1*y2,x1*y2+y1*x2)})"
+by (simp add: real_mult_def UN_UN_split_split_eq
+ UN_equiv_class2 [OF equiv_realrel equiv_realrel real_mult_congruent2])
+
+lemma real_mult_commute: "(z::real) * w = w * z"
+by (cases z, cases w, simp add: real_mult add_ac mult_ac)
+
+lemma real_mult_assoc: "((z1::real) * z2) * z3 = z1 * (z2 * z3)"
+apply (cases z1, cases z2, cases z3)
+apply (simp add: real_mult algebra_simps)
+done
+
+lemma real_mult_1: "(1::real) * z = z"
+apply (cases z)
+apply (simp add: real_mult real_one_def algebra_simps)
+done
+
+lemma real_add_mult_distrib: "((z1::real) + z2) * w = (z1 * w) + (z2 * w)"
+apply (cases z1, cases z2, cases w)
+apply (simp add: real_add real_mult algebra_simps)
+done
+
+text{*one and zero are distinct*}
+lemma real_zero_not_eq_one: "0 \<noteq> (1::real)"
+proof -
+ have "(1::preal) < 1 + 1"
+ by (simp add: preal_self_less_add_left)
+ thus ?thesis
+ by (simp add: real_zero_def real_one_def)
+qed
+
+instance real :: comm_ring_1
+proof
+ fix x y z :: real
+ show "(x * y) * z = x * (y * z)" by (rule real_mult_assoc)
+ show "x * y = y * x" by (rule real_mult_commute)
+ show "1 * x = x" by (rule real_mult_1)
+ show "(x + y) * z = x * z + y * z" by (rule real_add_mult_distrib)
+ show "0 \<noteq> (1::real)" by (rule real_zero_not_eq_one)
+qed
+
+subsection {* Inverse and Division *}
+
+lemma real_zero_iff: "Abs_Real (realrel `` {(x, x)}) = 0"
+by (simp add: real_zero_def add_commute)
+
+text{*Instead of using an existential quantifier and constructing the inverse
+within the proof, we could define the inverse explicitly.*}
+
+lemma real_mult_inverse_left_ex: "x \<noteq> 0 ==> \<exists>y. y*x = (1::real)"
+apply (simp add: real_zero_def real_one_def, cases x)
+apply (cut_tac x = xa and y = y in linorder_less_linear)
+apply (auto dest!: less_add_left_Ex simp add: real_zero_iff)
+apply (rule_tac
+ x = "Abs_Real (realrel``{(1, inverse (D) + 1)})"
+ in exI)
+apply (rule_tac [2]
+ x = "Abs_Real (realrel``{(inverse (D) + 1, 1)})"
+ in exI)
+apply (auto simp add: real_mult preal_mult_inverse_right algebra_simps)
+done
+
+lemma real_mult_inverse_left: "x \<noteq> 0 ==> inverse(x)*x = (1::real)"
+apply (simp add: real_inverse_def)
+apply (drule real_mult_inverse_left_ex, safe)
+apply (rule theI, assumption, rename_tac z)
+apply (subgoal_tac "(z * x) * y = z * (x * y)")
+apply (simp add: mult_commute)
+apply (rule mult_assoc)
+done
+
+
+subsection{*The Real Numbers form a Field*}
+
+instance real :: field_inverse_zero
+proof
+ fix x y z :: real
+ show "x \<noteq> 0 ==> inverse x * x = 1" by (rule real_mult_inverse_left)
+ show "x / y = x * inverse y" by (simp add: real_divide_def)
+ show "inverse 0 = (0::real)" by (simp add: real_inverse_def)
+qed
+
+
+subsection{*The @{text "\<le>"} Ordering*}
+
+lemma real_le_refl: "w \<le> (w::real)"
+by (cases w, force simp add: real_le_def)
+
+text{*The arithmetic decision procedure is not set up for type preal.
+ This lemma is currently unused, but it could simplify the proofs of the
+ following two lemmas.*}
+lemma preal_eq_le_imp_le:
+ assumes eq: "a+b = c+d" and le: "c \<le> a"
+ shows "b \<le> (d::preal)"
+proof -
+ have "c+d \<le> a+d" by (simp add: prems)
+ hence "a+b \<le> a+d" by (simp add: prems)
+ thus "b \<le> d" by simp
+qed
+
+lemma real_le_lemma:
+ assumes l: "u1 + v2 \<le> u2 + v1"
+ and "x1 + v1 = u1 + y1"
+ and "x2 + v2 = u2 + y2"
+ shows "x1 + y2 \<le> x2 + (y1::preal)"
+proof -
+ have "(x1+v1) + (u2+y2) = (u1+y1) + (x2+v2)" by (simp add: prems)
+ hence "(x1+y2) + (u2+v1) = (x2+y1) + (u1+v2)" by (simp add: add_ac)
+ also have "... \<le> (x2+y1) + (u2+v1)" by (simp add: prems)
+ finally show ?thesis by simp
+qed
+
+lemma real_le:
+ "(Abs_Real(realrel``{(x1,y1)}) \<le> Abs_Real(realrel``{(x2,y2)})) =
+ (x1 + y2 \<le> x2 + y1)"
+apply (simp add: real_le_def)
+apply (auto intro: real_le_lemma)
+done
+
+lemma real_le_antisym: "[| z \<le> w; w \<le> z |] ==> z = (w::real)"
+by (cases z, cases w, simp add: real_le)
+
+lemma real_trans_lemma:
+ assumes "x + v \<le> u + y"
+ and "u + v' \<le> u' + v"
+ and "x2 + v2 = u2 + y2"
+ shows "x + v' \<le> u' + (y::preal)"
+proof -
+ have "(x+v') + (u+v) = (x+v) + (u+v')" by (simp add: add_ac)
+ also have "... \<le> (u+y) + (u+v')" by (simp add: prems)
+ also have "... \<le> (u+y) + (u'+v)" by (simp add: prems)
+ also have "... = (u'+y) + (u+v)" by (simp add: add_ac)
+ finally show ?thesis by simp
+qed
+
+lemma real_le_trans: "[| i \<le> j; j \<le> k |] ==> i \<le> (k::real)"
+apply (cases i, cases j, cases k)
+apply (simp add: real_le)
+apply (blast intro: real_trans_lemma)
+done
+
+instance real :: order
+proof
+ fix u v :: real
+ show "u < v \<longleftrightarrow> u \<le> v \<and> \<not> v \<le> u"
+ by (auto simp add: real_less_def intro: real_le_antisym)
+qed (assumption | rule real_le_refl real_le_trans real_le_antisym)+
+
+(* Axiom 'linorder_linear' of class 'linorder': *)
+lemma real_le_linear: "(z::real) \<le> w | w \<le> z"
+apply (cases z, cases w)
+apply (auto simp add: real_le real_zero_def add_ac)
+done
+
+instance real :: linorder
+ by (intro_classes, rule real_le_linear)
+
+
+lemma real_le_eq_diff: "(x \<le> y) = (x-y \<le> (0::real))"
+apply (cases x, cases y)
+apply (auto simp add: real_le real_zero_def real_diff_def real_add real_minus
+ add_ac)
+apply (simp_all add: add_assoc [symmetric])
+done
+
+lemma real_add_left_mono:
+ assumes le: "x \<le> y" shows "z + x \<le> z + (y::real)"
+proof -
+ have "z + x - (z + y) = (z + -z) + (x - y)"
+ by (simp add: algebra_simps)
+ with le show ?thesis
+ by (simp add: real_le_eq_diff[of x] real_le_eq_diff[of "z+x"] diff_minus)
+qed
+
+lemma real_sum_gt_zero_less: "(0 < S + (-W::real)) ==> (W < S)"
+by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S] diff_minus)
+
+lemma real_less_sum_gt_zero: "(W < S) ==> (0 < S + (-W::real))"
+by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S] diff_minus)
+
+lemma real_mult_order: "[| 0 < x; 0 < y |] ==> (0::real) < x * y"
+apply (cases x, cases y)
+apply (simp add: linorder_not_le [where 'a = real, symmetric]
+ linorder_not_le [where 'a = preal]
+ real_zero_def real_le real_mult)
+ --{*Reduce to the (simpler) @{text "\<le>"} relation *}
+apply (auto dest!: less_add_left_Ex
+ simp add: algebra_simps preal_self_less_add_left)
+done
+
+lemma real_mult_less_mono2: "[| (0::real) < z; x < y |] ==> z * x < z * y"
+apply (rule real_sum_gt_zero_less)
+apply (drule real_less_sum_gt_zero [of x y])
+apply (drule real_mult_order, assumption)
+apply (simp add: right_distrib)
+done
+
+instantiation real :: distrib_lattice
+begin
+
+definition
+ "(inf \<Colon> real \<Rightarrow> real \<Rightarrow> real) = min"
+
+definition
+ "(sup \<Colon> real \<Rightarrow> real \<Rightarrow> real) = max"
+
+instance
+ by default (auto simp add: inf_real_def sup_real_def min_max.sup_inf_distrib1)
+
+end
+
+
+subsection{*The Reals Form an Ordered Field*}
+
+instance real :: linordered_field_inverse_zero
+proof
+ fix x y z :: real
+ show "x \<le> y ==> z + x \<le> z + y" by (rule real_add_left_mono)
+ show "x < y ==> 0 < z ==> z * x < z * y" by (rule real_mult_less_mono2)
+ show "\<bar>x\<bar> = (if x < 0 then -x else x)" by (simp only: real_abs_def)
+ show "sgn x = (if x=0 then 0 else if 0<x then 1 else - 1)"
+ by (simp only: real_sgn_def)
+qed
+
+text{*The function @{term real_of_preal} requires many proofs, but it seems
+to be essential for proving completeness of the reals from that of the
+positive reals.*}
+
+lemma real_of_preal_add:
+ "real_of_preal ((x::preal) + y) = real_of_preal x + real_of_preal y"
+by (simp add: real_of_preal_def real_add algebra_simps)
+
+lemma real_of_preal_mult:
+ "real_of_preal ((x::preal) * y) = real_of_preal x* real_of_preal y"
+by (simp add: real_of_preal_def real_mult algebra_simps)
+
+
+text{*Gleason prop 9-4.4 p 127*}
+lemma real_of_preal_trichotomy:
+ "\<exists>m. (x::real) = real_of_preal m | x = 0 | x = -(real_of_preal m)"
+apply (simp add: real_of_preal_def real_zero_def, cases x)
+apply (auto simp add: real_minus add_ac)
+apply (cut_tac x = x and y = y in linorder_less_linear)
+apply (auto dest!: less_add_left_Ex simp add: add_assoc [symmetric])
+done
+
+lemma real_of_preal_leD:
+ "real_of_preal m1 \<le> real_of_preal m2 ==> m1 \<le> m2"
+by (simp add: real_of_preal_def real_le)
+
+lemma real_of_preal_lessI: "m1 < m2 ==> real_of_preal m1 < real_of_preal m2"
+by (auto simp add: real_of_preal_leD linorder_not_le [symmetric])
+
+lemma real_of_preal_lessD:
+ "real_of_preal m1 < real_of_preal m2 ==> m1 < m2"
+by (simp add: real_of_preal_def real_le linorder_not_le [symmetric])
+
+lemma real_of_preal_less_iff [simp]:
+ "(real_of_preal m1 < real_of_preal m2) = (m1 < m2)"
+by (blast intro: real_of_preal_lessI real_of_preal_lessD)
+
+lemma real_of_preal_le_iff:
+ "(real_of_preal m1 \<le> real_of_preal m2) = (m1 \<le> m2)"
+by (simp add: linorder_not_less [symmetric])
+
+lemma real_of_preal_zero_less: "0 < real_of_preal m"
+apply (insert preal_self_less_add_left [of 1 m])
+apply (auto simp add: real_zero_def real_of_preal_def
+ real_less_def real_le_def add_ac)
+apply (rule_tac x="m + 1" in exI, rule_tac x="1" in exI)
+apply (simp add: add_ac)
+done
+
+lemma real_of_preal_minus_less_zero: "- real_of_preal m < 0"
+by (simp add: real_of_preal_zero_less)
+
+lemma real_of_preal_not_minus_gt_zero: "~ 0 < - real_of_preal m"
+proof -
+ from real_of_preal_minus_less_zero
+ show ?thesis by (blast dest: order_less_trans)
+qed
+
+
+subsection{*Theorems About the Ordering*}
+
+lemma real_gt_zero_preal_Ex: "(0 < x) = (\<exists>y. x = real_of_preal y)"
+apply (auto simp add: real_of_preal_zero_less)
+apply (cut_tac x = x in real_of_preal_trichotomy)
+apply (blast elim!: real_of_preal_not_minus_gt_zero [THEN notE])
+done
+
+lemma real_gt_preal_preal_Ex:
+ "real_of_preal z < x ==> \<exists>y. x = real_of_preal y"
+by (blast dest!: real_of_preal_zero_less [THEN order_less_trans]
+ intro: real_gt_zero_preal_Ex [THEN iffD1])
+
+lemma real_ge_preal_preal_Ex:
+ "real_of_preal z \<le> x ==> \<exists>y. x = real_of_preal y"
+by (blast dest: order_le_imp_less_or_eq real_gt_preal_preal_Ex)
+
+lemma real_less_all_preal: "y \<le> 0 ==> \<forall>x. y < real_of_preal x"
+by (auto elim: order_le_imp_less_or_eq [THEN disjE]
+ intro: real_of_preal_zero_less [THEN [2] order_less_trans]
+ simp add: real_of_preal_zero_less)
+
+lemma real_less_all_real2: "~ 0 < y ==> \<forall>x. y < real_of_preal x"
+by (blast intro!: real_less_all_preal linorder_not_less [THEN iffD1])
+
+
+subsection{*Numerals and Arithmetic*}
+
+instantiation real :: number_ring
+begin
+
+definition
+ real_number_of_def [code del]: "(number_of w :: real) = of_int w"
+
+instance
+ by intro_classes (simp add: real_number_of_def)
+
+end
+
+subsection {* Completeness of Positive Reals *}
+
+text {*
+ Supremum property for the set of positive reals
+
+ Let @{text "P"} be a non-empty set of positive reals, with an upper
+ bound @{text "y"}. Then @{text "P"} has a least upper bound
+ (written @{text "S"}).
+
+ FIXME: Can the premise be weakened to @{text "\<forall>x \<in> P. x\<le> y"}?
+*}
+
+lemma posreal_complete:
+ assumes positive_P: "\<forall>x \<in> P. (0::real) < x"
+ and not_empty_P: "\<exists>x. x \<in> P"
+ and upper_bound_Ex: "\<exists>y. \<forall>x \<in> P. x<y"
+ shows "\<exists>S. \<forall>y. (\<exists>x \<in> P. y < x) = (y < S)"
+proof (rule exI, rule allI)
+ fix y
+ let ?pP = "{w. real_of_preal w \<in> P}"
+
+ show "(\<exists>x\<in>P. y < x) = (y < real_of_preal (psup ?pP))"
+ proof (cases "0 < y")
+ assume neg_y: "\<not> 0 < y"
+ show ?thesis
+ proof
+ assume "\<exists>x\<in>P. y < x"
+ have "\<forall>x. y < real_of_preal x"
+ using neg_y by (rule real_less_all_real2)
+ thus "y < real_of_preal (psup ?pP)" ..
+ next
+ assume "y < real_of_preal (psup ?pP)"
+ obtain "x" where x_in_P: "x \<in> P" using not_empty_P ..
+ hence "0 < x" using positive_P by simp
+ hence "y < x" using neg_y by simp
+ thus "\<exists>x \<in> P. y < x" using x_in_P ..
+ qed
+ next
+ assume pos_y: "0 < y"
+
+ then obtain py where y_is_py: "y = real_of_preal py"
+ by (auto simp add: real_gt_zero_preal_Ex)
+
+ obtain a where "a \<in> P" using not_empty_P ..
+ with positive_P have a_pos: "0 < a" ..
+ then obtain pa where "a = real_of_preal pa"
+ by (auto simp add: real_gt_zero_preal_Ex)
+ hence "pa \<in> ?pP" using `a \<in> P` by auto
+ hence pP_not_empty: "?pP \<noteq> {}" by auto
+
+ obtain sup where sup: "\<forall>x \<in> P. x < sup"
+ using upper_bound_Ex ..
+ from this and `a \<in> P` have "a < sup" ..
+ hence "0 < sup" using a_pos by arith
+ then obtain possup where "sup = real_of_preal possup"
+ by (auto simp add: real_gt_zero_preal_Ex)
+ hence "\<forall>X \<in> ?pP. X \<le> possup"
+ using sup by (auto simp add: real_of_preal_lessI)
+ with pP_not_empty have psup: "\<And>Z. (\<exists>X \<in> ?pP. Z < X) = (Z < psup ?pP)"
+ by (rule preal_complete)
+
+ show ?thesis
+ proof
+ assume "\<exists>x \<in> P. y < x"
+ then obtain x where x_in_P: "x \<in> P" and y_less_x: "y < x" ..
+ hence "0 < x" using pos_y by arith
+ then obtain px where x_is_px: "x = real_of_preal px"
+ by (auto simp add: real_gt_zero_preal_Ex)
+
+ have py_less_X: "\<exists>X \<in> ?pP. py < X"
+ proof
+ show "py < px" using y_is_py and x_is_px and y_less_x
+ by (simp add: real_of_preal_lessI)
+ show "px \<in> ?pP" using x_in_P and x_is_px by simp
+ qed
+
+ have "(\<exists>X \<in> ?pP. py < X) ==> (py < psup ?pP)"
+ using psup by simp
+ hence "py < psup ?pP" using py_less_X by simp
+ thus "y < real_of_preal (psup {w. real_of_preal w \<in> P})"
+ using y_is_py and pos_y by (simp add: real_of_preal_lessI)
+ next
+ assume y_less_psup: "y < real_of_preal (psup ?pP)"
+
+ hence "py < psup ?pP" using y_is_py
+ by (simp add: real_of_preal_lessI)
+ then obtain "X" where py_less_X: "py < X" and X_in_pP: "X \<in> ?pP"
+ using psup by auto
+ then obtain x where x_is_X: "x = real_of_preal X"
+ by (simp add: real_gt_zero_preal_Ex)
+ hence "y < x" using py_less_X and y_is_py
+ by (simp add: real_of_preal_lessI)
+
+ moreover have "x \<in> P" using x_is_X and X_in_pP by simp
+
+ ultimately show "\<exists> x \<in> P. y < x" ..
+ qed
+ qed
+qed
+
+text {*
+ \medskip Completeness properties using @{text "isUb"}, @{text "isLub"} etc.
+*}
+
+lemma posreals_complete:
+ assumes positive_S: "\<forall>x \<in> S. 0 < x"
+ and not_empty_S: "\<exists>x. x \<in> S"
+ and upper_bound_Ex: "\<exists>u. isUb (UNIV::real set) S u"
+ shows "\<exists>t. isLub (UNIV::real set) S t"
+proof
+ let ?pS = "{w. real_of_preal w \<in> S}"
+
+ obtain u where "isUb UNIV S u" using upper_bound_Ex ..
+ hence sup: "\<forall>x \<in> S. x \<le> u" by (simp add: isUb_def setle_def)
+
+ obtain x where x_in_S: "x \<in> S" using not_empty_S ..
+ hence x_gt_zero: "0 < x" using positive_S by simp
+ have "x \<le> u" using sup and x_in_S ..
+ hence "0 < u" using x_gt_zero by arith
+
+ then obtain pu where u_is_pu: "u = real_of_preal pu"
+ by (auto simp add: real_gt_zero_preal_Ex)
+
+ have pS_less_pu: "\<forall>pa \<in> ?pS. pa \<le> pu"
+ proof
+ fix pa
+ assume "pa \<in> ?pS"
+ then obtain a where "a \<in> S" and "a = real_of_preal pa"
+ by simp
+ moreover hence "a \<le> u" using sup by simp
+ ultimately show "pa \<le> pu"
+ using sup and u_is_pu by (simp add: real_of_preal_le_iff)
+ qed
+
+ have "\<forall>y \<in> S. y \<le> real_of_preal (psup ?pS)"
+ proof
+ fix y
+ assume y_in_S: "y \<in> S"
+ hence "0 < y" using positive_S by simp
+ then obtain py where y_is_py: "y = real_of_preal py"
+ by (auto simp add: real_gt_zero_preal_Ex)
+ hence py_in_pS: "py \<in> ?pS" using y_in_S by simp
+ with pS_less_pu have "py \<le> psup ?pS"
+ by (rule preal_psup_le)
+ thus "y \<le> real_of_preal (psup ?pS)"
+ using y_is_py by (simp add: real_of_preal_le_iff)
+ qed
+
+ moreover {
+ fix x
+ assume x_ub_S: "\<forall>y\<in>S. y \<le> x"
+ have "real_of_preal (psup ?pS) \<le> x"
+ proof -
+ obtain "s" where s_in_S: "s \<in> S" using not_empty_S ..
+ hence s_pos: "0 < s" using positive_S by simp
+
+ hence "\<exists> ps. s = real_of_preal ps" by (simp add: real_gt_zero_preal_Ex)
+ then obtain "ps" where s_is_ps: "s = real_of_preal ps" ..
+ hence ps_in_pS: "ps \<in> {w. real_of_preal w \<in> S}" using s_in_S by simp
+
+ from x_ub_S have "s \<le> x" using s_in_S ..
+ hence "0 < x" using s_pos by simp
+ hence "\<exists> px. x = real_of_preal px" by (simp add: real_gt_zero_preal_Ex)
+ then obtain "px" where x_is_px: "x = real_of_preal px" ..
+
+ have "\<forall>pe \<in> ?pS. pe \<le> px"
+ proof
+ fix pe
+ assume "pe \<in> ?pS"
+ hence "real_of_preal pe \<in> S" by simp
+ hence "real_of_preal pe \<le> x" using x_ub_S by simp
+ thus "pe \<le> px" using x_is_px by (simp add: real_of_preal_le_iff)
+ qed
+
+ moreover have "?pS \<noteq> {}" using ps_in_pS by auto
+ ultimately have "(psup ?pS) \<le> px" by (simp add: psup_le_ub)
+ thus "real_of_preal (psup ?pS) \<le> x" using x_is_px by (simp add: real_of_preal_le_iff)
+ qed
+ }
+ ultimately show "isLub UNIV S (real_of_preal (psup ?pS))"
+ by (simp add: isLub_def leastP_def isUb_def setle_def setge_def)
+qed
+
+text {*
+ \medskip reals Completeness (again!)
+*}
+
+lemma reals_complete:
+ assumes notempty_S: "\<exists>X. X \<in> S"
+ and exists_Ub: "\<exists>Y. isUb (UNIV::real set) S Y"
+ shows "\<exists>t. isLub (UNIV :: real set) S t"
+proof -
+ obtain X where X_in_S: "X \<in> S" using notempty_S ..
+ obtain Y where Y_isUb: "isUb (UNIV::real set) S Y"
+ using exists_Ub ..
+ let ?SHIFT = "{z. \<exists>x \<in>S. z = x + (-X) + 1} \<inter> {x. 0 < x}"
+
+ {
+ fix x
+ assume "isUb (UNIV::real set) S x"
+ hence S_le_x: "\<forall> y \<in> S. y <= x"
+ by (simp add: isUb_def setle_def)
+ {
+ fix s
+ assume "s \<in> {z. \<exists>x\<in>S. z = x + - X + 1}"
+ hence "\<exists> x \<in> S. s = x + -X + 1" ..
+ then obtain x1 where "x1 \<in> S" and "s = x1 + (-X) + 1" ..
+ moreover hence "x1 \<le> x" using S_le_x by simp
+ ultimately have "s \<le> x + - X + 1" by arith
+ }
+ then have "isUb (UNIV::real set) ?SHIFT (x + (-X) + 1)"
+ by (auto simp add: isUb_def setle_def)
+ } note S_Ub_is_SHIFT_Ub = this
+
+ hence "isUb UNIV ?SHIFT (Y + (-X) + 1)" using Y_isUb by simp
+ hence "\<exists>Z. isUb UNIV ?SHIFT Z" ..
+ moreover have "\<forall>y \<in> ?SHIFT. 0 < y" by auto
+ moreover have shifted_not_empty: "\<exists>u. u \<in> ?SHIFT"
+ using X_in_S and Y_isUb by auto
+ ultimately obtain t where t_is_Lub: "isLub UNIV ?SHIFT t"
+ using posreals_complete [of ?SHIFT] by blast
+
+ show ?thesis
+ proof
+ show "isLub UNIV S (t + X + (-1))"
+ proof (rule isLubI2)
+ {
+ fix x
+ assume "isUb (UNIV::real set) S x"
+ hence "isUb (UNIV::real set) (?SHIFT) (x + (-X) + 1)"
+ using S_Ub_is_SHIFT_Ub by simp
+ hence "t \<le> (x + (-X) + 1)"
+ using t_is_Lub by (simp add: isLub_le_isUb)
+ hence "t + X + -1 \<le> x" by arith
+ }
+ then show "(t + X + -1) <=* Collect (isUb UNIV S)"
+ by (simp add: setgeI)
+ next
+ show "isUb UNIV S (t + X + -1)"
+ proof -
+ {
+ fix y
+ assume y_in_S: "y \<in> S"
+ have "y \<le> t + X + -1"
+ proof -
+ obtain "u" where u_in_shift: "u \<in> ?SHIFT" using shifted_not_empty ..
+ hence "\<exists> x \<in> S. u = x + - X + 1" by simp
+ then obtain "x" where x_and_u: "u = x + - X + 1" ..
+ have u_le_t: "u \<le> t" using u_in_shift and t_is_Lub by (simp add: isLubD2)
+
+ show ?thesis
+ proof cases
+ assume "y \<le> x"
+ moreover have "x = u + X + - 1" using x_and_u by arith
+ moreover have "u + X + - 1 \<le> t + X + -1" using u_le_t by arith
+ ultimately show "y \<le> t + X + -1" by arith
+ next
+ assume "~(y \<le> x)"
+ hence x_less_y: "x < y" by arith
+
+ have "x + (-X) + 1 \<in> ?SHIFT" using x_and_u and u_in_shift by simp
+ hence "0 < x + (-X) + 1" by simp
+ hence "0 < y + (-X) + 1" using x_less_y by arith
+ hence "y + (-X) + 1 \<in> ?SHIFT" using y_in_S by simp
+ hence "y + (-X) + 1 \<le> t" using t_is_Lub by (simp add: isLubD2)
+ thus ?thesis by simp
+ qed
+ qed
+ }
+ then show ?thesis by (simp add: isUb_def setle_def)
+ qed
+ qed
+ qed
+qed
+
+text{*A version of the same theorem without all those predicates!*}
+lemma reals_complete2:
+ fixes S :: "(real set)"
+ assumes "\<exists>y. y\<in>S" and "\<exists>(x::real). \<forall>y\<in>S. y \<le> x"
+ shows "\<exists>x. (\<forall>y\<in>S. y \<le> x) &
+ (\<forall>z. ((\<forall>y\<in>S. y \<le> z) --> x \<le> z))"
+proof -
+ have "\<exists>x. isLub UNIV S x"
+ by (rule reals_complete)
+ (auto simp add: isLub_def isUb_def leastP_def setle_def setge_def prems)
+ thus ?thesis
+ by (metis UNIV_I isLub_isUb isLub_le_isUb isUbD isUb_def setleI)
+qed
+
+
+subsection {* The Archimedean Property of the Reals *}
+
+theorem reals_Archimedean:
+ fixes x :: real
+ assumes x_pos: "0 < x"
+ shows "\<exists>n. inverse (of_nat (Suc n)) < x"
+proof (rule ccontr)
+ assume contr: "\<not> ?thesis"
+ have "\<forall>n. x * of_nat (Suc n) <= 1"
+ proof
+ fix n
+ from contr have "x \<le> inverse (of_nat (Suc n))"
+ by (simp add: linorder_not_less)
+ hence "x \<le> (1 / (of_nat (Suc n)))"
+ by (simp add: inverse_eq_divide)
+ moreover have "(0::real) \<le> of_nat (Suc n)"
+ by (rule of_nat_0_le_iff)
+ ultimately have "x * of_nat (Suc n) \<le> (1 / of_nat (Suc n)) * of_nat (Suc n)"
+ by (rule mult_right_mono)
+ thus "x * of_nat (Suc n) \<le> 1" by (simp del: of_nat_Suc)
+ qed
+ hence "{z. \<exists>n. z = x * (of_nat (Suc n))} *<= 1"
+ by (simp add: setle_def del: of_nat_Suc, safe, rule spec)
+ hence "isUb (UNIV::real set) {z. \<exists>n. z = x * (of_nat (Suc n))} 1"
+ by (simp add: isUbI)
+ hence "\<exists>Y. isUb (UNIV::real set) {z. \<exists>n. z = x* (of_nat (Suc n))} Y" ..
+ moreover have "\<exists>X. X \<in> {z. \<exists>n. z = x* (of_nat (Suc n))}" by auto
+ ultimately have "\<exists>t. isLub UNIV {z. \<exists>n. z = x * of_nat (Suc n)} t"
+ by (simp add: reals_complete)
+ then obtain "t" where
+ t_is_Lub: "isLub UNIV {z. \<exists>n. z = x * of_nat (Suc n)} t" ..
+
+ have "\<forall>n::nat. x * of_nat n \<le> t + - x"
+ proof
+ fix n
+ from t_is_Lub have "x * of_nat (Suc n) \<le> t"
+ by (simp add: isLubD2)
+ hence "x * (of_nat n) + x \<le> t"
+ by (simp add: right_distrib)
+ thus "x * (of_nat n) \<le> t + - x" by arith
+ qed
+
+ hence "\<forall>m. x * of_nat (Suc m) \<le> t + - x" by (simp del: of_nat_Suc)
+ hence "{z. \<exists>n. z = x * (of_nat (Suc n))} *<= (t + - x)"
+ by (auto simp add: setle_def)
+ hence "isUb (UNIV::real set) {z. \<exists>n. z = x * (of_nat (Suc n))} (t + (-x))"
+ by (simp add: isUbI)
+ hence "t \<le> t + - x"
+ using t_is_Lub by (simp add: isLub_le_isUb)
+ thus False using x_pos by arith
+qed
+
+text {*
+ There must be other proofs, e.g. @{text "Suc"} of the largest
+ integer in the cut representing @{text "x"}.
+*}
+
+lemma reals_Archimedean2: "\<exists>n. (x::real) < of_nat (n::nat)"
+proof cases
+ assume "x \<le> 0"
+ hence "x < of_nat (1::nat)" by simp
+ thus ?thesis ..
+next
+ assume "\<not> x \<le> 0"
+ hence x_greater_zero: "0 < x" by simp
+ hence "0 < inverse x" by simp
+ then obtain n where "inverse (of_nat (Suc n)) < inverse x"
+ using reals_Archimedean by blast
+ hence "inverse (of_nat (Suc n)) * x < inverse x * x"
+ using x_greater_zero by (rule mult_strict_right_mono)
+ hence "inverse (of_nat (Suc n)) * x < 1"
+ using x_greater_zero by simp
+ hence "of_nat (Suc n) * (inverse (of_nat (Suc n)) * x) < of_nat (Suc n) * 1"
+ by (rule mult_strict_left_mono) (simp del: of_nat_Suc)
+ hence "x < of_nat (Suc n)"
+ by (simp add: algebra_simps del: of_nat_Suc)
+ thus "\<exists>(n::nat). x < of_nat n" ..
+qed
+
+instance real :: archimedean_field
+proof
+ fix r :: real
+ obtain n :: nat where "r < of_nat n"
+ using reals_Archimedean2 ..
+ then have "r \<le> of_int (int n)"
+ by simp
+ then show "\<exists>z. r \<le> of_int z" ..
+qed
+
+end
--- a/src/HOL/ex/ROOT.ML Tue May 11 08:29:42 2010 +0200
+++ b/src/HOL/ex/ROOT.ML Tue May 11 08:30:02 2010 +0200
@@ -65,7 +65,8 @@
"Landau",
"Execute_Choice",
"Summation",
- "Gauge_Integration"
+ "Gauge_Integration",
+ "Dedekind_Real"
];
HTML.with_charset "utf-8" (no_document use_thys)
--- a/src/Pure/General/linear_set.scala Tue May 11 08:29:42 2010 +0200
+++ b/src/Pure/General/linear_set.scala Tue May 11 08:30:02 2010 +0200
@@ -8,10 +8,11 @@
package isabelle
import scala.collection.SetLike
-import scala.collection.generic.{SetFactory, CanBuildFrom, GenericSetTemplate, GenericCompanion}
+import scala.collection.generic.{ImmutableSetFactory, CanBuildFrom,
+ GenericSetTemplate, GenericCompanion}
-object Linear_Set extends SetFactory[Linear_Set]
+object Linear_Set extends ImmutableSetFactory[Linear_Set]
{
private case class Rep[A](
val start: Option[A], val end: Option[A], val nexts: Map[A, A], prevs: Map[A, A])
--- a/src/Pure/General/pretty.scala Tue May 11 08:29:42 2010 +0200
+++ b/src/Pure/General/pretty.scala Tue May 11 08:30:02 2010 +0200
@@ -30,7 +30,8 @@
object Break
{
def apply(width: Int): XML.Tree =
- XML.Elem(Markup.BREAK, List((Markup.WIDTH, width.toString)), List(XML.Text(" " * width)))
+ XML.Elem(Markup.BREAK, List((Markup.WIDTH, width.toString)),
+ List(XML.Text(Symbol.spaces(width))))
def unapply(tree: XML.Tree): Option[Int] =
tree match {
@@ -48,7 +49,7 @@
{
def newline: Text = copy(tx = FBreak :: tx, pos = 0, nl = nl + 1)
def string(s: String): Text = copy(tx = XML.Text(s) :: tx, pos = pos + s.length)
- def blanks(wd: Int): Text = string(" " * wd)
+ def blanks(wd: Int): Text = string(Symbol.spaces(wd))
def content: List[XML.Tree] = tx.reverse
}
@@ -126,7 +127,7 @@
def fmt(tree: XML.Tree): List[XML.Tree] =
tree match {
case Block(_, body) => body.flatMap(fmt)
- case Break(wd) => List(XML.Text(" " * wd))
+ case Break(wd) => List(XML.Text(Symbol.spaces(wd)))
case FBreak => List(XML.Text(" "))
case XML.Elem(name, atts, body) => List(XML.Elem(name, atts, body.flatMap(fmt)))
case XML.Text(_) => List(tree)
--- a/src/Pure/General/symbol.scala Tue May 11 08:29:42 2010 +0200
+++ b/src/Pure/General/symbol.scala Tue May 11 08:30:02 2010 +0200
@@ -13,6 +13,18 @@
object Symbol
{
+ /* spaces */
+
+ private val static_spaces = " " * 4000
+
+ def spaces(k: Int): String =
+ {
+ require(k >= 0)
+ if (k < static_spaces.length) static_spaces.substring(0, k)
+ else " " * k
+ }
+
+
/* Symbol regexps */
private val plain = new Regex("""(?xs)
--- a/src/Pure/System/gui_setup.scala Tue May 11 08:29:42 2010 +0200
+++ b/src/Pure/System/gui_setup.scala Tue May 11 08:30:02 2010 +0200
@@ -6,8 +6,6 @@
package isabelle
-import javax.swing.UIManager
-
import scala.swing._
import scala.swing.event._
@@ -16,7 +14,7 @@
{
def startup(args: Array[String]) =
{
- UIManager.setLookAndFeel(Platform.look_and_feel)
+ Platform.init_laf()
top.pack()
top.visible = true
}
--- a/src/Pure/System/isabelle_system.scala Tue May 11 08:29:42 2010 +0200
+++ b/src/Pure/System/isabelle_system.scala Tue May 11 08:30:02 2010 +0200
@@ -306,7 +306,7 @@
private def read_symbols(path: String): List[String] =
{
val file = platform_file(path)
- if (file.isFile) Source.fromFile(file).getLines().toList
+ if (file.isFile) Source.fromFile(file).getLines("\n").toList
else Nil
}
val symbols = new Symbol.Interpretation(
@@ -318,8 +318,8 @@
val font_family = "IsabelleText"
- def get_font(bold: Boolean): Font =
- new Font(font_family, if (bold) Font.BOLD else Font.PLAIN, 1)
+ def get_font(bold: Boolean = false, size: Int = 1): Font =
+ new Font(font_family, if (bold) Font.BOLD else Font.PLAIN, size)
def install_fonts()
{
--- a/src/Pure/System/platform.scala Tue May 11 08:29:42 2010 +0200
+++ b/src/Pure/System/platform.scala Tue May 11 08:30:02 2010 +0200
@@ -58,12 +58,14 @@
private def find_laf(name: String): Option[String] =
UIManager.getInstalledLookAndFeels().find(_.getName == name).map(_.getClassName)
- def look_and_feel(): String =
+ def get_laf(): String =
{
if (is_windows || is_macos) UIManager.getSystemLookAndFeelClassName()
else
find_laf("Nimbus") orElse find_laf("GTK+") getOrElse
UIManager.getCrossPlatformLookAndFeelClassName()
}
+
+ def init_laf(): Unit = UIManager.setLookAndFeel(get_laf())
}
--- a/src/Pure/System/session.scala Tue May 11 08:29:42 2010 +0200
+++ b/src/Pure/System/session.scala Tue May 11 08:30:02 2010 +0200
@@ -227,6 +227,8 @@
case result: Isabelle_Process.Result =>
handle_result(result.cache(xml_cache))
+ case TIMEOUT => // FIXME clarify!
+
case bad if prover != null =>
System.err.println("session_actor: ignoring bad message " + bad)
}
--- a/src/Pure/axclass.ML Tue May 11 08:29:42 2010 +0200
+++ b/src/Pure/axclass.ML Tue May 11 08:30:02 2010 +0200
@@ -423,7 +423,7 @@
val (c1, c2) = cert_classrel thy rel handle TYPE _ => err ();
val th' = th
|> Drule.instantiate' [SOME (ctyp_of thy (TVar ((Name.aT, 0), [c1])))] []
- |> Thm.unconstrain_allTs;
+ |> Thm.unconstrainT;
val _ = shyps_topped th' orelse raise Fail "add_classrel: nontop shyps after unconstrain";
in
thy
@@ -449,7 +449,7 @@
|> (map o apsnd o map_atyps) (K T);
val th' = th
|> Drule.instantiate' std_vars []
- |> Thm.unconstrain_allTs;
+ |> Thm.unconstrainT;
val _ = shyps_topped th' orelse raise Fail "add_arity: nontop shyps after unconstrain";
in
thy
--- a/src/Pure/logic.ML Tue May 11 08:29:42 2010 +0200
+++ b/src/Pure/logic.ML Tue May 11 08:30:02 2010 +0200
@@ -46,6 +46,7 @@
val name_arity: string * sort list * class -> string
val mk_arities: arity -> term list
val dest_arity: term -> string * sort list * class
+ val unconstrainT: sort list -> term -> ((typ -> typ) * ((typ * class) * term) list) * term
val protectC: term
val protect: term -> term
val unprotect: term -> term
@@ -269,6 +270,42 @@
in (t, Ss, c) end;
+(* internalized sort constraints *)
+
+fun unconstrainT shyps prop =
+ let
+ val present = rev ((fold_types o fold_atyps_sorts) (insert (eq_fst op =)) prop []);
+ val extra = fold (Sorts.remove_sort o #2) present shyps;
+
+ val n = length present;
+ val (names1, names2) = Name.invents Name.context Name.aT (n + length extra) |> chop n;
+
+ val present_map =
+ map2 (fn (T, S) => fn a => (T, TVar ((a, 0), S))) present names1;
+ val constraints_map =
+ map2 (fn (_, S) => fn a => (S, TVar ((a, 0), S))) present names1 @
+ map2 (fn S => fn a => (S, TVar ((a, 0), S))) extra names2;
+
+ fun atyp_map T =
+ (case AList.lookup (op =) present_map T of
+ SOME U => U
+ | NONE =>
+ (case AList.lookup (op =) constraints_map (Type.sort_of_atyp T) of
+ SOME U => U
+ | NONE => raise TYPE ("Dangling type variable", [T], [])));
+
+ val constraints =
+ maps (fn (_, T as TVar (ai, S)) =>
+ map (fn c => ((T, c), mk_of_class (TVar (ai, []), c))) S)
+ constraints_map;
+
+ val prop' =
+ prop
+ |> (Term.map_types o Term.map_atyps) (Type.strip_sorts o atyp_map)
+ |> curry list_implies (map snd constraints);
+ in ((atyp_map, constraints), prop') end;
+
+
(** protected propositions and embedded terms **)
--- a/src/Pure/term_subst.ML Tue May 11 08:29:42 2010 +0200
+++ b/src/Pure/term_subst.ML Tue May 11 08:30:02 2010 +0200
@@ -9,26 +9,20 @@
val map_atypsT_same: typ Same.operation -> typ Same.operation
val map_types_same: typ Same.operation -> term Same.operation
val map_aterms_same: term Same.operation -> term Same.operation
- val map_atypsT_option: (typ -> typ option) -> typ -> typ option
- val map_atyps_option: (typ -> typ option) -> term -> term option
- val map_types_option: (typ -> typ option) -> term -> term option
- val map_aterms_option: (term -> term option) -> term -> term option
val generalizeT_same: string list -> int -> typ Same.operation
val generalize_same: string list * string list -> int -> term Same.operation
+ val generalizeT: string list -> int -> typ -> typ
val generalize: string list * string list -> int -> term -> term
- val generalizeT: string list -> int -> typ -> typ
- val generalize_option: string list * string list -> int -> term -> term option
- val generalizeT_option: string list -> int -> typ -> typ option
val instantiateT_maxidx: ((indexname * sort) * (typ * int)) list -> typ -> int -> typ * int
val instantiate_maxidx:
((indexname * sort) * (typ * int)) list * ((indexname * typ) * (term * int)) list ->
term -> int -> term * int
+ val instantiateT_same: ((indexname * sort) * typ) list -> typ Same.operation
+ val instantiate_same: ((indexname * sort) * typ) list * ((indexname * typ) * term) list ->
+ term Same.operation
val instantiateT: ((indexname * sort) * typ) list -> typ -> typ
val instantiate: ((indexname * sort) * typ) list * ((indexname * typ) * term) list ->
term -> term
- val instantiateT_same: ((indexname * sort) * typ) list -> typ Same.operation
- val instantiate_same: ((indexname * sort) * typ) list * ((indexname * typ) * term) list ->
- term Same.operation
val zero_var_indexes: term -> term
val zero_var_indexes_inst: term list ->
((indexname * sort) * typ) list * ((indexname * typ) * term) list
@@ -64,11 +58,6 @@
| term a = f a;
in term end;
-val map_atypsT_option = Same.capture o map_atypsT_same o Same.function;
-val map_atyps_option = Same.capture o map_types_same o map_atypsT_same o Same.function;
-val map_types_option = Same.capture o map_types_same o Same.function;
-val map_aterms_option = Same.capture o map_aterms_same o Same.function;
-
(* generalization of fixed variables *)
@@ -99,11 +88,8 @@
| gen (t $ u) = (gen t $ Same.commit gen u handle Same.SAME => t $ gen u);
in gen tm end;
+fun generalizeT names i ty = Same.commit (generalizeT_same names i) ty;
fun generalize names i tm = Same.commit (generalize_same names i) tm;
-fun generalizeT names i ty = Same.commit (generalizeT_same names i) ty;
-
-fun generalize_option names i tm = SOME (generalize_same names i tm) handle Same.SAME => NONE;
-fun generalizeT_option names i ty = SOME (generalizeT_same names i ty) handle Same.SAME => NONE;
(* instantiation of schematic variables (types before terms) -- recomputes maxidx *)
@@ -160,18 +146,15 @@
let val maxidx = Unsynchronized.ref i
in (Same.commit (inst_same maxidx insts) tm, ! maxidx) end;
-fun instantiateT [] ty = ty
- | instantiateT instT ty = Same.commit (instT_same (Unsynchronized.ref ~1) (no_indexes1 instT)) ty;
-
-fun instantiate ([], []) tm = tm
- | instantiate insts tm = Same.commit (inst_same (Unsynchronized.ref ~1) (no_indexes2 insts)) tm;
-
fun instantiateT_same [] _ = raise Same.SAME
| instantiateT_same instT ty = instT_same (Unsynchronized.ref ~1) (no_indexes1 instT) ty;
fun instantiate_same ([], []) _ = raise Same.SAME
| instantiate_same insts tm = inst_same (Unsynchronized.ref ~1) (no_indexes2 insts) tm;
+fun instantiateT instT ty = Same.commit (instantiateT_same instT) ty;
+fun instantiate inst tm = Same.commit (instantiate_same inst) tm;
+
end;
--- a/src/Pure/thm.ML Tue May 11 08:29:42 2010 +0200
+++ b/src/Pure/thm.ML Tue May 11 08:30:02 2010 +0200
@@ -138,8 +138,8 @@
val varifyT_global': (string * sort) list -> thm -> ((string * sort) * indexname) list * thm
val of_class: ctyp * class -> thm
val strip_shyps: thm -> thm
- val unconstrainT: ctyp -> thm -> thm
- val unconstrain_allTs: thm -> thm
+ val unconstrainT: thm -> thm
+ val legacy_unconstrainT: ctyp -> thm -> thm
val legacy_freezeT: thm -> thm
val assumption: int -> thm -> thm Seq.seq
val eq_assumption: int -> thm -> thm
@@ -1221,8 +1221,29 @@
shyps = shyps', hyps = hyps, tpairs = tpairs, prop = prop})
end;
-(*Internalize sort constraints of type variable*)
-fun unconstrainT
+(*Internalize sort constraints of type variables*)
+fun unconstrainT (thm as Thm (_, {thy_ref, maxidx, shyps, hyps, tpairs, prop, ...})) =
+ let
+ fun err msg = raise THM ("unconstrainT: " ^ msg, 0, [thm]);
+
+ val _ = null hyps orelse err "illegal hyps";
+ val _ = null tpairs orelse err "unsolved flex-flex constraints";
+ val tfrees = rev (Term.add_tfree_names prop []);
+ val _ = null tfrees orelse err ("illegal free type variables " ^ commas_quote tfrees);
+
+ val (_, prop') = Logic.unconstrainT shyps prop;
+ in
+ Thm (deriv_rule0 (Pt.PAxm ("Pure.unconstrainT", prop, SOME [])),
+ {thy_ref = thy_ref,
+ tags = [],
+ maxidx = maxidx_of_term prop',
+ shyps = [[]], (*potentially redundant*)
+ hyps = hyps,
+ tpairs = tpairs,
+ prop = prop'})
+ end;
+
+fun legacy_unconstrainT
(Ctyp {thy_ref = thy_ref1, T, ...})
(th as Thm (_, {thy_ref = thy_ref2, maxidx, shyps, hyps, tpairs, prop, ...})) =
let
@@ -1242,10 +1263,6 @@
prop = Logic.list_implies (constraints, unconstrain prop)})
end;
-fun unconstrain_allTs th =
- fold (unconstrainT o ctyp_of (theory_of_thm th) o TVar)
- (fold_terms Term.add_tvars th []) th;
-
(* Replace all TFrees not fixed or in the hyps by new TVars *)
fun varifyT_global' fixed (Thm (der, {thy_ref, maxidx, shyps, hyps, tpairs, prop, ...})) =
--- a/src/Tools/jEdit/README_BUILD Tue May 11 08:29:42 2010 +0200
+++ b/src/Tools/jEdit/README_BUILD Tue May 11 08:30:02 2010 +0200
@@ -13,7 +13,7 @@
http://blogtrader.net/dcaoyuan/category/NetBeans
http://wiki.netbeans.org/Scala
-* jEdit 4.3.1 (final)
+* jEdit 4.3.1 or 4.3.2
http://www.jedit.org/
Netbeans Project "jEdit": install official sources as ./contrib/jEdit/.
--- a/src/Tools/jEdit/src/jedit/isabelle_sidekick.scala Tue May 11 08:29:42 2010 +0200
+++ b/src/Tools/jEdit/src/jedit/isabelle_sidekick.scala Tue May 11 08:30:02 2010 +0200
@@ -132,7 +132,7 @@
for ((command, command_start) <- document.command_range(0) if !stopped) {
root.add(document.current_state(command).get.markup_root.swing_tree((node: Markup_Node) =>
{
- val content = Pretty.str_of(List(XML.Text(command.source(node.start, node.stop))))
+ val content = command.source(node.start, node.stop).replace('\n', ' ')
val id = command.id
new DefaultMutableTreeNode(new IAsset {
--- a/src/Tools/nbe.ML Tue May 11 08:29:42 2010 +0200
+++ b/src/Tools/nbe.ML Tue May 11 08:30:02 2010 +0200
@@ -101,14 +101,16 @@
val prem_thms = map (the o AList.lookup (op =) of_classes
o apfst (fst o fst o dest_TVar) o Logic.dest_of_class) prem_props;
in Drule.implies_elim_list thm prem_thms end;
- in ct
+ in
+ (* FIXME avoid legacy operations *)
+ ct
|> Drule.cterm_rule Thm.varifyT_global
|> Thm.instantiate_cterm (Thm.certify_inst thy (map (fn (v, ((sort, _), sort')) =>
(((v, 0), sort), TFree (v, sort'))) vs, []))
|> Drule.cterm_rule Thm.legacy_freezeT
|> conv
|> Thm.varifyT_global
- |> fold (fn (v, (_, sort')) => Thm.unconstrainT (certT (TVar ((v, 0), sort')))) vs
+ |> fold (fn (v, (_, sort')) => Thm.legacy_unconstrainT (certT (TVar ((v, 0), sort')))) vs
|> Thm.certify_instantiate (map (fn (v, ((sort, _), _)) =>
(((v, 0), []), TVar ((v, 0), sort))) vs, [])
|> strip_of_class