Theory Int_ZF_2

(*    This file is a part of IsarMathLib - 
    a library of formalized mathematics for Isabelle/Isar.

    Copyright (C) 2005 - 2009  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.

THIS 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.

*)

section Integers 2

theory Int_ZF_2 imports func_ZF_1 Int_ZF_1 IntDiv_ZF_IML Group_ZF_3

begin

textIn this theory file we consider the properties of integers that are 
  needed for the real numbers construction in Real_ZF› series.

subsectionSlopes

textIn this section we study basic properties of slopes - the integer 
  almost homomorphisms. 
  The general definition of an almost homomorphism $f$ on a group $G$ 
  written in additive notation requires the set 
  $\{f(m+n) - f(m) - f(n): m,n\in G\}$ to be finite.
  In this section we establish a definition that is equivalent for integers: 
  that for all integer $m,n$ we have $|f(m+n) - f(m) - f(n)| \leq L$ for
  some $L$.

textFirst we extend the standard notation for integers with notation related
  to slopes. We define slopes as almost homomorphisms on the additive group of
  integers. The set of slopes is denoted 𝒮›. We also define "positive" 
  slopes as those that take infinite number of positive values on positive integers.
  We write δ(s,m,n)› to denote the homomorphism 
  difference of $s$ at $m,n$ (i.e the expression $s(m+n) - s(m) - s(n)$).
  We denote maxδ(s)› the maximum absolute value of homomorphism 
  difference of $s$ as $m,n$ range over integers.
  If $s$ is a slope, 
  then the set of homomorphism differences is finite and 
  this maximum exists.
  In Group_ZF_3› we define the equivalence relation on 
  almost homomorphisms using the notion of a quotient group relation
  and use "≈›" to denote it. As here this symbol seems to be hogged
  by the standard Isabelle, we will use "∼›" instead "≈›".
  We show in this section that $s\sim r$ iff for some $L$ we 
  have $|s(m)-r(m)| \leq L$ for all integer $m$.
  The "\<fp>›" denotes the first operation on almost homomorphisms. 
  For slopes this is addition of functions defined in the natural way.
  The "∘›" symbol denotes the second operation on almost homomorphisms
  (see Group_ZF_3› for definition), 
  defined for the group of integers.  In short "∘›" 
  is the composition of slopes.
  The "¯›" symbol acts as an infix operator that assigns the value 
  $\min\{n\in Z_+: p\leq f(n)\}$ to a pair (of sets) $f$ and $p$. 
  In application
  $f$ represents a function defined on $Z_+$ and $p$ is a positive integer.
  We choose this notation because we use it to construct the right inverse 
  in the ring of classes of slopes and show that this ring is in fact a field.
  To study the homomorphism difference of the function defined by $p\mapsto f^{-1}(p)$
  we introduce the symbol ε› defined as 
  $\varepsilon(f,\langle m,n \rangle ) = f^{-1}(m+n)-f^{-1}(m)-f^{-1}(n)$. Of course the intention is
  to use the fact that $\varepsilon(f,\langle m,n \rangle )$ is the homomorphism difference
  of the function $g$ defined as $g(m) = f^{-1}(m)$. We also define $\gamma (s,m,n)$ as the 
  expression $\delta(f,m,-n)+s(0)-\delta (f,n,-n)$. This is useful because of the identity 
  $f(m-n) = \gamma (m,n) + f(m)-f(n)$ that allows to obtain bounds on the value of a slope
  at the difference of of two integers.
  For every integer $m$ we introduce notation $m^S$ defined by $m^E(n)=m\cdot n$. The mapping
  $q\mapsto q^S$ embeds integers into 𝒮› preserving the order, (that is, 
  maps positive integers into 𝒮+).

locale int1 = int0 +

  fixes slopes ("𝒮" )
  defines slopes_def[simp]: "𝒮  AlmostHoms(,IntegerAddition)"

  fixes posslopes ("𝒮+")
  defines posslopes_def[simp]: "𝒮+  {s𝒮. s``(+)  +  Fin()}"
 
  fixes δ 
  defines δ_def[simp]: "δ(s,m,n)  s`(m\<ra>n)\<rs>s`(m)\<rs>s`(n)"

  fixes maxhomdiff ("maxδ" )
  defines maxhomdiff_def[simp]: 
  "maxδ(s)  Maximum(IntegerOrder,{abs(δ(s,m,n)).  m,n  ×})"

  fixes AlEqRel
  defines AlEqRel_def[simp]: 
  "AlEqRel  QuotientGroupRel(𝒮,AlHomOp1(,IntegerAddition),FinRangeFunctions(,))"

  fixes AlEq (infix "" 68)
  defines AlEq_def[simp]: "s  r   s,r  AlEqRel"

  fixes slope_add (infix "\<fp>" 70)
  defines slope_add_def[simp]: "s \<fp> r   AlHomOp1(,IntegerAddition)` s,r"

  fixes slope_comp (infix "" 70)
  defines slope_comp_def[simp]: "s  r   AlHomOp2(,IntegerAddition)` s,r"

  fixes neg ("\<fm>_" [90] 91)
  defines neg_def[simp]: "\<fm>s  GroupInv(,IntegerAddition) O s"

  fixes slope_inv (infix "¯" 71)
  defines slope_inv_def[simp]: 
  "f¯(p)  Minimum(IntegerOrder,{n+. p \<lsq> f`(n)})"
  fixes ε
  defines ε_def[simp]: 
  "ε(f,p)  f¯(fst(p)\<ra>snd(p)) \<rs> f¯(fst(p)) \<rs> f¯(snd(p))"

  fixes γ 
  defines γ_def[simp]:
  "γ(s,m,n)  δ(s,m,\<rm>n) \<rs> δ(s,n,\<rm>n) \<ra> s`(𝟬)"

  fixes intembed ("_S")
  defines intembed_def[simp]: "mS  {n,mn. n}"

textWe can use theorems proven in the group1› context.

lemma (in int1) Int_ZF_2_1_L1: shows "group1(,IntegerAddition)"
  using Int_ZF_1_T2 group1_axioms.intro group1_def by simp

textType information related to the homomorphism difference expression.

lemma (in int1) Int_ZF_2_1_L2: assumes "f𝒮" and "n" "m"
  shows 
  "m\<ra>n  "  
  "f`(m\<ra>n)  "  
  "f`(m)  "   "f`(n)  "
  "f`(m) \<ra> f`(n)  "  
  "HomDiff(,IntegerAddition,f, m,n)  " 
  using assms Int_ZF_2_1_L1 group1.Group_ZF_3_2_L4A
  by auto

textType information related to the homomorphism difference expression.

lemma (in int1) Int_ZF_2_1_L2A: 
  assumes "f:" and "n"  "m"
  shows 
  "m\<ra>n  " 
  "f`(m\<ra>n)  "   "f`(m)  "   "f`(n)  "
  "f`(m) \<ra> f`(n)  " 
  "HomDiff(,IntegerAddition,f, m,n)  "
  using assms Int_ZF_2_1_L1 group1.Group_ZF_3_2_L4
  by auto

textSlopes map integers into integers.

lemma (in int1) Int_ZF_2_1_L2B: 
  assumes A1: "f𝒮" and A2: "m" 
  shows "f`(m)  "
proof -
  from A1 have "f:" using AlmostHoms_def by simp
  with A2 show "f`(m)  " using apply_funtype by simp
qed

textThe homomorphism difference in multiplicative notation is defined as
  the expression $s(m\cdot n)\cdot(s(m)\cdot s(n))^{-1}$. The next lemma 
  shows that 
  in the additive notation used for integers the homomorphism 
  difference is $f(m+n) - f(m) - f(n)$ which we denote as δ(f,m,n)›.

lemma (in int1) Int_ZF_2_1_L3: 
  assumes "f:" and "m"  "n"
  shows "HomDiff(,IntegerAddition,f, m,n) = δ(f,m,n)"
  using assms Int_ZF_2_1_L2A Int_ZF_1_T2 group0.group0_4_L4A 
    HomDiff_def by auto

textThe next formula restates the definition of the homomorphism 
  difference to express the value an almost homomorphism on a sum.

lemma (in int1) Int_ZF_2_1_L3A: 
  assumes A1: "f𝒮" and A2: "m"  "n"
  shows 
  "f`(m\<ra>n) = f`(m)\<ra>(f`(n)\<ra>δ(f,m,n))"
proof -
  from A1 A2 have
    T: "f`(m) "  "f`(n)  "  "δ(f,m,n)  " and
    "HomDiff(,IntegerAddition,f, m,n) = δ(f,m,n)"  
    using Int_ZF_2_1_L2 AlmostHoms_def Int_ZF_2_1_L3 by auto
  with A1 A2 show  "f`(m\<ra>n) = f`(m)\<ra>(f`(n)\<ra>δ(f,m,n))" 
    using Int_ZF_2_1_L3 Int_ZF_1_L3 
      Int_ZF_2_1_L1 group1.Group_ZF_3_4_L1 
    by simp
qed

textThe homomorphism difference of any integer function is integer.

lemma (in int1) Int_ZF_2_1_L3B: 
  assumes "f:" and "m"  "n"
  shows "δ(f,m,n)  "
  using assms Int_ZF_2_1_L2A Int_ZF_2_1_L3 by simp

textThe value of an integer function at a sum expressed in 
  terms of δ›.

lemma (in int1) Int_ZF_2_1_L3C: assumes A1: "f:" and A2: "m"  "n"
  shows "f`(m\<ra>n) = δ(f,m,n) \<ra> f`(n) \<ra> f`(m)"
proof -
  from A1 A2 have T:
    "δ(f,m,n)  "  "f`(m\<ra>n)  "  "f`(m)  "  "f`(n)  "
    using Int_ZF_1_1_L5 apply_funtype by auto
  then show "f`(m\<ra>n) = δ(f,m,n) \<ra> f`(n) \<ra> f`(m)"
    using Int_ZF_1_2_L15 by simp
qed


textThe next lemma presents two ways the set of homomorphism differences
  can be written.

lemma (in int1) Int_ZF_2_1_L4: assumes A1: "f:"
  shows "{abs(HomDiff(,IntegerAddition,f,x)). x  ×} =
  {abs(δ(f,m,n)).  m,n  ×}"
proof -
  from A1 have "m. n. 
    abs(HomDiff(,IntegerAddition,f, m,n)) = abs(δ(f,m,n))"
    using Int_ZF_2_1_L3 by simp
  then show ?thesis by (rule ZF1_1_L4A)
qed

textIf $f$ maps integers into integers and 
  for all $m,n\in Z$ we have $|f(m+n) - f(m) - f(n)| \leq L$ for some $L$,
  then $f$ is a slope.

lemma (in int1) Int_ZF_2_1_L5: assumes A1: "f:"
  and A2: "m.n. abs(δ(f,m,n)) \<lsq> L"
  shows "f𝒮"
