Theory Dedekind_Real
section ‹The Reals as Dedekind Sections of Positive Rationals›
text ‹Fundamentals of Abstract Analysis [Gleason, p. 121] provides some of the definitions.›
theory Dedekind_Real
imports Complex_Main
begin
lemma add_eq_exists: "∃x. a+x = (b::'a::ab_group_add)"
by (rule_tac x="b-a" in exI, simp)
subsection ‹Dedekind cuts or sections›
definition
cut :: "rat set ⇒ bool" where
"cut A ≡ {} ⊂ A ∧ A ⊂ {0<..} ∧
(∀y ∈ A. ((∀z. 0<z ∧ z < y ⟶ z ∈ A) ∧ (∃u ∈ A. y < u)))"
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 ⊂ {0<..}" by force
have nonempty: "{} ⊂ ?A"
proof
show "{} ⊆ ?A" by simp
show "{} ≠ ?A"
using field_lbound_gt_zero q by auto
qed
show ?thesis
by (simp add: cut_def pos nonempty,
blast dest: dense intro: order_less_trans)
qed
typedef preal = "Collect cut"
by (blast intro: cut_of_rat [OF zero_less_one])
lemma Abs_preal_induct [induct type: preal]:
"(⋀x. cut x ⟹ P (Abs_preal x)) ⟹ P x"
using Abs_preal_induct [of P x] by simp
lemma cut_Rep_preal [simp]: "cut (Rep_preal x)"
using Rep_preal [of x] by simp
definition
psup :: "preal set ⇒ preal" where
"psup P = Abs_preal (⋃X ∈ P. Rep_preal X)"
definition
add_set :: "[rat set,rat set] ⇒ rat set" where
"add_set A B = {w. ∃x ∈ A. ∃y ∈ B. w = x + y}"
definition
diff_set :: "[rat set,rat set] ⇒ rat set" where
"diff_set A B = {w. ∃x. 0 < w ∧ 0 < x ∧ x ∉ B ∧ x + w ∈ A}"
definition
mult_set :: "[rat set,rat set] ⇒ rat set" where
"mult_set A B = {w. ∃x ∈ A. ∃y ∈ B. w = x * y}"
definition
inverse_set :: "rat set ⇒ rat set" where
"inverse_set A ≡ {x. ∃y. 0 < x ∧ x < y ∧ inverse y ∉ A}"
instantiation preal :: "{ord, plus, minus, times, inverse, one}"
begin
definition
preal_less_def:
"r < s ≡ Rep_preal r < Rep_preal s"
definition
preal_le_def:
"r ≤ s ≡ Rep_preal r ⊆ 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 div s = r * inverse (s::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 ⟹ cut {r::rat. 0 < r ∧ r < q}"
by (simp add: cut_of_rat)
lemma preal_nonempty: "cut A ⟹ ∃x∈A. 0 < x"
unfolding cut_def [abs_def] by blast
lemma preal_Ex_mem: "cut A ⟹ ∃x. x ∈ A"
using preal_nonempty by blast
lemma preal_exists_bound: "cut A ⟹ ∃x. 0 < x ∧ x ∉ A"
using Dedekind_Real.cut_def by fastforce
lemma preal_exists_greater: "⟦cut A; y ∈ A⟧ ⟹ ∃u ∈ A. y < u"
unfolding cut_def [abs_def] by blast
lemma preal_downwards_closed: "⟦cut A; y ∈ A; 0 < z; z < y⟧ ⟹ z ∈ A"
unfolding cut_def [abs_def] by blast
text‹Relaxing the final premise›
lemma preal_downwards_closed': "⟦cut A; y ∈ A; 0 < z; z ≤ y⟧ ⟹ z ∈ A"
using less_eq_rat_def preal_downwards_closed by blast
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: "cut A"
and notx: "x ∉ A"
and y: "y ∈ 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: "∃x. x ∈ Rep_preal X"
thm preal_Ex_mem
by (rule preal_Ex_mem [OF cut_Rep_preal])
lemma Rep_preal_exists_bound: "∃x>0. x ∉ Rep_preal X"
by (rule preal_exists_bound [OF cut_Rep_preal])
lemmas not_in_Rep_preal_ub = not_in_preal_ub [OF cut_Rep_preal]
subsection‹Properties of Ordering›
instance preal :: order
proof
fix w :: preal
show "w ≤ w" by (simp add: preal_le_def)
next
fix i j k :: preal
assume "i ≤ j" and "j ≤ k"
then show "i ≤ k" by (simp add: preal_le_def)
next
fix z w :: preal
assume "z ≤ w" and "w ≤ z"
then show "z = w" by (simp add: preal_le_def Rep_preal_inject)
next
fix z w :: preal
show "z < w ⟷ z ≤ w ∧ ¬ w ≤ z"
by (auto simp: preal_le_def preal_less_def Rep_preal_inject)
qed
lemma preal_imp_pos: "⟦cut A; r ∈ A⟧ ⟹ 0 < r"
by (auto simp: cut_def)
instance preal :: linorder
proof
fix x y :: preal
show "x ≤ y ∨ y ≤ x"
unfolding preal_le_def
by (meson cut_Rep_preal not_in_preal_ub preal_downwards_closed preal_imp_pos subsetI)
qed
instantiation preal :: distrib_lattice
begin
definition
"(inf :: preal ⇒ preal ⇒ preal) = min"
definition
"(sup :: preal ⇒ preal ⇒ preal) = max"
instance
by intro_classes
(auto simp: inf_preal_def sup_preal_def max_min_distrib2)
end
subsection‹Properties of Addition›
lemma preal_add_commute: "(x::preal) + y = y + x"
unfolding preal_add_def add_set_def
by (metis (no_types, opaque_lifting) add.commute)
text‹Lemmas for proving that addition of two positive reals gives
a positive real›
lemma mem_add_set:
assumes "cut A" "cut B"
shows "cut (add_set A B)"
proof -
have "{} ⊂ add_set A B"
using assms by (force simp: add_set_def dest: preal_nonempty)
moreover
obtain q where "q > 0" "q ∉ add_set A B"
proof -
obtain a b where "a > 0" "a ∉ A" "b > 0" "b ∉ B" "⋀x. x ∈ A ⟹ x < a" "⋀y. y ∈ B ⟹ y < b"
by (meson assms preal_exists_bound not_in_preal_ub)
with assms have "a+b ∉ add_set A B"
by (fastforce simp add: add_set_def)
then show thesis
using ‹0 < a› ‹0 < b› add_pos_pos that by blast
qed
then have "add_set A B ⊂ {0<..}"
unfolding add_set_def
using preal_imp_pos [OF ‹cut A›] preal_imp_pos [OF ‹cut B›] by fastforce
moreover have "z ∈ add_set A B"
if u: "u ∈ add_set A B" and "0 < z" "z < u" for u z
using u unfolding add_set_def
proof (clarify)
fix x::rat and y::rat
assume ueq: "u = x + y" and x: "x ∈ A" and y:"y ∈ B"
have xpos [simp]: "x > 0" and ypos [simp]: "y > 0"
using assms preal_imp_pos x y by blast+
have xypos [simp]: "x+y > 0" by (simp add: pos_add_strict)
let ?f = "z/(x+y)"
have fless: "?f < 1"
using divide_less_eq_1_pos ‹z < u› ueq xypos by blast
show "∃x' ∈ A. ∃y'∈B. z = x' + y'"
proof (intro bexI)
show "z = x*?f + y*?f"
by (simp add: distrib_right [symmetric] divide_inverse ac_simps order_less_imp_not_eq2)
next
show "y * ?f ∈ B"
proof (rule preal_downwards_closed [OF ‹cut B› y])
show "0 < y * ?f"
by (simp add: ‹0 < z›)
next
show "y * ?f < y"
by (insert mult_strict_left_mono [OF fless ypos], simp)
qed
next
show "x * ?f ∈ A"
proof (rule preal_downwards_closed [OF ‹cut A› x])
show "0 < x * ?f"
by (simp add: ‹0 < z›)
next
show "x * ?f < x"
by (insert mult_strict_left_mono [OF fless xpos], simp)
qed
qed
qed
moreover
have "⋀y. y ∈ add_set A B ⟹ ∃u ∈ add_set A B. y < u"
unfolding add_set_def using preal_exists_greater assms by fastforce
ultimately show ?thesis
by (simp add: Dedekind_Real.cut_def)
qed
lemma preal_add_assoc: "((x::preal) + y) + z = x + (y + z)"
apply (simp add: preal_add_def mem_add_set)
apply (force simp: add_set_def ac_simps)
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"
unfolding preal_mult_def mult_set_def
by (metis (no_types, opaque_lifting) mult.commute)
text‹Multiplication of two positive reals gives a positive real.›
lemma mem_mult_set:
assumes "cut A" "cut B"
shows "cut (mult_set A B)"
proof -
have "{} ⊂ mult_set A B"
using assms
by (force simp: mult_set_def dest: preal_nonempty)
moreover
obtain q where "q > 0" "q ∉ mult_set A B"
proof -
obtain x y where x [simp]: "0 < x" "x ∉ A" and y [simp]: "0 < y" "y ∉ B"
using preal_exists_bound assms by blast
show thesis
proof
show "0 < x*y" by simp
show "x * y ∉ mult_set A B"
proof -
{
fix u::rat and v::rat
assume u: "u ∈ A" and v: "v ∈ B" and xy: "x*y = u*v"
moreover have "u<x" and "v<y" using assms x y u v by (blast dest: not_in_preal_ub)+
moreover have "0≤v"
using less_imp_le preal_imp_pos assms x y u v by blast
moreover have "u*v < x*y"
using assms x ‹u < x› ‹v < y› ‹0 ≤ v› by (blast intro: mult_strict_mono)
ultimately have False by force
}
thus ?thesis by (auto simp: mult_set_def)
qed
qed
qed
then have "mult_set A B ⊂ {0<..}"
unfolding mult_set_def
using preal_imp_pos [OF ‹cut A›] preal_imp_pos [OF ‹cut B›] by fastforce
moreover have "z ∈ mult_set A B"
if u: "u ∈ mult_set A B" and "0 < z" "z < u" for u z
using u unfolding mult_set_def
proof (clarify)
fix x::rat and y::rat
assume ueq: "u = x * y" and x: "x ∈ A" and y: "y ∈ B"
have [simp]: "y > 0"
using ‹cut B› preal_imp_pos y by blast
show "∃x' ∈ A. ∃y' ∈ B. z = x' * y'"
proof
have "z = (z/y)*y"
by (simp add: divide_inverse mult.commute [of y] mult.assoc order_less_imp_not_eq2)
then show "∃y'∈B. z = (z/y) * y'"
using y by blast
next
show "z/y ∈ A"
proof (rule preal_downwards_closed [OF ‹cut A› x])
show "0 < z/y"
by (simp add: ‹0 < z›)
show "z/y < x"
using ‹0 < y› pos_divide_less_eq ‹z < u› ueq by blast
qed
qed
qed
moreover have "⋀y. y ∈ mult_set A B ⟹ ∃u ∈ mult_set A B. y < u"
apply (simp add: mult_set_def)
by (metis preal_exists_greater mult_strict_right_mono preal_imp_pos assms)
ultimately show ?thesis
by (simp add: Dedekind_Real.cut_def)
qed
lemma preal_mult_assoc: "((x::preal) * y) * z = x * (y * z)"
apply (simp add: preal_mult_def mem_mult_set Rep_preal)
apply (simp add: mult_set_def)
apply (metis (no_types, opaque_lifting) ab_semigroup_mult_class.mult_ac(1))
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: "cut A"
have "{w. ∃u. 0 < u ∧ u < 1 ∧ (∃v ∈ A. w = u * v)} = A" (is "?lhs = A")
proof
show "?lhs ⊆ A"
proof clarify
fix x::rat and u::rat and v::rat
assume upos: "0<u" and "u<1" and v: "v ∈ 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 upos ‹u < 1› v)
thus "u * v ∈ A"
by (force intro: preal_downwards_closed [OF A v] mult_pos_pos upos vpos)
qed
next
show "A ⊆ ?lhs"
proof clarify
fix x::rat
assume x: "x ∈ 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 ∈ A" and xlessv: "x < v" ..
have vpos: "0<v" by (rule preal_imp_pos [OF A v])
show "∃u. 0 < u ∧ u < 1 ∧ (∃v∈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)
have "x = (x/v)*v"
by (simp add: divide_inverse mult.assoc vpos order_less_imp_not_eq2)
then show "∃v'∈A. x = (x / v) * v'"
using v by blast
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 ∈ Rep_preal(r+s)) = (∃x ∈ Rep_preal r. ∃y ∈ 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 ∈ Rep_preal(r*s)) = (∃x ∈ Rep_preal r. ∃y ∈ 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)) ⊆ Rep_preal (w * x + w * y)"
by (force simp: Bex_def mem_Rep_preal_add_iff mem_Rep_preal_mult_iff distrib_left)
lemma preal_add_mult_distrib_mean:
assumes a: "a ∈ Rep_preal w"
and b: "b ∈ Rep_preal w"
and d: "d ∈ Rep_preal x"
and e: "e ∈ Rep_preal y"
shows "∃c ∈ 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 cut_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 ∈ Rep_preal w"
proof (cases rule: linorder_le_cases)
assume "a ≤ b"
hence "?c ≤ b"
by (simp add: pos_divide_le_eq distrib_left mult_right_mono
order_less_imp_le)
thus ?thesis by (rule preal_downwards_closed' [OF cut_Rep_preal b cpos])
next
assume "b ≤ a"
hence "?c ≤ a"
by (simp add: pos_divide_le_eq distrib_left mult_right_mono
order_less_imp_le)
thus ?thesis by (rule preal_downwards_closed' [OF cut_Rep_preal a cpos])
qed
qed
lemma distrib_subset2:
"Rep_preal (w * x + w * y) ⊆ Rep_preal (w * (x + y))"
apply (clarsimp simp: mem_Rep_preal_add_iff mem_Rep_preal_mult_iff)
using mem_Rep_preal_add_iff preal_add_mult_distrib_mean by blast
lemma preal_add_mult_distrib2: "(w * ((x::preal) + y)) = (w * x) + (w * y)"
by (metis Rep_preal_inverse distrib_subset1 distrib_subset2 subset_antisym)
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_inverse_set:
assumes "cut A" shows "cut (inverse_set A)"
proof -
have "∃x y. 0 < x ∧ x < y ∧ inverse y ∉ A"
proof -
from preal_exists_bound [OF ‹cut A›]
obtain x where [simp]: "0<x" "x ∉ 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) ∉ A"
by (simp add: order_less_imp_not_eq2)
qed
qed
then have "{} ⊂ inverse_set A"
using inverse_set_def by fastforce
moreover obtain q where "q > 0" "q ∉ inverse_set A"
proof -
from preal_nonempty [OF ‹cut A›]
obtain x where x: "x ∈ A" and xpos [simp]: "0<x" ..
show ?thesis
proof
show "0 < inverse x" by simp
show "inverse x ∉ 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 ∈ A"
by (simp add: preal_downwards_closed [OF ‹cut A› x] iyless)}
thus ?thesis by (auto simp: inverse_set_def)
qed
qed
qed
moreover have "inverse_set A ⊂ {0<..}"
using calculation inverse_set_def by blast
moreover have "z ∈ inverse_set A"
if u: "u ∈ inverse_set A" and "0 < z" "z < u" for u z
using u that less_trans unfolding inverse_set_def by auto
moreover have "⋀y. y ∈ inverse_set A ⟹ ∃u ∈ inverse_set A. y < u"
by (simp add: inverse_set_def) (meson dense less_trans)
ultimately show ?thesis
by (simp add: Dedekind_Real.cut_def)
qed
subsection‹Gleason's Lemma 9-3.4, page 122›
lemma Gleason9_34_exists:
assumes A: "cut A"
and "∀x∈A. x + u ∈ A"
and "0 ≤ z"
shows "∃b∈A. b + (of_int z) * u ∈ A"
proof (cases z rule: int_cases)
case (nonneg n)
show ?thesis
proof (simp add: nonneg, induct n)
case 0
from preal_nonempty [OF A]
show ?case by force
next
case (Suc k)
then obtain b where b: "b ∈ A" "b + of_nat k * u ∈ A" ..
hence "b + of_int (int k)*u + u ∈ A" by (simp add: assms)
thus ?case by (force simp: algebra_simps b)
qed
next
case (neg n)
with assms show ?thesis by simp
qed
lemma Gleason9_34_contra:
assumes A: "cut A"
shows "⟦∀x∈A. x + u ∈ A; 0 < u; 0 < y; y ∉ 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: "∀x∈A. x + (Fract c d) ∈ A"
and upos: "0 < Fract c d"
and ypos: "0 < Fract a b"
and notin: "Fract a b ∉ 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 ≤ Fract ?k 1 * (Fract c d)"
proof -
have "?thesis = ((a * d * b * d) ≤ c * b * (a * d * b * d))"
by (simp add: order_less_imp_not_eq2 ac_simps)
moreover
have "(1 * (a * d * b * d)) ≤ 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 ≤ ?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 ∈ A"
and mem: "z + of_int ?k * Fract c d ∈ 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 "cut A" "0 < u"
shows "∃r ∈ A. r + u ∉ A"
using assms Gleason9_34_contra preal_exists_bound by blast
subsection‹Gleason's Lemma 9-3.6›
lemma lemma_gleason9_36:
assumes A: "cut A"
and x: "1 < x"
shows "∃r ∈ A. r*x ∉ A"
proof -
from preal_nonempty [OF A]
obtain y where y: "y ∈ A" and ypos: "0<y" ..
show ?thesis
proof (rule classical)
assume "~(∃r∈A. r * x ∉ A)"
with y have ymem: "y * x ∈ 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∈A" and notin: "r + ?d ∉ 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 ≤ y + ?d)"
proof
assume le: "r + ?d ≤ y + ?d"
from ymem have yd: "y + ?d ∈ A" by (simp add: eq)
have "r + ?d ∈ 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"
using dpos less_divide_eq_1 by fastforce
have "r + ?d < r*x"
proof -
have "r + ?d < r + (r * ?d)/y" by (simp add: dless)
also from 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 "∃r∈A. r * x ∉ A" by (blast dest: preal_downwards_closed [OF A])
qed
qed
subsection‹Existence of Inverse: Part 2›
lemma mem_Rep_preal_inverse_iff:
"(z ∈ Rep_preal(inverse r)) ⟷ (0 < z ∧ (∃y. z < y ∧ inverse y ∉ 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 ∧ 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 "∃v u y. 0 < v ∧ v < y ∧ inverse y ∉ Rep_preal R ∧
u ∈ Rep_preal R ∧ x = v * u"
proof -
from xpos and xless have "1 < inverse x" by (simp add: one_less_inverse_iff)
from lemma_gleason9_36 [OF cut_Rep_preal this]
obtain t where t: "t ∈ Rep_preal R"
and notin: "t * (inverse x) ∉ Rep_preal R" ..
have rpos: "0<t" by (rule preal_imp_pos [OF cut_Rep_preal t])
from preal_exists_greater [OF cut_Rep_preal t]
obtain u where u: "u ∈ Rep_preal R" and rless: "t < u" ..
have upos: "0<u" by (rule preal_imp_pos [OF cut_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/t" using xpos upos rpos
by (simp add: divide_inverse mult_less_cancel_left rless)
show "inverse (x / t) ∉ Rep_preal R" using notin
by (simp add: divide_inverse mult.commute)
show "u ∈ 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 ⊆ Rep_preal(inverse r * r)"
by (force simp: Rep_preal_one mem_Rep_preal_inverse_iff mem_Rep_preal_mult_iff dest: subset_inverse_mult_lemma)
lemma inverse_mult_subset: "Rep_preal(inverse r * r) ⊆ Rep_preal 1"
proof -
have "0 < u * v" if "v ∈ Rep_preal r" "0 < u" "u < t" for u v t :: rat
using that by (simp add: zero_less_mult_iff preal_imp_pos [OF cut_Rep_preal])
moreover have "t * q < 1"
if "q ∈ Rep_preal r" "0 < t" "t < y" "inverse y ∉ Rep_preal r"
for t q y :: rat
proof -
have "q < inverse y"
using not_in_Rep_preal_ub that by auto
hence "t * q < t/y"
using that by (simp add: divide_inverse mult_less_cancel_left)
also have "… ≤ 1"
using that by (simp add: pos_divide_le_eq)
finally show ?thesis .
qed
ultimately show ?thesis
by (auto simp: Rep_preal_one mem_Rep_preal_inverse_iff mem_Rep_preal_mult_iff)
qed
lemma preal_mult_inverse: "inverse r * r = (1::preal)"
by (meson Rep_preal_inject inverse_mult_subset subset_antisym subset_inverse_mult)
lemma preal_mult_inverse_right: "r * inverse r = (1::preal)"
using preal_mult_commute preal_mult_inverse by auto
text‹Theorems needing ‹Gleason9_34››
lemma Rep_preal_self_subset: "Rep_preal (r) ⊆ Rep_preal(r + s)"
proof
fix x
assume x: "x ∈ Rep_preal r"
obtain y where y: "y ∈ Rep_preal s" and "y > 0"
using Rep_preal preal_nonempty by blast
have ry: "x+y ∈ Rep_preal(r + s)" using x y
by (auto simp: mem_Rep_preal_add_iff)
then show "x ∈ Rep_preal(r + s)"
by (meson ‹0 < y› add_less_same_cancel1 not_in_Rep_preal_ub order.asym preal_imp_pos [OF cut_Rep_preal x])
qed
lemma Rep_preal_sum_not_subset: "~ Rep_preal (r + s) ⊆ Rep_preal(r)"
proof -
obtain y where y: "y ∈ Rep_preal s" and "y > 0"
using Rep_preal preal_nonempty by blast
obtain x where "x ∈ Rep_preal r" and notin: "x + y ∉ Rep_preal r"
using Dedekind_Real.Rep_preal Gleason9_34 ‹0 < y› by blast
then have "x + y ∈ Rep_preal (r + s)" using y
by (auto simp: mem_Rep_preal_add_iff)
thus ?thesis using notin by blast
qed
text‹at last, Gleason prop. 9-3.5(iii) page 123›
proposition preal_self_less_add_left: "(r::preal) < r + s"
by (meson Rep_preal_sum_not_subset not_less preal_le_def)
subsection‹Subtraction for Positive Reals›
text‹gleason prop. 9-3.5(iv), page 123: proving \<^prop>‹a < b ⟹ ∃d. a + d = b›.
We define the claimed \<^term>‹D› and show that it is a positive real›
lemma mem_diff_set:
assumes "r < s"
shows "cut (diff_set (Rep_preal s) (Rep_preal r))"
proof -
obtain p where "Rep_preal r ⊆ Rep_preal s" "p ∈ Rep_preal s" "p ∉ Rep_preal r"
using assms unfolding preal_less_def by auto
then have "{} ⊂ diff_set (Rep_preal s) (Rep_preal r)"
apply (simp add: diff_set_def psubset_eq)
by (metis cut_Rep_preal add_eq_exists less_add_same_cancel1 preal_exists_greater preal_imp_pos)
moreover
obtain q where "q > 0" "q ∉ Rep_preal s"
using Rep_preal_exists_bound by blast
then have qnot: "q ∉ diff_set (Rep_preal s) (Rep_preal r)"
by (auto simp: diff_set_def dest: cut_Rep_preal [THEN preal_downwards_closed])
moreover have "diff_set (Rep_preal s) (Rep_preal r) ⊂ {0<..}" (is "?lhs < ?rhs")
using ‹0 < q› diff_set_def qnot by blast
moreover have "z ∈ diff_set (Rep_preal s) (Rep_preal r)"
if u: "u ∈ diff_set (Rep_preal s) (Rep_preal r)" and "0 < z" "z < u" for u z
using u that less_trans Rep_preal unfolding diff_set_def Dedekind_Real.cut_def by auto
moreover have "∃u ∈ diff_set (Rep_preal s) (Rep_preal r). y < u"
if y: "y ∈ diff_set (Rep_preal s) (Rep_preal r)" for y
proof -
obtain a b where "0 < a" "0 < b" "a ∉ Rep_preal r" "a + y + b ∈ Rep_preal s"
using y
by (simp add: diff_set_def) (metis cut_Rep_preal add_eq_exists less_add_same_cancel1 preal_exists_greater)
then have "a + (y + b) ∈ Rep_preal s"
by (simp add: add.assoc)
then have "y + b ∈ diff_set (Rep_preal s) (Rep_preal r)"
using ‹0 < a› ‹0 < b› ‹a ∉ Rep_preal r› y
by (auto simp: diff_set_def)
then show ?thesis
using ‹0 < b› less_add_same_cancel1 by blast
qed
ultimately show ?thesis
by (simp add: Dedekind_Real.cut_def)
qed
lemma mem_Rep_preal_diff_iff:
"r < s ⟹
(z ∈ Rep_preal (s - r)) ⟷
(∃x. 0 < x ∧ 0 < z ∧ x ∉ Rep_preal r ∧ x + z ∈ Rep_preal s)"
apply (simp add: preal_diff_def mem_diff_set Rep_preal)
apply (force simp: diff_set_def)
done
proposition less_add_left:
fixes r::preal
assumes "r < s"
shows "r + (s-r) = s"
proof -
have "a + b ∈ Rep_preal s"
if "a ∈ Rep_preal r" "c + b ∈ Rep_preal s" "c ∉ Rep_preal r"
and "0 < b" "0 < c" for a b c
by (meson cut_Rep_preal add_less_imp_less_right add_pos_pos not_in_Rep_preal_ub preal_downwards_closed preal_imp_pos that)
then have "r + (s-r) ≤ s"
using assms mem_Rep_preal_add_iff mem_Rep_preal_diff_iff preal_le_def by auto
have "x ∈ Rep_preal (r + (s - r))" if "x ∈ Rep_preal s" for x
proof (cases "x ∈ Rep_preal r")
case True
then show ?thesis
using Rep_preal_self_subset by blast
next
case False
have "∃u v z. 0 < v ∧ 0 < z ∧ u ∈ Rep_preal r ∧ z ∉ Rep_preal r ∧ z + v ∈ Rep_preal s ∧ x = u + v"
if x: "x ∈ Rep_preal s"
proof -
have xpos: "x > 0"
using Rep_preal preal_imp_pos that by blast
obtain e where epos: "0 < e" and xe: "x + e ∈ Rep_preal s"
by (metis cut_Rep_preal x add_eq_exists less_add_same_cancel1 preal_exists_greater)
from Gleason9_34 [OF cut_Rep_preal epos]
obtain u where r: "u ∈ Rep_preal r" and notin: "u + e ∉ Rep_preal r" ..
with x False xpos have rless: "u < x" by (blast intro: not_in_Rep_preal_ub)
from add_eq_exists [of u x]
obtain y where eq: "x = u+y" by auto
show ?thesis
proof (intro exI conjI)
show "u + e ∉ Rep_preal r" by (rule notin)
show "u + e + y ∈ Rep_preal s" using xe eq by (simp add: ac_simps)
show "0 < u + e"
using epos preal_imp_pos [OF cut_Rep_preal r] by simp
qed (use r rless eq in auto)
qed
then show ?thesis
using assms mem_Rep_preal_add_iff mem_Rep_preal_diff_iff that by blast
qed
then have "s ≤ r + (s-r)"
by (auto simp: preal_le_def)
then show ?thesis
by (simp add: ‹r + (s - r) ≤ s› antisym)
qed
lemma preal_add_less2_mono1: "r < (s::preal) ⟹ r + t < s + t"
by (metis add.assoc add.commute less_add_left preal_self_less_add_left)
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)"
by (metis linorder_cases order.asym preal_add_less2_mono1)
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 [simp]: "(t + (r::preal) < t + s) ⟷ (r < s)"
by (blast intro: preal_add_less2_mono2 preal_add_left_less_cancel)
lemma preal_add_less_cancel_right [simp]: "((r::preal) + t < s + t) = (r < s)"
using preal_add_less_cancel_left [symmetric, of r s t] by (simp add: ac_simps)
lemma preal_add_le_cancel_left [simp]: "(t + (r::preal) ≤ t + s) = (r ≤ s)"
by (simp add: linorder_not_less [symmetric])
lemma preal_add_le_cancel_right [simp]: "((r::preal) + t ≤ s + t) = (r ≤ s)"
using preal_add_le_cancel_left [symmetric, of r s t] by (simp add: ac_simps)
lemma preal_add_right_cancel: "(r::preal) + t = s + t ⟹ r = s"
by (metis less_irrefl linorder_cases preal_add_less_cancel_right)
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_ab_semigroup_add
proof
fix a b c :: preal
show "a ≤ b ⟹ c + a ≤ 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:
assumes le: "⋀X. X ∈ P ⟹ X ≤ Y" and "P ≠ {}"
shows "cut (⋃X ∈ P. Rep_preal(X))"
proof -
have "{} ⊂ (⋃X ∈ P. Rep_preal(X))"
using ‹P ≠ {}› mem_Rep_preal_Ex by fastforce
moreover
obtain q where "q > 0" and "q ∉ (⋃X ∈ P. Rep_preal(X))"
using Rep_preal_exists_bound [of Y] le by (auto simp: preal_le_def)
then have "(⋃X ∈ P. Rep_preal(X)) ⊂ {0<..}"
using cut_Rep_preal preal_imp_pos by force
moreover
have "⋀u z. ⟦u ∈ (⋃X ∈ P. Rep_preal(X)); 0 < z; z < u⟧ ⟹ z ∈ (⋃X ∈ P. Rep_preal(X))"
by (auto elim: cut_Rep_preal [THEN preal_downwards_closed])
moreover
have "⋀y. y ∈ (⋃X ∈ P. Rep_preal(X)) ⟹ ∃u ∈ (⋃X ∈ P. Rep_preal(X)). y < u"
by (blast dest: cut_Rep_preal [THEN preal_exists_greater])
ultimately show ?thesis
by (simp add: Dedekind_Real.cut_def)
qed
lemma preal_psup_le:
"⟦⋀X. X ∈ P ⟹ X ≤ Y; x ∈ P⟧ ⟹ x ≤ psup P"
using preal_sup [of P Y] unfolding preal_le_def psup_def by fastforce
lemma psup_le_ub: "⟦⋀X. X ∈ P ⟹ X ≤ Y; P ≠ {}⟧ ⟹ psup P ≤ Y"
using preal_sup [of P Y] by (simp add: SUP_least preal_le_def psup_def)
text‹Supremum property›
proposition preal_complete:
assumes le: "⋀X. X ∈ P ⟹ X ≤ Y" and "P ≠ {}"
shows "(∃X ∈ P. Z < X) ⟷ (Z < psup P)" (is "?lhs = ?rhs")
proof
assume ?lhs
then show ?rhs
using preal_sup [OF assms] preal_less_def psup_def by auto
next
assume ?rhs
then show ?lhs
by (meson ‹P ≠ {}› not_less psup_le_ub)
qed
subsection ‹Defining the Reals from the Positive Reals›
text ‹Here we do quotients the old-fashioned way›
definition
realrel :: "((preal * preal) * (preal * preal)) set" where
"realrel = {p. ∃x1 y1 x2 y2. p = ((x1,y1),(x2,y2)) ∧ x1+y2 = x2+y1}"
definition "Real = UNIV//realrel"
typedef real = Real
morphisms Rep_Real Abs_Real
unfolding Real_def by (auto simp: quotient_def)
text ‹This doesn't involve the overloaded "real" function: users don't see it›
definition
real_of_preal :: "preal ⇒ real" where
"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: "0 = Abs_Real(realrel``{(1, 1)})"
definition
real_one_def: "1 = Abs_Real(realrel``{(1 + 1, 1)})"
definition
real_add_def: "z + w =
the_elem (⋃(x,y) ∈ Rep_Real z. ⋃(u,v) ∈ Rep_Real w.
{ Abs_Real(realrel``{(x+u, y+v)}) })"
definition
real_minus_def: "- r = the_elem (⋃(x,y) ∈ Rep_Real r. { Abs_Real(realrel``{(y,x)}) })"
definition
real_diff_def: "r - (s::real) = r + - s"
definition
real_mult_def:
"z * w =
the_elem (⋃(x,y) ∈ Rep_Real z. ⋃(u,v) ∈ Rep_Real w.
{ Abs_Real(realrel``{(x*u + y*v, x*v + y*u)}) })"
definition
real_inverse_def: "inverse (r::real) ≡ (THE s. (r = 0 ∧ s = 0) ∨ s * r = 1)"
definition
real_divide_def: "r div (s::real) ≡ r * inverse s"
definition
real_le_def: "z ≤ (w::real) ≡
(∃x y u v. x+v ≤ u+y ∧ (x,y) ∈ Rep_Real z ∧ (u,v) ∈ Rep_Real w)"
definition
real_less_def: "x < (y::real) ≡ x ≤ y ∧ x ≠ y"
definition
real_abs_def: "¦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 realrel_iff [simp]: "(((x1,y1),(x2,y2)) ∈ realrel) = (x1 + y2 = x2 + y1)"
by (simp add: realrel_def)
lemma preal_trans_lemma:
assumes "x + y1 = x1 + y" and "x + y2 = x2 + y"
shows "x1 + y2 = x2 + (y1::preal)"
by (metis add.left_commute assms preal_add_left_cancel)
lemma equiv_realrel: "equiv UNIV realrel"
by (auto simp: equiv_def refl_on_def sym_def trans_def realrel_def intro: dest: preal_trans_lemma)
text‹Reduces equality of equivalence classes to the \<^term>‹realrel› relation:
\<^term>‹(realrel `` {x} = realrel `` {y}) = ((x,y) ∈ realrel)››
lemmas equiv_realrel_iff [simp] =
eq_equiv_class_iff [OF equiv_realrel UNIV_I UNIV_I]
lemma realrel_in_real [simp]: "realrel``{(x,y)} ∈ Real"
by (simp add: Real_def realrel_def quotient_def, blast)
declare Abs_Real_inject [simp] 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"
by (metis Rep_Real_inverse prod.exhaust Rep_Real [of z, unfolded Real_def, THEN quotientE])
subsection ‹Addition and Subtraction›
lemma real_add:
"Abs_Real (realrel``{(x,y)}) + Abs_Real (realrel``{(u,v)}) =
Abs_Real (realrel``{(x+u, y+v)})"
proof -
have "(λz w. (λ(x,y). (λ(u,v). {Abs_Real (realrel `` {(x+u, y+v)})}) w) z)
respects2 realrel"
by (clarsimp simp: congruent2_def) (metis add.left_commute preal_add_assoc)
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 "(λ(x,y). {Abs_Real (realrel``{(y,x)})}) respects realrel"
by (auto simp: 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 ac_simps)
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)"
by (metis (no_types, opaque_lifting) add.left_commute preal_add_commute preal_add_mult_distrib2)
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])
by (auto simp: mult.commute add.commute combine_common_factor preal_add_assoc preal_add_commute)
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 ac_simps)
lemma real_mult_assoc: "((z1::real) * z2) * z3 = z1 * (z2 * z3)"
by (cases z1, cases z2, cases z3) (simp add: real_mult algebra_simps)
lemma real_mult_1: "(1::real) * z = z"
by (cases z) (simp add: real_mult real_one_def algebra_simps)
lemma real_add_mult_distrib: "((z1::real) + z2) * w = (z1 * w) + (z2 * w)"
by (cases z1, cases z2, cases w) (simp add: real_add real_mult algebra_simps)
text‹one and zero are distinct›
lemma real_zero_not_eq_one: "0 ≠ (1::real)"
proof -
have "(1::preal) < 1 + 1"
by (simp add: preal_self_less_add_left)
then show ?thesis
by (simp add: real_zero_def real_one_def neq_iff)
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 ≠ (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)
lemma real_mult_inverse_left_ex:
assumes "x ≠ 0" obtains y::real where "y*x = 1"
proof (cases x)
case (Abs_Real u v)
show ?thesis
proof (cases u v rule: linorder_cases)
case less
then have "v * inverse (v - u) = 1 + u * inverse (v - u)"
using less_add_left [of u v]
by (metis preal_add_commute preal_add_mult_distrib preal_mult_inverse_right)
then have "Abs_Real (realrel``{(1, inverse (v-u) + 1)}) * x - 1 = 0"
by (simp add: Abs_Real real_mult preal_mult_inverse_right real_one_def) (simp add: algebra_simps)
with that show thesis by auto
next
case equal
then show ?thesis
using Abs_Real assms real_zero_iff by blast
next
case greater
then have "u * inverse (u - v) = 1 + v * inverse (u - v)"
using less_add_left [of v u] by (metis add.commute distrib_right preal_mult_inverse_right)
then have "Abs_Real (realrel``{(inverse (u-v) + 1, 1)}) * x - 1 = 0"
by (simp add: Abs_Real real_mult preal_mult_inverse_right real_one_def) (simp add: algebra_simps)
with that show thesis by auto
qed
qed
lemma real_mult_inverse_left:
fixes x :: real
assumes "x ≠ 0" shows "inverse x * x = 1"
proof -
obtain y where "y*x = 1"
using assms real_mult_inverse_left_ex by blast
then have "(THE s. s * x = 1) * x = 1"
proof (rule theI)
show "y' = y" if "y' * x = 1" for y'
by (metis ‹y * x = 1› mult.left_commute mult.right_neutral that)
qed
then show ?thesis
using assms real_inverse_def by auto
qed
subsection‹The Real Numbers form a Field›
instance real :: field
proof
fix x y z :: real
show "x ≠ 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 ‹≤› Ordering›
lemma real_le_refl: "w ≤ (w::real)"
by (cases w, force simp: 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 ≤ a"
shows "b ≤ (d::preal)"
proof -
from le have "c+d ≤ a+d" by simp
hence "a+b ≤ a+d" by (simp add: eq)
thus "b ≤ d" by simp
qed
lemma real_le_lemma:
assumes l: "u1 + v2 ≤ u2 + v1"
and "x1 + v1 = u1 + y1"
and "x2 + v2 = u2 + y2"
shows "x1 + y2 ≤ x2 + (y1::preal)"
proof -
have "(x1+v1) + (u2+y2) = (u1+y1) + (x2+v2)" by (simp add: assms)
hence "(x1+y2) + (u2+v1) = (x2+y1) + (u1+v2)" by (simp add: ac_simps)
also have "… ≤ (x2+y1) + (u2+v1)" by (simp add: assms)
finally show ?thesis by simp
qed
lemma real_le:
"Abs_Real(realrel``{(x1,y1)}) ≤ Abs_Real(realrel``{(x2,y2)}) ⟷ x1 + y2 ≤ x2 + y1"
unfolding real_le_def by (auto intro: real_le_lemma)
lemma real_le_antisym: "⟦z ≤ w; w ≤ z⟧ ⟹ z = (w::real)"
by (cases z, cases w, simp add: real_le)
lemma real_trans_lemma:
assumes "x + v ≤ u + y"
and "u + v' ≤ u' + v"
and "x2 + v2 = u2 + y2"
shows "x + v' ≤ u' + (y::preal)"
proof -
have "(x+v') + (u+v) = (x+v) + (u+v')" by (simp add: ac_simps)
also have "… ≤ (u+y) + (u+v')" by (simp add: assms)
also have "… ≤ (u+y) + (u'+v)" by (simp add: assms)
also have "… = (u'+y) + (u+v)" by (simp add: ac_simps)
finally show ?thesis by simp
qed
lemma real_le_trans: "⟦i ≤ j; j ≤ k⟧ ⟹ i ≤ (k::real)"
by (cases i, cases j, cases k) (auto simp: real_le intro: real_trans_lemma)
instance real :: order
proof
show "u < v ⟷ u ≤ v ∧ ¬ v ≤ u" for u v::real
by (auto simp: real_less_def intro: real_le_antisym)
qed (auto intro: real_le_refl real_le_trans real_le_antisym)
instance real :: linorder
proof
show "x ≤ y ∨ y ≤ x" for x y :: real
by (meson eq_refl le_cases real_le_def)
qed
instantiation real :: distrib_lattice
begin
definition
"(inf :: real ⇒ real ⇒ real) = min"
definition
"(sup :: real ⇒ real ⇒ real) = max"
instance
by standard (auto simp: inf_real_def sup_real_def max_min_distrib2)
end
subsection‹The Reals Form an Ordered Field›
lemma real_le_eq_diff: "(x ≤ y) ⟷ (x-y ≤ (0::real))"
by (cases x, cases y) (simp add: real_le real_zero_def real_diff_def real_add real_minus preal_add_commute)
lemma real_add_left_mono:
assumes le: "x ≤ y" shows "z + x ≤ 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"])
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])
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])
lemma real_mult_order:
fixes x y::real
assumes "0 < x" "0 < y"
shows "0 < x * y"
proof (cases x, cases y)
show "0 < x * y"
if x: "x = Abs_Real (Dedekind_Real.realrel `` {(x1, x2)})"
and y: "y = Abs_Real (Dedekind_Real.realrel `` {(y1, y2)})"
for x1 x2 y1 y2
proof -
have "x2 < x1" "y2 < y1"
using assms not_le real_zero_def real_le x y
by (metis preal_add_le_cancel_left real_zero_iff)+
then obtain xd yd where "x1 = x2 + xd" "y1 = y2 + yd"
using less_add_left by metis
then have "¬ (x * y ≤ 0)"
apply (simp add: x y real_mult real_zero_def real_le)
apply (simp add: not_le algebra_simps preal_self_less_add_left)
done
then show ?thesis
by auto
qed
qed
lemma real_mult_less_mono2: "⟦(0::real) < z; x < y⟧ ⟹ z * x < z * y"
by (metis add_uminus_conv_diff real_less_sum_gt_zero real_mult_order real_sum_gt_zero_less right_diff_distrib')
instance real :: linordered_field
proof
fix x y z :: real
show "x ≤ y ⟹ z + x ≤ z + y" by (rule real_add_left_mono)
show "¦x¦ = (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)
show "z * x < z * y" if "x < y" "0 < z"
by (simp add: real_mult_less_mono2 that)
qed
subsection ‹Completeness of the reals›
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:
"∃m. (x::real) = real_of_preal m ∨ x = 0 ∨ x = -(real_of_preal m)"
proof (cases x)
case (Abs_Real u v)
show ?thesis
proof (cases u v rule: linorder_cases)
case less
then show ?thesis
using less_add_left
apply (simp add: Abs_Real real_of_preal_def real_minus real_zero_def)
by (metis preal_add_assoc preal_add_commute)
next
case equal
then show ?thesis
using Abs_Real real_zero_iff by blast
next
case greater
then show ?thesis
using less_add_left
apply (simp add: Abs_Real real_of_preal_def real_minus real_zero_def)
by (metis preal_add_assoc preal_add_commute)
qed
qed
lemma real_of_preal_less_iff [simp]:
"(real_of_preal m1 < real_of_preal m2) = (m1 < m2)"
by (metis not_less preal_add_less_cancel_right real_le real_of_preal_def)
lemma real_of_preal_le_iff [simp]:
"(real_of_preal m1 ≤ real_of_preal m2) = (m1 ≤ m2)"
by (simp add: linorder_not_less [symmetric])
lemma real_of_preal_zero_less [simp]: "0 < real_of_preal m"
by (metis less_add_same_cancel2 preal_self_less_add_left real_of_preal_add real_of_preal_less_iff)
subsection‹Theorems About the Ordering›
lemma real_gt_zero_preal_Ex: "(0 < x) ⟷ (∃y. x = real_of_preal y)"
using order.asym real_of_preal_trichotomy by fastforce
subsection ‹Completeness of Positive Reals›
text ‹
Supremum property for the set of positive reals
Let ‹P› be a non-empty set of positive reals, with an upper
bound ‹y›. Then ‹P› has a least upper bound
(written ‹S›).
FIXME: Can the premise be weakened to ‹∀x ∈ P. x≤ y›?
›
lemma posreal_complete:
assumes positive_P: "∀x ∈ P. (0::real) < x"
and not_empty_P: "∃x. x ∈ P"
and upper_bound_Ex: "∃y. ∀x ∈ P. x<y"
shows "∃s. ∀y. (∃x ∈ P. y < x) = (y < s)"
proof (rule exI, rule allI)
fix y
let ?pP = "{w. real_of_preal w ∈ P}"
show "(∃x∈P. y < x) = (y < real_of_preal (psup ?pP))"
proof (cases "0 < y")
assume neg_y: "¬ 0 < y"
show ?thesis
proof
assume "∃x∈P. y < x"
thus "y < real_of_preal (psup ?pP)"
by (metis dual_order.strict_trans neg_y not_less_iff_gr_or_eq real_of_preal_zero_less)
next
assume "y < real_of_preal (psup ?pP)"
obtain "x" where x_in_P: "x ∈ P" using not_empty_P ..
thus "∃x ∈ P. y < x" using x_in_P
using neg_y not_less_iff_gr_or_eq positive_P by fastforce
qed
next
assume pos_y: "0 < y"
then obtain py where y_is_py: "y = real_of_preal py"
by (auto simp: real_gt_zero_preal_Ex)
obtain a where "a ∈ 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: real_gt_zero_preal_Ex)
hence "pa ∈ ?pP" using ‹a ∈ P› by auto
hence pP_not_empty: "?pP ≠ {}" by auto
obtain sup where sup: "∀x ∈ P. x < sup"
using upper_bound_Ex ..
from this and ‹a ∈ P› have "a < sup" ..
hence "0 < sup" using a_pos by arith
then obtain possup where "sup = real_of_preal possup"
by (auto simp: real_gt_zero_preal_Ex)
hence "∀X ∈ ?pP. X ≤ possup"
using sup by auto
with pP_not_empty have psup: "⋀Z. (∃X ∈ ?pP. Z < X) = (Z < psup ?pP)"
by (meson preal_complete)
show ?thesis
proof
assume "∃x ∈ P. y < x"
then obtain x where x_in_P: "x ∈ 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: real_gt_zero_preal_Ex)
have py_less_X: "∃X ∈ ?pP. py < X"
proof
show "py < px" using y_is_py and x_is_px and y_less_x
by simp
show "px ∈ ?pP" using x_in_P and x_is_px by simp
qed
have "(∃X ∈ ?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 ∈ P})"
using y_is_py and pos_y by simp
next
assume y_less_psup: "y < real_of_preal (psup ?pP)"
hence "py < psup ?pP" using y_is_py
by simp
then obtain "X" where py_less_X: "py < X" and X_in_pP: "X ∈ ?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
moreover have "x ∈ P"
using x_is_X and X_in_pP by simp
ultimately show "∃ x ∈ P. y < x" ..
qed
qed
qed
subsection ‹Completeness›
lemma reals_complete:
fixes S :: "real set"
assumes notempty_S: "∃X. X ∈ S"
and exists_Ub: "bdd_above S"
shows "∃x. (∀s∈S. s ≤ x) ∧ (∀y. (∀s∈S. s ≤ y) ⟶ x ≤ y)"
proof -
obtain X where X_in_S: "X ∈ S" using notempty_S ..
obtain Y where Y_isUb: "∀s∈S. s ≤ Y"
using exists_Ub by (auto simp: bdd_above_def)
let ?SHIFT = "{z. ∃x ∈S. z = x + (-X) + 1} ∩ {x. 0 < x}"
{
fix x
assume S_le_x: "∀s∈S. s ≤ x"
{
fix s
assume "s ∈ {z. ∃x∈S. z = x + - X + 1}"
hence "∃ x ∈ S. s = x + -X + 1" ..
then obtain x1 where x1: "x1 ∈ S" "s = x1 + (-X) + 1" ..
then have "x1 ≤ x" using S_le_x by simp
with x1 have "s ≤ x + - X + 1" by arith
}
then have "∀s∈?SHIFT. s ≤ x + (-X) + 1"
by auto
} note S_Ub_is_SHIFT_Ub = this
have *: "∀s∈?SHIFT. s ≤ Y + (-X) + 1" using Y_isUb by (rule S_Ub_is_SHIFT_Ub)
have "∀s∈?SHIFT. s < Y + (-X) + 2"
proof
fix s assume "s∈?SHIFT"
with * have "s ≤ Y + (-X) + 1" by simp
also have "… < Y + (-X) + 2" by simp
finally show "s < Y + (-X) + 2" .
qed
moreover have "∀y ∈ ?SHIFT. 0 < y" by auto
moreover have shifted_not_empty: "∃u. u ∈ ?SHIFT"
using X_in_S and Y_isUb by auto
ultimately obtain t where t_is_Lub: "∀y. (∃x∈?SHIFT. y < x) = (y < t)"
using posreal_complete [of ?SHIFT] unfolding bdd_above_def by blast
show ?thesis
proof
show "(∀s∈S. s ≤ (t + X + (-1))) ∧ (∀y. (∀s∈S. s ≤ y) ⟶ (t + X + (-1)) ≤ y)"
proof safe
fix x
assume "∀s∈S. s ≤ x"
hence "∀s∈?SHIFT. s ≤ x + (-X) + 1"
using S_Ub_is_SHIFT_Ub by simp
then have "¬ x + (-X) + 1 < t"
by (subst t_is_Lub[rule_format, symmetric]) (simp add: not_less)
thus "t + X + -1 ≤ x" by arith
next
fix y
assume y_in_S: "y ∈ S"
obtain "u" where u_in_shift: "u ∈ ?SHIFT" using shifted_not_empty ..
hence "∃ x ∈ S. u = x + - X + 1" by simp
then obtain "x" where x_and_u: "u = x + - X + 1" ..
have u_le_t: "u ≤ t"
proof (rule dense_le)
fix x assume "x < u" then have "x < t"
using u_in_shift t_is_Lub by auto
then show "x ≤ t" by simp
qed
show "y ≤ t + X + -1"
proof cases
assume "y ≤ x"
moreover have "x = u + X + - 1" using x_and_u by arith
moreover have "u + X + - 1 ≤ t + X + -1" using u_le_t by arith
ultimately show "y ≤ t + X + -1" by arith
next
assume "~(y ≤ x)"
hence x_less_y: "x < y" by arith
have "x + (-X) + 1 ∈ ?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 ∈ ?SHIFT" using y_in_S by simp
have "y + (-X) + 1 ≤ t"
proof (rule dense_le)
fix x assume "x < y + (-X) + 1" then have "x < t"
using * t_is_Lub by auto
then show "x ≤ t" by simp
qed
thus ?thesis by simp
qed
qed
qed
qed
subsection ‹The Archimedean Property of the Reals›
theorem reals_Archimedean:
fixes x :: real
assumes x_pos: "0 < x"
shows "∃n. inverse (of_nat (Suc n)) < x"
proof (rule ccontr)
assume contr: "¬ ?thesis"
have "∀n. x * of_nat (Suc n) ≤ 1"
proof
fix n
from contr have "x ≤ inverse (of_nat (Suc n))"
by (simp add: linorder_not_less)
hence "x ≤ (1 / (of_nat (Suc n)))"
by (simp add: inverse_eq_divide)
moreover have "(0::real) ≤ of_nat (Suc n)"
by (rule of_nat_0_le_iff)
ultimately have "x * of_nat (Suc n) ≤ (1 / of_nat (Suc n)) * of_nat (Suc n)"
by (rule mult_right_mono)
thus "x * of_nat (Suc n) ≤ 1" by (simp del: of_nat_Suc)
qed
hence 2: "bdd_above {z. ∃n. z = x * (of_nat (Suc n))}"
by (auto intro!: bdd_aboveI[of _ 1])
have 1: "∃X. X ∈ {z. ∃n. z = x* (of_nat (Suc n))}" by auto
obtain t where
upper: "⋀z. z ∈ {z. ∃n. z = x * of_nat (Suc n)} ⟹ z ≤ t" and
least: "⋀y. (⋀a. a ∈ {z. ∃n. z = x * of_nat (Suc n)} ⟹ a ≤ y) ⟹ t ≤ y"
using reals_complete[OF 1 2] by auto
have "t ≤ t + - x"
proof (rule least)
fix a assume a: "a ∈ {z. ∃n. z = x * (of_nat (Suc n))}"
have "∀n::nat. x * of_nat n ≤ t + - x"
proof
fix n
have "x * of_nat (Suc n) ≤ t"
by (simp add: upper)
hence "x * (of_nat n) + x ≤ t"
by (simp add: distrib_left)
thus "x * (of_nat n) ≤ t + - x" by arith
qed hence "∀m. x * of_nat (Suc m) ≤ t + - x" by (simp del: of_nat_Suc)
with a show "a ≤ t + - x"
by auto
qed
thus False using x_pos by arith
qed
text ‹
There must be other proofs, e.g. ‹Suc› of the largest
integer in the cut representing ‹x›.
›
lemma reals_Archimedean2: "∃n. (x::real) < of_nat (n::nat)"
proof cases
assume "x ≤ 0"
hence "x < of_nat (1::nat)" by simp
thus ?thesis ..
next
assume "¬ x ≤ 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 "∃(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 ≤ of_int (int n)"
by simp
then show "∃z. r ≤ of_int z" ..
qed
end