section ‹Construction of real numbers›
theory Real_ZF_1 imports Real_ZF Int_ZF_3 OrderedField_ZF
begin
text‹In this theory file we continue the construction of real numbers started
in ‹Real_ZF› to a succesful conclusion.
We put here those parts of the construction that
can not be done in the general settings of abelian groups and require
integers.›
subsection‹Definitions and notation›
text‹In this section we define notions and notation needed for the rest of
the construction.›
text‹We define positive slopes as those that take an infinite number of
posititive values on the positive integers (see ‹Int_ZF_2›
for properties of positive slopes).›
definition
"PositiveSlopes ≡ {s ∈ Slopes.
s``(PositiveIntegers) ∩ PositiveIntegers ∉ Fin(int)}"
text‹The order on the set of real numbers is constructed by
specifying the set of positive reals. This set is defined
as the projection of the set of positive slopes.›
definition
"PositiveReals ≡ {SlopeEquivalenceRel``{s}. s ∈ PositiveSlopes}"
text‹The order relation on real numbers is constructed from the set of
positive elements in a standard way (see section
"Alternative definitions" in ‹OrderedGroup_ZF›.)
›
definition
"OrderOnReals ≡ OrderFromPosSet(RealNumbers,RealAddition,PositiveReals)"
text‹The next locale extends the locale ‹real0› to define notation
specific to the construction of real numbers. The notation follows the
one defined in ‹Int_ZF_2.thy›.
If $m$ is an integer, then the real number which is the class of the slope
$n\mapsto m\cdot n$ is denoted ‹m⇧R›.
For a real number $a$ notation $\lfloor a \rfloor$ means the largest integer
$m$ such that the real version of it (that is, $m^R$) is not greater
than $a$.
For an integer $m$ and a subset
of reals $S$ the expression $\Gamma(S,m)$ is defined as
$\max \{\lfloor p^R\cdot x\rfloor : x\in S\}$. This is plays a role in
the proof of completeness of real numbers.
We also reuse some notation defined in the ‹int0› context, like
‹ℤ⇩+ › (the set of positive integers) and abs$(m)$ ( the absolute
value of an integer, and some defined in the ‹int1› context, like
the addition ( ‹\<fp>›) and composition (‹∘›
of slopes.›
locale real1 = real0 +
fixes AlEq (infix "∼" 68)
defines AlEq_def[simp]: "s ∼ r ≡ ⟨s,r⟩ ∈ SlopeEquivalenceRel"
fixes slope_add (infix "\<fp>" 70)
defines slope_add_def[simp]:
"s \<fp> r ≡ SlopeOp1`⟨s,r⟩"
fixes slope_comp (infix "∘" 71)
defines slope_comp_def[simp]: "s ∘ r ≡ SlopeOp2`⟨s,r⟩"
fixes slopes ("𝒮")
defines slopes_def[simp]: "𝒮 ≡ AlmostHoms(int,IntegerAddition)"
fixes posslopes ("𝒮⇩+")
defines posslopes_def[simp]: "𝒮⇩+ ≡ PositiveSlopes"
fixes slope_class ("[ _ ]")
defines slope_class_def[simp]: "[f] ≡ SlopeEquivalenceRel``{f}"
fixes slope_neg ("\<fm>_" [90] 91)
defines slope_neg_def[simp]: "\<fm>s ≡ GroupInv(int,IntegerAddition) O s"
fixes lesseqr (infix "\<lsq>" 60)
defines lesseqr_def[simp]: "a \<lsq> b ≡ ⟨a,b⟩ ∈ OrderOnReals"
fixes sless (infix "\<ls>" 60)
defines sless_def[simp]: "a \<ls> b ≡ a\<lsq>b ∧ a≠b"
fixes positivereals ("ℝ⇩+")
defines positivereals_def[simp]: "ℝ⇩+ ≡ PositiveSet(ℝ,RealAddition,OrderOnReals)"
fixes intembed ("_⇧R" [90] 91)
defines intembed_def[simp]:
"m⇧R ≡ [{⟨n,IntegerMultiplication`⟨m,n⟩ ⟩. n ∈ int}]"
fixes floor ("⌊ _ ⌋")
defines floor_def[simp]:
"⌊a⌋ ≡ Maximum(IntegerOrder,{m ∈ int. m⇧R \<lsq> a})"
fixes Γ
defines Γ_def[simp]: "Γ(S,p) ≡ Maximum(IntegerOrder,{⌊p⇧R⋅x⌋. x∈S})"
fixes ia (infixl "\<za>" 69)
defines ia_def[simp]: "a\<za>b ≡ IntegerAddition`⟨ a,b⟩"
fixes iminus ("\<zm> _" 72)
defines iminus_def[simp]: "\<zm>a ≡ GroupInv(int,IntegerAddition)`(a)"
fixes isub (infixl "\<zs>" 69)
defines isub_def[simp]: "a\<zs>b ≡ a\<za> (\<zm> b)"
fixes intpositives ("ℤ⇩+")
defines intpositives_def[simp]:
"ℤ⇩+ ≡ PositiveSet(int,IntegerAddition,IntegerOrder)"
fixes zlesseq (infix "\<zlq>" 60)
defines lesseq_def[simp]: "m \<zlq> n ≡ ⟨m,n⟩ ∈ IntegerOrder"
fixes imult (infixl "\<zmu>" 70)
defines imult_def[simp]: "a\<zmu>b ≡ IntegerMultiplication`⟨ a,b⟩"
fixes izero ("𝟬⇩Z")
defines izero_def[simp]: "𝟬⇩Z ≡ TheNeutralElement(int,IntegerAddition)"
fixes ione ("𝟭⇩Z")
defines ione_def[simp]: "𝟭⇩Z ≡ TheNeutralElement(int,IntegerMultiplication)"
fixes itwo ("𝟮⇩Z")
defines itwo_def[simp]: "𝟮⇩Z ≡ 𝟭⇩Z\<za>𝟭⇩Z"
fixes abs
defines abs_def[simp]:
"abs(m) ≡ AbsoluteValue(int,IntegerAddition,IntegerOrder)`(m)"
fixes δ
defines δ_def[simp]: "δ(s,m,n) ≡ s`(m\<za>n)\<zs>s`(m)\<zs>s`(n)"
subsection‹Multiplication of real numbers›
text‹Multiplication of real numbers is defined as a projection of
composition of slopes onto the space of equivalence classes of slopes.
Thus, the product of the real numbers given as classes of slopes $s$ and
$r$ is defined as the class of $s\circ r$. The goal of this section is to
show that multiplication defined this way is commutative.›
text‹Let's recall a theorem from ‹Int_ZF_2.thy› that states that
if $f,g$ are slopes,
then $f\circ g$ is equivalent to $g\circ f$.
Here we conclude from that that
the classes of $f\circ g$ and $g\circ f$ are the same.›
lemma (in real1) Real_ZF_1_1_L2: assumes A1: "f ∈ 𝒮" "g ∈ 𝒮"
shows "[f∘g] = [g∘f]"
proof -
from A1 have "f∘g ∼ g∘f"
using Slopes_def int1.Arthan_Th_9 SlopeOp1_def BoundedIntMaps_def
SlopeEquivalenceRel_def SlopeOp2_def by simp
then show ?thesis using Real_ZF_1_L11 equiv_class_eq
by simp
qed
text‹Classes of slopes are real numbers.›
lemma (in real1) Real_ZF_1_1_L3: assumes A1: "f ∈ 𝒮"
shows "[f] ∈ ℝ"
proof -
from A1 have "[f] ∈ Slopes//SlopeEquivalenceRel"
using Slopes_def quotientI by simp
then show "[f] ∈ ℝ" using RealNumbers_def by simp
qed
text‹Each real number is a class of a slope.›
lemma (in real1) Real_ZF_1_1_L3A: assumes A1: "a∈ℝ"
shows "∃f∈𝒮 . a = [f]"
proof -
from A1 have "a ∈ 𝒮//SlopeEquivalenceRel"
using RealNumbers_def Slopes_def by simp
then show ?thesis using quotient_def
by simp
qed
text‹It is useful to have the definition of addition and multiplication
in the ‹real1› context notation.›
lemma (in real1) Real_ZF_1_1_L4:
assumes A1: "f ∈ 𝒮" "g ∈ 𝒮"
shows
"[f] \<ra> [g] = [f\<fp>g]"
"[f] ⋅ [g] = [f∘g]"
proof -
let ?r = "SlopeEquivalenceRel"
have "[f]⋅[g] = ProjFun2(𝒮,?r,SlopeOp2)`⟨[f],[g]⟩"
using RealMultiplication_def Slopes_def by simp
also from A1 have "… = [f∘g]"
using Real_ZF_1_L11 EquivClass_1_L10 Slopes_def
by simp
finally show "[f] ⋅ [g] = [f∘g]" by simp
have "[f] \<ra> [g] = ProjFun2(𝒮,?r,SlopeOp1)`⟨[f],[g]⟩"
using RealAddition_def Slopes_def by simp
also from A1 have "… = [f\<fp>g]"
using Real_ZF_1_L11 EquivClass_1_L10 Slopes_def
by simp
finally show "[f] \<ra> [g] = [f\<fp>g]" by simp
qed
text‹The next lemma is essentially the same as ‹ Real_ZF_1_L12›, but
written in the notation defined in the ‹real1› context. It states
that if $f$ is a slope, then $-[f] = [-f]$.›
lemma (in real1) Real_ZF_1_1_L4A: assumes "f ∈ 𝒮"
shows "[\<fm>f] = \<rm>[f]"
using assms Slopes_def SlopeEquivalenceRel_def Real_ZF_1_L12
by simp
text‹Subtracting real numbers correspods to adding the opposite slope.›
lemma (in real1) Real_ZF_1_1_L4B: assumes A1: "f ∈ 𝒮" "g ∈ 𝒮"
shows "[f] \<rs> [g] = [f\<fp>(\<fm>g)]"
proof -
from A1 have "[f\<fp>(\<fm>g)] = [f] \<ra> [\<fm>g]"
using Slopes_def BoundedIntMaps_def int1.Int_ZF_2_1_L12
Real_ZF_1_1_L4 by simp
with A1 show "[f] \<rs> [g] = [f\<fp>(\<fm>g)]"
using Real_ZF_1_1_L4A by simp
qed
text‹Multiplication of real numbers is commutative.›
theorem (in real1) real_mult_commute: assumes A1: "a∈ℝ" "b∈ℝ"
shows "a⋅b = b⋅a"
proof -
from A1 have
"∃f∈𝒮 . a = [f]"
"∃g∈𝒮 . b = [g]"
using Real_ZF_1_1_L3A by auto
then obtain f g where
"f ∈ 𝒮" "g ∈ 𝒮" and "a = [f]" "b = [g]"
by auto
then show "a⋅b = b⋅a"
using Real_ZF_1_1_L4 Real_ZF_1_1_L2 by simp
qed
text‹Multiplication is commutative on reals.›
lemma real_mult_commutative: shows
"RealMultiplication {is commutative on} RealNumbers"
using real1.real_mult_commute IsCommutative_def
by simp
text‹The neutral element of multiplication of reals
(denoted as ‹𝟭› in the ‹real1› context)
is the class of identity function on integers. This is really shown
in ‹Real_ZF_1_L11›, here we only rewrite it in the notation used
in the ‹real1› context.›
lemma (in real1) real_one_cl_identity: shows "[id(int)] = 𝟭"
using Real_ZF_1_L11 by simp
text‹If $f$ is bounded, then its class is the neutral element of additive
operation on reals (denoted as ‹𝟬› in the ‹real1› context).›
lemma (in real1) real_zero_cl_bounded_map:
assumes "f ∈ BoundedIntMaps" shows "[f] = 𝟬"
using assms Real_ZF_1_L11A by simp
text‹Two real numbers are equal iff the slopes that represent them are
almost equal. This is proven in ‹Real_ZF_1_L13›, here we just
rewrite it in the notation used in the ‹real1› context.›
lemma (in real1) Real_ZF_1_1_L5:
assumes "f ∈ 𝒮" "g ∈ 𝒮"
shows "[f] = [g] ⟷ f ∼ g"
using assms Slopes_def Real_ZF_1_L13 by simp
text‹If the pair of function belongs to the slope equivalence relation, then
their classes are equal. This is convenient, because we don't need to assume
that $f,g$ are slopes (follows from the fact that $f\sim g$).›
lemma (in real1) Real_ZF_1_1_L5A: assumes "f ∼ g"
shows "[f] = [g]"
using assms Real_ZF_1_L11 Slopes_def Real_ZF_1_1_L5
by auto
text‹Identity function on integers is a slope.
This is proven in ‹Real_ZF_1_L13›, here we just
rewrite it in the notation used in the ‹real1› context.›
lemma (in real1) id_on_int_is_slope: shows "id(int) ∈ 𝒮"
using Real_ZF_1_L14 Slopes_def by simp
text‹A result from ‹Int_ZF_2.thy›: the identity function on integers
is not almost equal to any bounded function.›
lemma (in real1) Real_ZF_1_1_L7:
assumes A1: "f ∈ BoundedIntMaps"
shows "¬(id(int) ∼ f)"
using assms Slopes_def SlopeOp1_def BoundedIntMaps_def
SlopeEquivalenceRel_def BoundedIntMaps_def int1.Int_ZF_2_3_L12
by simp
text‹Zero is not one.›
lemma (in real1) real_zero_not_one: shows "𝟭≠𝟬"
proof -
{ assume A1: "𝟭=𝟬"
have "∃f ∈ 𝒮. 𝟬 = [f]"
using Real_ZF_1_L4 Real_ZF_1_1_L3A by simp
with A1 have
"∃f ∈ 𝒮. [id(int)] = [f] ∧ [f] = 𝟬"
using real_one_cl_identity by auto
then have False using Real_ZF_1_1_L5 Slopes_def
Real_ZF_1_L10 Real_ZF_1_1_L7 id_on_int_is_slope
by auto
} then show "𝟭≠𝟬" by auto
qed
text‹Negative of a real number is a real number. Property of groups.›
lemma (in real1) Real_ZF_1_1_L8: assumes "a∈ℝ" shows "(\<rm>a) ∈ ℝ"
using assms Real_ZF_1_L2 group0.inverse_in_group
by simp
text‹An identity with three real numbers.›
lemma (in real1) Real_ZF_1_1_L9: assumes "a∈ℝ" "b∈ℝ" "c∈ℝ"
shows "a⋅(b⋅c) = a⋅c⋅b"
using assms real_mult_commutative Real_ZF_1_L3 ring0.Ring_ZF_2_L4
by simp
subsection‹The order on reals›
text‹In this section we show that the order relation defined by prescribing
the set of positive reals as the projection of the set of positive
slopes makes the ring of real numbers into an ordered ring. We also
collect the facts about ordered groups and rings that we use in
the construction.›
text‹Positive slopes are slopes and positive reals are real.›
lemma Real_ZF_1_2_L1: shows
"PositiveSlopes ⊆ Slopes"
"PositiveReals ⊆ RealNumbers"
proof -
have "PositiveSlopes =
{s ∈ Slopes. s``(PositiveIntegers) ∩ PositiveIntegers ∉ Fin(int)}"
using PositiveSlopes_def by simp
then show "PositiveSlopes ⊆ Slopes" by (rule subset_with_property)
then have
"{SlopeEquivalenceRel``{s}. s ∈ PositiveSlopes } ⊆
Slopes//SlopeEquivalenceRel"
using EquivClass_1_L1A by simp
then show "PositiveReals ⊆ RealNumbers"
using PositiveReals_def RealNumbers_def by simp
qed
text‹Positive reals are the same as classes of a positive slopes.›
lemma (in real1) Real_ZF_1_2_L2:
shows "a ∈ PositiveReals ⟷ (∃f∈𝒮⇩+. a = [f])"
proof
assume "a ∈ PositiveReals"
then have "a ∈ {([s]). s ∈ 𝒮⇩+}" using PositiveReals_def
by simp
then show "∃f∈𝒮⇩+. a = [f]" by auto
next assume "∃f∈𝒮⇩+. a = [f]"
then have "a ∈ {([s]). s ∈ 𝒮⇩+}" by auto
then show "a ∈ PositiveReals" using PositiveReals_def
by simp
qed
text‹Let's recall from ‹Int_ZF_2.thy› that the sum and composition
of positive slopes is a positive slope.›
lemma (in real1) Real_ZF_1_2_L3:
assumes "f∈𝒮⇩+" "g∈𝒮⇩+"
shows
"f\<fp>g ∈ 𝒮⇩+"
"f∘g ∈ 𝒮⇩+"
using assms Slopes_def PositiveSlopes_def PositiveIntegers_def
SlopeOp1_def int1.sum_of_pos_sls_is_pos_sl
SlopeOp2_def int1.comp_of_pos_sls_is_pos_sl
by auto
text‹Bounded integer maps are not positive slopes.›
lemma (in real1) Real_ZF_1_2_L5:
assumes "f ∈ BoundedIntMaps"
shows "f ∉ 𝒮⇩+"
using assms BoundedIntMaps_def Slopes_def PositiveSlopes_def
PositiveIntegers_def int1.Int_ZF_2_3_L1B by simp
text‹The set of positive reals is closed under addition and multiplication.
Zero (the neutral element of addition) is not a positive number.›
lemma (in real1) Real_ZF_1_2_L6: shows
"PositiveReals {is closed under} RealAddition"
"PositiveReals {is closed under} RealMultiplication"
"𝟬 ∉ PositiveReals"
proof -
{ fix a fix b
assume "a ∈ PositiveReals" and "b ∈ PositiveReals"
then obtain f g where
I: "f ∈ 𝒮⇩+" "g ∈ 𝒮⇩+" and
II: "a = [f]" "b = [g]"
using Real_ZF_1_2_L2 by auto
then have "f ∈ 𝒮" "g ∈ 𝒮" using Real_ZF_1_2_L1 Slopes_def
by auto
with I II have
"a\<ra>b ∈ PositiveReals ∧ a⋅b ∈ PositiveReals"
using Real_ZF_1_1_L4 Real_ZF_1_2_L3 Real_ZF_1_2_L2
by auto
} then show
"PositiveReals {is closed under} RealAddition"
"PositiveReals {is closed under} RealMultiplication"
using IsOpClosed_def
by auto
{ assume "𝟬 ∈ PositiveReals"
then obtain f where "f ∈ 𝒮⇩+" and "𝟬 = [f]"
using Real_ZF_1_2_L2 by auto
then have False
using Real_ZF_1_2_L1 Slopes_def Real_ZF_1_L10 Real_ZF_1_2_L5
by auto
} then show "𝟬 ∉ PositiveReals" by auto
qed
text‹If a class of a slope $f$ is not zero, then either $f$ is
a positive slope or $-f$ is a positive slope. The real proof is in
‹Int_ZF_2.thy.››
lemma (in real1) Real_ZF_1_2_L7:
assumes A1: "f ∈ 𝒮" and A2: "[f] ≠ 𝟬"
shows "(f ∈ 𝒮⇩+) Xor ((\<fm>f) ∈ 𝒮⇩+)"
using assms Slopes_def SlopeEquivalenceRel_def BoundedIntMaps_def
PositiveSlopes_def PositiveIntegers_def
Real_ZF_1_L10 int1.Int_ZF_2_3_L8 by simp
text‹The next lemma rephrases ‹ Int_ZF_2_3_L10› in the notation
used in ‹real1› context.›
lemma (in real1) Real_ZF_1_2_L8:
assumes A1: "f ∈ 𝒮" "g ∈ 𝒮"
and A2: "(f ∈ 𝒮⇩+) Xor (g ∈ 𝒮⇩+)"
shows "([f] ∈ PositiveReals) Xor ([g] ∈ PositiveReals)"
using assms PositiveReals_def SlopeEquivalenceRel_def Slopes_def
SlopeOp1_def BoundedIntMaps_def PositiveSlopes_def PositiveIntegers_def
int1.Int_ZF_2_3_L10 by simp
text‹The trichotomy law for the (potential) order on reals: if $a\neq 0$,
then either $a$ is positive or $-a$ is potitive.›
lemma (in real1) Real_ZF_1_2_L9:
assumes A1: "a∈ℝ" and A2: "a≠𝟬"
shows "(a ∈ PositiveReals) Xor ((\<rm>a) ∈ PositiveReals)"
proof -
from A1 obtain f where I: "f ∈ 𝒮" "a = [f]"
using Real_ZF_1_1_L3A by auto
with A2 have "([f] ∈ PositiveReals) Xor ([\<fm>f] ∈ PositiveReals)"
using Slopes_def BoundedIntMaps_def int1.Int_ZF_2_1_L12
Real_ZF_1_2_L7 Real_ZF_1_2_L8 by simp
with I show "(a ∈ PositiveReals) Xor ((\<rm>a) ∈ PositiveReals)"
using Real_ZF_1_1_L4A by simp
qed
text‹Finally we are ready to prove that real numbers form an ordered ring
with no zero divisors.›
theorem reals_are_ord_ring: shows
"IsAnOrdRing(RealNumbers,RealAddition,RealMultiplication,OrderOnReals)"
"OrderOnReals {is total on} RealNumbers"
"PositiveSet(RealNumbers,RealAddition,OrderOnReals) = PositiveReals"
"HasNoZeroDivs(RealNumbers,RealAddition,RealMultiplication)"
proof -
let ?R = "RealNumbers"
let ?A = "RealAddition"
let ?M = "RealMultiplication"
let ?P = "PositiveReals"
let ?r = "OrderOnReals"
let ?z = "TheNeutralElement(?R, ?A)"
have I:
"ring0(?R, ?A, ?M)"
"?M {is commutative on} ?R"
"?P ⊆ ?R"
"?P {is closed under} ?A"
"TheNeutralElement(?R, ?A) ∉ ?P"
"∀a∈?R. a ≠ ?z ⟶ (a ∈ ?P) Xor (GroupInv(?R, ?A)`(a) ∈ ?P)"
"?P {is closed under} ?M"
"?r = OrderFromPosSet(?R, ?A, ?P)"
using real0.Real_ZF_1_L3 real_mult_commutative Real_ZF_1_2_L1
real1.Real_ZF_1_2_L6 real1.Real_ZF_1_2_L9 OrderOnReals_def
by auto
then show "IsAnOrdRing(?R, ?A, ?M, ?r)"
by (rule ring0.ring_ord_by_positive_set)
from I show "?r {is total on} ?R"
by (rule ring0.ring_ord_by_positive_set)
from I show "PositiveSet(?R,?A,?r) = ?P"
by (rule ring0.ring_ord_by_positive_set)
from I show "HasNoZeroDivs(?R,?A,?M)"
by (rule ring0.ring_ord_by_positive_set)
qed
text‹All theorems proven in the ‹ring1›
(about ordered rings), ‹group3› (about ordered groups) and
‹group1› (about groups)
contexts are valid as applied to ordered real numbers with addition
and (real) order.›
lemma Real_ZF_1_2_L10: shows
"ring1(RealNumbers,RealAddition,RealMultiplication,OrderOnReals)"
"IsAnOrdGroup(RealNumbers,RealAddition,OrderOnReals)"
"group3(RealNumbers,RealAddition,OrderOnReals)"
"OrderOnReals {is total on} RealNumbers"
proof -
show "ring1(RealNumbers,RealAddition,RealMultiplication,OrderOnReals)"
using reals_are_ord_ring OrdRing_ZF_1_L2 by simp
then show
"IsAnOrdGroup(RealNumbers,RealAddition,OrderOnReals)"
"group3(RealNumbers,RealAddition,OrderOnReals)"
"OrderOnReals {is total on} RealNumbers"
using ring1.OrdRing_ZF_1_L4 by auto
qed
text‹If $a=b$ or $b-a$ is positive, then $a$ is less or equal $b$.›
lemma (in real1) Real_ZF_1_2_L11: assumes A1: "a∈ℝ" "b∈ℝ" and
A3: "a=b ∨ b\<rs>a ∈ PositiveReals"
shows "a\<lsq>b"
using assms reals_are_ord_ring Real_ZF_1_2_L10
group3.OrderedGroup_ZF_1_L30 by simp
text‹A sufficient condition for two classes to be in the real order.›
lemma (in real1) Real_ZF_1_2_L12: assumes A1: "f ∈ 𝒮" "g ∈ 𝒮" and
A2: "f∼g ∨ (g \<fp> (\<fm>f)) ∈ 𝒮⇩+"
shows "[f] \<lsq> [g]"
proof -
from A1 A2 have "[f] = [g] ∨ [g]\<rs>[f] ∈ PositiveReals"
using Real_ZF_1_1_L5A Real_ZF_1_2_L2 Real_ZF_1_1_L4B
by auto
with A1 show "[f] \<lsq> [g]" using Real_ZF_1_1_L3 Real_ZF_1_2_L11
by simp
qed
text‹Taking negative on both sides reverses the inequality, a case with
an inverse on one side. Property of ordered groups.›
lemma (in real1) Real_ZF_1_2_L13:
assumes A1: "a∈ℝ" and A2: "(\<rm>a) \<lsq> b"
shows "(\<rm>b) \<lsq> a"
using assms Real_ZF_1_2_L10 group3.OrderedGroup_ZF_1_L5AG
by simp
text‹Real order is antisymmetric.›
lemma (in real1) real_ord_antisym:
assumes A1: "a\<lsq>b" "b\<lsq>a" shows "a=b"
proof -
from A1 have
"group3(RealNumbers,RealAddition,OrderOnReals)"
"⟨a,b⟩ ∈ OrderOnReals" "⟨b,a⟩ ∈ OrderOnReals"
using Real_ZF_1_2_L10 by auto
then show "a=b" by (rule group3.group_order_antisym)
qed
text‹Real order is transitive.›
lemma (in real1) real_ord_transitive: assumes A1: "a\<lsq>b" "b\<lsq>c"
shows "a\<lsq>c"
proof -
from A1 have
"group3(RealNumbers,RealAddition,OrderOnReals)"
"⟨a,b⟩ ∈ OrderOnReals" "⟨b,c⟩ ∈ OrderOnReals"
using Real_ZF_1_2_L10 by auto
then have "⟨a,c⟩ ∈ OrderOnReals"
by (rule group3.Group_order_transitive)
then show "a\<lsq>c" by simp
qed
text‹We can multiply both sides of an inequality
by a nonnegative real number.›
lemma (in real1) Real_ZF_1_2_L14:
assumes "a\<lsq>b" and "𝟬\<lsq>c"
shows
"a⋅c \<lsq> b⋅c"
"c⋅a \<lsq> c⋅b"
using assms Real_ZF_1_2_L10 ring1.OrdRing_ZF_1_L9
by auto
text‹A special case of ‹Real_ZF_1_2_L14›: we can multiply
an inequality by a real number.›
lemma (in real1) Real_ZF_1_2_L14A:
assumes A1: "a\<lsq>b" and A2: "c∈ℝ⇩+"
shows "c⋅a \<lsq> c⋅b"
using assms Real_ZF_1_2_L10 ring1.OrdRing_ZF_1_L9A
by simp
text‹In the ‹real1› context notation $a\leq b$
implies that $a$ and $b$ are real numbers.›
lemma (in real1) Real_ZF_1_2_L15: assumes "a\<lsq>b" shows "a∈ℝ" "b∈ℝ"
using assms Real_ZF_1_2_L10 group3.OrderedGroup_ZF_1_L4
by auto
text‹$a\leq b$ implies that $0 \leq b -a$.›
lemma (in real1) Real_ZF_1_2_L16: assumes "a\<lsq>b"
shows "𝟬 \<lsq> b\<rs>a"
using assms Real_ZF_1_2_L10 group3.OrderedGroup_ZF_1_L12A
by simp
text‹A sum of nonnegative elements is nonnegative.›
lemma (in real1) Real_ZF_1_2_L17: assumes "𝟬\<lsq>a" "𝟬\<lsq>b"
shows "𝟬 \<lsq> a\<ra>b"
using assms Real_ZF_1_2_L10 group3.OrderedGroup_ZF_1_L12
by simp
text‹We can add sides of two inequalities›
lemma (in real1) Real_ZF_1_2_L18: assumes "a\<lsq>b" "c\<lsq>d"
shows "a\<ra>c \<lsq> b\<ra>d"
using assms Real_ZF_1_2_L10 group3.OrderedGroup_ZF_1_L5B
by simp
text‹The order on real is reflexive.›
lemma (in real1) real_ord_refl: assumes "a∈ℝ" shows "a\<lsq>a"
using assms Real_ZF_1_2_L10 group3.OrderedGroup_ZF_1_L3
by simp
text‹We can add a real number to both sides of an inequality.›
lemma (in real1) add_num_to_ineq: assumes "a\<lsq>b" and "c∈ℝ"
shows "a\<ra>c \<lsq> b\<ra>c"
using assms Real_ZF_1_2_L10 IsAnOrdGroup_def by simp
text‹We can put a number on the other side of an inequality,
changing its sign.›
lemma (in real1) Real_ZF_1_2_L19:
assumes "a∈ℝ" "b∈ℝ" and "c \<lsq> a\<ra>b"
shows "c\<rs>b \<lsq> a"
using assms Real_ZF_1_2_L10 group3.OrderedGroup_ZF_1_L9C
by simp
text‹What happens when one real number is not greater or equal than another?›
lemma (in real1) Real_ZF_1_2_L20: assumes "a∈ℝ" "b∈ℝ" and "¬(a\<lsq>b)"
shows "b \<ls> a"
proof -
from assms have I:
"group3(ℝ,RealAddition,OrderOnReals)"
"OrderOnReals {is total on} ℝ"
"a∈ℝ" "b∈ℝ" "¬(⟨a,b⟩ ∈ OrderOnReals)"
using Real_ZF_1_2_L10 by auto
then have "⟨b,a⟩ ∈ OrderOnReals"
by (rule group3.OrderedGroup_ZF_1_L8)
then have "b \<lsq> a" by simp
moreover from I have "a≠b" by (rule group3.OrderedGroup_ZF_1_L8)
ultimately show "b \<ls> a" by auto
qed
text‹We can put a number on the other side of an inequality,
changing its sign, version with a minus.›
lemma (in real1) Real_ZF_1_2_L21:
assumes "a∈ℝ" "b∈ℝ" and "c \<lsq> a\<rs>b"
shows "c\<ra>b \<lsq> a"
using assms Real_ZF_1_2_L10 group3.OrderedGroup_ZF_1_L5J
by simp
text‹The order on reals is a relation on reals.›
lemma (in real1) Real_ZF_1_2_L22: shows "OrderOnReals ⊆ ℝ×ℝ"
using Real_ZF_1_2_L10 IsAnOrdGroup_def
by simp
text‹A set that is bounded above in the sense defined by order
on reals is a subset of real numbers.›
lemma (in real1) Real_ZF_1_2_L23:
assumes A1: "IsBoundedAbove(A,OrderOnReals)"
shows "A ⊆ ℝ"
using A1 Real_ZF_1_2_L22 Order_ZF_3_L1A
by blast
text‹Properties of the maximum of three real numbers.›
lemma (in real1) Real_ZF_1_2_L24:
assumes A1: "a∈ℝ" "b∈ℝ" "c∈ℝ"
shows
"Maximum(OrderOnReals,{a,b,c}) ∈ {a,b,c}"
"Maximum(OrderOnReals,{a,b,c}) ∈ ℝ"
"a \<lsq> Maximum(OrderOnReals,{a,b,c})"
"b \<lsq> Maximum(OrderOnReals,{a,b,c})"
"c \<lsq> Maximum(OrderOnReals,{a,b,c})"
proof -
have "IsLinOrder(ℝ,OrderOnReals)"
using Real_ZF_1_2_L10 group3.group_ord_total_is_lin
by simp
with A1 show
"Maximum(OrderOnReals,{a,b,c}) ∈ {a,b,c}"
"Maximum(OrderOnReals,{a,b,c}) ∈ ℝ"
"a \<lsq> Maximum(OrderOnReals,{a,b,c})"
"b \<lsq> Maximum(OrderOnReals,{a,b,c})"
"c \<lsq> Maximum(OrderOnReals,{a,b,c})"
using Finite_ZF_1_L2A by auto
qed
text‹A form of transitivity for the order on reals.›
lemma (in real1) real_strict_ord_transit:
assumes A1: "a\<lsq>b" and A2: "b\<ls>c"
shows "a\<ls>c"
proof -
from A1 A2 have I:
"group3(ℝ,RealAddition,OrderOnReals)"
"⟨a,b⟩ ∈ OrderOnReals" "⟨b,c⟩ ∈ OrderOnReals ∧ b≠c"
using Real_ZF_1_2_L10 by auto
then have "⟨a,c⟩ ∈ OrderOnReals ∧ a≠c" by (rule group3.group_strict_ord_transit)
then show "a\<ls>c" by simp
qed
text‹We can multiply a right hand side of an inequality between
positive real numbers by a number that is greater than one.›
lemma (in real1) Real_ZF_1_2_L25:
assumes "b ∈ ℝ⇩+" and "a\<lsq>b" and "𝟭\<ls>c"
shows "a\<ls>b⋅c"
using assms reals_are_ord_ring Real_ZF_1_2_L10 ring1.OrdRing_ZF_3_L17
by simp
text‹We can move a real number to the other side of a strict inequality,
changing its sign.›
lemma (in real1) Real_ZF_1_2_L26:
assumes "a∈ℝ" "b∈ℝ" and "a\<rs>b \<ls> c"
shows "a \<ls> c\<ra>b"
using assms Real_ZF_1_2_L10 group3.OrderedGroup_ZF_1_L12B
by simp
text‹Real order is translation invariant.›
lemma (in real1) real_ord_transl_inv:
assumes "a\<lsq>b" and "c∈ℝ"
shows "c\<ra>a \<lsq> c\<ra>b"
using assms Real_ZF_1_2_L10 IsAnOrdGroup_def
by simp
text‹It is convenient to have the transitivity of the order on integers
in the notation specific to ‹real1› context. This may be confusing
for the presentation readers: even though ‹\<zlq>› and
‹\<lsq>› are printed in the same way, they are different symbols
in the source. In the ‹real1› context the former denotes
inequality between integers, and the latter denotes inequality between real
numbers (classes of slopes). The next lemma is about transitivity of the
order relation on integers.›
lemma (in real1) int_order_transitive:
assumes A1: "a\<zlq>b" "b\<zlq>c"
shows "a\<zlq>c"
proof -
from A1 have
"⟨a,b⟩ ∈ IntegerOrder" and "⟨b,c⟩ ∈ IntegerOrder"
by auto
then have "⟨a,c⟩ ∈ IntegerOrder"
by (rule Int_ZF_2_L5)
then show "a\<zlq>c" by simp
qed
text‹A property of nonempty subsets of real numbers that don't
have a maximum: for any element we can find one that is (strictly) greater.›
lemma (in real1) Real_ZF_1_2_L27:
assumes "A⊆ℝ" and "¬HasAmaximum(OrderOnReals,A)" and "x∈A"
shows "∃y∈A. x\<ls>y"
using assms Real_ZF_1_2_L10 group3.OrderedGroup_ZF_2_L2B
by simp
text‹The next lemma shows what happens when one real number
is not greater or equal than another.›
lemma (in real1) Real_ZF_1_2_L28:
assumes "a∈ℝ" "b∈ℝ" and "¬(a\<lsq>b)"
shows "b\<ls>a"
proof -
from assms have
"group3(ℝ,RealAddition,OrderOnReals)"
"OrderOnReals {is total on} ℝ"
"a∈ℝ" "b∈ℝ" "⟨a,b⟩ ∉ OrderOnReals"
using Real_ZF_1_2_L10 by auto
then have "⟨b,a⟩ ∈ OrderOnReals ∧ b≠a"
by (rule group3.OrderedGroup_ZF_1_L8)
then show "b\<ls>a" by simp
qed
text‹If a real number is less than another, then the second one can not
be less or equal that the first.›
lemma (in real1) Real_ZF_1_2_L29:
assumes "a\<ls>b" shows "¬(b\<lsq>a)"
proof -
from assms have
"group3(ℝ,RealAddition,OrderOnReals)"
"⟨a,b⟩ ∈ OrderOnReals" "a≠b"
using Real_ZF_1_2_L10 by auto
then have "⟨b,a⟩ ∉ OrderOnReals"
by (rule group3.OrderedGroup_ZF_1_L8AA)
then show "¬(b\<lsq>a)" by simp
qed
subsection‹Inverting reals›
text‹In this section we tackle the issue of existence of (multiplicative)
inverses of real numbers and show that real numbers form an
ordered field. We also restate here some facts specific to ordered fields
that we need for the construction. The actual proofs of most of these facts
can be found in ‹Field_ZF.thy› and ‹OrderedField_ZF.thy››
text‹We rewrite the theorem from ‹Int_ZF_2.thy› that shows
that for every positive slope we can find one that is almost equal and
has an inverse.›
lemma (in real1) pos_slopes_have_inv: assumes "f ∈ 𝒮⇩+"
shows "∃g∈𝒮. f∼g ∧ (∃h∈𝒮. g∘h ∼ id(int))"
using assms PositiveSlopes_def Slopes_def PositiveIntegers_def
int1.pos_slope_has_inv SlopeOp1_def SlopeOp2_def
BoundedIntMaps_def SlopeEquivalenceRel_def
by simp
text‹The set of real numbers we are constructing is an ordered field.›
theorem (in real1) reals_are_ord_field: shows
"IsAnOrdField(RealNumbers,RealAddition,RealMultiplication,OrderOnReals)"
proof -
let ?R = "RealNumbers"
let ?A = "RealAddition"
let ?M = "RealMultiplication"
let ?r = "OrderOnReals"
have "ring1(?R,?A,?M,?r)" and "𝟬 ≠ 𝟭"
using reals_are_ord_ring OrdRing_ZF_1_L2 real_zero_not_one
by auto
moreover have "?M {is commutative on} ?R"
using real_mult_commutative by simp
moreover have
"∀a∈PositiveSet(?R,?A,?r). ∃b∈?R. a⋅b = 𝟭"
proof
fix a assume "a ∈ PositiveSet(?R,?A,?r)"
then obtain f where I: "f∈𝒮⇩+" and II: "a = [f]"
using reals_are_ord_ring Real_ZF_1_2_L2
by auto
then have "∃g∈𝒮. f∼g ∧ (∃h∈𝒮. g∘h ∼ id(int))"
using pos_slopes_have_inv by simp
then obtain g where
III: "g∈𝒮" and IV: "f∼g" and V: "∃h∈𝒮. g∘h ∼ id(int)"
by auto
from V obtain h where VII: "h∈𝒮" and VIII: "g∘h ∼ id(int)"
by auto
from I III IV have "[f] = [g]"
using Real_ZF_1_2_L1 Slopes_def Real_ZF_1_1_L5
by auto
with II III VII VIII have "a⋅[h] = 𝟭"
using Real_ZF_1_1_L4 Real_ZF_1_1_L5A real_one_cl_identity
by simp
with VII show "∃b∈?R. a⋅b = 𝟭" using Real_ZF_1_1_L3
by auto
qed
ultimately show ?thesis using ring1.OrdField_ZF_1_L4
by simp
qed
text‹Reals form a field.›
lemma reals_are_field:
shows "IsAfield(RealNumbers,RealAddition,RealMultiplication)"
using real1.reals_are_ord_field OrdField_ZF_1_L1A
by simp
text‹Theorem proven in ‹field0› and ‹field1› contexts
are valid as applied to real numbers.›
lemma field_cntxts_ok: shows
"field0(RealNumbers,RealAddition,RealMultiplication)"
"field1(RealNumbers,RealAddition,RealMultiplication,OrderOnReals)"
using reals_are_field real1.reals_are_ord_field
field_field0 OrdField_ZF_1_L2 by auto
text‹If $a$ is positive, then $a^{-1}$ is also positive.›
lemma (in real1) Real_ZF_1_3_L1: assumes "a ∈ ℝ⇩+"
shows "a¯ ∈ ℝ⇩+" "a¯ ∈ ℝ"
using assms field_cntxts_ok field1.OrdField_ZF_1_L8 PositiveSet_def
by auto
text‹A technical fact about multiplying strict inequality by the inverse
of one of the sides.›
lemma (in real1) Real_ZF_1_3_L2:
assumes "a ∈ ℝ⇩+" and "a¯ \<ls> b"
shows "𝟭 \<ls> b⋅a"
using assms field_cntxts_ok field1.OrdField_ZF_2_L2
by simp
text‹If $a$ is smaller than $b$, then $(b-a)^{-1}$ is positive.›
lemma (in real1) Real_ZF_1_3_L3: assumes "a\<ls>b"
shows "(b\<rs>a)¯ ∈ ℝ⇩+"
using assms field_cntxts_ok field1.OrdField_ZF_1_L9
by simp
text‹We can put a positive factor on the other side of a strict
inequality, changing it to its inverse.›
lemma (in real1) Real_ZF_1_3_L4:
assumes A1: "a∈ℝ" "b∈ℝ⇩+" and A2: "a⋅b \<ls> c"
shows "a \<ls> c⋅b¯"
using assms field_cntxts_ok field1.OrdField_ZF_2_L6
by simp
text‹We can put a positive factor on the other side of a strict
inequality, changing it to its inverse, version with the product
initially on the right hand side.›
lemma (in real1) Real_ZF_1_3_L4A:
assumes A1: "b∈ℝ" "c∈ℝ⇩+" and A2: "a \<ls> b⋅c"
shows "a⋅c¯ \<ls> b"
using assms field_cntxts_ok field1.OrdField_ZF_2_L6A
by simp
text‹We can put a positive factor on the other side of an inequality,
changing it to its inverse, version with the product
initially on the right hand side.›
lemma (in real1) Real_ZF_1_3_L4B:
assumes A1: "b∈ℝ" "c∈ℝ⇩+" and A2: "a \<lsq> b⋅c"
shows "a⋅c¯ \<lsq> b"
using assms field_cntxts_ok field1.OrdField_ZF_2_L5A
by simp
text‹We can put a positive factor on the other side of an inequality,
changing it to its inverse, version with the product
initially on the left hand side.›
lemma (in real1) Real_ZF_1_3_L4C:
assumes A1: "a∈ℝ" "b∈ℝ⇩+" and A2: "a⋅b \<lsq> c"
shows "a \<lsq> c⋅b¯"
using assms field_cntxts_ok field1.OrdField_ZF_2_L5
by simp
text‹A technical lemma about solving a strict inequality with three
real numbers and inverse of a difference.›
lemma (in real1) Real_ZF_1_3_L5:
assumes "a\<ls>b" and "(b\<rs>a)¯ \<ls> c"
shows "𝟭 \<ra> a⋅c \<ls> b⋅c"
using assms field_cntxts_ok field1.OrdField_ZF_2_L9
by simp
text‹We can multiply an inequality by the inverse of a positive number.›
lemma (in real1) Real_ZF_1_3_L6:
assumes "a\<lsq>b" and "c∈ℝ⇩+" shows "a⋅c¯ \<lsq> b⋅c¯"
using assms field_cntxts_ok field1.OrdField_ZF_2_L3
by simp
text‹We can multiply a strict inequality by a positive number or its inverse.
›
lemma (in real1) Real_ZF_1_3_L7:
assumes "a\<ls>b" and "c∈ℝ⇩+" shows
"a⋅c \<ls> b⋅c"
"c⋅a \<ls> c⋅b"
"a⋅c¯ \<ls> b⋅c¯"
using assms field_cntxts_ok field1.OrdField_ZF_2_L4
by auto
text‹An identity with three real numbers, inverse and cancelling.›
lemma (in real1) Real_ZF_1_3_L8: assumes"a∈ℝ" "b∈ℝ" "b≠𝟬" "c∈ℝ"
shows "a⋅b⋅(c⋅b¯) = a⋅c"
using assms field_cntxts_ok field0.Field_ZF_2_L6
by simp
subsection‹Completeness›
text‹This goal of this section is to show that the order on real numbers
is complete, that is every subset of reals that is bounded above
has a smallest upper bound.›
text‹If $m$ is an integer, then ‹m⇧R› is a real number.
Recall that in ‹real1› context ‹m⇧R› denotes the class
of the slope $n\mapsto m\cdot n$.›
lemma (in real1) real_int_is_real: assumes "m ∈ int"
shows "m⇧R ∈ ℝ"
using assms int1.Int_ZF_2_5_L1 Real_ZF_1_1_L3 by simp
text‹The negative of the real embedding of an integer is the embedding
of the negative of the integer.›
lemma (in real1) Real_ZF_1_4_L1: assumes "m ∈ int"
shows "(\<zm>m)⇧R = \<rm>(m⇧R)"
using assms int1.Int_ZF_2_5_L3 int1.Int_ZF_2_5_L1 Real_ZF_1_1_L4A
by simp
text‹The embedding of sum of integers is the sum of embeddings.›
lemma (in real1) Real_ZF_1_4_L1A: assumes "m ∈ int" "k ∈ int"
shows "m⇧R \<ra> k⇧R = ((m\<za>k)⇧R)"
using assms int1.Int_ZF_2_5_L1 SlopeOp1_def int1.Int_ZF_2_5_L3A
Real_ZF_1_1_L4 by simp
text‹The embedding of a difference of integers is the difference
of embeddings.›
lemma (in real1) Real_ZF_1_4_L1B: assumes A1: "m ∈ int" "k ∈ int"
shows "m⇧R \<rs> k⇧R = (m\<zs>k)⇧R"
proof -
from A1 have "(\<zm>k) ∈ int" using int0.Int_ZF_1_1_L4
by simp
with A1 have "(m\<zs>k)⇧R = m⇧R \<ra> (\<zm>k)⇧R"
using Real_ZF_1_4_L1A by simp
with A1 show "m⇧R \<rs> k⇧R = (m\<zs>k)⇧R"
using Real_ZF_1_4_L1 by simp
qed
text‹The embedding of the product of integers is the product of embeddings.›
lemma (in real1) Real_ZF_1_4_L1C: assumes "m ∈ int" "k ∈ int"
shows "m⇧R ⋅ k⇧R = (m\<zmu>k)⇧R"
using assms int1.Int_ZF_2_5_L1 SlopeOp2_def int1.Int_ZF_2_5_L3B
Real_ZF_1_1_L4 by simp
text‹For any real numbers there is an integer whose real version is
greater or equal.›
lemma (in real1) Real_ZF_1_4_L2: assumes A1: "a∈ℝ"
shows "∃m∈int. a \<lsq> m⇧R"
proof -
from A1 obtain f where I: "f∈𝒮" and II: "a = [f]"
using Real_ZF_1_1_L3A by auto
then have "∃m∈int. ∃g∈𝒮.
{⟨n,m\<zmu>n⟩ . n ∈ int} ∼ g ∧ (f∼g ∨ (g \<fp> (\<fm>f)) ∈ 𝒮⇩+)"
using int1.Int_ZF_2_5_L2 Slopes_def SlopeOp1_def
BoundedIntMaps_def SlopeEquivalenceRel_def
PositiveIntegers_def PositiveSlopes_def
by simp
then obtain m g where III: "m∈int" and IV: "g∈𝒮" and
"{⟨n,m\<zmu>n⟩ . n ∈ int} ∼ g ∧ (f∼g ∨ (g \<fp> (\<fm>f)) ∈ 𝒮⇩+)"
by auto
then have "m⇧R = [g]" and "f ∼ g ∨ (g \<fp> (\<fm>f)) ∈ 𝒮⇩+"
using Real_ZF_1_1_L5A by auto
with I II IV have "a \<lsq> m⇧R" using Real_ZF_1_2_L12
by simp
with III show "∃m∈int. a \<lsq> m⇧R" by auto
qed
text‹For any real numbers there is an integer whose real version (embedding)
is less or equal.›
lemma (in real1) Real_ZF_1_4_L3: assumes A1: "a∈ℝ"
shows "{m ∈ int. m⇧R \<lsq> a} ≠ 0"
proof -
from A1 have "(\<rm>a) ∈ ℝ" using Real_ZF_1_1_L8
by simp
then obtain m where I: "m∈int" and II: "(\<rm>a) \<lsq> m⇧R"
using Real_ZF_1_4_L2 by auto
let ?k = "GroupInv(int,IntegerAddition)`(m)"
from A1 I II have "?k ∈ int" and "?k⇧R \<lsq> a"
using Real_ZF_1_2_L13 Real_ZF_1_4_L1 int0.Int_ZF_1_1_L4
by auto
then show ?thesis by auto
qed
text‹Embeddings of two integers are equal only if the integers are equal.›
lemma (in real1) Real_ZF_1_4_L4:
assumes A1: "m ∈ int" "k ∈ int" and A2: "m⇧R = k⇧R"
shows "m=k"
proof -
let ?r = "{⟨n, IntegerMultiplication ` ⟨m, n⟩⟩ . n ∈ int}"
let ?s = "{⟨n, IntegerMultiplication ` ⟨k, n⟩⟩ . n ∈ int}"
from A1 A2 have "?r ∼ ?s"
using int1.Int_ZF_2_5_L1 AlmostHoms_def Real_ZF_1_1_L5
by simp
with A1 have
"m ∈ int" "k ∈ int"
"⟨?r,?s⟩ ∈ QuotientGroupRel(AlmostHoms(int, IntegerAddition),
AlHomOp1(int, IntegerAddition),FinRangeFunctions(int, int))"
using SlopeEquivalenceRel_def Slopes_def SlopeOp1_def
BoundedIntMaps_def by auto
then show "m=k" by (rule int1.Int_ZF_2_5_L6)
qed
text‹The embedding of integers preserves the order.›
lemma (in real1) Real_ZF_1_4_L5: assumes A1: "m\<zlq>k"
shows "m⇧R \<lsq> k⇧R"
proof -
let ?r = "{⟨n, m\<zmu>n⟩ . n ∈ int}"
let ?s = "{⟨n, k\<zmu>n⟩ . n ∈ int}"
from A1 have "?r ∈ 𝒮" "?s ∈ 𝒮"
using int0.Int_ZF_2_L1A int1.Int_ZF_2_5_L1 by auto
moreover from A1 have "?r ∼ ?s ∨ ?s \<fp> (\<fm>?r) ∈ 𝒮⇩+"
using Slopes_def SlopeOp1_def BoundedIntMaps_def SlopeEquivalenceRel_def
PositiveIntegers_def PositiveSlopes_def
int1.Int_ZF_2_5_L4 by simp
ultimately show "m⇧R \<lsq> k⇧R" using Real_ZF_1_2_L12
by simp
qed
text‹The embedding of integers preserves the strict order.›
lemma (in real1) Real_ZF_1_4_L5A: assumes A1: "m\<zlq>k" "m≠k"
shows "m⇧R \<ls> k⇧R"
proof -
from A1 have "m⇧R \<lsq> k⇧R" using Real_ZF_1_4_L5
by simp
moreover
from A1 have T: "m ∈ int" "k ∈ int"
using int0.Int_ZF_2_L1A by auto
with A1 have "m⇧R ≠ k⇧R" using Real_ZF_1_4_L4
by auto
ultimately show "m⇧R \<ls> k⇧R" by simp
qed
text‹For any real number there is a positive integer
whose real version is (strictly) greater.
This is Lemma 14 i) in \cite{Arthan2004}.›
lemma (in real1) Arthan_Lemma14i: assumes A1: "a∈ℝ"
shows "∃n∈ℤ⇩+. a \<ls> n⇧R"
proof -
from A1 obtain m where I: "m∈int" and II: "a \<lsq> m⇧R"
using Real_ZF_1_4_L2 by auto
let ?n = "GreaterOf(IntegerOrder,𝟭⇩Z,m) \<za> 𝟭⇩Z"
from I have T: "?n ∈ℤ⇩+" and "m \<zlq> ?n" "m≠?n"
using int0.Int_ZF_1_5_L7B by auto
then have III: "m⇧R \<ls> ?n⇧R"
using Real_ZF_1_4_L5A by simp
with II have "a \<ls> ?n⇧R" by (rule real_strict_ord_transit)
with T show ?thesis by auto
qed
text‹If one embedding is less or equal than another, then the integers
are also less or equal.›
lemma (in real1) Real_ZF_1_4_L6:
assumes A1: "k ∈ int" "m ∈ int" and A2: "m⇧R \<lsq> k⇧R"
shows "m\<zlq>k"
proof -
{ assume A3: "⟨m,k⟩ ∉ IntegerOrder"
with A1 have "⟨k,m⟩ ∈ IntegerOrder"
by (rule int0.Int_ZF_2_L19)
then have "k⇧R \<lsq> m⇧R" using Real_ZF_1_4_L5
by simp
with A2 have "m⇧R = k⇧R" by (rule real_ord_antisym)
with A1 have "k = m" using Real_ZF_1_4_L4
by auto
moreover from A1 A3 have "k≠m" by (rule int0.Int_ZF_2_L19)
ultimately have False by simp
} then show "m\<zlq>k" by auto
qed
text‹The floor function is well defined and has expected properties.›
lemma (in real1) Real_ZF_1_4_L7: assumes A1: "a∈ℝ"
shows
"IsBoundedAbove({m ∈ int. m⇧R \<lsq> a},IntegerOrder)"
"{m ∈ int. m⇧R \<lsq> a} ≠ 0"
"⌊a⌋ ∈ int"
"⌊a⌋⇧R \<lsq> a"
proof -
let ?A = "{m ∈ int. m⇧R \<lsq> a}"
from A1 obtain K where I: "K∈int" and II: "a \<lsq> (K⇧R)"
using Real_ZF_1_4_L2 by auto
{ fix n assume "n ∈ ?A"
then have III: "n ∈ int" and IV: "n⇧R \<lsq> a"
by auto
from IV II have "(n⇧R) \<lsq> (K⇧R)"
by (rule real_ord_transitive)
with I III have "n\<zlq>K" using Real_ZF_1_4_L6
by simp
} then have "∀n∈?A. ⟨n,K⟩ ∈ IntegerOrder"
by simp
then show "IsBoundedAbove(?A,IntegerOrder)"
by (rule Order_ZF_3_L10)
moreover from A1 show "?A ≠ 0" using Real_ZF_1_4_L3
by simp
ultimately have "Maximum(IntegerOrder,?A) ∈ ?A"
by (rule int0.int_bounded_above_has_max)
then show "⌊a⌋ ∈ int" "⌊a⌋⇧R \<lsq> a" by auto
qed
text‹Every integer whose embedding is less or equal a real number $a$
is less or equal than the floor of $a$.›
lemma (in real1) Real_ZF_1_4_L8:
assumes A1: "m ∈ int" and A2: "m⇧R \<lsq> a"
shows "m \<zlq> ⌊a⌋"
proof -
let ?A = "{m ∈ int. m⇧R \<lsq> a}"
from A2 have "IsBoundedAbove(?A,IntegerOrder)" and "?A≠0"
using Real_ZF_1_2_L15 Real_ZF_1_4_L7 by auto
then have "∀x∈?A. ⟨x,Maximum(IntegerOrder,?A)⟩ ∈ IntegerOrder"
by (rule int0.int_bounded_above_has_max)
with A1 A2 show "m \<zlq> ⌊a⌋" by simp
qed
text‹Integer zero and one embed as real zero and one.›
lemma (in real1) int_0_1_are_real_zero_one:
shows "𝟬⇩Z⇧R = 𝟬" "𝟭⇩Z⇧R = 𝟭"
using int1.Int_ZF_2_5_L7 BoundedIntMaps_def
real_one_cl_identity real_zero_cl_bounded_map
by auto
text‹Integer two embeds as the real two.›
lemma (in real1) int_two_is_real_two: shows "𝟮⇩Z⇧R = 𝟮"
proof -
have "𝟮⇩Z⇧R = 𝟭⇩Z⇧R \<ra> 𝟭⇩Z⇧R"
using int0.int_zero_one_are_int Real_ZF_1_4_L1A
by simp
also have "… = 𝟮" using int_0_1_are_real_zero_one
by simp
finally show "𝟮⇩Z⇧R = 𝟮" by simp
qed
text‹A positive integer embeds as a positive (hence nonnegative) real.›
lemma (in real1) int_pos_is_real_pos: assumes A1: "p∈ℤ⇩+"
shows
"p⇧R ∈ ℝ"
"𝟬 \<lsq> p⇧R"
"p⇧R ∈ ℝ⇩+"
proof -
from A1 have I: "p ∈ int" "𝟬⇩Z \<zlq> p" "𝟬⇩Z ≠ p"
using PositiveSet_def by auto
then have "p⇧R ∈ ℝ" "𝟬⇩Z⇧R \<lsq> p⇧R"
using real_int_is_real Real_ZF_1_4_L5 by auto
then show "p⇧R ∈ ℝ" "𝟬 \<lsq> p⇧R"
using int_0_1_are_real_zero_one by auto
moreover have "𝟬 ≠ p⇧R"
proof -
{ assume "𝟬 = p⇧R"
with I have False using int_0_1_are_real_zero_one
int0.int_zero_one_are_int Real_ZF_1_4_L4 by auto
} then show "𝟬 ≠ p⇧R" by auto
qed
ultimately show "p⇧R ∈ ℝ⇩+" using PositiveSet_def
by simp
qed
text‹The ordered field of reals we are constructing is archimedean, i.e.,
if $x,y$ are its elements with $y$ positive, then there is a positive
integer $M$ such that $x$ is smaller than $M^R y$. This is Lemma 14 ii) in \cite{Arthan2004}.›
lemma (in real1) Arthan_Lemma14ii: assumes A1: "x∈ℝ" "y ∈ ℝ⇩+"
shows "∃M∈ℤ⇩+. x \<ls> M⇧R⋅y"
proof -
from A1 have
"∃C∈ℤ⇩+. x \<ls> C⇧R" and "∃D∈ℤ⇩+. y¯ \<ls> D⇧R"
using Real_ZF_1_3_L1 Arthan_Lemma14i by auto
then obtain C D where
I: "C∈ℤ⇩+" and II: "x \<ls> C⇧R" and
III: "D∈ℤ⇩+" and IV: "y¯ \<ls> D⇧R"
by auto
let ?M = "C\<zmu>D"
from I III have
T: "?M ∈ ℤ⇩+" "C⇧R ∈ ℝ" "D⇧R ∈ ℝ"
using int0.pos_int_closed_mul_unfold PositiveSet_def real_int_is_real
by auto
with A1 I III have "C⇧R⋅(D⇧R⋅y) = ?M⇧R⋅y"
using PositiveSet_def Real_ZF_1_L6A Real_ZF_1_4_L1C
by simp
moreover from A1 I II IV have
"x \<ls> C⇧R⋅(D⇧R⋅y)"
using int_pos_is_real_pos Real_ZF_1_3_L2 Real_ZF_1_2_L25
by auto
ultimately have "x \<ls> ?M⇧R⋅y"
by auto
with T show ?thesis by auto
qed
text‹Taking the floor function preserves the order.›
lemma (in real1) Real_ZF_1_4_L9: assumes A1: "a\<lsq>b"
shows "⌊a⌋ \<zlq> ⌊b⌋"
proof -
from A1 have T: "a∈ℝ" using Real_ZF_1_2_L15
by simp
with A1 have "⌊a⌋⇧R \<lsq> a" and "a\<lsq>b"
using Real_ZF_1_4_L7 by auto
then have "⌊a⌋⇧R \<lsq> b" by (rule real_ord_transitive)
moreover from T have "⌊a⌋ ∈ int" using Real_ZF_1_4_L7
by simp
ultimately show "⌊a⌋ \<zlq> ⌊b⌋" using Real_ZF_1_4_L8
by simp
qed
text‹If $S$ is bounded above and $p$ is a positive intereger, then
$\Gamma(S,p)$ is well defined.›
lemma (in real1) Real_ZF_1_4_L10:
assumes A1: "IsBoundedAbove(S,OrderOnReals)" "S≠0" and A2: "p∈ℤ⇩+"
shows
"IsBoundedAbove({⌊p⇧R⋅x⌋. x∈S},IntegerOrder)"
"Γ(S,p) ∈ {⌊p⇧R⋅x⌋. x∈S}"
"Γ(S,p) ∈ int"
proof -
let ?A = "{⌊p⇧R⋅x⌋. x∈S}"
from A1 obtain X where I: "∀x∈S. x\<lsq>X"
using IsBoundedAbove_def by auto
{ fix m assume "m ∈ ?A"
then obtain x where "x∈S" and II: "m = ⌊p⇧R⋅x⌋"
by auto
with I have "x\<lsq>X" by simp
moreover from A2 have "𝟬 \<lsq> p⇧R" using int_pos_is_real_pos
by simp
ultimately have "p⇧R⋅x \<lsq> p⇧R⋅X" using Real_ZF_1_2_L14
by simp
with II have "m \<zlq> ⌊p⇧R⋅X⌋" using Real_ZF_1_4_L9
by simp
} then have "∀m∈?A. ⟨m,⌊p⇧R⋅X⌋⟩ ∈ IntegerOrder"
by auto
then show II: "IsBoundedAbove(?A,IntegerOrder)"
by (rule Order_ZF_3_L10)
moreover from A1 have III: "?A ≠ 0" by simp
ultimately have "Maximum(IntegerOrder,?A) ∈ ?A"
by (rule int0.int_bounded_above_has_max)
moreover from II III have "Maximum(IntegerOrder,?A) ∈ int"
by (rule int0.int_bounded_above_has_max)
ultimately show "Γ(S,p) ∈ {⌊p⇧R⋅x⌋. x∈S}" and "Γ(S,p) ∈ int"
by auto
qed
text‹If $p$ is a positive integer, then
for all $s\in S$ the floor of $p\cdot x$ is not greater that $\Gamma(S,p)$.›
lemma (in real1) Real_ZF_1_4_L11:
assumes A1: "IsBoundedAbove(S,OrderOnReals)" and A2: "x∈S" and A3: "p∈ℤ⇩+"
shows "⌊p⇧R⋅x⌋ \<zlq> Γ(S,p)"
proof -
let ?A = "{⌊p⇧R⋅x⌋. x∈S}"
from A2 have "S≠0" by auto
with A1 A3 have "IsBoundedAbove(?A,IntegerOrder)" "?A ≠ 0"
using Real_ZF_1_4_L10 by auto
then have "∀x∈?A. ⟨x,Maximum(IntegerOrder,?A)⟩ ∈ IntegerOrder"
by (rule int0.int_bounded_above_has_max)
with A2 show "⌊p⇧R⋅x⌋ \<zlq> Γ(S,p)" by simp
qed
text‹The candidate for supremum is an integer mapping with values
given by $\Gamma$.›
lemma (in real1) Real_ZF_1_4_L12:
assumes A1: "IsBoundedAbove(S,OrderOnReals)" "S≠0" and
A2: "g = {⟨p,Γ(S,p)⟩. p∈ℤ⇩+}"
shows
"g : ℤ⇩+→int"
"∀n∈ℤ⇩+. g`(n) = Γ(S,n)"
proof -
from A1 have "∀n∈ℤ⇩+. Γ(S,n) ∈ int" using Real_ZF_1_4_L10
by simp
with A2 show I: "g : ℤ⇩+→int" using ZF_fun_from_total by simp
{ fix n assume "n∈ℤ⇩+"
with A2 I have "g`(n) = Γ(S,n)" using ZF_fun_from_tot_val
by simp
} then show "∀n∈ℤ⇩+. g`(n) = Γ(S,n)" by simp
qed
text‹Every integer is equal to the floor of its embedding.›
lemma (in real1) Real_ZF_1_4_L14: assumes A1: "m ∈ int"
shows "⌊m⇧R⌋ = m"
proof -
let ?A = "{n ∈ int. n⇧R \<lsq> m⇧R }"
have "antisym(IntegerOrder)" using int0.Int_ZF_2_L4
by simp
moreover from A1 have "m ∈ ?A"
using real_int_is_real real_ord_refl by auto
moreover from A1 have "∀n ∈ ?A. ⟨n,m⟩ ∈ IntegerOrder"
using Real_ZF_1_4_L6 by auto
ultimately show "⌊m⇧R⌋ = m" using Order_ZF_4_L14
by auto
qed
text‹Floor of (real) zero is (integer) zero.›
lemma (in real1) floor_01_is_zero_one: shows
"⌊𝟬⌋ = 𝟬⇩Z" "⌊𝟭⌋ = 𝟭⇩Z"
proof -
have "⌊(𝟬⇩Z)⇧R⌋ = 𝟬⇩Z" and "⌊(𝟭⇩Z)⇧R⌋ = 𝟭⇩Z"
using int0.int_zero_one_are_int Real_ZF_1_4_L14
by auto
then show "⌊𝟬⌋ = 𝟬⇩Z" and "⌊𝟭⌋ = 𝟭⇩Z"
using int_0_1_are_real_zero_one
by auto
qed
text‹Floor of (real) two is (integer) two.›
lemma (in real1) floor_2_is_two: shows "⌊𝟮⌋ = 𝟮⇩Z"
proof -
have "⌊(𝟮⇩Z)⇧R⌋ = 𝟮⇩Z"
using int0.int_two_three_are_int Real_ZF_1_4_L14
by simp
then show "⌊𝟮⌋ = 𝟮⇩Z" using int_two_is_real_two
by simp
qed
text‹Floor of a product of embeddings of integers is equal to the
product of integers.›
lemma (in real1) Real_ZF_1_4_L14A: assumes A1: "m ∈ int" "k ∈ int"
shows "⌊m⇧R⋅k⇧R⌋ = m\<zmu>k"
proof -
from A1 have T: "m\<zmu>k ∈ int"
using int0.Int_ZF_1_1_L5 by simp
from A1 have "⌊m⇧R⋅k⇧R⌋ = ⌊(m\<zmu>k)⇧R⌋" using Real_ZF_1_4_L1C
by simp
with T show "⌊m⇧R⋅k⇧R⌋ = m\<zmu>k" using Real_ZF_1_4_L14
by simp
qed
text‹Floor of the sum of a number and the embedding of an
integer is the floor of the number plus the integer.›
lemma (in real1) Real_ZF_1_4_L15: assumes A1: "x∈ℝ" and A2: "p ∈ int"
shows "⌊x \<ra> p⇧R⌋ = ⌊x⌋ \<za> p"
proof -
let ?A = "{n ∈ int. n⇧R \<lsq> x \<ra> p⇧R}"
have "antisym(IntegerOrder)" using int0.Int_ZF_2_L4
by simp
moreover have "⌊x⌋ \<za> p ∈ ?A"
proof -
from A1 A2 have "⌊x⌋⇧R \<lsq> x" and "p⇧R ∈ ℝ"
using Real_ZF_1_4_L7 real_int_is_real by auto
then have "⌊x⌋⇧R \<ra> p⇧R \<lsq> x \<ra> p⇧R"
using add_num_to_ineq by simp
moreover from A1 A2 have "(⌊x⌋ \<za> p)⇧R = ⌊x⌋⇧R \<ra> p⇧R"
using Real_ZF_1_4_L7 Real_ZF_1_4_L1A by simp
ultimately have "(⌊x⌋ \<za> p)⇧R \<lsq> x \<ra> p⇧R"
by simp
moreover from A1 A2 have "⌊x⌋ \<za> p ∈ int"
using Real_ZF_1_4_L7 int0.Int_ZF_1_1_L5 by simp
ultimately show "⌊x⌋ \<za> p ∈ ?A" by auto
qed
moreover have "∀n∈?A. n \<zlq> ⌊x⌋ \<za> p"
proof
fix n assume "n∈?A"
then have I: "n ∈ int" and "n⇧R \<lsq> x \<ra> p⇧R"
by auto
with A1 A2 have "n⇧R \<rs> p⇧R \<lsq> x"
using real_int_is_real Real_ZF_1_2_L19
by simp
with A2 I have "⌊(n\<zs>p)⇧R⌋ \<zlq> ⌊x⌋"
using Real_ZF_1_4_L1B Real_ZF_1_4_L9
by simp
moreover
from A2 I have "n\<zs>p ∈ int"
using int0.Int_ZF_1_1_L5 by simp
then have "⌊(n\<zs>p)⇧R⌋ = n\<zs>p"
using Real_ZF_1_4_L14 by simp
ultimately have "n\<zs>p \<zlq> ⌊x⌋"
by simp
with A2 I show "n \<zlq> ⌊x⌋ \<za> p"
using int0.Int_ZF_2_L9C by simp
qed
ultimately show "⌊x \<ra> p⇧R⌋ = ⌊x⌋ \<za> p"
using Order_ZF_4_L14 by auto
qed
text‹Floor of the difference of a number and the embedding of an
integer is the floor of the number minus the integer.›
lemma (in real1) Real_ZF_1_4_L16: assumes A1: "x∈ℝ" and A2: "p ∈ int"
shows "⌊x \<rs> p⇧R⌋ = ⌊x⌋ \<zs> p"
proof -
from A2 have "⌊x \<rs> p⇧R⌋ = ⌊x \<ra> (\<zm>p)⇧R⌋"
using Real_ZF_1_4_L1 by simp
with A1 A2 show "⌊x \<rs> p⇧R⌋ = ⌊x⌋ \<zs> p"
using int0.Int_ZF_1_1_L4 Real_ZF_1_4_L15 by simp
qed
text‹The floor of sum of embeddings is the sum of the integers.›
lemma (in real1) Real_ZF_1_4_L17: assumes "m ∈ int" "n ∈ int"
shows "⌊(m⇧R) \<ra> n⇧R⌋ = m \<za> n"
using assms real_int_is_real Real_ZF_1_4_L15 Real_ZF_1_4_L14
by simp
text‹A lemma about adding one to floor.›
lemma (in real1) Real_ZF_1_4_L17A: assumes A1: "a∈ℝ"
shows "𝟭 \<ra> ⌊a⌋⇧R = (𝟭⇩Z \<za> ⌊a⌋)⇧R"
proof -
have "𝟭 \<ra> ⌊a⌋⇧R = 𝟭⇩Z⇧R \<ra> ⌊a⌋⇧R"
using int_0_1_are_real_zero_one by simp
with A1 show "𝟭 \<ra> ⌊a⌋⇧R = (𝟭⇩Z \<za> ⌊a⌋)⇧R"
using int0.int_zero_one_are_int Real_ZF_1_4_L7 Real_ZF_1_4_L1A
by simp
qed
text‹The difference between the a number and the embedding of its floor
is (strictly) less than one.›
lemma (in real1) Real_ZF_1_4_L17B: assumes A1: "a∈ℝ"
shows
"a \<rs> ⌊a⌋⇧R \<ls> 𝟭"
"a \<ls> (𝟭⇩Z \<za> ⌊a⌋)⇧R"
proof -
from A1 have T1: "⌊a⌋ ∈ int" "⌊a⌋⇧R ∈ ℝ" and
T2: "𝟭 ∈ ℝ" "a \<rs> ⌊a⌋⇧R ∈ ℝ"
using Real_ZF_1_4_L7 real_int_is_real Real_ZF_1_L6 Real_ZF_1_L4
by auto
{ assume "𝟭 \<lsq> a \<rs> ⌊a⌋⇧R"
with A1 T1 have "⌊𝟭⇩Z⇧R \<ra> ⌊a⌋⇧R⌋ \<zlq> ⌊a⌋"
using Real_ZF_1_2_L21 Real_ZF_1_4_L9 int_0_1_are_real_zero_one
by simp
with T1 have False
using int0.int_zero_one_are_int Real_ZF_1_4_L17
int0.Int_ZF_1_2_L3AA by simp
} then have I: "¬(𝟭 \<lsq> a \<rs> ⌊a⌋⇧R)" by auto
with T2 show II: "a \<rs> ⌊a⌋⇧R \<ls> 𝟭"
by (rule Real_ZF_1_2_L20)
with A1 T1 I II have
"a \<ls> 𝟭 \<ra> ⌊a⌋⇧R"
using Real_ZF_1_2_L26 by simp
with A1 show "a \<ls> (𝟭⇩Z \<za> ⌊a⌋)⇧R"
using Real_ZF_1_4_L17A by simp
qed
text‹The next lemma corresponds to Lemma 14 iii) in \cite{Arthan2004}.
It says that we can find a rational number between any two different
real numbers.›
lemma (in real1) Arthan_Lemma14iii: assumes A1: "x\<ls>y"
shows "∃M∈int. ∃N∈ℤ⇩+. x⋅N⇧R \<ls> M⇧R ∧ M⇧R \<ls> y⋅N⇧R"
proof -
from A1 have "(y\<rs>x)¯ ∈ ℝ⇩+" using Real_ZF_1_3_L3
by simp
then have
"∃N∈ℤ⇩+. (y\<rs>x)¯ \<ls> N⇧R"
using Arthan_Lemma14i PositiveSet_def by simp
then obtain N where I: "N∈ℤ⇩+" and II: "(y\<rs>x)¯ \<ls> N⇧R"
by auto
let ?M = "𝟭⇩Z \<za> ⌊x⋅N⇧R⌋"
from A1 I have
T1: "x∈ℝ" "N⇧R ∈ ℝ" "N⇧R ∈ ℝ⇩+" "x⋅N⇧R ∈ ℝ"
using Real_ZF_1_2_L15 PositiveSet_def real_int_is_real
Real_ZF_1_L6 int_pos_is_real_pos by auto
then have T2: "?M ∈ int" using
int0.int_zero_one_are_int Real_ZF_1_4_L7 int0.Int_ZF_1_1_L5
by simp
from T1 have III: "x⋅N⇧R \<ls> ?M⇧R"
using Real_ZF_1_4_L17B by simp
from T1 have "(𝟭 \<ra> ⌊x⋅N⇧R⌋⇧R) \<lsq> (𝟭 \<ra> x⋅N⇧R)"
using Real_ZF_1_4_L7 Real_ZF_1_L4 real_ord_transl_inv
by simp
with T1 have "?M⇧R \<lsq> (𝟭 \<ra> x⋅N⇧R)"
using Real_ZF_1_4_L17A by simp
moreover from A1 II have "(𝟭 \<ra> x⋅N⇧R) \<ls> y⋅N⇧R"
using Real_ZF_1_3_L5 by simp
ultimately have "?M⇧R \<ls> y⋅N⇧R"
by (rule real_strict_ord_transit)
with I T2 III show ?thesis by auto
qed
text‹Some estimates for the homomorphism difference of the floor
function.›
lemma (in real1) Real_ZF_1_4_L18: assumes A1: "x∈ℝ" "y∈ℝ"
shows
"abs(⌊x\<ra>y⌋ \<zs> ⌊x⌋ \<zs> ⌊y⌋) \<zlq> 𝟮⇩Z"
proof -
from A1 have T:
"⌊x⌋⇧R ∈ ℝ" "⌊y⌋⇧R ∈ ℝ"
"x\<ra>y \<rs> (⌊x⌋⇧R) ∈ ℝ"
using Real_ZF_1_4_L7 real_int_is_real Real_ZF_1_L6
by auto
from A1 have
"𝟬 \<lsq> x \<rs> (⌊x⌋⇧R) \<ra> (y \<rs> (⌊y⌋⇧R))"
"x \<rs> (⌊x⌋⇧R) \<ra> (y \<rs> (⌊y⌋⇧R)) \<lsq> 𝟮"
using Real_ZF_1_4_L7 Real_ZF_1_2_L16 Real_ZF_1_2_L17
Real_ZF_1_4_L17B Real_ZF_1_2_L18 by auto
moreover from A1 T have
"x \<rs> (⌊x⌋⇧R) \<ra> (y \<rs> (⌊y⌋⇧R)) = x\<ra>y \<rs> (⌊x⌋⇧R) \<rs> (⌊y⌋⇧R)"
using Real_ZF_1_L7A by simp
ultimately have
"𝟬 \<lsq> x\<ra>y \<rs> (⌊x⌋⇧R) \<rs> (⌊y⌋⇧R)"
"x\<ra>y \<rs> (⌊x⌋⇧R) \<rs> (⌊y⌋⇧R) \<lsq> 𝟮"
by auto
then have
"⌊𝟬⌋ \<zlq> ⌊x\<ra>y \<rs> (⌊x⌋⇧R) \<rs> (⌊y⌋⇧R)⌋"
"⌊x\<ra>y \<rs> (⌊x⌋⇧R) \<rs> (⌊y⌋⇧R)⌋ \<zlq> ⌊𝟮⌋"
using Real_ZF_1_4_L9 by auto
then have
"𝟬⇩Z \<zlq> ⌊x\<ra>y \<rs> (⌊x⌋⇧R) \<rs> (⌊y⌋⇧R)⌋"
"⌊x\<ra>y \<rs> (⌊x⌋⇧R) \<rs> (⌊y⌋⇧R)⌋ \<zlq> 𝟮⇩Z"
using floor_01_is_zero_one floor_2_is_two by auto
moreover from A1 have
"⌊x\<ra>y \<rs> (⌊x⌋⇧R) \<rs> (⌊y⌋⇧R)⌋ = ⌊x\<ra>y⌋ \<zs> ⌊x⌋ \<zs> ⌊y⌋"
using Real_ZF_1_L6 Real_ZF_1_4_L7 real_int_is_real Real_ZF_1_4_L16
by simp
ultimately have
"𝟬⇩Z \<zlq> ⌊x\<ra>y⌋ \<zs> ⌊x⌋ \<zs> ⌊y⌋"
"⌊x\<ra>y⌋ \<zs> ⌊x⌋ \<zs> ⌊y⌋ \<zlq> 𝟮⇩Z"
by auto
then show "abs(⌊x\<ra>y⌋ \<zs> ⌊x⌋ \<zs> ⌊y⌋) \<zlq> 𝟮⇩Z"
using int0.Int_ZF_2_L16 by simp
qed
text‹Suppose $S\neq \emptyset$ is bounded above and
$\Gamma(S,m) =\lfloor m^R\cdot x\rfloor$ for some positive integer $m$
and $x\in S$. Then if $y\in S, x\leq y$ we also have
$\Gamma(S,m) =\lfloor m^R\cdot y\rfloor$.›
lemma (in real1) Real_ZF_1_4_L20:
assumes A1: "IsBoundedAbove(S,OrderOnReals)" "S≠0" and
A2: "n∈ℤ⇩+" "x∈S" and
A3: "Γ(S,n) = ⌊n⇧R⋅x⌋" and
A4: "y∈S" "x\<lsq>y"
shows "Γ(S,n) = ⌊n⇧R⋅y⌋"
proof -
from A2 A4 have "⌊n⇧R⋅x⌋ \<zlq> ⌊(n⇧R)⋅y⌋"
using int_pos_is_real_pos Real_ZF_1_2_L14 Real_ZF_1_4_L9
by simp
with A3 have "⟨Γ(S,n),⌊(n⇧R)⋅y⌋⟩ ∈ IntegerOrder"
by simp
moreover from A1 A2 A4 have "⟨⌊n⇧R⋅y⌋,Γ(S,n)⟩ ∈ IntegerOrder"
using Real_ZF_1_4_L11 by simp
ultimately show "Γ(S,n) = ⌊n⇧R⋅y⌋"
by (rule int0.Int_ZF_2_L3)
qed
text‹The homomorphism difference of $n\mapsto \Gamma(S,n)$ is bounded
by $2$ on positive integers.›
lemma (in real1) Real_ZF_1_4_L21:
assumes A1: "IsBoundedAbove(S,OrderOnReals)" "S≠0" and
A2: "m∈ℤ⇩+" "n∈ℤ⇩+"
shows "abs(Γ(S,m\<za>n) \<zs> Γ(S,m) \<zs> Γ(S,n)) \<zlq> 𝟮⇩Z"
proof -
from A2 have T: "m\<za>n ∈ ℤ⇩+" using int0.pos_int_closed_add_unfolded
by simp
with A1 A2 have
"Γ(S,m) ∈ {⌊m⇧R⋅x⌋. x∈S}" and
"Γ(S,n) ∈ {⌊n⇧R⋅x⌋. x∈S}" and
"Γ(S,m\<za>n) ∈ {⌊(m\<za>n)⇧R⋅x⌋. x∈S}"
using Real_ZF_1_4_L10 by auto
then obtain a b c where I: "a∈S" "b∈S" "c∈S"
and II:
"Γ(S,m) = ⌊m⇧R⋅a⌋"
"Γ(S,n) = ⌊n⇧R⋅b⌋"
"Γ(S,m\<za>n) = ⌊(m\<za>n)⇧R⋅c⌋"
by auto
let ?d = "Maximum(OrderOnReals,{a,b,c})"
from A1 I have "a∈ℝ" "b∈ℝ" "c∈ℝ"
using Real_ZF_1_2_L23 by auto
then have IV:
"?d ∈ {a,b,c}"
"?d ∈ ℝ"
"a \<lsq> ?d"
"b \<lsq> ?d"
"c \<lsq> ?d"
using Real_ZF_1_2_L24 by auto
with I have V: "?d ∈ S" by auto
from A1 T I II IV V have "Γ(S,m\<za>n) = ⌊(m\<za>n)⇧R⋅?d⌋"
using Real_ZF_1_4_L20 by blast
also from A2 have "… = ⌊((m⇧R)\<ra>(n⇧R))⋅?d⌋"
using Real_ZF_1_4_L1A PositiveSet_def by simp
also from A2 IV have "… = ⌊(m⇧R)⋅?d \<ra> (n⇧R)⋅?d⌋"
using PositiveSet_def real_int_is_real Real_ZF_1_L7
by simp
finally have "Γ(S,m\<za>n) = ⌊(m⇧R)⋅?d \<ra> (n⇧R)⋅?d⌋"
by simp
moreover from A1 A2 I II IV V have "Γ(S,m) = ⌊m⇧R⋅?d⌋"
using Real_ZF_1_4_L20 by blast
moreover from A1 A2 I II IV V have "Γ(S,n) = ⌊n⇧R⋅?d⌋"
using Real_ZF_1_4_L20 by blast
moreover from A1 T I II IV V have "Γ(S,m\<za>n) = ⌊(m\<za>n)⇧R⋅?d⌋"
using Real_ZF_1_4_L20 by blast
ultimately have "abs(Γ(S,m\<za>n) \<zs> Γ(S,m) \<zs> Γ(S,n)) =
abs(⌊(m⇧R)⋅?d \<ra> (n⇧R)⋅?d⌋ \<zs> ⌊m⇧R⋅?d⌋ \<zs> ⌊n⇧R⋅?d⌋)"
by simp
with A2 IV show
"abs(Γ(S,m\<za>n) \<zs> Γ(S,m) \<zs> Γ(S,n)) \<zlq> 𝟮⇩Z"
using PositiveSet_def real_int_is_real Real_ZF_1_L6
Real_ZF_1_4_L18 by simp
qed
text‹The next lemma provides sufficient condition for an odd function
to be an almost homomorphism.
It says for odd functions we only need to check that
the homomorphism difference
(denoted ‹δ› in the ‹real1› context) is bounded on positive
integers. This is really proven in ‹Int_ZF_2.thy›, but we
restate it here for convenience. Recall from ‹Group_ZF_3.thy› that
‹OddExtension› of a function defined on the set of positive elements
(of an ordered group) is the only odd function that is equal to the given
one when restricted to positive elements.›
lemma (in real1) Real_ZF_1_4_L21A:
assumes A1: "f:ℤ⇩+→int" "∀a∈ℤ⇩+. ∀b∈ℤ⇩+. abs(δ(f,a,b)) \<zlq> L"
shows "OddExtension(int,IntegerAddition,IntegerOrder,f) ∈ 𝒮"
using A1 int1.Int_ZF_2_1_L24 by auto
text‹The candidate for (a representant of) the supremum of a
nonempty bounded above set is a slope.›
lemma (in real1) Real_ZF_1_4_L22:
assumes A1: "IsBoundedAbove(S,OrderOnReals)" "S≠0" and
A2: "g = {⟨p,Γ(S,p)⟩. p∈ℤ⇩+}"
shows "OddExtension(int,IntegerAddition,IntegerOrder,g) ∈ 𝒮"
proof -
from A1 A2 have "g: ℤ⇩+→int" by (rule Real_ZF_1_4_L12)
moreover have "∀m∈ℤ⇩+. ∀n∈ℤ⇩+. abs(δ(g,m,n)) \<zlq> 𝟮⇩Z"
proof -
{ fix m n assume A3: "m∈ℤ⇩+" "n∈ℤ⇩+"
then have "m\<za>n ∈ ℤ⇩+" "m∈ℤ⇩+" "n∈ℤ⇩+"
using int0.pos_int_closed_add_unfolded
by auto
moreover from A1 A2 have "∀n∈ℤ⇩+. g`(n) = Γ(S,n)"
by (rule Real_ZF_1_4_L12)
ultimately have "δ(g,m,n) = Γ(S,m\<za>n) \<zs> Γ(S,m) \<zs> Γ(S,n)"
by simp
moreover from A1 A3 have
"abs(Γ(S,m\<za>n) \<zs> Γ(S,m) \<zs> Γ(S,n)) \<zlq> 𝟮⇩Z"
by (rule Real_ZF_1_4_L21)
ultimately have "abs(δ(g,m,n)) \<zlq> 𝟮⇩Z"
by simp
} then show "∀m∈ℤ⇩+. ∀n∈ℤ⇩+. abs(δ(g,m,n)) \<zlq> 𝟮⇩Z"
by simp
qed
ultimately show ?thesis by (rule Real_ZF_1_4_L21A)
qed
text‹A technical lemma used in the proof that all elements
of $S$ are less or equal than the candidate for supremum of $S$.›
lemma (in real1) Real_ZF_1_4_L23:
assumes A1: "f ∈ 𝒮" and A2: "N ∈ int" "M ∈ int" and
A3: "∀n∈ℤ⇩+. M\<zmu>n \<zlq> f`(N\<zmu>n)"
shows "M⇧R \<lsq> [f]⋅(N⇧R)"
proof -
let ?M⇩S = "{⟨n, M\<zmu>n⟩ . n ∈ int}"
let ?N⇩S = "{⟨n, N\<zmu>n⟩ . n ∈ int}"
from A1 A2 have T: "?M⇩S ∈ 𝒮" "?N⇩S ∈ 𝒮" "f∘?N⇩S ∈ 𝒮"
using int1.Int_ZF_2_5_L1 int1.Int_ZF_2_1_L11 SlopeOp2_def
by auto
moreover from A1 A2 A3 have "?M⇩S ∼ f∘?N⇩S ∨ f∘?N⇩S \<fp> (\<fm>?M⇩S) ∈ 𝒮⇩+"
using int1.Int_ZF_2_5_L8 SlopeOp2_def SlopeOp1_def Slopes_def
BoundedIntMaps_def SlopeEquivalenceRel_def PositiveIntegers_def
PositiveSlopes_def by simp
ultimately have "[?M⇩S] \<lsq> [f∘?N⇩S]" using Real_ZF_1_2_L12
by simp
with A1 T show "M⇧R \<lsq> [f]⋅(N⇧R)" using Real_ZF_1_1_L4
by simp
qed
text‹A technical lemma aimed used in the proof the candidate for supremum of
$S$ is less or equal than any upper bound for $S$.›
lemma (in real1) Real_ZF_1_4_L23A:
assumes A1: "f ∈ 𝒮" and A2: "N ∈ int" "M ∈ int" and
A3: "∀n∈ℤ⇩+. f`(N\<zmu>n) \<zlq> M\<zmu>n "
shows "[f]⋅(N⇧R) \<lsq> M⇧R"
proof -
let ?M⇩S = "{⟨n, M\<zmu>n⟩ . n ∈ int}"
let ?N⇩S = "{⟨n, N\<zmu>n⟩ . n ∈ int}"
from A1 A2 have T: "?M⇩S ∈ 𝒮" "?N⇩S ∈ 𝒮" "f∘?N⇩S ∈ 𝒮"
using int1.Int_ZF_2_5_L1 int1.Int_ZF_2_1_L11 SlopeOp2_def
by auto
moreover from A1 A2 A3 have
"f∘?N⇩S ∼ ?M⇩S ∨ ?M⇩S \<fp> (\<fm>(f∘?N⇩S)) ∈ 𝒮⇩+"
using int1.Int_ZF_2_5_L9 SlopeOp2_def SlopeOp1_def Slopes_def
BoundedIntMaps_def SlopeEquivalenceRel_def PositiveIntegers_def
PositiveSlopes_def by simp
ultimately have "[f∘?N⇩S] \<lsq> [?M⇩S]" using Real_ZF_1_2_L12
by simp
with A1 T show " [f]⋅(N⇧R)\<lsq> M⇧R" using Real_ZF_1_1_L4
by simp
qed
text‹The essential condition to claim that the candidate for supremum
of $S$ is greater or equal than all elements of $S$.›
lemma (in real1) Real_ZF_1_4_L24:
assumes A1: "IsBoundedAbove(S,OrderOnReals)" and
A2: "x\<ls>y" "y∈S" and
A4: "N ∈ ℤ⇩+" "M ∈ int" and
A5: "M⇧R \<ls> y⋅N⇧R" and A6: "p ∈ ℤ⇩+"
shows "p\<zmu>M \<zlq> Γ(S,p\<zmu>N)"
proof -
from A2 A4 A6 have T1:
"N⇧R ∈ ℝ⇩+" "y∈ℝ" "p⇧R ∈ ℝ⇩+"
"p\<zmu>N ∈ ℤ⇩+" "(p\<zmu>N)⇧R ∈ ℝ⇩+"
using int_pos_is_real_pos Real_ZF_1_2_L15
int0.pos_int_closed_mul_unfold by auto
with A4 A6 have T2:
"p ∈ int" "p⇧R ∈ ℝ" "N⇧R ∈ ℝ" "N⇧R ≠ 𝟬" "M⇧R ∈ ℝ"
using real_int_is_real PositiveSet_def by auto
from T1 A5 have "⌊(p\<zmu>N)⇧R⋅(M⇧R⋅(N⇧R)¯)⌋ \<zlq> ⌊(p\<zmu>N)⇧R⋅y⌋"
using Real_ZF_1_3_L4A Real_ZF_1_3_L7 Real_ZF_1_4_L9
by simp
moreover from A1 A2 T1 have "⌊(p\<zmu>N)⇧R⋅y⌋ \<zlq> Γ(S,p\<zmu>N)"
using Real_ZF_1_4_L11 by simp
ultimately have I: "⌊(p\<zmu>N)⇧R⋅(M⇧R⋅(N⇧R)¯)⌋ \<zlq> Γ(S,p\<zmu>N)"
by (rule int_order_transitive)
from A4 A6 have "(p\<zmu>N)⇧R⋅(M⇧R⋅(N⇧R)¯) = p⇧R⋅N⇧R⋅(M⇧R⋅(N⇧R)¯)"
using PositiveSet_def Real_ZF_1_4_L1C by simp
with A4 T2 have "⌊(p\<zmu>N)⇧R⋅(M⇧R⋅(N⇧R)¯)⌋ = p\<zmu>M"
using Real_ZF_1_3_L8 Real_ZF_1_4_L14A by simp
with I show "p\<zmu>M \<zlq> Γ(S,p\<zmu>N)" by simp
qed
text‹An obvious fact about odd extension
of a function $p\mapsto \Gamma(s,p)$ that is used a couple of times
in proofs.›
lemma (in real1) Real_ZF_1_4_L24A:
assumes A1: "IsBoundedAbove(S,OrderOnReals)" "S≠0" and A2: "p ∈ ℤ⇩+"
and A3:
"h = OddExtension(int,IntegerAddition,IntegerOrder,{⟨p,Γ(S,p)⟩. p∈ℤ⇩+})"
shows "h`(p) = Γ(S,p)"
proof -
let ?g = "{⟨p,Γ(S,p)⟩. p∈ℤ⇩+}"
from A1 have I: "?g : ℤ⇩+→int" using Real_ZF_1_4_L12
by blast
with A2 A3 show "h`(p) = Γ(S,p)"
using int0.Int_ZF_1_5_L11 ZF_fun_from_tot_val
by simp
qed
text‹The candidate for the supremum of $S$ is not smaller than
any element of $S$.›
lemma (in real1) Real_ZF_1_4_L25:
assumes A1: "IsBoundedAbove(S,OrderOnReals)" and
A2: "¬HasAmaximum(OrderOnReals,S)" and
A3: "x∈S" and A4:
"h = OddExtension(int,IntegerAddition,IntegerOrder,{⟨p,Γ(S,p)⟩. p∈ℤ⇩+})"
shows "x \<lsq> [h]"
proof -
from A1 A2 A3 have
"S ⊆ ℝ" "¬HasAmaximum(OrderOnReals,S)" "x∈S"
using Real_ZF_1_2_L23 by auto
then have "∃y∈S. x\<ls>y" by (rule Real_ZF_1_2_L27)
then obtain y where I: "y∈S" and II: "x\<ls>y"
by auto
from II have
"∃M∈int. ∃N∈ℤ⇩+. x⋅N⇧R \<ls> M⇧R ∧ M⇧R \<ls> y⋅N⇧R"
using Arthan_Lemma14iii by simp
then obtain M N where III: "M ∈ int" "N∈ℤ⇩+" and
IV: "x⋅N⇧R \<ls> M⇧R" "M⇧R \<ls> y⋅N⇧R"
by auto
from II III IV have V: "x \<lsq> M⇧R⋅(N⇧R)¯"
using int_pos_is_real_pos Real_ZF_1_2_L15 Real_ZF_1_3_L4
by auto
from A3 have VI: "S≠0" by auto
with A1 A4 have T1: "h ∈ 𝒮" using Real_ZF_1_4_L22
by simp
moreover from III have "N ∈ int" "M ∈ int"
using PositiveSet_def by auto
moreover have "∀n∈ℤ⇩+. M\<zmu>n \<zlq> h`(N\<zmu>n)"
proof
let ?g = "{⟨p,Γ(S,p)⟩. p∈ℤ⇩+}"
fix n assume A5: "n∈ℤ⇩+"
with III have T2: "N\<zmu>n ∈ ℤ⇩+"
using int0.pos_int_closed_mul_unfold by simp
from III A5 have
"N\<zmu>n = n\<zmu>N" and "n\<zmu>M = M\<zmu>n"
using PositiveSet_def int0.Int_ZF_1_1_L5 by auto
moreover
from A1 I II III IV A5 have
"IsBoundedAbove(S,OrderOnReals)"
"x\<ls>y" "y∈S"
"N ∈ ℤ⇩+" "M ∈ int"
"M⇧R \<ls> y⋅N⇧R" "n ∈ ℤ⇩+"
by auto
then have "n\<zmu>M \<zlq> Γ(S,n\<zmu>N)" by (rule Real_ZF_1_4_L24)
moreover from A1 A4 VI T2 have "h`(N\<zmu>n) = Γ(S,N\<zmu>n)"
using Real_ZF_1_4_L24A by simp
ultimately show "M\<zmu>n \<zlq> h`(N\<zmu>n)" by auto
qed
ultimately have "M⇧R \<lsq> [h]⋅N⇧R" using Real_ZF_1_4_L23
by simp
with III T1 have "M⇧R⋅(N⇧R)¯ \<lsq> [h]"
using int_pos_is_real_pos Real_ZF_1_1_L3 Real_ZF_1_3_L4B
by simp
with V show "x \<lsq> [h]" by (rule real_ord_transitive)
qed
text‹The essential condition to claim that the candidate for supremum
of $S$ is less or equal than any upper bound of $S$.›
lemma (in real1) Real_ZF_1_4_L26:
assumes A1: "IsBoundedAbove(S,OrderOnReals)" and
A2: "x\<lsq>y" "x∈S" and
A4: "N ∈ ℤ⇩+" "M ∈ int" and
A5: "y⋅N⇧R \<ls> M⇧R " and A6: "p ∈ ℤ⇩+"
shows "⌊(N\<zmu>p)⇧R⋅x⌋ \<zlq> M\<zmu>p"
proof -
from A2 A4 A6 have T:
"p\<zmu>N ∈ ℤ⇩+" "p ∈ int" "N ∈ int"
"p⇧R ∈ ℝ⇩+" "p⇧R ∈ ℝ" "N⇧R ∈ ℝ" "x ∈ ℝ" "y ∈ ℝ"
using int0.pos_int_closed_mul_unfold PositiveSet_def
real_int_is_real Real_ZF_1_2_L15 int_pos_is_real_pos
by auto
with A2 have "(p\<zmu>N)⇧R⋅x \<lsq> (p\<zmu>N)⇧R⋅y"
using int_pos_is_real_pos Real_ZF_1_2_L14A
by simp
moreover from A4 T have I:
"(p\<zmu>N)⇧R = p⇧R⋅N⇧R"
"(p\<zmu>M)⇧R = p⇧R⋅M⇧R"
using Real_ZF_1_4_L1C by auto
ultimately have "(p\<zmu>N)⇧R⋅x \<lsq> p⇧R⋅N⇧R⋅y"
by simp
moreover
from A5 T I have "p⇧R⋅(y⋅N⇧R) \<ls> (p\<zmu>M)⇧R"
using Real_ZF_1_3_L7 by simp
with T have "p⇧R⋅N⇧R⋅y \<ls> (p\<zmu>M)⇧R" using Real_ZF_1_1_L9
by simp
ultimately have "(p\<zmu>N)⇧R⋅x \<ls> (p\<zmu>M)⇧R"
by (rule real_strict_ord_transit)
then have "⌊(p\<zmu>N)⇧R⋅x⌋ \<zlq> ⌊(p\<zmu>M)⇧R⌋"
using Real_ZF_1_4_L9 by simp
moreover
from A4 T have "p\<zmu>M ∈ int" using int0.Int_ZF_1_1_L5
by simp
then have "⌊(p\<zmu>M)⇧R⌋ = p\<zmu>M" using Real_ZF_1_4_L14
by simp
moreover from A4 A6 have "p\<zmu>N = N\<zmu>p" and "p\<zmu>M = M\<zmu>p"
using PositiveSet_def int0.Int_ZF_1_1_L5 by auto
ultimately show "⌊(N\<zmu>p)⇧R⋅x⌋ \<zlq> M\<zmu>p" by simp
qed
text‹A piece of the proof of the fact that the candidate for the supremum
of $S$ is not greater than any upper bound of $S$, done separately for
clarity (of mind).›
lemma (in real1) Real_ZF_1_4_L27:
assumes "IsBoundedAbove(S,OrderOnReals)" "S≠0" and
"h = OddExtension(int,IntegerAddition,IntegerOrder,{⟨p,Γ(S,p)⟩. p∈ℤ⇩+})"
and "p ∈ ℤ⇩+"
shows "∃x∈S. h`(p) = ⌊p⇧R⋅x⌋"
using assms Real_ZF_1_4_L10 Real_ZF_1_4_L24A by auto
text‹The candidate for the supremum of $S$ is not greater than
any upper bound of $S$.›
lemma (in real1) Real_ZF_1_4_L28:
assumes A1: "IsBoundedAbove(S,OrderOnReals)" "S≠0"
and A2: "∀x∈S. x\<lsq>y" and A3:
"h = OddExtension(int,IntegerAddition,IntegerOrder,{⟨p,Γ(S,p)⟩. p∈ℤ⇩+})"
shows "[h] \<lsq> y"
proof -
from A1 obtain a where "a∈S" by auto
with A1 A2 A3 have T: "y∈ℝ" "h ∈ 𝒮" "[h] ∈ ℝ"
using Real_ZF_1_2_L15 Real_ZF_1_4_L22 Real_ZF_1_1_L3
by auto
{ assume "¬([h] \<lsq> y)"
with T have "y \<ls> [h]" using Real_ZF_1_2_L28
by blast
then have "∃M∈int. ∃N∈ℤ⇩+. y⋅N⇧R \<ls> M⇧R ∧ M⇧R \<ls> [h]⋅N⇧R"
using Arthan_Lemma14iii by simp
then obtain M N where I: "M∈int" "N∈ℤ⇩+" and
II: "y⋅N⇧R \<ls> M⇧R" "M⇧R \<ls> [h]⋅N⇧R"
by auto
from I have III: "N⇧R ∈ ℝ⇩+" using int_pos_is_real_pos
by simp
have "∀p∈ℤ⇩+. h`(N\<zmu>p) \<zlq> M\<zmu>p"
proof
fix p assume A4: "p∈ℤ⇩+"
with A1 A3 I have "∃x∈S. h`(N\<zmu>p) = ⌊(N\<zmu>p)⇧R⋅x⌋"
using int0.pos_int_closed_mul_unfold Real_ZF_1_4_L27
by simp
with A1 A2 I II A4 show "h`(N\<zmu>p) \<zlq> M\<zmu>p"
using Real_ZF_1_4_L26 by auto
qed
with T I have "[h]⋅N⇧R \<lsq> M⇧R"
using PositiveSet_def Real_ZF_1_4_L23A
by simp
with T III have "[h] \<lsq> M⇧R⋅(N⇧R)¯"
using Real_ZF_1_3_L4C by simp
moreover from T II III have "M⇧R⋅(N⇧R)¯ \<ls> [h]"
using Real_ZF_1_3_L4A by simp
ultimately have False using Real_ZF_1_2_L29 by blast
} then show "[h] \<lsq> y" by auto
qed
text‹Now we can prove that every nonempty subset of reals that is bounded
above has a supremum. Proof by considering two cases: when the set has a
maximum and when it does not.›
lemma (in real1) real_order_complete:
assumes A1: "IsBoundedAbove(S,OrderOnReals)" "S≠0"
shows "HasAminimum(OrderOnReals,⋂a∈S. OrderOnReals``{a})"
proof -
{ assume "HasAmaximum(OrderOnReals,S)"
with A1 have "HasAminimum(OrderOnReals,⋂a∈S. OrderOnReals``{a})"
using Real_ZF_1_2_L10 IsAnOrdGroup_def IsPartOrder_def
Order_ZF_5_L6 by simp }
moreover
{ assume A2: "¬HasAmaximum(OrderOnReals,S)"
let ?h = "OddExtension(int,IntegerAddition,IntegerOrder,{⟨p,Γ(S,p)⟩. p∈ℤ⇩+})"
let ?r = "OrderOnReals"
from A1 have "antisym(OrderOnReals)" "S≠0"
using Real_ZF_1_2_L10 IsAnOrdGroup_def IsPartOrder_def by auto
moreover from A1 A2 have "∀x∈S. ⟨x,[?h]⟩ ∈ ?r"
using Real_ZF_1_4_L25 by simp
moreover from A1 have "∀y. (∀x∈S. ⟨x,y⟩ ∈ ?r) ⟶ ⟨[?h],y⟩ ∈ ?r"
using Real_ZF_1_4_L28 by simp
ultimately have "HasAminimum(OrderOnReals,⋂a∈S. OrderOnReals``{a})"
by (rule Order_ZF_5_L5) }
ultimately show ?thesis by blast
qed
text‹Finally, we are ready to formulate the main result: that the
construction of real numbers from the additive group of integers
results in a complete ordered field.
This theorem completes the construction. It was fun.›
theorem eudoxus_reals_are_reals: shows
"IsAmodelOfReals(RealNumbers,RealAddition,RealMultiplication,OrderOnReals)"
using real1.reals_are_ord_field real1.real_order_complete
IsComplete_def IsAmodelOfReals_def by simp
end