proof -
  let ?Abs = "AbsoluteValue(,IntegerAddition,IntegerOrder)"
  have "group3(,IntegerAddition,IntegerOrder)" 
    "IntegerOrder {is total on} "
    using Int_ZF_2_T1 by auto
  moreover from A1 A2 have 
    "x×. HomDiff(,IntegerAddition,f,x)   
    ?Abs`(HomDiff(,IntegerAddition,f,x)),L   IntegerOrder"
    using Int_ZF_2_1_L2A Int_ZF_2_1_L3 by auto
  ultimately have 
    "IsBounded({HomDiff(,IntegerAddition,f,x). x×},IntegerOrder)"
    by (rule group3.OrderedGroup_ZF_3_L9A)
  with A1 show "f  𝒮" using Int_bounded_iff_fin AlmostHoms_def
    by simp
qed

textThe absolute value of homomorphism difference 
  of a slope $s$ does not exceed maxδ(s)›.

lemma (in int1) Int_ZF_2_1_L7: 
  assumes A1: "s𝒮" and A2: "n"  "m"
  shows 
  "abs(δ(s,m,n)) \<lsq> maxδ(s)"   
  "δ(s,m,n)  "   "maxδ(s)  "
  "(\<rm>maxδ(s)) \<lsq> δ(s,m,n)"
proof -
  from A1 A2 show T: "δ(s,m,n)  "
    using Int_ZF_2_1_L2 Int_ZF_1_1_L5 by simp
  let ?A = "{abs(HomDiff(,IntegerAddition,s,x)). x×}"
  let ?B = "{abs(δ(s,m,n)).  m,n  ×}"
  let ?d = "abs(δ(s,m,n))"
  have "IsLinOrder(,IntegerOrder)" using Int_ZF_2_T1
    by simp
  moreover have "?A  Fin()" 
  proof -
    have "k. abs(k)  " using Int_ZF_2_L14 by simp
    moreover from A1 have 
      "{HomDiff(,IntegerAddition,s,x). x  ×}  Fin()"
      using AlmostHoms_def by simp
    ultimately show "?A  Fin()" by (rule Finite1_L6C)
  qed
  moreover have "?A0" by auto
  ultimately have "k?A. k,Maximum(IntegerOrder,?A)  IntegerOrder"
    by (rule Finite_ZF_1_T2)
  moreover from A1 A2 have "?d?A" using AlmostHoms_def Int_ZF_2_1_L4
    by auto
  ultimately have "?d \<lsq> Maximum(IntegerOrder,?A)" by auto 
  with A1 show "?d \<lsq> maxδ(s)"  "maxδ(s)  "
    using AlmostHoms_def Int_ZF_2_1_L4 Int_ZF_2_L1A 
    by auto
  with T show "(\<rm>maxδ(s)) \<lsq> δ(s,m,n)"
    using Int_ZF_1_3_L19 by simp
qed

textA useful estimate for the value of a slope at $0$, plus some type information
  for slopes.

lemma (in int1) Int_ZF_2_1_L8: assumes A1: "s𝒮"
  shows 
  "abs(s`(𝟬)) \<lsq> maxδ(s)"
  "𝟬 \<lsq> maxδ(s)"  
  "abs(s`(𝟬))  "   "maxδ(s)  "
  "abs(s`(𝟬)) \<ra> maxδ(s)  "
proof -
  from A1 have "s`(𝟬)  " 
    using int_zero_one_are_int Int_ZF_2_1_L2B by simp
  then have I: "𝟬 \<lsq> abs(s`(𝟬))"  
    and "abs(δ(s,𝟬,𝟬)) = abs(s`(𝟬))" 
    using int_abs_nonneg int_zero_one_are_int Int_ZF_1_1_L4 
      Int_ZF_2_L17 by auto
  moreover from A1 have "abs(δ(s,𝟬,𝟬)) \<lsq> maxδ(s)"
    using int_zero_one_are_int Int_ZF_2_1_L7 by simp
  ultimately show II: "abs(s`(𝟬)) \<lsq> maxδ(s)"
    by simp
  with I show "𝟬\<lsq>maxδ(s)" by (rule Int_order_transitive)
  with II show 
    "maxδ(s)  "   "abs(s`(𝟬))  " 
    "abs(s`(𝟬)) \<ra> maxδ(s)  "
    using Int_ZF_2_L1A Int_ZF_1_1_L5 by auto
qed

textInt Group_ZF_3.thy› we show that finite range functions 
  valued in an abelian group 
  form a normal subgroup of almost homomorphisms. 
  This allows to define the equivalence relation 
  between almost homomorphisms as the relation resulting from dividing 
  by that normal subgroup. 
  Then we show in Group_ZF_3_4_L12› that if the difference of $f$ and $g$ 
  has finite range (actually $f(n)\cdot g(n)^{-1}$ as we use multiplicative 
  notation 
  in Group_ZF_3.thy›), then $f$ and $g$ are equivalent.
  The next lemma translates that fact into the notation used in int1› 
  context.

lemma (in int1) Int_ZF_2_1_L9: assumes A1: "s𝒮"  "r𝒮"
  and A2: "m. abs(s`(m)\<rs>r`(m)) \<lsq> L"
  shows "s  r"
proof -
  from A1 A2 have 
    "m. s`(m)\<rs>r`(m)    abs(s`(m)\<rs>r`(m)) \<lsq> L"
    using Int_ZF_2_1_L2B Int_ZF_1_1_L5 by simp
  then have
    "IsBounded({s`(n)\<rs>r`(n). n}, IntegerOrder)"
    by (rule Int_ZF_1_3_L20)
  with A1 show "s  r" using Int_bounded_iff_fin 
    Int_ZF_2_1_L1 group1.Group_ZF_3_4_L12 by simp
qed

textA neccessary condition for two slopes to be almost equal. 
  For slopes the definition postulates the 
  set $\{f(m)-g(m): m\in Z\}$ to be finite. 
  This lemma shows that this implies that
  $|f(m)-g(m)|$ is bounded (by some integer) as $m$ varies over integers.
  We also mention here that in this context s ∼ r› implies that both
  $s$ and $r$ are slopes.

lemma (in int1) Int_ZF_2_1_L9A: assumes "s  r" 
  shows 
  "L. m. abs(s`(m)\<rs>r`(m)) \<lsq> L"
  "s𝒮"  "r𝒮"
  using assms Int_ZF_2_1_L1 group1.Group_ZF_3_4_L11 
    Int_ZF_1_3_L20AA QuotientGroupRel_def by auto

textLet's recall that the relation of almost equality is an equivalence relation
  on the set of slopes.

lemma (in int1) Int_ZF_2_1_L9B: shows
  "AlEqRel  𝒮×𝒮"
  "equiv(𝒮,AlEqRel)"
  using Int_ZF_2_1_L1 group1.Group_ZF_3_3_L3 by auto

textAnother version of sufficient condition for two slopes to be almost
  equal: if the difference of two slopes is a finite range function, then
  they are almost equal.

lemma (in int1) Int_ZF_2_1_L9C: assumes "s𝒮"  "r𝒮" and 
  "s \<fp> (\<fm>r)  FinRangeFunctions(,)"
  shows  
  "s  r"  
  "r  s"
  using assms Int_ZF_2_1_L1 
    group1.Group_ZF_3_2_L13 group1.Group_ZF_3_4_L12A
  by auto


textIf two slopes are almost equal, then the difference has finite range.
  This is the inverse of Int_ZF_2_1_L9C›.

lemma (in int1) Int_ZF_2_1_L9D: assumes A1: "s  r"
  shows "s \<fp> (\<fm>r)  FinRangeFunctions(,)"
proof -
  let ?G = ""
  let ?f = "IntegerAddition"
  from A1 have "AlHomOp1(?G, ?f)`s,GroupInv(AlmostHoms(?G, ?f),AlHomOp1(?G, ?f))`(r) 
     FinRangeFunctions(?G, ?G)"
    using Int_ZF_2_1_L1 group1.Group_ZF_3_4_L12B by auto
  with A1 show "s \<fp> (\<fm>r)  FinRangeFunctions(,)"
    using Int_ZF_2_1_L9A Int_ZF_2_1_L1 group1.Group_ZF_3_2_L13
    by simp
qed
  
textWhat is the value of a composition of slopes?

lemma (in int1) Int_ZF_2_1_L10: 
  assumes "s𝒮"  "r𝒮" and "m"
  shows "(sr)`(m) = s`(r`(m))"  "s`(r`(m))  "
  using assms Int_ZF_2_1_L1 group1.Group_ZF_3_4_L2 by auto

textComposition of slopes is a slope.

lemma (in int1) Int_ZF_2_1_L11:
  assumes "s𝒮"  "r𝒮"
  shows "sr  𝒮"
  using assms Int_ZF_2_1_L1 group1.Group_ZF_3_4_T1 by simp

textNegative of a slope is a slope.

lemma (in int1) Int_ZF_2_1_L12: assumes "s𝒮" shows "\<fm>s  𝒮"
  using assms Int_ZF_1_T2 Int_ZF_2_1_L1 group1.Group_ZF_3_2_L13 
  by simp

textWhat is the value of a negative of a slope?

lemma (in int1) Int_ZF_2_1_L12A: 
  assumes "s𝒮" and "m" shows "(\<fm>s)`(m) = \<rm>(s`(m))"
  using assms Int_ZF_2_1_L1 group1.Group_ZF_3_2_L5
  by simp

textWhat are the values of a sum of slopes?

lemma (in int1) Int_ZF_2_1_L12B: assumes "s𝒮"  "r𝒮" and "m"
  shows "(s\<fp>r)`(m) = s`(m) \<ra> r`(m)"
  using assms Int_ZF_2_1_L1 group1.Group_ZF_3_2_L12
  by simp

textSum of slopes is a slope.

lemma (in int1) Int_ZF_2_1_L12C: assumes "s𝒮"  "r𝒮"
  shows "s\<fp>r  𝒮"
  using assms Int_ZF_2_1_L1 group1.Group_ZF_3_2_L16
  by simp

textA simple but useful identity.

lemma (in int1) Int_ZF_2_1_L13: 
  assumes "s𝒮" and "n"  "m"
  shows "s`(nm) \<ra> (s`(m) \<ra> δ(s,nm,m)) = s`((n\<ra>𝟭)m)"
  using assms Int_ZF_1_1_L5 Int_ZF_2_1_L2B Int_ZF_1_2_L9 Int_ZF_1_2_L7
  by simp

textSome estimates for the absolute value of a slope at the opposite 
  integer.

lemma (in int1) Int_ZF_2_1_L14: assumes A1: "s𝒮" and A2: "m"
  shows 
  "s`(\<rm>m) = s`(𝟬) \<rs> δ(s,m,\<rm>m) \<rs> s`(m)"
  "abs(s`(m)\<ra>s`(\<rm>m)) \<lsq> 𝟮maxδ(s)"
  "abs(s`(\<rm>m)) \<lsq> 𝟮maxδ(s) \<ra> abs(s`(m))"
  "s`(\<rm>m) \<lsq> abs(s`(𝟬)) \<ra> maxδ(s) \<rs> s`(m)"
