Up to index of Isabelle/ZF/IsarMathLib
theory Real_ZF_1(* This file is a part of IsarMathLib - a library of formalized mathematics for Isabelle/Isar. Copyright (C) 2005, 2006 Slawomir Kolodynski This program is free software; Redistribution and use in source and binary forms, with or without modification, are permitted provided that the following conditions are met: 1. Redistributions of source code must retain the above copyright notice, this list of conditions and the following disclaimer. 2. Redistributions in binary form must reproduce the above copyright notice, this list of conditions and the following disclaimer in the documentation and/or other materials provided with the distribution. 3. The name of the author may not be used to endorse or promote products derived from this software without specific prior written permission. HIS SOFTWARE IS PROVIDED BY THE AUTHOR ``AS IS'' AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE AUTHOR BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. *) header {*\isaheader{Real\_ZF\_1.thy}*} theory Real_ZF_1 imports Real_ZF Int_ZF_2 OrderedField_ZF begin text{*In this theory file we continue the construction of real numbers started in @{text "Real_ZF.thy"} 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.*} section{*Definitions and notation*} text{*In this section we define notions and notation needed for the rest of the construction.*} text{* The order on the set of real numbers is constructed by specifying the set of positive reals. This is defined as the projection of the set of positive slopes. A slope is positive if it takes an infinite number of posititive values on the positive integers (see @{text "Int_ZF_2.thy"} for properties of positive slopes). The order relation on real numbers is defined by prescribing the set of positive numbers (see section "Alternative definitions" in @{text "OrderedGroup_ZF.thy."}). *} constdefs "PositiveSlopes ≡ {s ∈ Slopes. s``(PositiveIntegers) ∩ PositiveIntegers ∉ Fin(int)}" "PositiveReals ≡ {SlopeEquivalenceRel``{s}. s ∈ PositiveSlopes}" "OrderOnReals ≡ OrderFromPosSet(RealNumbers,RealAddition,PositiveReals)" text{*The next locale extends the locale @{text "real0"} to define notation specific to the construction of real numbers. The notation follows the one defined in @{text "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 @{text "mR"}. 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 @{text "int0"} context, like @{text "\<int>+ "} (the set of positive integers) and abs$(m)$ ( the absolute value of an integer, and some defined in the @{text "int1"} context, like the addition ( @{text "\<fp>"}) and composition (@{text "o"} 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 "o" 71) defines slope_comp_def [simp]: "s o r ≡ SlopeOp2`〈s,r〉" fixes slopes ("\<S>") defines slopes_def [simp]: "\<S> ≡ AlmostHoms(int,IntegerAddition)" fixes posslopes ("\<S>+") defines posslopes_def [simp]: "\<S>+ ≡ PositiveSlopes" fixes slope_class ("[ _ ]") defines slope_class_def [simp]: "[f] ≡ SlopeEquivalenceRel``{f}" fixes slope_neg :: "i=>i" ("\<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 ("\<real>+") defines positivereals_def [simp]: "\<real>+ ≡ PositiveSet(\<real>,RealAddition,OrderOnReals)" fixes intembed ("_R" [90] 91) defines intembed_def [simp]: "mR ≡ [{〈n,IntegerMultiplication`〈m,n〉 〉. n ∈ int}]" fixes floor ("⌊ _ ⌋") defines floor_def [simp]: "⌊a⌋ ≡ Maximum(IntegerOrder,{m ∈ int. mR \<lsq> a})" fixes Γ defines Γ_def [simp]: "Γ(S,p) ≡ Maximum(IntegerOrder,{⌊pR·x⌋. x∈S})" fixes ia (infixl "\<za>" 69) defines ia_def [simp]: "a\<za>b ≡ IntegerAddition`<a,b>" fixes iminus :: "i=>i" ("\<zm> _" 72) defines rminus_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 ("\<int>+") defines intpositives_def [simp]: "\<int>+ ≡ 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 ("\<zero>Z") defines izero_def [simp]: "\<zero>Z ≡ TheNeutralElement(int,IntegerAddition)" fixes ione ("\<one>Z") defines ione_def [simp]: "\<one>Z ≡ TheNeutralElement(int,IntegerMultiplication)" fixes itwo ("\<two>Z") defines itwo_def [simp]: "\<two>Z ≡ \<one>Z\<za>\<one>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)"; section{*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 @{text "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 ∈ \<S>" "g ∈ \<S>" shows "[fog] = [gof]" proof - from A1 have "fog ∼ gof" using Slopes_def int1.Arthan_Th_9 SlopeOp1_def BoundedIntMaps_def SlopeEquivalenceRel_def SlopeOp2_def by simp; then show ?thesis using (*Real_ZF_1_1_L1*) 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 ∈ \<S>" shows "[f] ∈ \<real>" proof - from A1 have "[f] ∈ Slopes//SlopeEquivalenceRel" using Slopes_def quotientI by simp; then show "[f] ∈ \<real>" 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∈\<real>" shows "∃f∈\<S> . a = [f]" proof - from A1 have "a ∈ \<S>//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 @{text "real1"} context notation.*} lemma (in real1) Real_ZF_1_1_L4: assumes A1: "f ∈ \<S>" "g ∈ \<S>" shows "[f] \<ra> [g] = [f\<fp>g]" "[f] · [g] = [fog]" proof - let ?r = "SlopeEquivalenceRel" have "[f]·[g] = ProjFun2(\<S>,?r,SlopeOp2)`〈[f],[g]〉" using RealMultiplication_def Slopes_def by simp; also from A1 have "… = [fog]" using Real_ZF_1_L11 EquivClass_1_L10 Slopes_def by simp; finally show "[f] · [g] = [fog]" by simp; have "[f] \<ra> [g] = ProjFun2(\<S>,?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 @{text " Real_ZF_1_L12"}, but written in the notation defined in the @{text "real1"} context. It states that if $f$ is a slope, then $-[f] = [-f]$.*} lemma (in real1) Real_ZF_1_1_L4A: assumes "f ∈ \<S>" shows "[\<fm>f] = \<rm>[f]" using prems 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 ∈ \<S>" "g ∈ \<S>" 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∈\<real>" "b∈\<real>" shows "a·b = b·a" proof - from A1 have "∃f∈\<S> . a = [f]" "∃g∈\<S> . b = [g]" using Real_ZF_1_1_L3A by auto; then obtain f g where "f ∈ \<S>" "g ∈ \<S>" 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 @{text "\<one>"} in the @{text "real1"} context) is the class of identity function on integers. This is really shown in @{text "Real_ZF_1_L11"}, here we only rewrite it in the notation used in the @{text "real1"} context.*} lemma (in real1) real_one_cl_identity: shows "[id(int)] = \<one>" 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 @{text "\<zero>"} in the @{text "real1"} context).*} lemma (in real1) real_zero_cl_bounded_map: assumes "f ∈ BoundedIntMaps" shows "[f] = \<zero>" using prems 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 @{text "Real_ZF_1_L13"}, here we just rewrite it in the notation used in the @{text "real1"} context.*} lemma (in real1) Real_ZF_1_1_L5: assumes "f ∈ \<S>" "g ∈ \<S>" shows "[f] = [g] <-> f ∼ g" using prems 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 prems 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 @{text "Real_ZF_1_L13"}, here we just rewrite it in the notation used in the @{text "real1"} context.*} lemma (in real1) id_on_int_is_slope: shows "id(int) ∈ \<S>" using Real_ZF_1_L14 Slopes_def by simp; text{*A result from @{text "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 prems 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 "\<one>≠\<zero>" proof - { assume A1: "\<one>=\<zero>" have "∃f ∈ \<S>. \<zero> = [f]" using Real_ZF_1_L4 Real_ZF_1_1_L3A by simp; with A1 have "∃f ∈ \<S>. [id(int)] = [f] ∧ [f] = \<zero>" 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 "\<one>≠\<zero>" 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∈\<real>" shows "(\<rm>a) ∈ \<real>" using prems 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∈\<real>" "b∈\<real>" "c∈\<real>" shows "a·(b·c) = a·c·b" using prems real_mult_commutative Real_ZF_1_L3 ring0.Ring_ZF_2_L4 by simp section{*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∈\<S>+. a = [f])" proof assume "a ∈ PositiveReals" then have "a ∈ {([s]). s ∈ \<S>+}" using PositiveReals_def by simp; (* it has to be ([a]), otherwise fails*) then show "∃f∈\<S>+. a = [f]" by auto; next assume "∃f∈\<S>+. a = [f]" then have "a ∈ {([s]). s ∈ \<S>+}" by auto; then show "a ∈ PositiveReals" using PositiveReals_def by simp; qed; text{*Let's recall from @{text "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∈\<S>+" "g∈\<S>+" shows "f\<fp>g ∈ \<S>+" "fog ∈ \<S>+" using prems 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 ∉ \<S>+" using prems 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" "\<zero> ∉ PositiveReals" proof - { fix a fix b assume "a ∈ PositiveReals" and "b ∈ PositiveReals" then obtain f g where I: "f ∈ \<S>+" "g ∈ \<S>+" and II: "a = [f]" "b = [g]" using Real_ZF_1_2_L2 by auto; then have "f ∈ \<S>" "g ∈ \<S>" 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 "\<zero> ∈ PositiveReals" then obtain f where "f ∈ \<S>+" and "\<zero> = [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 "\<zero> ∉ 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 @{text "Int_ZF_2.thy."}*} lemma (in real1) Real_ZF_1_2_L7: assumes A1: "f ∈ \<S>" and A2: "[f] ≠ \<zero>" shows "(f ∈ \<S>+) Xor ((\<fm>f) ∈ \<S>+)" using prems 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 @{text " Int_ZF_2_3_L10"} in the notation used in @{text "real1"} context.*} lemma (in real1) Real_ZF_1_2_L8: assumes A1: "f ∈ \<S>" "g ∈ \<S>" and A2: "(f ∈ \<S>+) Xor (g ∈ \<S>+)" shows "([f] ∈ PositiveReals) Xor ([g] ∈ PositiveReals)" using prems 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∈\<real>" and A2: "a≠\<zero>" shows "(a ∈ PositiveReals) Xor ((\<rm>a) ∈ PositiveReals)" proof - from A1 obtain f where I: "f ∈ \<S>" "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 @{text "ring1"} (about ordered rings), @{text "group3"} (about ordered groups) and @{text "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∈\<real>" "b∈\<real>" and A3: "a=b ∨ b\<rs>a ∈ PositiveReals" shows "a\<lsq>b" using prems 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 ∈ \<S>" "g ∈ \<S>" and A2: "f∼g ∨ (g \<fp> (\<fm>f)) ∈ \<S>+" 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∈\<real>" and A2: "(\<rm>a) \<lsq> b" shows "(\<rm>b) \<lsq> a" using prems 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 "\<zero>\<lsq>c" shows "a·c \<lsq> b·c" "c·a \<lsq> c·b" using prems Real_ZF_1_2_L10 ring1.OrdRing_ZF_1_L9 by auto; text{*A special case of @{text "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∈\<real>+" shows "c·a \<lsq> c·b" using prems Real_ZF_1_2_L10 ring1.OrdRing_ZF_1_L9A by simp text{*In the @{text "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∈\<real>" "b∈\<real>" using prems 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 "\<zero> \<lsq> b\<rs>a" using prems 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 "\<zero>\<lsq>a" "\<zero>\<lsq>b" shows "\<zero> \<lsq> a\<ra>b" using prems 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 prems 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∈\<real>" shows "a\<lsq>a" using prems 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∈\<real>" shows "a\<ra>c \<lsq> b\<ra>c" using prems 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∈\<real>" "b∈\<real>" and "c \<lsq> a\<ra>b" shows "c\<rs>b \<lsq> a" using prems 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∈\<real>" "b∈\<real>" and "¬(a\<lsq>b)" shows "b \<ls> a" proof - from prems have I: "group3(\<real>,RealAddition,OrderOnReals)" "OrderOnReals {is total on} \<real>" "a∈\<real>" "b∈\<real>" "¬(〈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∈\<real>" "b∈\<real>" and "c \<lsq> a\<rs>b" shows "c\<ra>b \<lsq> a" using prems 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 ⊆ \<real>×\<real>" 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 ⊆ \<real>" using A1 Real_ZF_1_2_L22 Order_ZF_3_L1A by blast; (* A1 has to be here *) text{*Properties of the maximum of three real numbers.*} lemma (in real1) Real_ZF_1_2_L24: assumes A1: "a∈\<real>" "b∈\<real>" "c∈\<real>" shows "Maximum(OrderOnReals,{a,b,c}) ∈ {a,b,c}" "Maximum(OrderOnReals,{a,b,c}) ∈ \<real>" "a \<lsq> Maximum(OrderOnReals,{a,b,c})" "b \<lsq> Maximum(OrderOnReals,{a,b,c})" "c \<lsq> Maximum(OrderOnReals,{a,b,c})" proof - have "IsLinOrder(\<real>,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}) ∈ \<real>" "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; 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(\<real>,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 ∈ \<real>+" and "a\<lsq>b" and "\<one>\<ls>c" shows "a\<ls>b·c" using prems 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∈\<real>" "b∈\<real>" and "a\<rs>b \<ls> c" shows "a \<ls> c\<ra>b" using prems 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∈\<real>" shows "c\<ra>a \<lsq> c\<ra>b" using prems 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 @{text "real1"} context. This may be confusing for the presentation readers: even though @{text "\<zlq>"} and @{text "\<lsq>"} are printed in the same way, they are different symbols in the source. In the @{text "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⊆\<real>" and "¬HasAmaximum(OrderOnReals,A)" and "x∈A" shows "∃y∈A. x\<ls>y" using prems 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∈\<real>" "b∈\<real>" and "¬(a\<lsq>b)" shows "b\<ls>a" proof - from prems have "group3(\<real>,RealAddition,OrderOnReals)" "OrderOnReals {is total on} \<real>" "a∈\<real>" "b∈\<real>" "〈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 secon 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 prems have "group3(\<real>,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; section{*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 @{text "Field_ZF.thy"} and @{text "OrderedField_ZF.thy"}*} text{*We rewrite the theorem from @{text "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 ∈ \<S>+" shows "∃g∈\<S>. f∼g ∧ (∃h∈\<S>. goh ∼ id(int))" using prems 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 "\<zero> ≠ \<one>" 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 = \<one>" proof fix a assume "a ∈ PositiveSet(?R,?A,?r)" then obtain f where I: "f∈\<S>+" and II: "a = [f]" using reals_are_ord_ring Real_ZF_1_2_L2 by auto; then have "∃g∈\<S>. f∼g ∧ (∃h∈\<S>. goh ∼ id(int))" using pos_slopes_have_inv by simp; then obtain g where III: "g∈\<S>" and IV: "f∼g" and V: "∃h∈\<S>. goh ∼ id(int)" by auto; from V obtain h where VII: "h∈\<S>" and VIII: "goh ∼ 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] = \<one>" using Real_ZF_1_1_L4 Real_ZF_1_1_L5A real_one_cl_identity by simp; with VII show "∃b∈?R. a·b = \<one>" 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 @{text "field0"} and @{text "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_ZF_1_L2 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 ∈ \<real>+" shows "a¯ ∈ \<real>+" "a¯ ∈ \<real>" using prems 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 ∈ \<real>+" and "a¯ \<ls> b" shows "\<one> \<ls> b·a" using prems field_cntxts_ok field1.OrdField_ZF_2_L2 by simp; text{*If $a<b$, then $(b-a)^{-1}$ is positive.*} lemma (in real1) Real_ZF_1_3_L3: assumes "a\<ls>b" shows "(b\<rs>a)¯ ∈ \<real>+" using prems 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∈\<real>" "b∈\<real>+" and A2: "a·b \<ls> c" shows "a \<ls> c·b¯" using prems 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∈\<real>" "c∈\<real>+" and A2: "a \<ls> b·c" shows "a·c¯ \<ls> b" using prems 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∈\<real>" "c∈\<real>+" and A2: "a \<lsq> b·c" shows "a·c¯ \<lsq> b" using prems 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∈\<real>" "b∈\<real>+" and A2: "a·b \<lsq> c" shows "a \<lsq> c·b¯" using prems 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 "\<one> \<ra> a·c \<ls> b·c" using prems 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∈\<real>+" shows "a·c¯ \<lsq> b·c¯" using prems 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∈\<real>+" shows "a·c \<ls> b·c" "c·a \<ls> c·b" "a·c¯ \<ls> b·c¯" using prems 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∈\<real>" "b∈\<real>" "b≠\<zero>" "c∈\<real>" shows "a·b·(c·b¯) = a·c" using prems field_cntxts_ok field0.Field_ZF_2_L6 by simp; section{*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 @{text "mR"} is a real number. Recall that in @{text "real1"} context @{text "mR"} denotes the class of the slope $n\mapsto m\cdot n$.*} lemma (in real1) real_int_is_real: assumes "m ∈ int" shows "mR ∈ \<real>" using prems 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>(mR)" using prems 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 "mR \<ra> kR = ((m\<za>k)R)" using prems 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 "mR \<rs> kR = (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 = mR \<ra> (\<zm>k)R" using Real_ZF_1_4_L1A by simp; with A1 show "mR \<rs> kR = (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 "mR · kR = (m\<zmu>k)R" using prems 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∈\<real>" shows "∃m∈int. a \<lsq> mR" proof - from A1 obtain f where I: "f∈\<S>" and II: "a = [f]" using Real_ZF_1_1_L3A by auto; then have "∃m∈int. ∃g∈\<S>. {〈n,m\<zmu>n〉 . n ∈ int} ∼ g ∧ (f∼g ∨ (g \<fp> (\<fm>f)) ∈ \<S>+)" 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∈\<S>" and "{〈n,m\<zmu>n〉 . n ∈ int} ∼ g ∧ (f∼g ∨ (g \<fp> (\<fm>f)) ∈ \<S>+)" by auto; then have "mR = [g]" and "f ∼ g ∨ (g \<fp> (\<fm>f)) ∈ \<S>+" using Real_ZF_1_1_L5A by auto; with I II IV have "a \<lsq> mR" using Real_ZF_1_2_L12 by simp; with III show "∃m∈int. a \<lsq> mR" 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∈\<real>" shows "{m ∈ int. mR \<lsq> a} ≠ 0" proof - from A1 have "(\<rm>a) ∈ \<real>" using Real_ZF_1_1_L8 by simp; then obtain m where I: "m∈int" and II: "(\<rm>a) \<lsq> mR" using Real_ZF_1_4_L2 by auto; let ?k = "GroupInv(int,IntegerAddition)`(m)" from A1 I II have "?k ∈ int" and "?kR \<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: "mR = kR" 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 "mR \<lsq> kR" proof - let ?r = "{〈n, m\<zmu>n〉 . n ∈ int}" let ?s = "{〈n, k\<zmu>n〉 . n ∈ int}" from A1 have "?r ∈ \<S>" "?s ∈ \<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) ∈ \<S>+" using Slopes_def SlopeOp1_def BoundedIntMaps_def SlopeEquivalenceRel_def PositiveIntegers_def PositiveSlopes_def int1.Int_ZF_2_5_L4 by simp; ultimately show "mR \<lsq> kR" 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 "mR \<ls> kR" proof - from A1 have "mR \<lsq> kR" 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 "mR ≠ kR" using Real_ZF_1_4_L4 by auto; ultimately show "mR \<ls> kR" 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∈\<real>" shows "∃n∈\<int>+. a \<ls> nR" proof - from A1 obtain m where I: "m∈int" and II: "a \<lsq> mR" using Real_ZF_1_4_L2 by auto; let ?n = "GreaterOf(IntegerOrder,\<one>Z,m) \<za> \<one>Z" from I have T: "?n ∈\<int>+" and "m \<zlq> ?n" "m≠?n" using int0.Int_ZF_1_5_L7B by auto; then have III: "mR \<ls> ?nR" using Real_ZF_1_4_L5A by simp; with II have "a \<ls> ?nR" 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: "mR \<lsq> kR" 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 "kR \<lsq> mR" using Real_ZF_1_4_L5 by simp; with A2 have "mR = kR" 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∈\<real>" shows "IsBoundedAbove({m ∈ int. mR \<lsq> a},IntegerOrder)" "{m ∈ int. mR \<lsq> a} ≠ 0" "⌊a⌋ ∈ int" "⌊a⌋R \<lsq> a" proof - let ?A = "{m ∈ int. mR \<lsq> a}" from A1 obtain K where I: "K∈int" and II: "a \<lsq> (KR)" using Real_ZF_1_4_L2 by auto; { fix n assume "n ∈ ?A" then have III: "n ∈ int" and IV: "nR \<lsq> a" by auto from IV II have "(nR) \<lsq> (KR)" 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: "mR \<lsq> a" shows "m \<zlq> ⌊a⌋" proof - let ?A = "{m ∈ int. mR \<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 "\<zero>ZR = \<zero>" "\<one>ZR = \<one>" 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 "\<two>ZR = \<two>" proof - have "\<two>ZR = \<one>ZR \<ra> \<one>ZR" using int0.int_zero_one_are_int Real_ZF_1_4_L1A by simp; also have "… = \<two>" using int_0_1_are_real_zero_one by simp finally show "\<two>ZR = \<two>" 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∈\<int>+" shows "pR ∈ \<real>" "\<zero> \<lsq> pR" "pR ∈ \<real>+" proof - from A1 have I: "p ∈ int" "\<zero>Z \<zlq> p" "\<zero>Z ≠ p" using PositiveSet_def by auto; then have "pR ∈ \<real>" "\<zero>ZR \<lsq> pR" using real_int_is_real Real_ZF_1_4_L5 by auto; then show "pR ∈ \<real>" "\<zero> \<lsq> pR" using int_0_1_are_real_zero_one by auto; moreover have "\<zero> ≠ pR" proof - { assume "\<zero> = pR" 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 "\<zero> ≠ pR" by auto; qed; ultimately show "pR ∈ \<real>+" 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<M^Ry$. This is Lemma 14 ii) in \cite{Arthan2004}.*} lemma (in real1) Arthan_Lemma14ii: assumes A1: "x∈\<real>" "y ∈ \<real>+" shows "∃M∈\<int>+. x \<ls> MR·y" proof - from A1 have "∃C∈\<int>+. x \<ls> CR" and "∃D∈\<int>+. y¯ \<ls> DR" using Real_ZF_1_3_L1 Arthan_Lemma14i by auto; then obtain C D where I: "C∈\<int>+" and II: "x \<ls> CR" and III: "D∈\<int>+" and IV: "y¯ \<ls> DR" by auto; let ?M = "C\<zmu>D" from I III have T: "?M ∈ \<int>+" "CR ∈ \<real>" "DR ∈ \<real>" using int0.pos_int_closed_mul_unfold PositiveSet_def real_int_is_real by auto; with A1 I III have "CR·(DR·y) = ?MR·y" using PositiveSet_def Real_ZF_1_L6A Real_ZF_1_4_L1C by simp; moreover from A1 I II IV have "x \<ls> CR·(DR·y)" using int_pos_is_real_pos Real_ZF_1_3_L2 Real_ZF_1_2_L25 by auto; ultimately have "x \<ls> ?MR·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∈\<real>" 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∈\<int>+" shows "IsBoundedAbove({⌊pR·x⌋. x∈S},IntegerOrder)" "Γ(S,p) ∈ {⌊pR·x⌋. x∈S}" "Γ(S,p) ∈ int" proof - let ?A = "{⌊pR·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 = ⌊pR·x⌋" by auto; with I have "x\<lsq>X" by simp moreover from A2 have "\<zero> \<lsq> pR" using int_pos_is_real_pos by simp; ultimately have "pR·x \<lsq> pR·X" using Real_ZF_1_2_L14 by simp; with II have "m \<zlq> ⌊pR·X⌋" using Real_ZF_1_4_L9 by simp; } then have "∀m∈?A. 〈m,⌊pR·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) ∈ {⌊pR·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∈\<int>+" shows "⌊pR·x⌋ \<zlq> Γ(S,p)" proof - let ?A = "{⌊pR·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 "⌊pR·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∈\<int>+}" shows "g : \<int>+->int" "∀n∈\<int>+. g`(n) = Γ(S,n)" proof - from A1 have "∀n∈\<int>+. Γ(S,n) ∈ int" using Real_ZF_1_4_L10 by simp with A2 show I: "g : \<int>+->int" using ZF_fun_from_total by simp { fix n assume "n∈\<int>+" with A2 I have "g`(n) = Γ(S,n)" using ZF_fun_from_tot_val by simp } then show "∀n∈\<int>+. 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 "⌊mR⌋ = m" proof - let ?A = "{n ∈ int. nR \<lsq> mR }" 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 "⌊mR⌋ = 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 "⌊\<zero>⌋ = \<zero>Z" "⌊\<one>⌋ = \<one>Z" proof - have "⌊(\<zero>Z)R⌋ = \<zero>Z" and "⌊(\<one>Z)R⌋ = \<one>Z" using int0.int_zero_one_are_int Real_ZF_1_4_L14 by auto then show "⌊\<zero>⌋ = \<zero>Z" and "⌊\<one>⌋ = \<one>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 "⌊\<two>⌋ = \<two>Z" proof - have "⌊(\<two>Z)R⌋ = \<two>Z" using int0.int_two_three_are_int Real_ZF_1_4_L14 by simp then show "⌊\<two>⌋ = \<two>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 "⌊mR·kR⌋ = m\<zmu>k" proof - from A1 have T: "m\<zmu>k ∈ int" using int0.Int_ZF_1_1_L5 by simp; from A1 have "⌊mR·kR⌋ = ⌊(m\<zmu>k)R⌋" using Real_ZF_1_4_L1C by simp; with T show "⌊mR·kR⌋ = 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∈\<real>" and A2: "p ∈ int" shows "⌊x \<ra> pR⌋ = ⌊x⌋ \<za> p" proof - let ?A = "{n ∈ int. nR \<lsq> x \<ra> pR}" 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 "pR ∈ \<real>" using Real_ZF_1_4_L7 real_int_is_real by auto; then have "⌊x⌋R \<ra> pR \<lsq> x \<ra> pR" using add_num_to_ineq by simp; moreover from A1 A2 have "(⌊x⌋ \<za> p)R = ⌊x⌋R \<ra> pR" using Real_ZF_1_4_L7 Real_ZF_1_4_L1A by simp; ultimately have "(⌊x⌋ \<za> p)R \<lsq> x \<ra> pR" 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 "nR \<lsq> x \<ra> pR" by auto with A1 A2 have "nR \<rs> pR \<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> pR⌋ = ⌊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∈\<real>" and A2: "p ∈ int" shows "⌊x \<rs> pR⌋ = ⌊x⌋ \<zs> p" proof - from A2 have "⌊x \<rs> pR⌋ = ⌊x \<ra> (\<zm>p)R⌋" using Real_ZF_1_4_L1 by simp; with A1 A2 show "⌊x \<rs> pR⌋ = ⌊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 "⌊(mR) \<ra> nR⌋ = m \<za> n" using prems 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∈\<real>" shows "\<one> \<ra> ⌊a⌋R = (\<one>Z \<za> ⌊a⌋)R" proof - have "\<one> \<ra> ⌊a⌋R = \<one>ZR \<ra> ⌊a⌋R" using int_0_1_are_real_zero_one by simp; with A1 show "\<one> \<ra> ⌊a⌋R = (\<one>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∈\<real>" shows "a \<rs> ⌊a⌋R \<ls> \<one>" "a \<ls> (\<one>Z \<za> ⌊a⌋)R" proof - from A1 have T1: "⌊a⌋ ∈ int" "⌊a⌋R ∈ \<real>" and T2: "\<one> ∈ \<real>" "a \<rs> ⌊a⌋R ∈ \<real>" using Real_ZF_1_4_L7 real_int_is_real Real_ZF_1_L6 Real_ZF_1_L4 by auto; { assume "\<one> \<lsq> a \<rs> ⌊a⌋R" with A1 T1 have "⌊\<one>ZR \<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: "¬(\<one> \<lsq> a \<rs> ⌊a⌋R)" by auto; with T2 show II: "a \<rs> ⌊a⌋R \<ls> \<one>" by (rule Real_ZF_1_2_L20); with A1 T1 I II have "a \<ls> \<one> \<ra> ⌊a⌋R" using Real_ZF_1_2_L26 by simp; with A1 show "a \<ls> (\<one>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∈\<int>+. x·NR \<ls> MR ∧ MR \<ls> y·NR" proof -; from A1 have "(y\<rs>x)¯ ∈ \<real>+" using Real_ZF_1_3_L3 by simp then have "∃N∈\<int>+. (y\<rs>x)¯ \<ls> NR" using Arthan_Lemma14i PositiveSet_def by simp; then obtain N where I: "N∈\<int>+" and II: "(y\<rs>x)¯ \<ls> NR" by auto; let ?M = "\<one>Z \<za> ⌊x·NR⌋" from A1 I have T1: "x∈\<real>" "NR ∈ \<real>" "NR ∈ \<real>+" "x·NR ∈ \<real>" 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·NR \<ls> ?MR" using Real_ZF_1_4_L17B by simp; from T1 have "(\<one> \<ra> ⌊x·NR⌋R) \<lsq> (\<one> \<ra> x·NR)" using Real_ZF_1_4_L7 Real_ZF_1_L4 real_ord_transl_inv by simp; with T1 have "?MR \<lsq> (\<one> \<ra> x·NR)" using Real_ZF_1_4_L17A by simp; moreover from A1 II have "(\<one> \<ra> x·NR) \<ls> y·NR" using Real_ZF_1_3_L5 by simp; ultimately have "?MR \<ls> y·NR" 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∈\<real>" "y∈\<real>" shows "abs(⌊x\<ra>y⌋ \<zs> ⌊x⌋ \<zs> ⌊y⌋) \<zlq> \<two>Z" proof - from A1 have T: "⌊x⌋R ∈ \<real>" "⌊y⌋R ∈ \<real>" "x\<ra>y \<rs> (⌊x⌋R) ∈ \<real>" using Real_ZF_1_4_L7 real_int_is_real Real_ZF_1_L6 by auto; from A1 have "\<zero> \<lsq> x \<rs> (⌊x⌋R) \<ra> (y \<rs> (⌊y⌋R))" "x \<rs> (⌊x⌋R) \<ra> (y \<rs> (⌊y⌋R)) \<lsq> \<two>" 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 "\<zero> \<lsq> x\<ra>y \<rs> (⌊x⌋R) \<rs> (⌊y⌋R)" "x\<ra>y \<rs> (⌊x⌋R) \<rs> (⌊y⌋R) \<lsq> \<two>" by auto; then have "⌊\<zero>⌋ \<zlq> ⌊x\<ra>y \<rs> (⌊x⌋R) \<rs> (⌊y⌋R)⌋" "⌊x\<ra>y \<rs> (⌊x⌋R) \<rs> (⌊y⌋R)⌋ \<zlq> ⌊\<two>⌋" using Real_ZF_1_4_L9 by auto; then have "\<zero>Z \<zlq> ⌊x\<ra>y \<rs> (⌊x⌋R) \<rs> (⌊y⌋R)⌋" "⌊x\<ra>y \<rs> (⌊x⌋R) \<rs> (⌊y⌋R)⌋ \<zlq> \<two>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 "\<zero>Z \<zlq> ⌊x\<ra>y⌋ \<zs> ⌊x⌋ \<zs> ⌊y⌋" "⌊x\<ra>y⌋ \<zs> ⌊x⌋ \<zs> ⌊y⌋ \<zlq> \<two>Z" by auto; then show "abs(⌊x\<ra>y⌋ \<zs> ⌊x⌋ \<zs> ⌊y⌋) \<zlq> \<two>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∈\<int>+" "x∈S" and A3: "Γ(S,n) = ⌊nR·x⌋" and A4: "y∈S" "x\<lsq>y" shows "Γ(S,n) = ⌊nR·y⌋" proof - from A2 A4 have "⌊nR·x⌋ \<zlq> ⌊(nR)·y⌋" using int_pos_is_real_pos Real_ZF_1_2_L14 Real_ZF_1_4_L9 by simp; with A3 have "〈Γ(S,n),⌊(nR)·y⌋〉 ∈ IntegerOrder" by simp moreover from A1 A2 A4 have "〈⌊nR·y⌋,Γ(S,n)〉 ∈ IntegerOrder" using Real_ZF_1_4_L11 by simp; ultimately show "Γ(S,n) = ⌊nR·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∈\<int>+" "n∈\<int>+" shows "abs(Γ(S,m\<za>n) \<zs> Γ(S,m) \<zs> Γ(S,n)) \<zlq> \<two>Z" proof - from A2 have T: "m\<za>n ∈ \<int>+" using int0.pos_int_closed_add_unfolded by simp; with A1 A2 have "Γ(S,m) ∈ {⌊mR·x⌋. x∈S}" and "Γ(S,n) ∈ {⌊nR·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) = ⌊mR·a⌋" "Γ(S,n) = ⌊nR·b⌋" "Γ(S,m\<za>n) = ⌊(m\<za>n)R·c⌋" by auto let ?d = "Maximum(OrderOnReals,{a,b,c})" from A1 I have "a∈\<real>" "b∈\<real>" "c∈\<real>" using Real_ZF_1_2_L23 by auto; then have IV: "?d ∈ {a,b,c}" "?d ∈ \<real>" "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 "… = ⌊((mR)\<ra>(nR))·?d⌋" using Real_ZF_1_4_L1A PositiveSet_def by simp; also from A2 IV have "… = ⌊(mR)·?d \<ra> (nR)·?d⌋" using PositiveSet_def real_int_is_real Real_ZF_1_L7 by simp; finally have "Γ(S,m\<za>n) = ⌊(mR)·?d \<ra> (nR)·?d⌋" by simp moreover from A1 A2 I II IV V have "Γ(S,m) = ⌊mR·?d⌋" using Real_ZF_1_4_L20 by blast; moreover from A1 A2 I II IV V have "Γ(S,n) = ⌊nR·?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(⌊(mR)·?d \<ra> (nR)·?d⌋ \<zs> ⌊mR·?d⌋ \<zs> ⌊nR·?d⌋)" by simp; with A2 IV show "abs(Γ(S,m\<za>n) \<zs> Γ(S,m) \<zs> Γ(S,n)) \<zlq> \<two>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 @{text "δ"} in the @{text "real1"} context is bounded on positive integers. This is really proven in @{text "Int_ZF_2.thy"}, but we restate it here for convenience. Recall from @{text "Group_ZF_3.thy"} that @{text "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>+->int" "∀a∈\<int>+. ∀b∈\<int>+. abs(δ(f,a,b)) \<zlq> L" shows "OddExtension(int,IntegerAddition,IntegerOrder,f) ∈ \<S>" using A1 int1.Int_ZF_2_1_L24 by auto; (*A1 has to be here *) 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∈\<int>+}" shows "OddExtension(int,IntegerAddition,IntegerOrder,g) ∈ \<S>" proof - from A1 A2 have "g: \<int>+->int" by (rule Real_ZF_1_4_L12); moreover have "∀m∈\<int>+. ∀n∈\<int>+. abs(δ(g,m,n)) \<zlq> \<two>Z" proof - { fix m n assume A3: "m∈\<int>+" "n∈\<int>+" then have "m\<za>n ∈ \<int>+" "m∈\<int>+" "n∈\<int>+" using int0.pos_int_closed_add_unfolded by auto moreover from A1 A2 have "∀n∈\<int>+. 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> \<two>Z" by (rule Real_ZF_1_4_L21) ultimately have "abs(δ(g,m,n)) \<zlq> \<two>Z" by simp } then show "∀m∈\<int>+. ∀n∈\<int>+. abs(δ(g,m,n)) \<zlq> \<two>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 ∈ \<S>" and A2: "N ∈ int" "M ∈ int" and A3: "∀n∈\<int>+. M\<zmu>n \<zlq> f`(N\<zmu>n)" shows "MR \<lsq> [f]·(NR)" proof - let ?MS = "{〈n, M\<zmu>n〉 . n ∈ int}" let ?NS = "{〈n, N\<zmu>n〉 . n ∈ int}" from A1 A2 have T: "?MS ∈ \<S>" "?NS ∈ \<S>" "fo?NS ∈ \<S>" using int1.Int_ZF_2_5_L1 int1.Int_ZF_2_1_L11 SlopeOp2_def by auto; moreover from A1 A2 A3 have "?MS ∼ fo?NS ∨ fo?NS \<fp> (\<fm>?MS) ∈ \<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 "[?MS] \<lsq> [fo?NS]" using Real_ZF_1_2_L12 by simp; with A1 T show "MR \<lsq> [f]·(NR)" 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 ∈ \<S>" and A2: "N ∈ int" "M ∈ int" and A3: "∀n∈\<int>+. f`(N\<zmu>n) \<zlq> M\<zmu>n " shows "[f]·(NR) \<lsq> MR" proof - let ?MS = "{〈n, M\<zmu>n〉 . n ∈ int}" let ?NS = "{〈n, N\<zmu>n〉 . n ∈ int}" from A1 A2 have T: "?MS ∈ \<S>" "?NS ∈ \<S>" "fo?NS ∈ \<S>" using int1.Int_ZF_2_5_L1 int1.Int_ZF_2_1_L11 SlopeOp2_def by auto; moreover from A1 A2 A3 have "fo?NS ∼ ?MS ∨ ?MS \<fp> (\<fm>(fo?NS)) ∈ \<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 "[fo?NS] \<lsq> [?MS]" using Real_ZF_1_2_L12 by simp; with A1 T show " [f]·(NR)\<lsq> MR" 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 ∈ \<int>+" "M ∈ int" and A5: "MR \<ls> y·NR" and A6: "p ∈ \<int>+" shows "p\<zmu>M \<zlq> Γ(S,p\<zmu>N)" proof - from A2 A4 A6 have T1: "NR ∈ \<real>+" "y∈\<real>" "pR ∈ \<real>+" "p\<zmu>N ∈ \<int>+" "(p\<zmu>N)R ∈ \<real>+" 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" "pR ∈ \<real>" "NR ∈ \<real>" "NR ≠ \<zero>" "MR ∈ \<real>" using real_int_is_real PositiveSet_def by auto; from T1 A5 have "⌊(p\<zmu>N)R·(MR·(NR)¯)⌋ \<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·(MR·(NR)¯)⌋ \<zlq> Γ(S,p\<zmu>N)" by (rule int_order_transitive); from A4 A6 have "(p\<zmu>N)R·(MR·(NR)¯) = pR·NR·(MR·(NR)¯)" using PositiveSet_def Real_ZF_1_4_L1C by simp; with A4 T2 have "⌊(p\<zmu>N)R·(MR·(NR)¯)⌋ = 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 ∈ \<int>+" and A3: "h = OddExtension(int,IntegerAddition,IntegerOrder,{〈p,Γ(S,p)〉. p∈\<int>+})" shows "h`(p) = Γ(S,p)" proof - let ?g = "{〈p,Γ(S,p)〉. p∈\<int>+}" from A1 have I: "?g : \<int>+->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∈\<int>+})" shows "x \<lsq> [h]" proof - from A1 A2 A3 have "S ⊆ \<real>" "¬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∈\<int>+. x·NR \<ls> MR ∧ MR \<ls> y·NR" using Arthan_Lemma14iii by simp; then obtain M N where III: "M ∈ int" "N∈\<int>+" and IV: "x·NR \<ls> MR" "MR \<ls> y·NR" by auto; from II III IV have V: "x \<lsq> MR·(NR)¯" 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 ∈ \<S>" using Real_ZF_1_4_L22 by simp; moreover from III have "N ∈ int" "M ∈ int" using PositiveSet_def by auto; moreover have "∀n∈\<int>+. M\<zmu>n \<zlq> h`(N\<zmu>n)" proof let ?g = "{〈p,Γ(S,p)〉. p∈\<int>+}" fix n assume A5: "n∈\<int>+" with III have T2: "N\<zmu>n ∈ \<int>+" 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 ∈ \<int>+" "M ∈ int" "MR \<ls> y·NR" "n ∈ \<int>+" 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 "MR \<lsq> [h]·NR" using Real_ZF_1_4_L23 by simp; with III T1 have "MR·(NR)¯ \<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 ∈ \<int>+" "M ∈ int" and A5: "y·NR \<ls> MR " and A6: "p ∈ \<int>+" shows "⌊(N\<zmu>p)R·x⌋ \<zlq> M\<zmu>p" proof - from A2 A4 A6 have T: "p\<zmu>N ∈ \<int>+" "p ∈ int" "N ∈ int" "pR ∈ \<real>+" "pR ∈ \<real>" "NR ∈ \<real>" "x ∈ \<real>" "y ∈ \<real>" 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 = pR·NR" "(p\<zmu>M)R = pR·MR" using Real_ZF_1_4_L1C by auto; ultimately have "(p\<zmu>N)R·x \<lsq> pR·NR·y" by simp; moreover from A5 T I have "pR·(y·NR) \<ls> (p\<zmu>M)R" using Real_ZF_1_3_L7 by simp; with T have "pR·NR·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∈\<int>+})" and "p ∈ \<int>+" shows "∃x∈S. h`(p) = ⌊pR·x⌋" using prems 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∈\<int>+})" shows "[h] \<lsq> y" proof - from A1 obtain a where "a∈S" by auto; with A1 A2 A3 have T: "y∈\<real>" "h ∈ \<S>" "[h] ∈ \<real>" 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∈\<int>+. y·NR \<ls> MR ∧ MR \<ls> [h]·NR" using Arthan_Lemma14iii by simp; then obtain M N where I: "M∈int" "N∈\<int>+" and II: "y·NR \<ls> MR" "MR \<ls> [h]·NR" by auto; from I have III: "NR ∈ \<real>+" using int_pos_is_real_pos by simp; have "∀p∈\<int>+. h`(N\<zmu>p) \<zlq> M\<zmu>p" proof fix p assume A4: "p∈\<int>+" 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]·NR \<lsq> MR" using PositiveSet_def Real_ZF_1_4_L23A by simp; with T III have "[h] \<lsq> MR·(NR)¯" using Real_ZF_1_3_L4C by simp; moreover from T II III have "MR·(NR)¯ \<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.*} lemma (in real1) real_order_complete: assumes A1: "IsBoundedAbove(S,OrderOnReals)" "S≠0" shows "HasAminimum(OrderOnReals,\<Inter>a∈S. OrderOnReals``{a})" proof (cases "HasAmaximum(OrderOnReals,S)") assume "HasAmaximum(OrderOnReals,S)" with A1 show "HasAminimum(OrderOnReals,\<Inter>a∈S. OrderOnReals``{a})" using Real_ZF_1_2_L10 IsAnOrdGroup_def IsPartOrder_def Order_ZF_5_L6 by simp; next assume A2: "¬HasAmaximum(OrderOnReals,S)" let ?h = "OddExtension(int,IntegerAddition,IntegerOrder,{〈p,Γ(S,p)〉. p∈\<int>+})" let ?r = "OrderOnReals" from A1 have "antisym(OrderOnReals)" "S≠0" using Real_ZF_1_2_L10 IsAnOrdGroup_def IsPartOrder_def by simp; 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 show "HasAminimum(OrderOnReals,\<Inter>a∈S. OrderOnReals``{a})" by (rule Order_ZF_5_L5) qed; (*lemma (in real1) real_order_complete: assumes A1: "IsBoundedAbove(S,OrderOnReals)" "S≠0" shows "∃s. s = Supremum(OrderOnReals,S)" proof (cases "HasAmaximum(OrderOnReals,S)") assume "HasAmaximum(OrderOnReals,S)" with A1 have "Maximum(OrderOnReals,S) = Supremum(OrderOnReals,S)" using Real_ZF_1_2_L10 IsAnOrdGroup_def IsPartOrder_def Order_ZF_5_L6 by simp; then show "∃s. s = Supremum(OrderOnReals,S)" by auto; next assume A2: "¬HasAmaximum(OrderOnReals,S)" let ?h = "OddExtension(int,IntegerAddition,IntegerOrder,{〈p,Γ(S,p)〉. p∈\<int>+})" let ?r = "OrderOnReals" from A1 have "antisym(OrderOnReals)" "S≠0" using Real_ZF_1_2_L10 IsAnOrdGroup_def IsPartOrder_def by simp; 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 "[?h] = Supremum(OrderOnReals,S)" by (rule Order_ZF_5_L5); then show "∃s. s = Supremum(OrderOnReals,S)" by auto; 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.*} 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; text{* This completes the construction. It was fun.*} end;
lemma Real_ZF_1_1_L2:
[| f ∈ AlmostHoms(int, IntegerAddition); g ∈ AlmostHoms(int, IntegerAddition) |] ==> SlopeEquivalenceRel `` {SlopeOp2 ` 〈f, g〉} = SlopeEquivalenceRel `` {SlopeOp2 ` 〈g, f〉}
lemma Real_ZF_1_1_L3:
f ∈ AlmostHoms(int, IntegerAddition) ==> SlopeEquivalenceRel `` {f} ∈ RealNumbers
lemma Real_ZF_1_1_L3A:
a ∈ RealNumbers ==> ∃f∈AlmostHoms(int, IntegerAddition). a = SlopeEquivalenceRel `` {f}
lemma Real_ZF_1_1_L4:
[| f ∈ AlmostHoms(int, IntegerAddition); g ∈ AlmostHoms(int, IntegerAddition) |] ==> RealAddition ` 〈SlopeEquivalenceRel `` {f}, SlopeEquivalenceRel `` {g}〉 = SlopeEquivalenceRel `` {SlopeOp1 ` 〈f, g〉}
[| f ∈ AlmostHoms(int, IntegerAddition); g ∈ AlmostHoms(int, IntegerAddition) |] ==> RealMultiplication ` 〈SlopeEquivalenceRel `` {f}, SlopeEquivalenceRel `` {g}〉 = SlopeEquivalenceRel `` {SlopeOp2 ` 〈f, g〉}
lemma Real_ZF_1_1_L4A:
f ∈ AlmostHoms(int, IntegerAddition) ==> SlopeEquivalenceRel `` {GroupInv(int, IntegerAddition) O f} = GroupInv(RealNumbers, RealAddition) ` (SlopeEquivalenceRel `` {f})
lemma Real_ZF_1_1_L4B:
[| f ∈ AlmostHoms(int, IntegerAddition); g ∈ AlmostHoms(int, IntegerAddition) |] ==> RealAddition ` 〈SlopeEquivalenceRel `` {f}, GroupInv(RealNumbers, RealAddition) ` (SlopeEquivalenceRel `` {g})〉 = SlopeEquivalenceRel `` {SlopeOp1 ` 〈f, GroupInv(int, IntegerAddition) O g〉}
theorem real_mult_commute:
[| a ∈ RealNumbers; b ∈ RealNumbers |] ==> RealMultiplication ` 〈a, b〉 = RealMultiplication ` 〈b, a〉
lemma real_mult_commutative:
RealMultiplication {is commutative on} RealNumbers
lemma real_one_cl_identity:
SlopeEquivalenceRel `` {id(int)} = TheNeutralElement(RealNumbers, RealMultiplication)
lemma real_zero_cl_bounded_map:
f ∈ BoundedIntMaps ==> SlopeEquivalenceRel `` {f} = TheNeutralElement(RealNumbers, RealAddition)
lemma Real_ZF_1_1_L5:
[| f ∈ AlmostHoms(int, IntegerAddition); g ∈ AlmostHoms(int, IntegerAddition) |] ==> SlopeEquivalenceRel `` {f} = SlopeEquivalenceRel `` {g} <-> 〈f, g〉 ∈ SlopeEquivalenceRel
lemma Real_ZF_1_1_L5A:
〈f, g〉 ∈ SlopeEquivalenceRel ==> SlopeEquivalenceRel `` {f} = SlopeEquivalenceRel `` {g}
lemma id_on_int_is_slope:
id(int) ∈ AlmostHoms(int, IntegerAddition)
lemma Real_ZF_1_1_L7:
f ∈ BoundedIntMaps ==> 〈id(int), f〉 ∉ SlopeEquivalenceRel
lemma real_zero_not_one:
TheNeutralElement(RealNumbers, RealMultiplication) ≠ TheNeutralElement(RealNumbers, RealAddition)
lemma Real_ZF_1_1_L8:
a ∈ RealNumbers ==> GroupInv(RealNumbers, RealAddition) ` a ∈ RealNumbers
lemma Real_ZF_1_1_L9:
[| a ∈ RealNumbers; b ∈ RealNumbers; c ∈ RealNumbers |] ==> RealMultiplication ` 〈a, RealMultiplication ` 〈b, c〉〉 = RealMultiplication ` 〈RealMultiplication ` 〈a, c〉, b〉
lemma Real_ZF_1_2_L1:
PositiveSlopes ⊆ Slopes
PositiveReals ⊆ RealNumbers
lemma Real_ZF_1_2_L2:
a ∈ PositiveReals <-> (∃f∈PositiveSlopes. a = SlopeEquivalenceRel `` {f})
lemma Real_ZF_1_2_L3:
[| f ∈ PositiveSlopes; g ∈ PositiveSlopes |] ==> SlopeOp1 ` 〈f, g〉 ∈ PositiveSlopes
[| f ∈ PositiveSlopes; g ∈ PositiveSlopes |] ==> SlopeOp2 ` 〈f, g〉 ∈ PositiveSlopes
lemma Real_ZF_1_2_L5:
f ∈ BoundedIntMaps ==> f ∉ PositiveSlopes
lemma Real_ZF_1_2_L6:
PositiveReals {is closed under} RealAddition
PositiveReals {is closed under} RealMultiplication
TheNeutralElement(RealNumbers, RealAddition) ∉ PositiveReals
lemma Real_ZF_1_2_L7:
[| f ∈ AlmostHoms(int, IntegerAddition); SlopeEquivalenceRel `` {f} ≠ TheNeutralElement(RealNumbers, RealAddition) |] ==> (f ∈ PositiveSlopes) Xor (GroupInv(int, IntegerAddition) O f ∈ PositiveSlopes)
lemma Real_ZF_1_2_L8:
[| f ∈ AlmostHoms(int, IntegerAddition); g ∈ AlmostHoms(int, IntegerAddition); (f ∈ PositiveSlopes) Xor (g ∈ PositiveSlopes) |] ==> (SlopeEquivalenceRel `` {f} ∈ PositiveReals) Xor (SlopeEquivalenceRel `` {g} ∈ PositiveReals)
lemma Real_ZF_1_2_L9:
[| a ∈ RealNumbers; a ≠ TheNeutralElement(RealNumbers, RealAddition) |] ==> (a ∈ PositiveReals) Xor (GroupInv(RealNumbers, RealAddition) ` a ∈ PositiveReals)
theorem reals_are_ord_ring:
IsAnOrdRing(RealNumbers, RealAddition, RealMultiplication, OrderOnReals)
OrderOnReals {is total on} RealNumbers
PositiveSet(RealNumbers, RealAddition, OrderOnReals) = PositiveReals
HasNoZeroDivs(RealNumbers, RealAddition, RealMultiplication)
lemma Real_ZF_1_2_L10:
ring1(RealNumbers, RealAddition, RealMultiplication, OrderOnReals)
IsAnOrdGroup(RealNumbers, RealAddition, OrderOnReals)
group3(RealNumbers, RealAddition, OrderOnReals)
OrderOnReals {is total on} RealNumbers
lemma Real_ZF_1_2_L11:
[| a ∈ RealNumbers; b ∈ RealNumbers; a = b ∨ RealAddition ` 〈b, GroupInv(RealNumbers, RealAddition) ` a〉 ∈ PositiveReals |] ==> 〈a, b〉 ∈ OrderOnReals
lemma Real_ZF_1_2_L12:
[| f ∈ AlmostHoms(int, IntegerAddition); g ∈ AlmostHoms(int, IntegerAddition); 〈f, g〉 ∈ SlopeEquivalenceRel ∨ SlopeOp1 ` 〈g, GroupInv(int, IntegerAddition) O f〉 ∈ PositiveSlopes |] ==> 〈SlopeEquivalenceRel `` {f}, SlopeEquivalenceRel `` {g}〉 ∈ OrderOnReals
lemma Real_ZF_1_2_L13:
[| a ∈ RealNumbers; 〈GroupInv(RealNumbers, RealAddition) ` a, b〉 ∈ OrderOnReals |] ==> 〈GroupInv(RealNumbers, RealAddition) ` b, a〉 ∈ OrderOnReals
lemma real_ord_antisym:
[| 〈a, b〉 ∈ OrderOnReals; 〈b, a〉 ∈ OrderOnReals |] ==> a = b
lemma real_ord_transitive:
[| 〈a, b〉 ∈ OrderOnReals; 〈b, c〉 ∈ OrderOnReals |] ==> 〈a, c〉 ∈ OrderOnReals
lemma Real_ZF_1_2_L14:
[| 〈a, b〉 ∈ OrderOnReals; 〈TheNeutralElement(RealNumbers, RealAddition), c〉 ∈ OrderOnReals |] ==> 〈RealMultiplication ` 〈a, c〉, RealMultiplication ` 〈b, c〉〉 ∈ OrderOnReals
[| 〈a, b〉 ∈ OrderOnReals; 〈TheNeutralElement(RealNumbers, RealAddition), c〉 ∈ OrderOnReals |] ==> 〈RealMultiplication ` 〈c, a〉, RealMultiplication ` 〈c, b〉〉 ∈ OrderOnReals
lemma Real_ZF_1_2_L14A:
[| 〈a, b〉 ∈ OrderOnReals; c ∈ PositiveSet(RealNumbers, RealAddition, OrderOnReals) |] ==> 〈RealMultiplication ` 〈c, a〉, RealMultiplication ` 〈c, b〉〉 ∈ OrderOnReals
lemma Real_ZF_1_2_L15:
〈a, b〉 ∈ OrderOnReals ==> a ∈ RealNumbers
〈a, b〉 ∈ OrderOnReals ==> b ∈ RealNumbers
lemma Real_ZF_1_2_L16:
〈a, b〉 ∈ OrderOnReals ==> 〈TheNeutralElement(RealNumbers, RealAddition), RealAddition ` 〈b, GroupInv(RealNumbers, RealAddition) ` a〉〉 ∈ OrderOnReals
lemma Real_ZF_1_2_L17:
[| 〈TheNeutralElement(RealNumbers, RealAddition), a〉 ∈ OrderOnReals; 〈TheNeutralElement(RealNumbers, RealAddition), b〉 ∈ OrderOnReals |] ==> 〈TheNeutralElement(RealNumbers, RealAddition), RealAddition ` 〈a, b〉〉 ∈ OrderOnReals
lemma Real_ZF_1_2_L18:
[| 〈a, b〉 ∈ OrderOnReals; 〈c, d〉 ∈ OrderOnReals |] ==> 〈RealAddition ` 〈a, c〉, RealAddition ` 〈b, d〉〉 ∈ OrderOnReals
lemma real_ord_refl:
a ∈ RealNumbers ==> 〈a, a〉 ∈ OrderOnReals
lemma add_num_to_ineq:
[| 〈a, b〉 ∈ OrderOnReals; c ∈ RealNumbers |] ==> 〈RealAddition ` 〈a, c〉, RealAddition ` 〈b, c〉〉 ∈ OrderOnReals
lemma Real_ZF_1_2_L19:
[| a ∈ RealNumbers; b ∈ RealNumbers; 〈c, RealAddition ` 〈a, b〉〉 ∈ OrderOnReals |] ==> 〈RealAddition ` 〈c, GroupInv(RealNumbers, RealAddition) ` b〉, a〉 ∈ OrderOnReals
lemma Real_ZF_1_2_L20:
[| a ∈ RealNumbers; b ∈ RealNumbers; 〈a, b〉 ∉ OrderOnReals |] ==> 〈b, a〉 ∈ OrderOnReals ∧ b ≠ a
lemma Real_ZF_1_2_L21:
[| a ∈ RealNumbers; b ∈ RealNumbers; 〈c, RealAddition ` 〈a, GroupInv(RealNumbers, RealAddition) ` b〉〉 ∈ OrderOnReals |] ==> 〈RealAddition ` 〈c, b〉, a〉 ∈ OrderOnReals
lemma Real_ZF_1_2_L22:
OrderOnReals ⊆ RealNumbers × RealNumbers
lemma Real_ZF_1_2_L23:
IsBoundedAbove(A, OrderOnReals) ==> A ⊆ RealNumbers
lemma Real_ZF_1_2_L24:
[| a ∈ RealNumbers; b ∈ RealNumbers; c ∈ RealNumbers |] ==> Maximum(OrderOnReals, {a, b, c}) ∈ {a, b, c}
[| a ∈ RealNumbers; b ∈ RealNumbers; c ∈ RealNumbers |] ==> Maximum(OrderOnReals, {a, b, c}) ∈ RealNumbers
[| a ∈ RealNumbers; b ∈ RealNumbers; c ∈ RealNumbers |] ==> 〈a, Maximum(OrderOnReals, {a, b, c})〉 ∈ OrderOnReals
[| a ∈ RealNumbers; b ∈ RealNumbers; c ∈ RealNumbers |] ==> 〈b, Maximum(OrderOnReals, {a, b, c})〉 ∈ OrderOnReals
[| a ∈ RealNumbers; b ∈ RealNumbers; c ∈ RealNumbers |] ==> 〈c, Maximum(OrderOnReals, {a, b, c})〉 ∈ OrderOnReals
lemma real_strict_ord_transit:
[| 〈a, b〉 ∈ OrderOnReals; 〈b, c〉 ∈ OrderOnReals ∧ b ≠ c |] ==> 〈a, c〉 ∈ OrderOnReals ∧ a ≠ c
lemma Real_ZF_1_2_L25:
[| b ∈ PositiveSet(RealNumbers, RealAddition, OrderOnReals); 〈a, b〉 ∈ OrderOnReals; 〈TheNeutralElement(RealNumbers, RealMultiplication), c〉 ∈ OrderOnReals ∧ TheNeutralElement(RealNumbers, RealMultiplication) ≠ c |] ==> 〈a, RealMultiplication ` 〈b, c〉〉 ∈ OrderOnReals ∧ a ≠ RealMultiplication ` 〈b, c〉
lemma Real_ZF_1_2_L26:
[| a ∈ RealNumbers; b ∈ RealNumbers; 〈RealAddition ` 〈a, GroupInv(RealNumbers, RealAddition) ` b〉, c〉 ∈ OrderOnReals ∧ RealAddition ` 〈a, GroupInv(RealNumbers, RealAddition) ` b〉 ≠ c |] ==> 〈a, RealAddition ` 〈c, b〉〉 ∈ OrderOnReals ∧ a ≠ RealAddition ` 〈c, b〉
lemma real_ord_transl_inv:
[| 〈a, b〉 ∈ OrderOnReals; c ∈ RealNumbers |] ==> 〈RealAddition ` 〈c, a〉, RealAddition ` 〈c, b〉〉 ∈ OrderOnReals
lemma int_order_transitive:
[| 〈a, b〉 ∈ IntegerOrder; 〈b, c〉 ∈ IntegerOrder |] ==> 〈a, c〉 ∈ IntegerOrder
lemma Real_ZF_1_2_L27:
[| A ⊆ RealNumbers; ¬ HasAmaximum(OrderOnReals, A); x ∈ A |] ==> ∃y∈A. 〈x, y〉 ∈ OrderOnReals ∧ x ≠ y
lemma Real_ZF_1_2_L28:
[| a ∈ RealNumbers; b ∈ RealNumbers; 〈a, b〉 ∉ OrderOnReals |] ==> 〈b, a〉 ∈ OrderOnReals ∧ b ≠ a
lemma Real_ZF_1_2_L29:
〈a, b〉 ∈ OrderOnReals ∧ a ≠ b ==> 〈b, a〉 ∉ OrderOnReals
lemma pos_slopes_have_inv:
f ∈ PositiveSlopes ==> ∃g∈AlmostHoms(int, IntegerAddition). 〈f, g〉 ∈ SlopeEquivalenceRel ∧ (∃h∈AlmostHoms(int, IntegerAddition). 〈SlopeOp2 ` 〈g, h〉, id(int)〉 ∈ SlopeEquivalenceRel)
theorem reals_are_ord_field:
IsAnOrdField(RealNumbers, RealAddition, RealMultiplication, OrderOnReals)
lemma reals_are_field:
IsAfield(RealNumbers, RealAddition, RealMultiplication)
lemma field_cntxts_ok:
field0(RealNumbers, RealAddition, RealMultiplication)
field1(RealNumbers, RealAddition, RealMultiplication, OrderOnReals)
lemma Real_ZF_1_3_L1:
a ∈ PositiveSet(RealNumbers, RealAddition, OrderOnReals) ==> GroupInv (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}, restrict (RealMultiplication, (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}) × (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}))) ` a ∈ PositiveSet(RealNumbers, RealAddition, OrderOnReals)
a ∈ PositiveSet(RealNumbers, RealAddition, OrderOnReals) ==> GroupInv (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}, restrict (RealMultiplication, (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}) × (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}))) ` a ∈ RealNumbers
lemma Real_ZF_1_3_L2:
[| a ∈ PositiveSet(RealNumbers, RealAddition, OrderOnReals); 〈GroupInv (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}, restrict (RealMultiplication, (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}) × (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}))) ` a, b〉 ∈ OrderOnReals ∧ GroupInv (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}, restrict (RealMultiplication, (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}) × (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}))) ` a ≠ b |] ==> 〈TheNeutralElement(RealNumbers, RealMultiplication), RealMultiplication ` 〈b, a〉〉 ∈ OrderOnReals ∧ TheNeutralElement(RealNumbers, RealMultiplication) ≠ RealMultiplication ` 〈b, a〉
lemma Real_ZF_1_3_L3:
〈a, b〉 ∈ OrderOnReals ∧ a ≠ b ==> GroupInv (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}, restrict (RealMultiplication, (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}) × (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}))) ` (RealAddition ` 〈b, GroupInv(RealNumbers, RealAddition) ` a〉) ∈ PositiveSet(RealNumbers, RealAddition, OrderOnReals)
lemma Real_ZF_1_3_L4:
[| a ∈ RealNumbers; b ∈ PositiveSet(RealNumbers, RealAddition, OrderOnReals); 〈RealMultiplication ` 〈a, b〉, c〉 ∈ OrderOnReals ∧ RealMultiplication ` 〈a, b〉 ≠ c |] ==> 〈a, RealMultiplication ` 〈c, GroupInv (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}, restrict (RealMultiplication, (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}) × (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}))) ` b〉〉 ∈ OrderOnReals ∧ a ≠ RealMultiplication ` 〈c, GroupInv (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}, restrict (RealMultiplication, (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}) × (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}))) ` b〉
lemma Real_ZF_1_3_L4A:
[| b ∈ RealNumbers; c ∈ PositiveSet(RealNumbers, RealAddition, OrderOnReals); 〈a, RealMultiplication ` 〈b, c〉〉 ∈ OrderOnReals ∧ a ≠ RealMultiplication ` 〈b, c〉 |] ==> 〈RealMultiplication ` 〈a, GroupInv (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}, restrict (RealMultiplication, (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}) × (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}))) ` c〉, b〉 ∈ OrderOnReals ∧ RealMultiplication ` 〈a, GroupInv (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}, restrict (RealMultiplication, (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}) × (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}))) ` c〉 ≠ b
lemma Real_ZF_1_3_L4B:
[| b ∈ RealNumbers; c ∈ PositiveSet(RealNumbers, RealAddition, OrderOnReals); 〈a, RealMultiplication ` 〈b, c〉〉 ∈ OrderOnReals |] ==> 〈RealMultiplication ` 〈a, GroupInv (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}, restrict (RealMultiplication, (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}) × (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}))) ` c〉, b〉 ∈ OrderOnReals
lemma Real_ZF_1_3_L4C:
[| a ∈ RealNumbers; b ∈ PositiveSet(RealNumbers, RealAddition, OrderOnReals); 〈RealMultiplication ` 〈a, b〉, c〉 ∈ OrderOnReals |] ==> 〈a, RealMultiplication ` 〈c, GroupInv (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}, restrict (RealMultiplication, (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}) × (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}))) ` b〉〉 ∈ OrderOnReals
lemma Real_ZF_1_3_L5:
[| 〈a, b〉 ∈ OrderOnReals ∧ a ≠ b; 〈GroupInv (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}, restrict (RealMultiplication, (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}) × (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}))) ` (RealAddition ` 〈b, GroupInv(RealNumbers, RealAddition) ` a〉), c〉 ∈ OrderOnReals ∧ GroupInv (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}, restrict (RealMultiplication, (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}) × (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}))) ` (RealAddition ` 〈b, GroupInv(RealNumbers, RealAddition) ` a〉) ≠ c |] ==> 〈RealAddition ` 〈TheNeutralElement(RealNumbers, RealMultiplication), RealMultiplication ` 〈a, c〉〉, RealMultiplication ` 〈b, c〉〉 ∈ OrderOnReals ∧ RealAddition ` 〈TheNeutralElement(RealNumbers, RealMultiplication), RealMultiplication ` 〈a, c〉〉 ≠ RealMultiplication ` 〈b, c〉
lemma Real_ZF_1_3_L6:
[| 〈a, b〉 ∈ OrderOnReals; c ∈ PositiveSet(RealNumbers, RealAddition, OrderOnReals) |] ==> 〈RealMultiplication ` 〈a, GroupInv (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}, restrict (RealMultiplication, (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}) × (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}))) ` c〉, RealMultiplication ` 〈b, GroupInv (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}, restrict (RealMultiplication, (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}) × (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}))) ` c〉〉 ∈ OrderOnReals
lemma Real_ZF_1_3_L7:
[| 〈a, b〉 ∈ OrderOnReals ∧ a ≠ b; c ∈ PositiveSet(RealNumbers, RealAddition, OrderOnReals) |] ==> 〈RealMultiplication ` 〈a, c〉, RealMultiplication ` 〈b, c〉〉 ∈ OrderOnReals ∧ RealMultiplication ` 〈a, c〉 ≠ RealMultiplication ` 〈b, c〉
[| 〈a, b〉 ∈ OrderOnReals ∧ a ≠ b; c ∈ PositiveSet(RealNumbers, RealAddition, OrderOnReals) |] ==> 〈RealMultiplication ` 〈c, a〉, RealMultiplication ` 〈c, b〉〉 ∈ OrderOnReals ∧ RealMultiplication ` 〈c, a〉 ≠ RealMultiplication ` 〈c, b〉
[| 〈a, b〉 ∈ OrderOnReals ∧ a ≠ b; c ∈ PositiveSet(RealNumbers, RealAddition, OrderOnReals) |] ==> 〈RealMultiplication ` 〈a, GroupInv (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}, restrict (RealMultiplication, (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}) × (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}))) ` c〉, RealMultiplication ` 〈b, GroupInv (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}, restrict (RealMultiplication, (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}) × (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}))) ` c〉〉 ∈ OrderOnReals ∧ RealMultiplication ` 〈a, GroupInv (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}, restrict (RealMultiplication, (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}) × (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}))) ` c〉 ≠ RealMultiplication ` 〈b, GroupInv (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}, restrict (RealMultiplication, (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}) × (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}))) ` c〉
lemma Real_ZF_1_3_L8:
[| a ∈ RealNumbers; b ∈ RealNumbers; b ≠ TheNeutralElement(RealNumbers, RealAddition); c ∈ RealNumbers |] ==> RealMultiplication ` 〈RealMultiplication ` 〈a, b〉, RealMultiplication ` 〈c, GroupInv (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}, restrict (RealMultiplication, (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}) × (RealNumbers - {TheNeutralElement(RealNumbers, RealAddition)}))) ` b〉〉 = RealMultiplication ` 〈a, c〉
lemma real_int_is_real:
m ∈ int ==> SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}} ∈ RealNumbers
lemma Real_ZF_1_4_L1:
m ∈ int ==> SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈GroupInv(int, IntegerAddition) ` m, n〉〉 . n ∈ int}} = GroupInv(RealNumbers, RealAddition) ` (SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}})
lemma Real_ZF_1_4_L1A:
[| m ∈ int; k ∈ int |] ==> RealAddition ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈k, n〉〉 . n ∈ int}}〉 = SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈IntegerAddition ` 〈m, k〉, n〉〉 . n ∈ int}}
lemma Real_ZF_1_4_L1B:
[| m ∈ int; k ∈ int |] ==> RealAddition ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, GroupInv(RealNumbers, RealAddition) ` (SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈k, n〉〉 . n ∈ int}})〉 = SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈IntegerAddition ` 〈m, GroupInv(int, IntegerAddition) ` k〉, n〉〉 . n ∈ int}}
lemma Real_ZF_1_4_L1C:
[| m ∈ int; k ∈ int |] ==> RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈k, n〉〉 . n ∈ int}}〉 = SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈IntegerMultiplication ` 〈m, k〉, n〉〉 . n ∈ int}}
lemma Real_ZF_1_4_L2:
a ∈ RealNumbers ==> ∃m∈int. 〈a, SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}〉 ∈ OrderOnReals
lemma Real_ZF_1_4_L3:
a ∈ RealNumbers ==> {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, a〉 ∈ OrderOnReals} ≠ 0
lemma Real_ZF_1_4_L4:
[| m ∈ int; k ∈ int; SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}} = SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈k, n〉〉 . n ∈ int}} |] ==> m = k
lemma Real_ZF_1_4_L5:
〈m, k〉 ∈ IntegerOrder ==> 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈k, n〉〉 . n ∈ int}}〉 ∈ OrderOnReals
lemma Real_ZF_1_4_L5A:
[| 〈m, k〉 ∈ IntegerOrder; m ≠ k |] ==> 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈k, n〉〉 . n ∈ int}}〉 ∈ OrderOnReals ∧ SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}} ≠ SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈k, n〉〉 . n ∈ int}}
lemma Arthan_Lemma14i:
a ∈ RealNumbers ==> ∃n∈PositiveSet(int, IntegerAddition, IntegerOrder). 〈a, SlopeEquivalenceRel `` {{〈na, IntegerMultiplication ` 〈n, na〉〉 . na ∈ int}}〉 ∈ OrderOnReals ∧ a ≠ SlopeEquivalenceRel `` {{〈na, IntegerMultiplication ` 〈n, na〉〉 . na ∈ int}}
lemma Real_ZF_1_4_L6:
[| k ∈ int; m ∈ int; 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈k, n〉〉 . n ∈ int}}〉 ∈ OrderOnReals |] ==> 〈m, k〉 ∈ IntegerOrder
lemma Real_ZF_1_4_L7:
a ∈ RealNumbers ==> IsBoundedAbove ({m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, a〉 ∈ OrderOnReals}, IntegerOrder)
a ∈ RealNumbers ==> {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, a〉 ∈ OrderOnReals} ≠ 0
a ∈ RealNumbers ==> Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, a〉 ∈ OrderOnReals}) ∈ int
a ∈ RealNumbers ==> 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, a〉 ∈ OrderOnReals}), n〉〉 . n ∈ int}}, a〉 ∈ OrderOnReals
lemma Real_ZF_1_4_L8:
[| m ∈ int; 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, a〉 ∈ OrderOnReals |] ==> 〈m, Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, a〉 ∈ OrderOnReals})〉 ∈ IntegerOrder
lemma int_0_1_are_real_zero_one:
SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈TheNeutralElement(int, IntegerAddition), n〉〉 . n ∈ int}} = TheNeutralElement(RealNumbers, RealAddition)
SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈TheNeutralElement(int, IntegerMultiplication), n〉〉 . n ∈ int}} = TheNeutralElement(RealNumbers, RealMultiplication)
lemma int_two_is_real_two:
SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈IntegerAddition ` 〈TheNeutralElement(int, IntegerMultiplication), TheNeutralElement(int, IntegerMultiplication)〉, n〉〉 . n ∈ int}} = RealAddition ` 〈TheNeutralElement(RealNumbers, RealMultiplication), TheNeutralElement(RealNumbers, RealMultiplication)〉
lemma int_pos_is_real_pos:
p ∈ PositiveSet(int, IntegerAddition, IntegerOrder) ==> SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈p, n〉〉 . n ∈ int}} ∈ RealNumbers
p ∈ PositiveSet(int, IntegerAddition, IntegerOrder) ==> 〈TheNeutralElement(RealNumbers, RealAddition), SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈p, n〉〉 . n ∈ int}}〉 ∈ OrderOnReals
p ∈ PositiveSet(int, IntegerAddition, IntegerOrder) ==> SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈p, n〉〉 . n ∈ int}} ∈ PositiveSet(RealNumbers, RealAddition, OrderOnReals)
lemma Arthan_Lemma14ii:
[| x ∈ RealNumbers; y ∈ PositiveSet(RealNumbers, RealAddition, OrderOnReals) |] ==> ∃M∈PositiveSet(int, IntegerAddition, IntegerOrder). 〈x, RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈M, n〉〉 . n ∈ int}}, y〉〉 ∈ OrderOnReals ∧ x ≠ RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈M, n〉〉 . n ∈ int}}, y〉
lemma Real_ZF_1_4_L9:
〈a, b〉 ∈ OrderOnReals ==> 〈Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, a〉 ∈ OrderOnReals}), Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, b〉 ∈ OrderOnReals})〉 ∈ IntegerOrder
lemma Real_ZF_1_4_L10:
[| IsBoundedAbove(S, OrderOnReals); S ≠ 0; p ∈ PositiveSet(int, IntegerAddition, IntegerOrder) |] ==> IsBoundedAbove ({Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈p, n〉〉 . n ∈ int}}, x〉〉 ∈ OrderOnReals}) . x ∈ S}, IntegerOrder)
[| IsBoundedAbove(S, OrderOnReals); S ≠ 0; p ∈ PositiveSet(int, IntegerAddition, IntegerOrder) |] ==> Maximum (IntegerOrder, {Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈p, n〉〉 . n ∈ int}}, x〉〉 ∈ OrderOnReals}) . x ∈ S}) ∈ {Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈p, n〉〉 . n ∈ int}}, x〉〉 ∈ OrderOnReals}) . x ∈ S}
[| IsBoundedAbove(S, OrderOnReals); S ≠ 0; p ∈ PositiveSet(int, IntegerAddition, IntegerOrder) |] ==> Maximum (IntegerOrder, {Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈p, n〉〉 . n ∈ int}}, x〉〉 ∈ OrderOnReals}) . x ∈ S}) ∈ int
lemma Real_ZF_1_4_L11:
[| IsBoundedAbove(S, OrderOnReals); x ∈ S; p ∈ PositiveSet(int, IntegerAddition, IntegerOrder) |] ==> 〈Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈p, n〉〉 . n ∈ int}}, x〉〉 ∈ OrderOnReals}), Maximum (IntegerOrder, {Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈p, n〉〉 . n ∈ int}}, x〉〉 ∈ OrderOnReals}) . x ∈ S})〉 ∈ IntegerOrder
lemma Real_ZF_1_4_L12:
[| IsBoundedAbove(S, OrderOnReals); S ≠ 0; g = {〈p, Maximum (IntegerOrder, {Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈p, n〉〉 . n ∈ int}}, x〉〉 ∈ OrderOnReals}) . x ∈ S})〉 . p ∈ PositiveSet(int, IntegerAddition, IntegerOrder)} |] ==> g ∈ PositiveSet(int, IntegerAddition, IntegerOrder) -> int
[| IsBoundedAbove(S, OrderOnReals); S ≠ 0; g = {〈p, Maximum (IntegerOrder, {Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈p, n〉〉 . n ∈ int}}, x〉〉 ∈ OrderOnReals}) . x ∈ S})〉 . p ∈ PositiveSet(int, IntegerAddition, IntegerOrder)} |] ==> ∀n∈PositiveSet(int, IntegerAddition, IntegerOrder). g ` n = Maximum (IntegerOrder, {Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈na, IntegerMultiplication ` 〈n, na〉〉 . na ∈ int}}, x〉〉 ∈ OrderOnReals}) . x ∈ S})
lemma Real_ZF_1_4_L14:
m ∈ int ==> Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}〉 ∈ OrderOnReals}) = m
lemma floor_01_is_zero_one:
Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, TheNeutralElement(RealNumbers, RealAddition)〉 ∈ OrderOnReals}) = TheNeutralElement(int, IntegerAddition)
Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, TheNeutralElement(RealNumbers, RealMultiplication)〉 ∈ OrderOnReals}) = TheNeutralElement(int, IntegerMultiplication)
lemma floor_2_is_two:
Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealAddition ` 〈TheNeutralElement(RealNumbers, RealMultiplication), TheNeutralElement(RealNumbers, RealMultiplication)〉〉 ∈ OrderOnReals}) = IntegerAddition ` 〈TheNeutralElement(int, IntegerMultiplication), TheNeutralElement(int, IntegerMultiplication)〉
lemma Real_ZF_1_4_L14A:
[| m ∈ int; k ∈ int |] ==> Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈k, n〉〉 . n ∈ int}}〉〉 ∈ OrderOnReals}) = IntegerMultiplication ` 〈m, k〉
lemma Real_ZF_1_4_L15:
[| x ∈ RealNumbers; p ∈ int |] ==> Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealAddition ` 〈x, SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈p, n〉〉 . n ∈ int}}〉〉 ∈ OrderOnReals}) = IntegerAddition ` 〈Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, x〉 ∈ OrderOnReals}), p〉
lemma Real_ZF_1_4_L16:
[| x ∈ RealNumbers; p ∈ int |] ==> Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealAddition ` 〈x, GroupInv(RealNumbers, RealAddition) ` (SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈p, n〉〉 . n ∈ int}})〉〉 ∈ OrderOnReals}) = IntegerAddition ` 〈Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, x〉 ∈ OrderOnReals}), GroupInv(int, IntegerAddition) ` p〉
lemma Real_ZF_1_4_L17:
[| m ∈ int; n ∈ int |] ==> Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealAddition ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈n, n〉〉 . n ∈ int}}〉〉 ∈ OrderOnReals}) = IntegerAddition ` 〈m, n〉
lemma Real_ZF_1_4_L17A:
a ∈ RealNumbers ==> RealAddition ` 〈TheNeutralElement(RealNumbers, RealMultiplication), SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, a〉 ∈ OrderOnReals}), n〉〉 . n ∈ int}}〉 = SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈IntegerAddition ` 〈TheNeutralElement(int, IntegerMultiplication), Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, a〉 ∈ OrderOnReals})〉, n〉〉 . n ∈ int}}
lemma Real_ZF_1_4_L17B:
a ∈ RealNumbers ==> 〈RealAddition ` 〈a, GroupInv(RealNumbers, RealAddition) ` (SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, a〉 ∈ OrderOnReals}), n〉〉 . n ∈ int}})〉, TheNeutralElement(RealNumbers, RealMultiplication)〉 ∈ OrderOnReals ∧ RealAddition ` 〈a, GroupInv(RealNumbers, RealAddition) ` (SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, a〉 ∈ OrderOnReals}), n〉〉 . n ∈ int}})〉 ≠ TheNeutralElement(RealNumbers, RealMultiplication)
a ∈ RealNumbers ==> 〈a, SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈IntegerAddition ` 〈TheNeutralElement(int, IntegerMultiplication), Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, a〉 ∈ OrderOnReals})〉, n〉〉 . n ∈ int}}〉 ∈ OrderOnReals ∧ a ≠ SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈IntegerAddition ` 〈TheNeutralElement(int, IntegerMultiplication), Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, a〉 ∈ OrderOnReals})〉, n〉〉 . n ∈ int}}
lemma Arthan_Lemma14iii:
〈x, y〉 ∈ OrderOnReals ∧ x ≠ y ==> ∃M∈int. ∃N∈PositiveSet(int, IntegerAddition, IntegerOrder). (〈RealMultiplication ` 〈x, SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈N, n〉〉 . n ∈ int}}〉, SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈M, n〉〉 . n ∈ int}}〉 ∈ OrderOnReals ∧ RealMultiplication ` 〈x, SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈N, n〉〉 . n ∈ int}}〉 ≠ SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈M, n〉〉 . n ∈ int}}) ∧ 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈M, n〉〉 . n ∈ int}}, RealMultiplication ` 〈y, SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈N, n〉〉 . n ∈ int}}〉〉 ∈ OrderOnReals ∧ SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈M, n〉〉 . n ∈ int}} ≠ RealMultiplication ` 〈y, SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈N, n〉〉 . n ∈ int}}〉
lemma Real_ZF_1_4_L18:
[| x ∈ RealNumbers; y ∈ RealNumbers |] ==> 〈AbsoluteValue(int, IntegerAddition, IntegerOrder) ` (IntegerAddition ` 〈IntegerAddition ` 〈Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealAddition ` 〈x, y〉〉 ∈ OrderOnReals}), GroupInv(int, IntegerAddition) ` Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, x〉 ∈ OrderOnReals})〉, GroupInv(int, IntegerAddition) ` Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, y〉 ∈ OrderOnReals})〉), IntegerAddition ` 〈TheNeutralElement(int, IntegerMultiplication), TheNeutralElement(int, IntegerMultiplication)〉〉 ∈ IntegerOrder
lemma Real_ZF_1_4_L20:
[| IsBoundedAbove(S, OrderOnReals); S ≠ 0; n ∈ PositiveSet(int, IntegerAddition, IntegerOrder); x ∈ S; Maximum (IntegerOrder, {Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈n, n〉〉 . n ∈ int}}, x〉〉 ∈ OrderOnReals}) . x ∈ S}) = Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈n, n〉〉 . n ∈ int}}, x〉〉 ∈ OrderOnReals}); y ∈ S; 〈x, y〉 ∈ OrderOnReals |] ==> Maximum (IntegerOrder, {Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈n, n〉〉 . n ∈ int}}, x〉〉 ∈ OrderOnReals}) . x ∈ S}) = Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈n, n〉〉 . n ∈ int}}, y〉〉 ∈ OrderOnReals})
lemma Real_ZF_1_4_L21:
[| IsBoundedAbove(S, OrderOnReals); S ≠ 0; m ∈ PositiveSet(int, IntegerAddition, IntegerOrder); n ∈ PositiveSet(int, IntegerAddition, IntegerOrder) |] ==> 〈AbsoluteValue(int, IntegerAddition, IntegerOrder) ` (IntegerAddition ` 〈IntegerAddition ` 〈Maximum (IntegerOrder, {Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈IntegerAddition ` 〈m, n〉, n〉〉 . n ∈ int}}, x〉〉 ∈ OrderOnReals}) . x ∈ S}), GroupInv(int, IntegerAddition) ` Maximum (IntegerOrder, {Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, x〉〉 ∈ OrderOnReals}) . x ∈ S})〉, GroupInv(int, IntegerAddition) ` Maximum (IntegerOrder, {Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈n, n〉〉 . n ∈ int}}, x〉〉 ∈ OrderOnReals}) . x ∈ S})〉), IntegerAddition ` 〈TheNeutralElement(int, IntegerMultiplication), TheNeutralElement(int, IntegerMultiplication)〉〉 ∈ IntegerOrder
lemma Real_ZF_1_4_L21A:
[| f ∈ PositiveSet(int, IntegerAddition, IntegerOrder) -> int; ∀a∈PositiveSet(int, IntegerAddition, IntegerOrder). ∀b∈PositiveSet(int, IntegerAddition, IntegerOrder). 〈AbsoluteValue(int, IntegerAddition, IntegerOrder) ` (IntegerAddition ` 〈IntegerAddition ` 〈f ` (IntegerAddition ` 〈a, b〉), GroupInv(int, IntegerAddition) ` (f ` a)〉, GroupInv(int, IntegerAddition) ` (f ` b)〉), L〉 ∈ IntegerOrder |] ==> OddExtension(int, IntegerAddition, IntegerOrder, f) ∈ AlmostHoms(int, IntegerAddition)
lemma Real_ZF_1_4_L22:
[| IsBoundedAbove(S, OrderOnReals); S ≠ 0; g = {〈p, Maximum (IntegerOrder, {Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈p, n〉〉 . n ∈ int}}, x〉〉 ∈ OrderOnReals}) . x ∈ S})〉 . p ∈ PositiveSet(int, IntegerAddition, IntegerOrder)} |] ==> OddExtension(int, IntegerAddition, IntegerOrder, g) ∈ AlmostHoms(int, IntegerAddition)
lemma Real_ZF_1_4_L23:
[| f ∈ AlmostHoms(int, IntegerAddition); N ∈ int; M ∈ int; ∀n∈PositiveSet(int, IntegerAddition, IntegerOrder). 〈IntegerMultiplication ` 〈M, n〉, f ` (IntegerMultiplication ` 〈N, n〉)〉 ∈ IntegerOrder |] ==> 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈M, n〉〉 . n ∈ int}}, RealMultiplication ` 〈SlopeEquivalenceRel `` {f}, SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈N, n〉〉 . n ∈ int}}〉〉 ∈ OrderOnReals
lemma Real_ZF_1_4_L23A:
[| f ∈ AlmostHoms(int, IntegerAddition); N ∈ int; M ∈ int; ∀n∈PositiveSet(int, IntegerAddition, IntegerOrder). 〈f ` (IntegerMultiplication ` 〈N, n〉), IntegerMultiplication ` 〈M, n〉〉 ∈ IntegerOrder |] ==> 〈RealMultiplication ` 〈SlopeEquivalenceRel `` {f}, SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈N, n〉〉 . n ∈ int}}〉, SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈M, n〉〉 . n ∈ int}}〉 ∈ OrderOnReals
lemma Real_ZF_1_4_L24:
[| IsBoundedAbove(S, OrderOnReals); 〈x, y〉 ∈ OrderOnReals ∧ x ≠ y; y ∈ S; N ∈ PositiveSet(int, IntegerAddition, IntegerOrder); M ∈ int; 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈M, n〉〉 . n ∈ int}}, RealMultiplication ` 〈y, SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈N, n〉〉 . n ∈ int}}〉〉 ∈ OrderOnReals ∧ SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈M, n〉〉 . n ∈ int}} ≠ RealMultiplication ` 〈y, SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈N, n〉〉 . n ∈ int}}〉; p ∈ PositiveSet(int, IntegerAddition, IntegerOrder) |] ==> 〈IntegerMultiplication ` 〈p, M〉, Maximum (IntegerOrder, {Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈IntegerMultiplication ` 〈p, N〉, n〉〉 . n ∈ int}}, x〉〉 ∈ OrderOnReals}) . x ∈ S})〉 ∈ IntegerOrder
lemma Real_ZF_1_4_L24A:
[| IsBoundedAbove(S, OrderOnReals); S ≠ 0; p ∈ PositiveSet(int, IntegerAddition, IntegerOrder); h = OddExtension (int, IntegerAddition, IntegerOrder, {〈p, Maximum (IntegerOrder, {Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈p, n〉〉 . n ∈ int}}, x〉〉 ∈ OrderOnReals}) . x ∈ S})〉 . p ∈ PositiveSet(int, IntegerAddition, IntegerOrder)}) |] ==> h ` p = Maximum (IntegerOrder, {Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈p, n〉〉 . n ∈ int}}, x〉〉 ∈ OrderOnReals}) . x ∈ S})
lemma Real_ZF_1_4_L25:
[| IsBoundedAbove(S, OrderOnReals); ¬ HasAmaximum(OrderOnReals, S); x ∈ S; h = OddExtension (int, IntegerAddition, IntegerOrder, {〈p, Maximum (IntegerOrder, {Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈p, n〉〉 . n ∈ int}}, x〉〉 ∈ OrderOnReals}) . x ∈ S})〉 . p ∈ PositiveSet(int, IntegerAddition, IntegerOrder)}) |] ==> 〈x, SlopeEquivalenceRel `` {h}〉 ∈ OrderOnReals
lemma Real_ZF_1_4_L26:
[| IsBoundedAbove(S, OrderOnReals); 〈x, y〉 ∈ OrderOnReals; x ∈ S; N ∈ PositiveSet(int, IntegerAddition, IntegerOrder); M ∈ int; 〈RealMultiplication ` 〈y, SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈N, n〉〉 . n ∈ int}}〉, SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈M, n〉〉 . n ∈ int}}〉 ∈ OrderOnReals ∧ RealMultiplication ` 〈y, SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈N, n〉〉 . n ∈ int}}〉 ≠ SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈M, n〉〉 . n ∈ int}}; p ∈ PositiveSet(int, IntegerAddition, IntegerOrder) |] ==> 〈Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈IntegerMultiplication ` 〈N, p〉, n〉〉 . n ∈ int}}, x〉〉 ∈ OrderOnReals}), IntegerMultiplication ` 〈M, p〉〉 ∈ IntegerOrder
lemma Real_ZF_1_4_L27:
[| IsBoundedAbove(S, OrderOnReals); S ≠ 0; h = OddExtension (int, IntegerAddition, IntegerOrder, {〈p, Maximum (IntegerOrder, {Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈p, n〉〉 . n ∈ int}}, x〉〉 ∈ OrderOnReals}) . x ∈ S})〉 . p ∈ PositiveSet(int, IntegerAddition, IntegerOrder)}); p ∈ PositiveSet(int, IntegerAddition, IntegerOrder) |] ==> ∃x∈S. h ` p = Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈p, n〉〉 . n ∈ int}}, x〉〉 ∈ OrderOnReals})
lemma Real_ZF_1_4_L28:
[| IsBoundedAbove(S, OrderOnReals); S ≠ 0; ∀x∈S. 〈x, y〉 ∈ OrderOnReals; h = OddExtension (int, IntegerAddition, IntegerOrder, {〈p, Maximum (IntegerOrder, {Maximum (IntegerOrder, {m ∈ int . 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈m, n〉〉 . n ∈ int}}, RealMultiplication ` 〈SlopeEquivalenceRel `` {{〈n, IntegerMultiplication ` 〈p, n〉〉 . n ∈ int}}, x〉〉 ∈ OrderOnReals}) . x ∈ S})〉 . p ∈ PositiveSet(int, IntegerAddition, IntegerOrder)}) |] ==> 〈SlopeEquivalenceRel `` {h}, y〉 ∈ OrderOnReals
lemma real_order_complete:
[| IsBoundedAbove(S, OrderOnReals); S ≠ 0 |] ==> HasAminimum(OrderOnReals, \<Inter>a∈S. OrderOnReals `` {a})
theorem eudoxus_reals_are_reals:
IsAmodelOfReals(RealNumbers, RealAddition, RealMultiplication, OrderOnReals)