proof -
  from A1 A2 have T:
    "(\<rm>m)  "  "abs(s`(m))  "  "s`(𝟬)  "  "abs(s`(𝟬))  "  
    "δ(s,m,\<rm>m)  "   "s`(m)  "   "s`(\<rm>m)  "  
    "(\<rm>(s`(m)))  "  "s`(𝟬) \<rs> δ(s,m,\<rm>m)  "
    using Int_ZF_1_1_L4 Int_ZF_2_1_L2B Int_ZF_2_L14 Int_ZF_2_1_L2 
      Int_ZF_1_1_L5 int_zero_one_are_int by auto
  with A2 show I: "s`(\<rm>m) = s`(𝟬) \<rs> δ(s,m,\<rm>m) \<rs> s`(m)"
    using Int_ZF_1_1_L4 Int_ZF_1_2_L15 by simp
  from T have "abs(s`(𝟬) \<rs> δ(s,m,\<rm>m)) \<lsq> abs(s`(𝟬)) \<ra> abs(δ(s,m,\<rm>m))"
    using Int_triangle_ineq1 by simp
  moreover from A1 A2 T have "abs(s`(𝟬)) \<ra> abs(δ(s,m,\<rm>m)) \<lsq>  𝟮maxδ(s)"
    using Int_ZF_2_1_L7 Int_ZF_2_1_L8 Int_ZF_1_3_L21 by simp
  ultimately have "abs(s`(𝟬) \<rs> δ(s,m,\<rm>m)) \<lsq> 𝟮maxδ(s)"
    by (rule Int_order_transitive)
  moreover
  from I have "s`(m) \<ra> s`(\<rm>m) = s`(m) \<ra> (s`(𝟬) \<rs> δ(s,m,\<rm>m) \<rs> s`(m))"
    by simp
  with T have "abs(s`(m) \<ra> s`(\<rm>m)) = abs(s`(𝟬) \<rs> δ(s,m,\<rm>m))"
    using Int_ZF_1_2_L3 by simp
  ultimately show "abs(s`(m)\<ra>s`(\<rm>m)) \<lsq> 𝟮maxδ(s)"
    by simp
  from I have "abs(s`(\<rm>m)) = abs(s`(𝟬) \<rs> δ(s,m,\<rm>m) \<rs> s`(m))"
    by simp
  with T have 
    "abs(s`(\<rm>m)) \<lsq> abs(s`(𝟬)) \<ra> abs(δ(s,m,\<rm>m)) \<ra> abs(s`(m))"
    using int_triangle_ineq3 by simp
  moreover from A1 A2 T have
    "abs(s`(𝟬)) \<ra> abs(δ(s,m,\<rm>m)) \<ra> abs(s`(m)) \<lsq> 𝟮maxδ(s) \<ra> abs(s`(m))"
    using Int_ZF_2_1_L7 Int_ZF_2_1_L8 Int_ZF_1_3_L21 int_ord_transl_inv by simp
  ultimately show "abs(s`(\<rm>m)) \<lsq> 𝟮maxδ(s) \<ra> abs(s`(m))"
    by (rule Int_order_transitive)
  from T have "s`(𝟬) \<rs> δ(s,m,\<rm>m) \<lsq> abs(s`(𝟬)) \<ra> abs(δ(s,m,\<rm>m))"
    using Int_ZF_2_L15E by simp
  moreover from A1 A2 T have 
    "abs(s`(𝟬)) \<ra> abs(δ(s,m,\<rm>m)) \<lsq> abs(s`(𝟬)) \<ra> maxδ(s)"
    using Int_ZF_2_1_L7 int_ord_transl_inv by simp
  ultimately have "s`(𝟬) \<rs> δ(s,m,\<rm>m) \<lsq> abs(s`(𝟬)) \<ra> maxδ(s)"
    by (rule Int_order_transitive)
  with T have 
    "s`(𝟬) \<rs> δ(s,m,\<rm>m) \<rs> s`(m) \<lsq> abs(s`(𝟬)) \<ra> maxδ(s) \<rs> s`(m)"
    using int_ord_transl_inv by simp
  with I show "s`(\<rm>m) \<lsq> abs(s`(𝟬)) \<ra> maxδ(s) \<rs> s`(m)"
    by simp
qed

textAn identity that expresses the value of an integer function at the opposite
  integer in terms of the value of that function at the integer, zero, and the 
  homomorphism difference. We have a similar identity in Int_ZF_2_1_L14›, but
  over there we assume that $f$ is a slope.

lemma (in int1) Int_ZF_2_1_L14A: assumes A1: "f:" and A2: "m"
  shows "f`(\<rm>m) = (\<rm>δ(f,m,\<rm>m)) \<ra> f`(𝟬) \<rs> f`(m)"
proof -
  from A1 A2 have T:
    "f`(\<rm>m)  "  "δ(f,m,\<rm>m)  "  "f`(𝟬)  "  "f`(m)  "
    using Int_ZF_1_1_L4 Int_ZF_1_1_L5 int_zero_one_are_int apply_funtype 
    by auto
   with A2 show "f`(\<rm>m) = (\<rm>δ(f,m,\<rm>m)) \<ra> f`(𝟬) \<rs> f`(m)"
     using Int_ZF_1_1_L4 Int_ZF_1_2_L15 by simp
qed

textThe next lemma allows to use the expression maxf(f,𝟬..M-1)›. 
  Recall that maxf(f,A)› is the maximum of (function) $f$ on 
  (the set) $A$.

lemma (in int1) Int_ZF_2_1_L15:
  assumes "s𝒮" and "M  +"
  shows 
  "maxf(s,𝟬..(M\<rs>𝟭))  "
  "n  𝟬..(M\<rs>𝟭). s`(n) \<lsq> maxf(s,𝟬..(M\<rs>𝟭))"
  "minf(s,𝟬..(M\<rs>𝟭))  "
  "n  𝟬..(M\<rs>𝟭). minf(s,𝟬..(M\<rs>𝟭)) \<lsq> s`(n)"
  using assms AlmostHoms_def Int_ZF_1_5_L6 Int_ZF_1_4_L2
  by auto

textA lower estimate for the value of a slope at $nM+k$.

lemma (in int1) Int_ZF_2_1_L16:
  assumes A1: "s𝒮"  and A2: "m" and A3: "M  +" and A4: "k  𝟬..(M\<rs>𝟭)"
  shows "s`(mM) \<ra> (minf(s,𝟬..(M\<rs>𝟭))\<rs> maxδ(s)) \<lsq> s`(mM\<ra>k)"
proof -
  from A3 have "𝟬..(M\<rs>𝟭)  "
    using Int_ZF_1_5_L6 by simp
  with A1 A2 A3 A4 have T: "mM  "   "k  "  "s`(mM)  "
    using PositiveSet_def Int_ZF_1_1_L5  Int_ZF_2_1_L2B
    by auto
  with A1 A3 A4 have 
    "s`(mM) \<ra> (minf(s,𝟬..(M\<rs>𝟭)) \<rs> maxδ(s)) \<lsq> s`(mM) \<ra> (s`(k) \<ra> δ(s,mM,k))"
    using Int_ZF_2_1_L15 Int_ZF_2_1_L7 int_ineq_add_sides int_ord_transl_inv
    by simp
  with A1 T show ?thesis using Int_ZF_2_1_L3A by simp
qed

textIdentity is a slope.

lemma (in int1) Int_ZF_2_1_L17: shows "id()  𝒮"
  using Int_ZF_2_1_L1 group1.Group_ZF_3_4_L15 by simp

textSimple identities about (absolute value of) homomorphism differences.

lemma (in int1) Int_ZF_2_1_L18:  
  assumes A1: "f:" and A2: "m"  "n"
  shows 
  "abs(f`(n) \<ra> f`(m) \<rs> f`(m\<ra>n)) = abs(δ(f,m,n))"
  "abs(f`(m) \<ra> f`(n) \<rs> f`(m\<ra>n)) = abs(δ(f,m,n))"
  "(\<rm>(f`(m))) \<rs> f`(n) \<ra> f`(m\<ra>n) = δ(f,m,n)"
  "(\<rm>(f`(n))) \<rs> f`(m) \<ra> f`(m\<ra>n) = δ(f,m,n)"
  "abs((\<rm>f`(m\<ra>n)) \<ra> f`(m) \<ra> f`(n)) = abs(δ(f,m,n))"
proof -
  from A1 A2 have T: 
    "f`(m\<ra>n)  "  "f`(m)  "  "f`(n)  "
    "f`(m\<ra>n) \<rs> f`(m) \<rs>  f`(n)   "
    "(\<rm>(f`(m)))  "
    "(\<rm>f`(m\<ra>n)) \<ra> f`(m) \<ra> f`(n)  "
    using apply_funtype Int_ZF_1_1_L4 Int_ZF_1_1_L5 by auto
  then have 
    "abs(\<rm>(f`(m\<ra>n) \<rs> f`(m) \<rs>  f`(n))) = abs(f`(m\<ra>n) \<rs> f`(m) \<rs>  f`(n))"
    using Int_ZF_2_L17 by simp
  moreover from T have 
    "(\<rm>(f`(m\<ra>n) \<rs> f`(m) \<rs>  f`(n))) = f`(n) \<ra> f`(m) \<rs> f`(m\<ra>n)"
    using Int_ZF_1_2_L9A by simp
  ultimately show "abs(f`(n) \<ra> f`(m) \<rs> f`(m\<ra>n)) = abs(δ(f,m,n))"
    by simp
  moreover from T have "f`(n) \<ra> f`(m) = f`(m) \<ra> f`(n)"
    using Int_ZF_1_1_L5 by simp
  ultimately show "abs(f`(m) \<ra> f`(n) \<rs> f`(m\<ra>n)) = abs(δ(f,m,n))"
    by simp
  from T show 
    "(\<rm>(f`(m))) \<rs> f`(n) \<ra> f`(m\<ra>n) = δ(f,m,n)"
    "(\<rm>(f`(n))) \<rs> f`(m) \<ra> f`(m\<ra>n) = δ(f,m,n)"
    using Int_ZF_1_2_L9 by auto
  from T have 
    "abs((\<rm>f`(m\<ra>n)) \<ra> f`(m) \<ra> f`(n)) =
    abs(\<rm>((\<rm>f`(m\<ra>n)) \<ra> f`(m) \<ra> f`(n)))"
    using Int_ZF_2_L17 by simp
  also from T have 
    "abs(\<rm>((\<rm>f`(m\<ra>n)) \<ra> f`(m) \<ra> f`(n))) = abs(δ(f,m,n))"
    using Int_ZF_1_2_L9 by simp
  finally show "abs((\<rm>f`(m\<ra>n)) \<ra> f`(m) \<ra> f`(n)) = abs(δ(f,m,n))"
    by simp
qed
  
textSome identities about the homomorphism difference of odd functions.

lemma (in int1) Int_ZF_2_1_L19: 
  assumes A1: "f:" and A2: "x. (\<rm>f`(\<rm>x)) = f`(x)"
  and A3: "m"  "n"
  shows
  "abs(δ(f,\<rm>m,m\<ra>n)) = abs(δ(f,m,n))"
  "abs(δ(f,\<rm>n,m\<ra>n)) = abs(δ(f,m,n))"
  "δ(f,n,\<rm>(m\<ra>n)) = δ(f,m,n)"
  "δ(f,m,\<rm>(m\<ra>n)) = δ(f,m,n)"
  "abs(δ(f,\<rm>m,\<rm>n)) = abs(δ(f,m,n))"
proof -
  from A1 A2 A3 show 
    "abs(δ(f,\<rm>m,m\<ra>n)) = abs(δ(f,m,n))"
    "abs(δ(f,\<rm>n,m\<ra>n)) = abs(δ(f,m,n))"
    using Int_ZF_1_2_L3 Int_ZF_2_1_L18 by auto
  from A3 have T: "m\<ra>n  " using Int_ZF_1_1_L5 by simp
  from A1 A2 have I: "x. f`(\<rm>x) = (\<rm>f`(x))"
    using Int_ZF_1_5_L13 by simp
  with A1 A2 A3 T show 
    "δ(f,n,\<rm>(m\<ra>n)) = δ(f,m,n)"
    "δ(f,m,\<rm>(m\<ra>n)) = δ(f,m,n)"
    using Int_ZF_1_2_L3 Int_ZF_2_1_L18 by auto
  from A3 have 
    "abs(δ(f,\<rm>m,\<rm>n)) = abs(f`(\<rm>(m\<ra>n)) \<rs> f`(\<rm>m) \<rs> f`(\<rm>n))"
    using Int_ZF_1_1_L5 by simp
  also from A1 A2 A3 T I have " = abs(δ(f,m,n))"
    using Int_ZF_2_1_L18 by simp
  finally show "abs(δ(f,\<rm>m,\<rm>n)) = abs(δ(f,m,n))" by simp
qed

textRecall that $f$ is a slope iff $f(m+n)-f(m)-f(n)$ is bounded
  as $m,n$ ranges over integers. The next lemma is the first 
  step in showing that we only need to check this condition as $m,n$ ranges
  over positive intergers. Namely we show that if the condition holds for
  positive integers, then it holds if one integer is positive and the second 
  one is nonnegative.

lemma (in int1) Int_ZF_2_1_L20: assumes A1: "f:" and
  A2: "a+. b+. abs(δ(f,a,b)) \<lsq> L" and
  A3:  "m+"  "n+"
  shows 
  "𝟬 \<lsq> L"
  "abs(δ(f,m,n)) \<lsq> L \<ra> abs(f`(𝟬))"
proof -       
  from A1 A2 have 
    "δ(f,𝟭,𝟭)  "  and "abs(δ(f,𝟭,𝟭)) \<lsq> L" 
    using int_one_two_are_pos PositiveSet_def Int_ZF_2_1_L3B
    by auto
  then show I: "𝟬 \<lsq> L" using Int_ZF_1_3_L19 by simp
  from A1 A3 have T: 
    "n  "  "f`(n)  "  "f`(𝟬)  "  
    "δ(f,m,n)  "  "abs(δ(f,m,n))  "
    using PositiveSet_def int_zero_one_are_int apply_funtype
      Nonnegative_def Int_ZF_2_1_L3B Int_ZF_2_L14 by auto
  from A3 have "m=𝟬  m+" using Int_ZF_1_5_L3A by auto
  moreover
  { assume "m = 𝟬"
    with T I have "abs(δ(f,m,n)) \<lsq> L \<ra> abs(f`(𝟬))"
      using Int_ZF_1_1_L4 Int_ZF_1_2_L3 Int_ZF_2_L17 
	int_ord_is_refl refl_def Int_ZF_2_L15F by simp }
  moreover
  { assume "m+"
    with A2 A3 T have "abs(δ(f,m,n)) \<lsq> L \<ra> abs(f`(𝟬))"
       using int_abs_nonneg Int_ZF_2_L15F by simp }
   ultimately show "abs(δ(f,m,n)) \<lsq> L \<ra> abs(f`(𝟬))"
     by auto
qed

textIf the slope condition holds for all pairs of integers such that one integer is 
  positive and the second one is nonnegative, then it holds when both integers are 
  nonnegative.

lemma (in int1) Int_ZF_2_1_L21: assumes A1: "f:" and
  A2: "a+. b+. abs(δ(f,a,b)) \<lsq> L" and
  A3: "n+"  "m+"
  shows "abs(δ(f,m,n)) \<lsq> L \<ra> abs(f`(𝟬))"
proof -
  from A1 A2 have 
    "δ(f,𝟭,𝟭)  "  and "abs(δ(f,𝟭,𝟭)) \<lsq> L" 
    using int_one_two_are_pos PositiveSet_def Nonnegative_def Int_ZF_2_1_L3B
    by auto
  then have I: "𝟬 \<lsq> L" using Int_ZF_1_3_L19 by simp
  from A1 A3 have T: 
    "m  "  "f`(m)  "  "f`(𝟬)  "  "(\<rm>f`(𝟬))  "  
    "δ(f,m,n)  "  "abs(δ(f,m,n))  "
    using int_zero_one_are_int apply_funtype Nonnegative_def 
      Int_ZF_2_1_L3B Int_ZF_2_L14 Int_ZF_1_1_L4 by auto
  from A3 have "n=𝟬  n+" using Int_ZF_1_5_L3A by auto
  moreover
  { assume "n=𝟬"
     with T have "δ(f,m,n) = \<rm>f`(𝟬)"
      using Int_ZF_1_1_L4 by simp
    with T have "abs(δ(f,m,n)) = abs(f`(𝟬))"
      using Int_ZF_2_L17 by simp
    with T have "abs(δ(f,m,n)) \<lsq> abs(f`(𝟬))"
      using int_ord_is_refl refl_def by simp
    with T I have "abs(δ(f,m,n)) \<lsq> L \<ra> abs(f`(𝟬))"
      using Int_ZF_2_L15F by simp }
  moreover
  { assume "n+"
    with A2 A3 T have "abs(δ(f,m,n)) \<lsq> L \<ra> abs(f`(𝟬))"
      using int_abs_nonneg Int_ZF_2_L15F by simp }
  ultimately show  "abs(δ(f,m,n)) \<lsq> L \<ra> abs(f`(𝟬))"
    by auto
qed

textIf the homomorphism difference is bounded on +×ℤ+, 
  then it is bounded on +×ℤ+.

lemma (in int1) Int_ZF_2_1_L22: assumes A1: "f:" and
  A2: "a+. b+. abs(δ(f,a,b)) \<lsq> L"
  shows "M. m+. n+. abs(δ(f,m,n)) \<lsq> M"
proof -
  from A1 A2 have 
    "m+. n+. abs(δ(f,m,n)) \<lsq> L \<ra> abs(f`(𝟬)) \<ra> abs(f`(𝟬))"
    using Int_ZF_2_1_L20 Int_ZF_2_1_L21 by simp
  then show ?thesis by auto
qed

textFor odd functions we can do better than in Int_ZF_2_1_L22›: 
  if the homomorphism 
  difference of $f$ is bounded on +×ℤ+, then it is bounded 
  on ℤ×ℤ›, hence $f$ is a slope. 
  Loong prof by splitting the ℤ×ℤ› into six subsets.

lemma (in int1) Int_ZF_2_1_L23: assumes A1: "f:" and
  A2: "a+. b+. abs(δ(f,a,b)) \<lsq> L"
  and A3: "x. (\<rm>f`(\<rm>x)) = f`(x)"
  shows (*"∃M. ∀m∈ℤ. ∀n∈ℤ. abs(δ(f,m,n)) \<lsq> M"*) "f𝒮"
proof -
  from A1 A2 have
    "M.a+. b+. abs(δ(f,a,b)) \<lsq> M"
    by (rule Int_ZF_2_1_L22)
  then obtain M where I: "m+. n+. abs(δ(f,m,n)) \<lsq> M"
    by auto
  { fix a b assume A4: "a"  "b"
    then have 
      "𝟬\<lsq>a  𝟬\<lsq>b    a\<lsq>𝟬  b\<lsq>𝟬    
      a\<lsq>𝟬  𝟬\<lsq>b  𝟬 \<lsq> a\<ra>b   a\<lsq>𝟬  𝟬\<lsq>b  a\<ra>b \<lsq> 𝟬    
      𝟬\<lsq>a  b\<lsq>𝟬  𝟬 \<lsq> a\<ra>b    𝟬\<lsq>a  b\<lsq>𝟬  a\<ra>b \<lsq> 𝟬"
      using int_plane_split_in6 by simp
    moreover
    { assume "𝟬\<lsq>a  𝟬\<lsq>b" 
      then have "a+"  "b+"
	using Int_ZF_2_L16 by auto
      with I have "abs(δ(f,a,b)) \<lsq> M" by simp }
    moreover
    { assume "a\<lsq>𝟬  b\<lsq>𝟬"
      with I have "abs(δ(f,\<rm>a,\<rm>b)) \<lsq> M"
	using Int_ZF_2_L10A Int_ZF_2_L16 by simp
      with A1 A3 A4 have "abs(δ(f,a,b)) \<lsq> M"
	using Int_ZF_2_1_L19 by simp }
    moreover
    { assume "a\<lsq>𝟬  𝟬\<lsq>b  𝟬 \<lsq> a\<ra>b"
      with I have "abs(δ(f,\<rm>a,a\<ra>b)) \<lsq> M"
	using Int_ZF_2_L10A Int_ZF_2_L16 by simp
      with A1 A3 A4 have "abs(δ(f,a,b)) \<lsq> M"
	using Int_ZF_2_1_L19 by simp } 
    moreover
    { assume "a\<lsq>𝟬  𝟬\<lsq>b  a\<ra>b \<lsq> 𝟬"
      with I have "abs(δ(f,b,\<rm>(a\<ra>b))) \<lsq> M"
	using Int_ZF_2_L10A Int_ZF_2_L16 by simp
      with A1 A3 A4 have "abs(δ(f,a,b)) \<lsq> M"
	using Int_ZF_2_1_L19 by simp }
    moreover
    { assume "𝟬\<lsq>a  b\<lsq>𝟬  𝟬 \<lsq> a\<ra>b"
      with I have "abs(δ(f,\<rm>b,a\<ra>b)) \<lsq> M"
	using Int_ZF_2_L10A Int_ZF_2_L16 by simp
      with A1 A3 A4 have "abs(δ(f,a,b)) \<lsq> M"
	using Int_ZF_2_1_L19 by simp }
    moreover
    { assume "𝟬\<lsq>a  b\<lsq>𝟬  a\<ra>b \<lsq> 𝟬" 
      with I have "abs(δ(f,a,\<rm>(a\<ra>b))) \<lsq> M"
	using Int_ZF_2_L10A Int_ZF_2_L16 by simp
      with A1 A3 A4 have "abs(δ(f,a,b)) \<lsq> M"
	using Int_ZF_2_1_L19 by simp }
    ultimately have "abs(δ(f,a,b)) \<lsq> M" by auto } 
  then have "m. n. abs(δ(f,m,n)) \<lsq> M" by simp
  with A1 show "f𝒮" by (rule Int_ZF_2_1_L5)
qed

textIf the homomorphism difference of a function defined 
  on positive integers is bounded, then the odd extension
  of this function is a slope.

lemma (in int1) Int_ZF_2_1_L24: 
  assumes A1: "f:+" and A2: "a+. b+. abs(δ(f,a,b)) \<lsq> L"
  shows "OddExtension(,IntegerAddition,IntegerOrder,f)  𝒮"
proof -
  let ?g = "OddExtension(,IntegerAddition,IntegerOrder,f)"
  from A1 have "?g : "
    using Int_ZF_1_5_L10 by simp
  moreover have "a+. b+. abs(δ(?g,a,b)) \<lsq> L"
  proof -
    { fix a b assume A3: "a+"  "b+"
      with A1 have "abs(δ(f,a,b)) =  abs(δ(?g,a,b))"
	using pos_int_closed_add_unfolded Int_ZF_1_5_L11 
	by simp
      moreover from A2 A3 have "abs(δ(f,a,b)) \<lsq> L" by simp
      ultimately have "abs(δ(?g,a,b)) \<lsq> L" by simp
    } then show ?thesis by simp
  qed
  moreover from A1 have "x. (\<rm>?g`(\<rm>x)) = ?g`(x)"
    using int_oddext_is_odd_alt by simp
  ultimately show "?g  𝒮" by (rule Int_ZF_2_1_L23)
qed

textType information related to $\gamma$.

lemma (in int1) Int_ZF_2_1_L25: 
  assumes A1: "f:" and A2: "m"  "n"
  shows 
  "δ(f,m,\<rm>n)  "
  "δ(f,n,\<rm>n)  "
  "(\<rm>δ(f,n,\<rm>n))  "
  "f`(𝟬)  "
  "γ(f,m,n)   "
proof -
  from A1 A2 show T1:
    "δ(f,m,\<rm>n)  "  "f`(𝟬)  "
    using Int_ZF_1_1_L4 Int_ZF_2_1_L3B int_zero_one_are_int apply_funtype
    by auto
  from A2 have "(\<rm>n)  "
    using Int_ZF_1_1_L4 by simp
  with A1 A2 show "δ(f,n,\<rm>n)  "
    using Int_ZF_2_1_L3B by simp
  then show "(\<rm>δ(f,n,\<rm>n))  "
    using Int_ZF_1_1_L4 by simp
  with T1 show "γ(f,m,n)   "
    using Int_ZF_1_1_L5 by simp
qed 

textA couple of formulae involving $f(m-n)$ and $\gamma(f,m,n)$.

lemma (in int1) Int_ZF_2_1_L26: 
  assumes A1: "f:" and A2: "m"  "n"
  shows 
  "f`(m\<rs>n) = γ(f,m,n) \<ra> f`(m) \<rs> f`(n)"
  "f`(m\<rs>n) = γ(f,m,n) \<ra> (f`(m) \<rs> f`(n))"
  "f`(m\<rs>n) \<ra> (f`(n) \<rs> γ(f,m,n)) = f`(m)"
proof -
  from A1 A2 have T:
    "(\<rm>n)  "  "δ(f,m,\<rm>n)  "  
    "f`(𝟬)  "  "f`(m)  "  "f`(n)  "  "(\<rm>f`(n))  "
    "(\<rm>δ(f,n,\<rm>n))  "  
    "(\<rm>δ(f,n,\<rm>n))  \<ra> f`(𝟬)  "
    "γ(f,m,n)  "
    using  Int_ZF_1_1_L4 Int_ZF_2_1_L25 apply_funtype Int_ZF_1_1_L5 
    by auto
   with A1 A2 have "f`(m\<rs>n) = 
    δ(f,m,\<rm>n) \<ra> ((\<rm>δ(f,n,\<rm>n)) \<ra> f`(𝟬) \<rs> f`(n)) \<ra> f`(m)"
    using Int_ZF_2_1_L3C Int_ZF_2_1_L14A by simp
  with T have "f`(m\<rs>n) =
    δ(f,m,\<rm>n) \<ra> ((\<rm>δ(f,n,\<rm>n)) \<ra> f`(𝟬)) \<ra> f`(m) \<rs> f`(n)"
    using Int_ZF_1_2_L16 by simp
  moreover from T have 
    "δ(f,m,\<rm>n) \<ra> ((\<rm>δ(f,n,\<rm>n)) \<ra> f`(𝟬)) = γ(f,m,n)"
    using Int_ZF_1_1_L7 by simp
  ultimately show  I: "f`(m\<rs>n) = γ(f,m,n) \<ra> f`(m) \<rs> f`(n)"
    by simp
  then have "f`(m\<rs>n) \<ra> (f`(n) \<rs> γ(f,m,n)) = 
    (γ(f,m,n) \<ra> f`(m) \<rs> f`(n)) \<ra> (f`(n) \<rs> γ(f,m,n))"
    by simp
  moreover from T have " = f`(m)" using Int_ZF_1_2_L18 
    by simp
  ultimately show "f`(m\<rs>n) \<ra> (f`(n) \<rs> γ(f,m,n)) = f`(m)"
    by simp
  from T have "γ(f,m,n)  "  "f`(m)  "  "(\<rm>f`(n))  "
    by auto
  then have 
    "γ(f,m,n) \<ra> f`(m) \<ra> (\<rm>f`(n)) =  γ(f,m,n) \<ra> (f`(m) \<ra> (\<rm>f`(n)))"
    by (rule Int_ZF_1_1_L7)
  with I show  "f`(m\<rs>n) = γ(f,m,n) \<ra> (f`(m) \<rs> f`(n))" by simp
qed

textA formula expressing the difference between $f(m-n-k)$ and
  $f(m)-f(n)-f(k)$ in terms of $\gamma$.

lemma (in int1) Int_ZF_2_1_L26A: 
  assumes A1: "f:" and A2: "m"  "n"  "k"
  shows 
  "f`(m\<rs>n\<rs>k) \<rs> (f`(m)\<rs> f`(n) \<rs> f`(k)) = γ(f,m\<rs>n,k) \<ra> γ(f,m,n)"
proof -
  from A1 A2 have 
    T: "m\<rs>n  " "γ(f,m\<rs>n,k)  "  "f`(m) \<rs> f`(n) \<rs> f`(k)  " and
    T1: "γ(f,m,n)  "  "f`(m) \<rs> f`(n)  "  "(\<rm>f`(k))  "
    using Int_ZF_1_1_L4 Int_ZF_1_1_L5 Int_ZF_2_1_L25 apply_funtype 
    by auto
  from A1 A2 have 
    "f`(m\<rs>n) \<rs> f`(k) = γ(f,m,n) \<ra> (f`(m) \<rs> f`(n)) \<ra> (\<rm>f`(k))"
    using Int_ZF_2_1_L26 by simp
  also from T1 have " = γ(f,m,n) \<ra> (f`(m) \<rs> f`(n) \<ra> (\<rm>f`(k)))"
    by (rule Int_ZF_1_1_L7)
  finally have 
    "f`(m\<rs>n) \<rs> f`(k) = γ(f,m,n) \<ra> (f`(m) \<rs> f`(n) \<rs> f`(k))"
    by simp
  moreover from A1 A2 T have
    "f`(m\<rs>n\<rs>k) =  γ(f,m\<rs>n,k) \<ra> (f`(m\<rs>n)\<rs>f`(k))"
    using Int_ZF_2_1_L26 by simp
  ultimately have 
    "f`(m\<rs>n\<rs>k) \<rs> (f`(m)\<rs> f`(n) \<rs> f`(k)) = 
    γ(f,m\<rs>n,k) \<ra> ( γ(f,m,n) \<ra> (f`(m) \<rs> f`(n) \<rs> f`(k))) 
    \<rs> (f`(m)\<rs> f`(n) \<rs> f`(k))"
    by simp
  with T T1 show ?thesis 
    using Int_ZF_1_2_L17 by simp
qed


textIf $s$ is a slope, then $\gamma (s,m,n)$ is uniformly bounded.

lemma (in int1) Int_ZF_2_1_L27: assumes A1: "s𝒮"
  shows "L. m.n. abs(γ(s,m,n)) \<lsq> L"
proof -
  let ?L = "maxδ(s) \<ra> maxδ(s) \<ra> abs(s`(𝟬))"
  from A1 have T: 
    "maxδ(s)  "  "abs(s`(𝟬))  "  "?L  "
    using Int_ZF_2_1_L8 int_zero_one_are_int Int_ZF_2_1_L2B 
      Int_ZF_2_L14 Int_ZF_1_1_L5 by auto
  moreover
  { fix m 
    fix n
    assume A2: "m"  "n"
    with A1 have T: 
      "(\<rm>n)  "
      "δ(s,m,\<rm>n)  "
      "δ(s,n,\<rm>n)  "
      "(\<rm>δ(s,n,\<rm>n))  "
      "s`(𝟬)  "  "abs(s`(𝟬))  "
      using Int_ZF_1_1_L4 AlmostHoms_def Int_ZF_2_1_L25 Int_ZF_2_L14
      by auto
    with T have
      "abs(δ(s,m,\<rm>n) \<rs> δ(s,n,\<rm>n) \<ra> s`(𝟬)) \<lsq>
      abs(δ(s,m,\<rm>n)) \<ra> abs(\<rm>δ(s,n,\<rm>n)) \<ra> abs(s`(𝟬))"
      using Int_triangle_ineq3 by simp
    moreover from A1 A2 T have 
      "abs(δ(s,m,\<rm>n)) \<ra> abs(\<rm>δ(s,n,\<rm>n)) \<ra> abs(s`(𝟬)) \<lsq> ?L"
      using Int_ZF_2_1_L7 int_ineq_add_sides int_ord_transl_inv Int_ZF_2_L17
      by simp
   ultimately have "abs(δ(s,m,\<rm>n) \<rs> δ(s,n,\<rm>n) \<ra> s`(𝟬)) \<lsq> ?L"
      by (rule Int_order_transitive)    
    then have "abs(γ(s,m,n)) \<lsq> ?L" by simp }
  ultimately show "L. m.n. abs(γ(s,m,n)) \<lsq> L"
    by auto
qed


textIf $s$ is a slope, then $s(m) \leq s(m-1) + M$, where $L$ does not depend
  on $m$.

lemma (in int1) Int_ZF_2_1_L28: assumes A1: "s𝒮"
  shows "M. m. s`(m) \<lsq> s`(m\<rs>𝟭) \<ra> M"
proof -
  from A1 have
    "L. m.n.abs(γ(s,m,n)) \<lsq> L"
    using Int_ZF_2_1_L27 by simp
  then obtain L where T: "L" and "m.n.abs(γ(s,m,n)) \<lsq> L"
    using Int_ZF_2_1_L27 by auto
  then have I: "m.abs(γ(s,m,𝟭)) \<lsq> L"
    using int_zero_one_are_int by simp
  let ?M = "s`(𝟭) \<ra> L"
  from A1 T have "?M  "
    using int_zero_one_are_int Int_ZF_2_1_L2B Int_ZF_1_1_L5
    by simp
  moreover
  { fix m assume A2: "m"
    with A1 have 
      T1: "s:"  "m"  "𝟭" and
      T2: "γ(s,m,𝟭)  "  "s`(𝟭)  "
      using int_zero_one_are_int AlmostHoms_def 
	Int_ZF_2_1_L25 by auto
    from A2 T1 have T3: "s`(m\<rs>𝟭)  "
      using Int_ZF_1_1_L5 apply_funtype by simp
    from I A2 T2 have 
      "(\<rm>γ(s,m,𝟭)) \<lsq> abs(γ(s,m,𝟭))"
      "abs(γ(s,m,𝟭)) \<lsq> L"
      using Int_ZF_2_L19C by auto
    then have "(\<rm>γ(s,m,𝟭)) \<lsq> L"
      by (rule Int_order_transitive)
    with T2 T3 have 
      "s`(m\<rs>𝟭) \<ra> (s`(𝟭) \<rs> γ(s,m,𝟭)) \<lsq> s`(m\<rs>𝟭) \<ra> ?M"
      using int_ord_transl_inv by simp
    moreover from T1 have
      "s`(m\<rs>𝟭) \<ra> (s`(𝟭) \<rs> γ(s,m,𝟭)) = s`(m)"
      by (rule Int_ZF_2_1_L26)
    ultimately have "s`(m) \<lsq> s`(m\<rs>𝟭) \<ra> ?M"  by simp  }
  ultimately show "M. m. s`(m) \<lsq> s`(m\<rs>𝟭) \<ra> M"
    by auto
qed

textIf $s$ is a slope, then the difference between 
  $s(m-n-k)$ and $s(m)-s(n)-s(k)$ is uniformly bounded.

lemma (in int1) Int_ZF_2_1_L29: assumes A1: "s𝒮"
  shows 
  "M. m.n.k. abs(s`(m\<rs>n\<rs>k) \<rs> (s`(m)\<rs>s`(n)\<rs>s`(k))) \<lsq>M"
proof -
  from A1 have "L. m.n. abs(γ(s,m,n)) \<lsq> L"
    using Int_ZF_2_1_L27 by simp
  then obtain L where I: "L" and 
    II: "m.n. abs(γ(s,m,n)) \<lsq> L"
    by auto
  from I have "L\<ra>L  "
    using Int_ZF_1_1_L5 by simp
  moreover
  { fix m n k assume A2: "m"  "n"  "k"
    with A1 have T: 
      "m\<rs>n  "  "γ(s,m\<rs>n,k)  "  "γ(s,m,n)  "
      using Int_ZF_1_1_L5 AlmostHoms_def Int_ZF_2_1_L25 
      by auto
    then have 
      I: "abs(γ(s,m\<rs>n,k) \<ra> γ(s,m,n)) \<lsq> abs(γ(s,m\<rs>n,k)) \<ra> abs(γ(s,m,n))"
      using Int_triangle_ineq by simp
    from II A2 T have 
      "abs(γ(s,m\<rs>n,k)) \<lsq> L"
      "abs(γ(s,m,n)) \<lsq> L"
      by auto
    then have "abs(γ(s,m\<rs>n,k)) \<ra> abs(γ(s,m,n)) \<lsq> L\<ra>L"
      using int_ineq_add_sides by simp
    with I have "abs(γ(s,m\<rs>n,k) \<ra> γ(s,m,n)) \<lsq> L\<ra>L"
      by (rule Int_order_transitive)
    moreover from A1 A2 have 
      "s`(m\<rs>n\<rs>k) \<rs> (s`(m)\<rs> s`(n) \<rs> s`(k)) = γ(s,m\<rs>n,k) \<ra> γ(s,m,n)"
      using AlmostHoms_def Int_ZF_2_1_L26A by simp
    ultimately have 
      "abs(s`(m\<rs>n\<rs>k) \<rs> (s`(m)\<rs> s`(n) \<rs> s`(k))) \<lsq> L\<ra>L"
      by simp }
  ultimately show ?thesis by auto
qed

textIf $s$ is a slope, then we can find integers $M,K$ such that
  $s(m-n-k) \leq s(m)-s(n)-s(k) + M$ and $s(m)-s(n)-s(k) + K \leq s(m-n-k)$, 
  for all integer $m,n,k$.

lemma (in int1) Int_ZF_2_1_L30: assumes A1: "s𝒮"
  shows 
  "M. m.n.k. s`(m\<rs>n\<rs>k) \<lsq> s`(m)\<rs>s`(n)\<rs>s`(k)\<ra>M"
  "K. m.n.k. s`(m)\<rs>s`(n)\<rs>s`(k)\<ra>K \<lsq> s`(m\<rs>n\<rs>k)"
proof -
  from A1 have
    "M. m.n.k. abs(s`(m\<rs>n\<rs>k) \<rs> (s`(m)\<rs>s`(n)\<rs>s`(k))) \<lsq>M"
    using Int_ZF_2_1_L29 by simp
  then obtain M where I: "M" and II:
    "m.n.k. abs(s`(m\<rs>n\<rs>k) \<rs> (s`(m)\<rs>s`(n)\<rs>s`(k))) \<lsq>M"
    by auto
  from I have III: "(\<rm>M)  " using Int_ZF_1_1_L4 by simp
  { fix m n k assume A2: "m"  "n"  "k"
    with A1 have "s`(m\<rs>n\<rs>k)  "  and "s`(m)\<rs>s`(n)\<rs>s`(k)  "
      using Int_ZF_1_1_L5 Int_ZF_2_1_L2B by auto
    moreover from II A2 have
      "abs(s`(m\<rs>n\<rs>k) \<rs> (s`(m)\<rs>s`(n)\<rs>s`(k))) \<lsq>M"
      by simp
    ultimately have 
      "s`(m\<rs>n\<rs>k) \<lsq> s`(m)\<rs>s`(n)\<rs>s`(k)\<ra>M  
      s`(m)\<rs>s`(n)\<rs>s`(k) \<rs> M \<lsq> s`(m\<rs>n\<rs>k)"
      using Int_triangle_ineq2 by simp
  } then have 
      "m.n.k. s`(m\<rs>n\<rs>k) \<lsq> s`(m)\<rs>s`(n)\<rs>s`(k)\<ra>M"
      "m.n.k. s`(m)\<rs>s`(n)\<rs>s`(k) \<rs> M \<lsq> s`(m\<rs>n\<rs>k)"
    by auto
  with I III show  
    "M. m.n.k. s`(m\<rs>n\<rs>k) \<lsq> s`(m)\<rs>s`(n)\<rs>s`(k)\<ra>M"
    "K. m.n.k. s`(m)\<rs>s`(n)\<rs>s`(k)\<ra>K \<lsq> s`(m\<rs>n\<rs>k)"
    by auto
qed

textBy definition functions $f,g$ are almost equal if $f-g$* is bounded. 
  In the next lemma we show it is sufficient to check the boundedness on positive
  integers.

lemma (in int1) Int_ZF_2_1_L31: assumes A1: "s𝒮"  "r𝒮" 
  and A2: "m+. abs(s`(m)\<rs>r`(m)) \<lsq> L"
  shows "s  r"
proof -
  let ?a = "abs(s`(𝟬) \<rs> r`(𝟬))"
  let ?c = "𝟮maxδ(s) \<ra> 𝟮maxδ(r) \<ra> L"
  let ?M = "Maximum(IntegerOrder,{?a,L,?c})"
  from A2 have "abs(s`(𝟭)\<rs>r`(𝟭)) \<lsq> L"
    using int_one_two_are_pos by simp
  then have T: "L" using Int_ZF_2_L1A by simp
  moreover from A1 have "?a  "
    using int_zero_one_are_int Int_ZF_2_1_L2B 
      Int_ZF_1_1_L5 Int_ZF_2_L14 by simp
  moreover from A1 T have "?c  "
    using Int_ZF_2_1_L8 int_two_three_are_int Int_ZF_1_1_L5
    by simp
  ultimately have 
    I: "?a \<lsq> ?M" and
    II: "L \<lsq> ?M" and 
    III: "?c \<lsq> ?M"
    using Int_ZF_1_4_L1A by auto
  
  { fix m assume A5: "m"
    with A1 have T: 
      "s`(m)  "  "r`(m)  "  "s`(m) \<rs> r`(m)  "
      "s`(\<rm>m)  "  "r`(\<rm>m)  "
      using Int_ZF_2_1_L2B Int_ZF_1_1_L4 Int_ZF_1_1_L5 
      by auto
    from A5 have "m=𝟬  m+  (\<rm>m)  +"
      using int_decomp_cases by simp
    moreover
    { assume "m=𝟬"
      with I have "abs(s`(m) \<rs> r`(m)) \<lsq> ?M"
	by simp }
    moreover
    { assume "m+"
      with A2 II have 
	"abs(s`(m)\<rs>r`(m)) \<lsq> L" and "L\<lsq>?M"
	by auto
      then have "abs(s`(m)\<rs>r`(m)) \<lsq> ?M"
	by (rule Int_order_transitive) }
    moreover
    { assume A6: "(\<rm>m)  +"
      from T have "abs(s`(m)\<rs>r`(m)) \<lsq> 
	abs(s`(m)\<ra>s`(\<rm>m)) \<ra> abs(r`(m)\<ra>r`(\<rm>m)) \<ra> abs(s`(\<rm>m)\<rs>r`(\<rm>m))"
	using Int_ZF_1_3_L22A by simp
      moreover 
      from A1 A2 III A5 A6 have
	"abs(s`(m)\<ra>s`(\<rm>m)) \<ra> abs(r`(m)\<ra>r`(\<rm>m)) \<ra> abs(s`(\<rm>m)\<rs>r`(\<rm>m)) \<lsq> ?c"
	"?c \<lsq> ?M"
	using Int_ZF_2_1_L14 int_ineq_add_sides by auto
      then have 
	"abs(s`(m)\<ra>s`(\<rm>m)) \<ra> abs(r`(m)\<ra>r`(\<rm>m)) \<ra> abs(s`(\<rm>m)\<rs>r`(\<rm>m)) \<lsq> ?M"
	by (rule Int_order_transitive)
      ultimately have  "abs(s`(m)\<rs>r`(m)) \<lsq> ?M"
	by (rule Int_order_transitive) }
    ultimately have "abs(s`(m) \<rs> r`(m)) \<lsq> ?M"
      by auto
  } then have "m. abs(s`(m)\<rs>r`(m)) \<lsq> ?M"
    by simp
  with A1 show "s  r" by (rule Int_ZF_2_1_L9)
qed

textA sufficient condition for an odd slope to be almost equal to identity:
  If for all positive integers the value of the slope at $m$ is between $m$ and
  $m$ plus some constant independent of $m$, then the slope is almost identity.

lemma (in int1) Int_ZF_2_1_L32: assumes A1: "s𝒮"  "M"
  and A2: "m+. m \<lsq> s`(m)  s`(m) \<lsq> m\<ra>M"
  shows "s  id()"
proof -
  let ?r = "id()"
  from A1 have "s𝒮"  "?r  𝒮"
    using Int_ZF_2_1_L17 by auto
  moreover from A1 A2 have "m+. abs(s`(m)\<rs>?r`(m)) \<lsq> M"
    using Int_ZF_1_3_L23 PositiveSet_def id_conv by simp
  ultimately show "s  id()" by (rule Int_ZF_2_1_L31)
qed

textA lemma about adding a constant to slopes. This is actually proven in
  Group_ZF_3_5_L1›, in Group_ZF_3.thy› here we just refer to 
  that lemma to show it in notation used for integers. Unfortunately we have
  to use raw set notation in the proof.

lemma (in int1) Int_ZF_2_1_L33:
  assumes A1: "s𝒮" and A2: "c" and 
  A3: "r = {m,s`(m)\<ra>c. m}"
  shows
  "m. r`(m) = s`(m)\<ra>c"
  "r𝒮"
  "s  r"
proof -
  let ?G = ""
  let ?f = "IntegerAddition"
  let ?AH = "AlmostHoms(?G, ?f)"
  from assms have I:
    "group1(?G, ?f)"
    "s  AlmostHoms(?G, ?f)"
    "c  ?G"
    "r = {x, ?f`s`(x), c. x  ?G}"
    using Int_ZF_2_1_L1 by auto
  then have "x?G. r`(x) = ?f`s`(x),c"
    by (rule group1.Group_ZF_3_5_L1)
  moreover from I have "r  AlmostHoms(?G, ?f)"
    by (rule group1.Group_ZF_3_5_L1)
  moreover from I have 
    "s, r  QuotientGroupRel(AlmostHoms(?G, ?f), AlHomOp1(?G, ?f), FinRangeFunctions(?G, ?G))"
    by (rule group1.Group_ZF_3_5_L1)
  ultimately show 
    "m. r`(m) = s`(m)\<ra>c"
    "r𝒮"
    "s  r"
    by auto
qed  

subsectionComposing slopes

textComposition of slopes is not commutative. However, as we show in this 
  section if $f$ and $g$ are slopes then the range of $f\circ g - g\circ f$ 
  is bounded. This allows to show that the multiplication of real 
  numbers is commutative.

textTwo useful estimates.

lemma (in int1) Int_ZF_2_2_L1: 
  assumes A1: "f:" and A2: "p"  "q"
  shows 
  "abs(f`((p\<ra>𝟭)q)\<rs>(p\<ra>𝟭)f`(q)) \<lsq> abs(δ(f,pq,q))\<ra>abs(f`(pq)\<rs>pf`(q))"
  "abs(f`((p\<rs>𝟭)q)\<rs>(p\<rs>𝟭)f`(q)) \<lsq> abs(δ(f,(p\<rs>𝟭)q,q))\<ra>abs(f`(pq)\<rs>pf`(q))"
proof -
  let ?R = ""
  let ?A = "IntegerAddition"
  let ?M = "IntegerMultiplication"
  let ?I = "GroupInv(?R, ?A)"
  let ?a = "f`((p\<ra>𝟭)q)"
  let ?b = "p"
  let ?c = "f`(q)"
  let ?d = "f`(pq)"
  from A1 A2 have T1:
    "ring0(?R, ?A, ?M)"  "?a  ?R"  "?b  ?R"  "?c  ?R"  "?d  ?R"
    using  Int_ZF_1_1_L2 int_zero_one_are_int Int_ZF_1_1_L5 apply_funtype 
    by auto
  then have 
    "?A`?a,?I`(?M`?A`?b, TheNeutralElement(?R, ?M),?c) =
    ?A`?A`?A`?a,?I`(?d),?I`(?c),?A`?d, ?I`(?M`?b, ?c)"
    by (rule ring0.Ring_ZF_2_L2)
  with A2 have 
    "f`((p\<ra>𝟭)q)\<rs>(p\<ra>𝟭)f`(q) = δ(f,pq,q)\<ra>(f`(pq)\<rs>pf`(q))"
    using int_zero_one_are_int Int_ZF_1_1_L1 Int_ZF_1_1_L4 by simp
  moreover from A1 A2 T1 have "δ(f,pq,q)  " "f`(pq)\<rs>pf`(q)  "
    using Int_ZF_1_1_L5 apply_funtype by auto
  ultimately show 
    "abs(f`((p\<ra>𝟭)q)\<rs>(p\<ra>𝟭)f`(q)) \<lsq> abs(δ(f,pq,q))\<ra>abs(f`(pq)\<rs>pf`(q))"
    using Int_triangle_ineq by simp
  from A1 A2 have T1: 
    "f`((p\<rs>𝟭)q)  "   "p"  "f`(q)  "   "f`(pq)  " 
    using int_zero_one_are_int Int_ZF_1_1_L5 apply_funtype by auto
  then have
    "f`((p\<rs>𝟭)q)\<rs>(p\<rs>𝟭)f`(q) = (f`(pq)\<rs>pf`(q))\<rs>(f`(pq)\<rs>f`((p\<rs>𝟭)q)\<rs>f`(q))"
    by (rule Int_ZF_1_2_L6)
  with A2 have "f`((p\<rs>𝟭)q)\<rs>(p\<rs>𝟭)f`(q) = (f`(pq)\<rs>pf`(q))\<rs>δ(f,(p\<rs>𝟭)q,q)"
    using Int_ZF_1_2_L7 by simp
  moreover from A1 A2 have 
    "f`(pq)\<rs>pf`(q)  "   "δ(f,(p\<rs>𝟭)q,q)  " 
    using Int_ZF_1_1_L5 int_zero_one_are_int apply_funtype by auto
  ultimately show 
    "abs(f`((p\<rs>𝟭)q)\<rs>(p\<rs>𝟭)f`(q)) \<lsq> abs(δ(f,(p\<rs>𝟭)q,q))\<ra>abs(f`(pq)\<rs>pf`(q))"
    using Int_triangle_ineq1 by simp
qed
  
textIf $f$ is a slope, then 
  $|f(p\cdot q)-p\cdot f(q)|\leq (|p|+1)\cdot$maxδ(f)›. 
  The proof is by induction on $p$ and the next lemma is the induction step for the case when $0\leq p$.

lemma (in int1) Int_ZF_2_2_L2: 
  assumes A1: "f𝒮" and A2: "𝟬\<lsq>p"  "q"
  and A3: "abs(f`(pq)\<rs>pf`(q)) \<lsq> (abs(p)\<ra>𝟭)maxδ(f)"
  shows 
  "abs(f`((p\<ra>𝟭)q)\<rs>(p\<ra>𝟭)f`(q)) \<lsq> (abs(p\<ra>𝟭)\<ra> 𝟭)maxδ(f)"
proof -
  from A2 have "q"  "pq  " 
    using Int_ZF_2_L1A Int_ZF_1_1_L5 by auto
  with A1 have I: "abs(δ(f,pq,q)) \<lsq> maxδ(f)" by (rule Int_ZF_2_1_L7)
  moreover note A3
  moreover from A1 A2 have
    "abs(f`((p\<ra>𝟭)q)\<rs>(p\<ra>𝟭)f`(q)) \<lsq> abs(δ(f,pq,q))\<ra>abs(f`(pq)\<rs>pf`(q))"
    using AlmostHoms_def Int_ZF_2_L1A Int_ZF_2_2_L1 by simp
  ultimately have 
    "abs(f`((p\<ra>𝟭)q)\<rs>(p\<ra>𝟭)f`(q)) \<lsq> maxδ(f)\<ra>(abs(p)\<ra>𝟭)maxδ(f)"
    by (rule Int_ZF_2_L15)
  moreover from I A2 have 
    "maxδ(f)\<ra>(abs(p)\<ra>𝟭)maxδ(f) = (abs(p\<ra>𝟭)\<ra> 𝟭)maxδ(f)"
    using Int_ZF_2_L1A Int_ZF_1_2_L2 by simp
  ultimately show
    "abs(f`((p\<ra>𝟭)q)\<rs>(p\<ra>𝟭)f`(q)) \<lsq> (abs(p\<ra>𝟭)\<ra> 𝟭)maxδ(f)"
    by simp
qed

textIf $f$ is a slope, then 
  $|f(p\cdot q)-p\cdot f(q)|\leq (|p|+1)\cdot$maxδ›. 
  The proof is by induction on $p$ and the next lemma is the induction step for the case when $p\leq 0$.

lemma (in int1) Int_ZF_2_2_L3: 
  assumes A1: "f𝒮" and A2: "p\<lsq>𝟬"  "q"
  and A3: "abs(f`(pq)\<rs>pf`(q)) \<lsq> (abs(p)\<ra>𝟭)maxδ(f)"
  shows  "abs(f`((p\<rs>𝟭)q)\<rs>(p\<rs>𝟭)f`(q)) \<lsq> (abs(p\<rs>𝟭)\<ra> 𝟭)maxδ(f)"
proof -
  from A2 have "q"  "(p\<rs>𝟭)q  " 
    using Int_ZF_2_L1A int_zero_one_are_int Int_ZF_1_1_L5 by auto
  with A1 have I: "abs(δ(f,(p\<rs>𝟭)q,q)) \<lsq> maxδ(f)" by (rule Int_ZF_2_1_L7)
  moreover note A3
  moreover from A1 A2 have 
    "abs(f`((p\<rs>𝟭)q)\<rs>(p\<rs>𝟭)f`(q)) \<lsq> abs(δ(f,(p\<rs>𝟭)q,q))\<ra>abs(f`(pq)\<rs>pf`(q))"
    using AlmostHoms_def Int_ZF_2_L1A Int_ZF_2_2_L1 by simp
  ultimately have 
    "abs(f`((p\<rs>𝟭)q)\<rs>(p\<rs>𝟭)f`(q)) \<lsq> maxδ(f)\<ra>(abs(p)\<ra>𝟭)maxδ(f)"
    by (rule Int_ZF_2_L15)
  with I A2 show ?thesis using Int_ZF_2_L1A Int_ZF_1_2_L5 by simp
qed

textIf $f$ is a slope, then 
  $|f(p\cdot q)-p\cdot f(q)|\leq (|p|+1)\cdot$maxδ›$(f)$.
  Proof by cases on $0 \leq p$. 

lemma (in int1) Int_ZF_2_2_L4: 
  assumes A1: "f𝒮" and A2: "p" "q"
  shows "abs(f`(pq)\<rs>pf`(q)) \<lsq> (abs(p)\<ra>𝟭)maxδ(f)"
proof -
  { assume "𝟬\<lsq>p"
    moreover from A1 A2 have "abs(f`(𝟬q)\<rs>𝟬f`(q)) \<lsq> (abs(𝟬)\<ra>𝟭)maxδ(f)"
      using int_zero_one_are_int Int_ZF_2_1_L2B Int_ZF_1_1_L4 
	Int_ZF_2_1_L8 Int_ZF_2_L18 by simp
    moreover from A1 A2 have 
      "p. 𝟬\<lsq>p  abs(f`(pq)\<rs>pf`(q)) \<lsq> (abs(p)\<ra>𝟭)maxδ(f) 
      abs(f`((p\<ra>𝟭)q)\<rs>(p\<ra>𝟭)f`(q)) \<lsq> (abs(p\<ra>𝟭)\<ra> 𝟭)maxδ(f)"
      using Int_ZF_2_2_L2 by simp
    ultimately have "abs(f`(pq)\<rs>pf`(q)) \<lsq> (abs(p)\<ra>𝟭)maxδ(f)" 
      by (rule Induction_on_int) }
  moreover
  { assume "¬(𝟬\<lsq>p)"
    with A2 have "p\<lsq>𝟬" using Int_ZF_2_L19A by simp
    moreover from A1 A2 have "abs(f`(𝟬q)\<rs>𝟬f`(q)) \<lsq> (abs(𝟬)\<ra>𝟭)maxδ(f)"
      using int_zero_one_are_int Int_ZF_2_1_L2B Int_ZF_1_1_L4
	Int_ZF_2_1_L8 Int_ZF_2_L18 by simp
    moreover from A1 A2 have 
      "p. p\<lsq>𝟬  abs(f`(pq)\<rs>pf`(q)) \<lsq> (abs(p)\<ra>𝟭)maxδ(f) 
      abs(f`((p\<rs>𝟭)q)\<rs>(p\<rs>𝟭)f`(q)) \<lsq> (abs(p\<rs>𝟭)\<ra> 𝟭)maxδ(f)"
      using Int_ZF_2_2_L3 by simp
    ultimately have "abs(f`(pq)\<rs>pf`(q)) \<lsq> (abs(p)\<ra>𝟭)maxδ(f)" 
      by (rule Back_induct_on_int) }
  ultimately show ?thesis by blast
qed

textThe next elegant result is Lemma 7 in the Arthan's paper \cite{Arthan2004}.

lemma (in int1) Arthan_Lem_7: 
 assumes A1: "f𝒮" and A2: "p"  "q"
  shows "abs(qf`(p)\<rs>pf`(q)) \<lsq> (abs(p)\<ra>abs(q)\<ra>𝟮)maxδ(f)"
proof -
  from A1 A2 have T:
    "qf`(p)\<rs>f`(pq)  " 
    "f`(pq)\<rs>pf`(q)  "
    "f`(qp)  "  "f`(pq)  "
    "qf`(p)  "  "pf`(q)  " 
    "maxδ(f)  "
    "abs(q)  "  "abs(p)  "
    using Int_ZF_1_1_L5 Int_ZF_2_1_L2B Int_ZF_2_1_L7 Int_ZF_2_L14 by auto
  moreover have "abs(qf`(p)\<rs>f`(pq)) \<lsq> (abs(q)\<ra>𝟭)maxδ(f)"
  proof -
    from A1 A2 have "abs(f`(qp)\<rs>qf`(p)) \<lsq> (abs(q)\<ra>𝟭)maxδ(f)"
      using Int_ZF_2_2_L4 by simp
    with T A2 show ?thesis
      using Int_ZF_2_L20 Int_ZF_1_1_L5 by simp
  qed
  moreover from A1 A2 have "abs(f`(pq)\<rs>pf`(q)) \<lsq> (abs(p)\<ra>𝟭)maxδ(f)"
    using Int_ZF_2_2_L4 by simp
  ultimately have 
    "abs(qf`(p)\<rs>f`(pq)\<ra>(f`(pq)\<rs>pf`(q))) \<lsq> (abs(q)\<ra>𝟭)maxδ(f)\<ra>(abs(p)\<ra>𝟭)maxδ(f)"
    using Int_ZF_2_L21 by simp
  with T show ?thesis using Int_ZF_1_2_L9 int_zero_one_are_int Int_ZF_1_2_L10
    by simp
qed

textThis is Lemma 8 in the Arthan's paper.

lemma (in int1) Arthan_Lem_8: assumes A1: "f𝒮"
  shows "A B. A  B  (p. abs(f`(p)) \<lsq> Aabs(p)\<ra>B)"
proof -
  let ?A = "maxδ(f) \<ra> abs(f`(𝟭))"
  let ?B = "𝟯maxδ(f)"
  from A1 have "?A" "?B"
    using int_zero_one_are_int Int_ZF_1_1_L5 Int_ZF_2_1_L2B 
      Int_ZF_2_1_L7 Int_ZF_2_L14 by auto
  moreover have "p. abs(f`(p)) \<lsq> ?Aabs(p)\<ra>?B"
  proof
    fix p assume A2: "p" 
    with A1 have T: 
      "f`(p)  "  "abs(p)  "  "f`(𝟭)  " 
      "pf`(𝟭)  "  "𝟯"  "maxδ(f)  "
      using Int_ZF_2_1_L2B Int_ZF_2_L14 int_zero_one_are_int 
	Int_ZF_1_1_L5 Int_ZF_2_1_L7 by auto
    from A1 A2 have 
      "abs(𝟭f`(p)\<rs>pf`(𝟭)) \<lsq> (abs(p)\<ra>abs(𝟭)\<ra>𝟮)maxδ(f)"
      using int_zero_one_are_int Arthan_Lem_7 by simp
    with T have "abs(f`(p)) \<lsq> abs(pf`(𝟭))\<ra>(abs(p)\<ra>𝟯)maxδ(f)"
      using Int_ZF_2_L16A Int_ZF_1_1_L4 Int_ZF_1_2_L11 
	Int_triangle_ineq2 by simp
    with A2 T show "abs(f`(p)) \<lsq> ?Aabs(p)\<ra>?B"
      using Int_ZF_1_3_L14 by simp
  qed
  ultimately show ?thesis by auto
qed

textIf $f$ and $g$ are slopes, then $f\circ g$ is equivalent 
  (almost equal) to $g\circ f$. This is Theorem 9 in Arthan's paper \cite{Arthan2004}.

theorem (in int1) Arthan_Th_9: assumes A1: "f𝒮"  "g𝒮"
  shows "fg  gf"
proof -
   from A1 have 
      "A B. A  B  (p. abs(f`(p)) \<lsq> Aabs(p)\<ra>B)"
      "C D. C  D  (p. abs(g`(p)) \<lsq> Cabs(p)\<ra>D)"
      using Arthan_Lem_8 by auto
    then obtain A B C D where D1: "A" "B" "C" "D" and D2: 
      "p. abs(f`(p)) \<lsq> Aabs(p)\<ra>B"
      "p. abs(g`(p)) \<lsq> Cabs(p)\<ra>D"
      by auto
    let ?E = "maxδ(g)(A\<ra>𝟭) \<ra> maxδ(f)(C\<ra>𝟭)"
    let ?F = "(Bmaxδ(g) \<ra> 𝟮maxδ(g)) \<ra> (Dmaxδ(f) \<ra> 𝟮maxδ(f))"
  { fix p assume A2: "p"
    with A1 have T1:
      "g`(p)  "  "f`(p)  "  "abs(p)  "  "𝟮  "
      "f`(g`(p))  "  "g`(f`(p))  "  "f`(g`(p)) \<rs> g`(f`(p))  "
      "pf`(g`(p))  "  "pg`(f`(p))  "
      "abs(f`(g`(p))\<rs>g`(f`(p)))  "
      using Int_ZF_2_1_L2B Int_ZF_2_1_L10 Int_ZF_1_1_L5 Int_ZF_2_L14 int_two_three_are_int
      by auto
    with A1 A2 have
      "abs((f`(g`(p))\<rs>g`(f`(p)))p) \<lsq> 
      (abs(p)\<ra>abs(f`(p))\<ra>𝟮)maxδ(g) \<ra> (abs(p)\<ra>abs(g`(p))\<ra>𝟮)maxδ(f)"
      using Arthan_Lem_7 Int_ZF_1_2_L10A Int_ZF_1_2_L12 by simp
    moreover have 
      "(abs(p)\<ra>abs(f`(p))\<ra>𝟮)maxδ(g) \<ra> (abs(p)\<ra>abs(g`(p))\<ra>𝟮)maxδ(f) \<lsq>
      ((maxδ(g)(A\<ra>𝟭) \<ra> maxδ(f)(C\<ra>𝟭)))abs(p) \<ra>
      ((Bmaxδ(g) \<ra> 𝟮maxδ(g)) \<ra> (Dmaxδ(f) \<ra> 𝟮maxδ(f)))"
    proof -
      from D2 A2 T1 have 
	"abs(p)\<ra>abs(f`(p))\<ra>𝟮 \<lsq> abs(p)\<ra>(Aabs(p)\<ra>B)\<ra>𝟮"
	"abs(p)\<ra>abs(g`(p))\<ra>𝟮 \<lsq> abs(p)\<ra>(Cabs(p)\<ra>D)\<ra>𝟮"
	using Int_ZF_2_L15C by auto
      with A1 have 
	"(abs(p)\<ra>abs(f`(p))\<ra>𝟮)maxδ(g) \<lsq> (abs(p)\<ra>(Aabs(p)\<ra>B)\<ra>𝟮)maxδ(g)"
	"(abs(p)\<ra>abs(g`(p))\<ra>𝟮)maxδ(f) \<lsq> (abs(p)\<ra>(Cabs(p)\<ra>D)\<ra>𝟮)maxδ(f)"
	using Int_ZF_2_1_L8 Int_ZF_1_3_L13 by auto
      moreover from A1 D1 T1 have 
	"(abs(p)\<ra>(Aabs(p)\<ra>B)\<ra>𝟮)maxδ(g) = 
	maxδ(g)(A\<ra>𝟭)abs(p) \<ra> (Bmaxδ(g) \<ra> 𝟮maxδ(g))"
	"(abs(p)\<ra>(Cabs(p)\<ra>D)\<ra>𝟮)maxδ(f) = 
	maxδ(f)(C\<ra>𝟭)abs(p) \<ra> (Dmaxδ(f) \<ra> 𝟮maxδ(f))"
	using Int_ZF_2_1_L8 Int_ZF_1_2_L13 by auto
      ultimately have 
	"(abs(p)\<ra>abs(f`(p))\<ra>𝟮)maxδ(g) \<ra> (abs(p)\<ra>abs(g`(p))\<ra>𝟮)maxδ(f) \<lsq>
	(maxδ(g)(A\<ra>𝟭)abs(p) \<ra> (Bmaxδ(g) \<ra> 𝟮maxδ(g))) \<ra> 
	(maxδ(f)(C\<ra>𝟭)abs(p) \<ra> (Dmaxδ(f) \<ra> 𝟮maxδ(f)))"
	using int_ineq_add_sides by simp
      moreover from A1 A2 D1 have "abs(p)  "
	"maxδ(g)(A\<ra>𝟭)  "  "Bmaxδ(g) \<ra> 𝟮maxδ(g)  "
	"maxδ(f)(C\<ra>𝟭)  "  "Dmaxδ(f) \<ra> 𝟮maxδ(f)  " 
	using Int_ZF_2_L14 Int_ZF_2_1_L8 int_zero_one_are_int 
	  Int_ZF_1_1_L5 int_two_three_are_int by auto
      ultimately show ?thesis using Int_ZF_1_2_L14 by simp
    qed
    ultimately have
      "abs((f`(g`(p))\<rs>g`(f`(p)))p) \<lsq> ?Eabs(p) \<ra> ?F"
      by (rule Int_order_transitive)
    with A2 T1 have
      "abs(f`(g`(p))\<rs>g`(f`(p)))abs(p) \<lsq> ?Eabs(p) \<ra> ?F"
      "abs(f`(g`(p))\<rs>g`(f`(p)))  "
      using Int_ZF_1_3_L5 by auto
  } then have 
      "p. abs(f`(g`(p))\<rs>g`(f`(p)))  "
      "p. abs(f`(g`(p))\<rs>g`(f`(p)))abs(p) \<lsq> ?Eabs(p) \<ra> ?F"
    by auto
  moreover from A1 D1 have "?E  "  "?F  "
    using int_zero_one_are_int int_two_three_are_int Int_ZF_2_1_L8 Int_ZF_1_1_L5
    by auto
  ultimately have
    "L. p. abs(f`(g`(p))\<rs>g`(f`(p))) \<lsq> L"
    by (rule Int_ZF_1_7_L1)
  with A1 obtain L where "p. abs((fg)`(p)\<rs>(gf)`(p)) \<lsq> L"
    using Int_ZF_2_1_L10 by auto
  moreover from A1 have "fg  𝒮"  "gf  𝒮"
    using Int_ZF_2_1_L11 by auto
  ultimately show "fg  gf" using Int_ZF_2_1_L9 by auto
qed
 
end