Theory Int_ZF_3

(*    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 3

theory Int_ZF_3 imports Int_ZF_2

begin

textThis theory is a continuation of Int_ZF_2›. We consider 
  here the properties of slopes (almost homomorphisms on integers)
  that allow to define the order relation and multiplicative
  inverse on real numbers. We also prove theorems that allow to show 
  completeness of the order relation of real numbers we define in Real_ZF›.


subsectionPositive slopes

textThis section provides background material for defining the order relation on real numbers.

textPositive slopes are functions (of course.)

lemma (in int1) Int_ZF_2_3_L1: assumes A1: "f𝒮+" shows "f:"
  using assms AlmostHoms_def PositiveSet_def by simp

textA small technical lemma to simplify the proof of the next theorem.

lemma (in int1) Int_ZF_2_3_L1A: 
  assumes A1: "f𝒮+" and A2: "n  f``(+)  +. a\<lsq>n"
  shows "M+. a \<lsq> f`(M)"
proof -
 from A1 have "f:"  "+  " 
    using AlmostHoms_def PositiveSet_def by auto
 with A2 show ?thesis using func_imagedef by auto
qed

textThe next lemma is Lemma 3 in the Arthan's paper.

lemma (in int1) Arthan_Lem_3: 
  assumes A1: "f𝒮+" and A2: "D  +"
  shows "M+. m+. (m\<ra>𝟭)D \<lsq> f`(mM)" 
proof -
  let ?E = "maxδ(f) \<ra> D"
  let ?A = "f``(+)  +"
  from A1 A2 have I: "D\<lsq>?E"
    using Int_ZF_1_5_L3 Int_ZF_2_1_L8 Int_ZF_2_L1A Int_ZF_2_L15D
    by simp
  from A1 A2 have "?A  +"  "?A  Fin()"  "𝟮?E  " 
    using int_two_three_are_int Int_ZF_2_1_L8 PositiveSet_def Int_ZF_1_1_L5
    by auto
  with A1 have "M+.  𝟮?E \<lsq> f`(M)"
    using Int_ZF_1_5_L2A Int_ZF_2_3_L1A by simp
  then obtain M where II: "M+"  and III: "𝟮?E \<lsq> f`(M)"
    by auto
  { fix m assume "m+" then have A4: "𝟭\<lsq>m"
      using Int_ZF_1_5_L3 by simp
    moreover from II III have "(𝟭\<ra>𝟭) ?E \<lsq> f`(𝟭M)"
      using PositiveSet_def Int_ZF_1_1_L4 by simp
    moreover have "k. 
      𝟭\<lsq>k  (k\<ra>𝟭)?E \<lsq> f`(kM)  (k\<ra>𝟭\<ra>𝟭)?E \<lsq> f`((k\<ra>𝟭)M)"
    proof -
      { fix k assume A5: "𝟭\<lsq>k"  and A6: "(k\<ra>𝟭)?E \<lsq> f`(kM)"
	with A1 A2 II have T:
	  "k"  "M"  "k\<ra>𝟭  "  "?E"  "(k\<ra>𝟭)?E  "  "𝟮?E  "
	  using Int_ZF_2_L1A PositiveSet_def int_zero_one_are_int 
	    Int_ZF_1_1_L5 Int_ZF_2_1_L8 by auto
	from A1 A2 A5 II have 
	  "δ(f,kM,M)  "   "abs(δ(f,kM,M)) \<lsq> maxδ(f)"   "𝟬\<lsq>D"
	  using Int_ZF_2_L1A PositiveSet_def Int_ZF_1_1_L5 
	    Int_ZF_2_1_L7 Int_ZF_2_L16C by auto
	with III A6 have 
	  "(k\<ra>𝟭)?E \<ra> (𝟮?E \<rs> ?E) \<lsq> f`(kM) \<ra> (f`(M) \<ra> δ(f,kM,M))"
	  using Int_ZF_1_3_L19A int_ineq_add_sides by simp
	with A1 T have "(k\<ra>𝟭\<ra>𝟭)?E \<lsq> f`((k\<ra>𝟭)M)"
	  using Int_ZF_1_1_L1 int_zero_one_are_int Int_ZF_1_1_L4 
	    Int_ZF_1_2_L11 Int_ZF_2_1_L13 by simp
      } then show ?thesis by simp
    qed
    ultimately have "(m\<ra>𝟭)?E \<lsq> f`(mM)" by (rule Induction_on_int)
    with A4 I have "(m\<ra>𝟭)D \<lsq> f`(mM)" using Int_ZF_1_3_L13A
      by simp
  } then have "m+.(m\<ra>𝟭)D \<lsq> f`(mM)" by simp
  with II show ?thesis by auto
qed

textA special case of  Arthan_Lem_3› when $D=1$.

corollary (in int1) Arthan_L_3_spec: assumes A1: "f  𝒮+"
  shows "M+.n+. n\<ra>𝟭 \<lsq> f`(nM)"
proof -
  have "n+. n\<ra>𝟭  "
    using PositiveSet_def int_zero_one_are_int Int_ZF_1_1_L5
    by simp
  then have "n+. (n\<ra>𝟭)𝟭 = n\<ra>𝟭"
    using Int_ZF_1_1_L4 by simp
  moreover from A1 have "M+. n+. (n\<ra>𝟭)𝟭 \<lsq> f`(nM)" 
    using int_one_two_are_pos Arthan_Lem_3 by simp
  ultimately show ?thesis by simp
qed

textWe know  from Group_ZF_3.thy› that finite range functions are almost homomorphisms. 
  Besides reminding that fact for slopes the next lemma shows 
  that finite range functions do not belong to 𝒮+. 
  This is important, because the projection
  of the set of finite range functions defines zero in the real number construction in Real_ZF_x.thy› 
  series, while the projection of 𝒮+ becomes the set of (strictly) positive reals. 
  We don't want zero to be positive, do we? The next lemma is a part of Lemma 5 in the Arthan's paper 
  \cite{Arthan2004}.

lemma (in int1) Int_ZF_2_3_L1B: 
  assumes A1: "f  FinRangeFunctions(,)"
  shows "f𝒮"   "f  𝒮+"
proof -
  from A1 show "f𝒮" using Int_ZF_2_1_L1 group1.Group_ZF_3_3_L1
    by auto
  have "+  " using PositiveSet_def by auto
  with A1 have "f``(+)  Fin()"
    using Finite1_L21 by simp
  then have "f``(+)  +  Fin()"
    using Fin_subset_lemma by blast
  thus "f  𝒮+" by auto
qed

textWe want to show that if $f$ is a slope and neither $f$ nor $-f$ are in 𝒮+, 
  then $f$ is bounded. The next lemma is the first step towards that goal and 
  shows that if slope is not in 𝒮+ then $f($+$)$ is bounded above.

lemma (in int1) Int_ZF_2_3_L2: assumes A1: "f𝒮" and A2: "f  𝒮+"
  shows "IsBoundedAbove(f``(+), IntegerOrder)"
proof -
  from A1 have "f:" using AlmostHoms_def by simp
  then have "f``(+)  " using func1_1_L6 by simp
  moreover from A1 A2 have "f``(+)  +  Fin()" by auto
  ultimately show ?thesis using Int_ZF_2_T1 group3.OrderedGroup_ZF_2_L4
    by simp
qed

textIf $f$ is a slope and $-f\notin$ 𝒮+, then 
  $f($+$)$ is bounded below.

lemma (in int1) Int_ZF_2_3_L3: assumes A1: "f𝒮" and A2: "\<fm>f  𝒮+"
  shows "IsBoundedBelow(f``(+), IntegerOrder)"
proof -
  from A1 have T: "f:" using AlmostHoms_def by simp
  then have "(\<sm>(f``(+))) = (\<fm>f)``(+)"
    using Int_ZF_1_T2 group0_2_T2 PositiveSet_def func1_1_L15C
    by auto
  with A1 A2 T show "IsBoundedBelow(f``(+), IntegerOrder)"
    using Int_ZF_2_1_L12 Int_ZF_2_3_L2 PositiveSet_def func1_1_L6 
      Int_ZF_2_T1 group3.OrderedGroup_ZF_2_L5 by simp
qed

textA slope that is bounded on + is bounded everywhere.

lemma (in int1) Int_ZF_2_3_L4: 
  assumes A1: "f𝒮" and A2: "m" 
  and A3: "n+. abs(f`(n)) \<lsq> L"
  shows "abs(f`(m)) \<lsq> 𝟮maxδ(f) \<ra> L"
proof -
  from A1 A3 have 
    "𝟬 \<lsq> abs(f`(𝟭))"  "abs(f`(𝟭)) \<lsq> L"
    using int_zero_one_are_int Int_ZF_2_1_L2B int_abs_nonneg int_one_two_are_pos
    by auto
  then have II: "𝟬\<lsq>L" by (rule Int_order_transitive)
  note A2
  moreover have "abs(f`(𝟬)) \<lsq> 𝟮maxδ(f) \<ra> L"
  proof -
    from A1 have 
      "abs(f`(𝟬)) \<lsq> maxδ(f)"  "𝟬 \<lsq> maxδ(f)" 
      and T: "maxδ(f)  "
      using Int_ZF_2_1_L8 by auto
    with II have "abs(f`(𝟬)) \<lsq> maxδ(f) \<ra> maxδ(f) \<ra> L"
      using Int_ZF_2_L15F by simp
    with T show ?thesis using Int_ZF_1_1_L4 by simp
  qed
  moreover from A1 A3 II have 
    "n+. abs(f`(n)) \<lsq> 𝟮maxδ(f) \<ra> L"
    using Int_ZF_2_1_L8 Int_ZF_1_3_L5A Int_ZF_2_L15F 
    by simp
  moreover have "n+. abs(f`(\<rm>n)) \<lsq> 𝟮maxδ(f) \<ra> L"
  proof
    fix n assume "n+"
    with A1 A3 have
      "𝟮maxδ(f)  "
      "abs(f`(\<rm>n)) \<lsq> 𝟮maxδ(f) \<ra> abs(f`(n))"
      "abs(f`(n)) \<lsq> L"
      using int_two_three_are_int Int_ZF_2_1_L8 Int_ZF_1_1_L5
	PositiveSet_def Int_ZF_2_1_L14 by auto
    then show "abs(f`(\<rm>n)) \<lsq> 𝟮maxδ(f) \<ra> L"
      using Int_ZF_2_L15A by blast
  qed    
  ultimately show ?thesis by (rule Int_ZF_2_L19B)
qed

textA slope whose image of the set of positive integers is bounded
  is a finite range function.

lemma (in int1) Int_ZF_2_3_L4A: 
  assumes A1: "f𝒮" and A2: "IsBounded(f``(+), IntegerOrder)"
  shows "f  FinRangeFunctions(,)"
proof -
  have T1: "+  " using PositiveSet_def by auto
  from A1 have T2: "f:" using AlmostHoms_def by simp
  from A2 obtain L where "af``(+). abs(a) \<lsq> L"
    using Int_ZF_1_3_L20A by auto
  with T2 T1 have "n+. abs(f`(n)) \<lsq> L"
    by (rule func1_1_L15B)
  with A1 have "m. abs(f`(m)) \<lsq> 𝟮maxδ(f) \<ra> L"
    using Int_ZF_2_3_L4 by simp
  with T2 have "f``()  Fin()"
    by (rule Int_ZF_1_3_L20C)
  with T2 show "f  FinRangeFunctions(,)"
    using FinRangeFunctions_def by simp
qed

textA slope whose image of the set of positive integers is bounded
  below is a finite range function or a positive slope.

lemma (in int1) Int_ZF_2_3_L4B: 
  assumes "f𝒮" and "IsBoundedBelow(f``(+), IntegerOrder)"
  shows "f  FinRangeFunctions(,)  f𝒮+"
  using assms Int_ZF_2_3_L2 IsBounded_def Int_ZF_2_3_L4A
  by auto

textIf one slope is not greater then another on positive integers,
  then they are almost equal or the difference is a positive slope.

lemma (in int1) Int_ZF_2_3_L4C: assumes A1: "f𝒮"  "g𝒮" and
  A2: "n+. f`(n) \<lsq> g`(n)"
  shows "fg  g \<fp> (\<fm>f)  𝒮+"
proof -
  let ?h = "g \<fp> (\<fm>f)"
  from A1 have "(\<fm>f)  𝒮" using Int_ZF_2_1_L12 
    by simp
  with A1 have I: "?h  𝒮" using Int_ZF_2_1_L12C 
    by simp
  moreover have "IsBoundedBelow(?h``(+), IntegerOrder)"
  proof -
    from I have 
      "?h:" and "+" using AlmostHoms_def PositiveSet_def
      by auto
    moreover from A1 A2 have "n+. 𝟬, ?h`(n)  IntegerOrder"
      using Int_ZF_2_1_L2B PositiveSet_def Int_ZF_1_3_L10A 
	Int_ZF_2_1_L12 Int_ZF_2_1_L12B Int_ZF_2_1_L12A
      by simp
    ultimately show "IsBoundedBelow(?h``(+), IntegerOrder)"
      by (rule func_ZF_8_L1)
  qed
  ultimately have "?h  FinRangeFunctions(,)  ?h𝒮+"
    using Int_ZF_2_3_L4B by simp
  with A1 show "fg  g \<fp> (\<fm>f)  𝒮+"
    using Int_ZF_2_1_L9C by auto
qed
  
textPositive slopes are arbitrarily large for large enough arguments.

lemma (in int1) Int_ZF_2_3_L5: 
  assumes A1: "f𝒮+" and A2: "K"
  shows "N+. m. N\<lsq>m  K \<lsq> f`(m)"
proof -
  from A1 obtain M where I: "M+" and II: "n+. n\<ra>𝟭 \<lsq> f`(nM)"
    using Arthan_L_3_spec by auto
  let ?j = "GreaterOf(IntegerOrder,M,K \<rs> (minf(f,𝟬..(M\<rs>𝟭)) \<rs> maxδ(f)) \<rs> 𝟭)"
  from A1 I have T1: 
    "minf(f,𝟬..(M\<rs>𝟭)) \<rs> maxδ(f)  "  "M"
    using Int_ZF_2_1_L15 Int_ZF_2_1_L8 Int_ZF_1_1_L5 PositiveSet_def
    by auto
  with A2 I have T2: 
    "K \<rs> (minf(f,𝟬..(M\<rs>𝟭)) \<rs> maxδ(f))  "
    "K \<rs> (minf(f,𝟬..(M\<rs>𝟭)) \<rs> maxδ(f)) \<rs> 𝟭  "
    using Int_ZF_1_1_L5 int_zero_one_are_int by auto
  with T1 have III: "M \<lsq> ?j"  and 
    "K \<rs> (minf(f,𝟬..(M\<rs>𝟭)) \<rs> maxδ(f)) \<rs> 𝟭 \<lsq> ?j"
    using Int_ZF_1_3_L18 by auto
  with A2 T1 T2 have 
    IV: "K \<lsq> ?j\<ra>𝟭 \<ra> (minf(f,𝟬..(M\<rs>𝟭)) \<rs> maxδ(f))"
    using int_zero_one_are_int Int_ZF_2_L9C by simp
  let ?N = "GreaterOf(IntegerOrder,𝟭,?jM)"
  from T1 III have T3: "?j  "  "?jM  "
    using Int_ZF_2_L1A Int_ZF_1_1_L5 by auto
  then have V: "?N  +" and VI: "?jM \<lsq> ?N"
    using int_zero_one_are_int Int_ZF_1_5_L3 Int_ZF_1_3_L18 
    by auto
  { fix m
    let ?n = "m zdiv M"
    let ?k = "m zmod M"
    assume "?N\<lsq>m"
    with VI have "?jM \<lsq> m" by (rule Int_order_transitive)
    with I III have 
      VII: "m = ?nM\<ra>?k" 
      "?j \<lsq> ?n"  and 
      VIII: "?n  +"  "?k  𝟬..(M\<rs>𝟭)"
      using IntDiv_ZF_1_L5 by auto
    with II have 
      "?j \<ra> 𝟭 \<lsq> ?n \<ra> 𝟭"  "?n\<ra>𝟭 \<lsq> f`(?nM)"
      using int_zero_one_are_int int_ord_transl_inv by auto
    then have "?j \<ra> 𝟭 \<lsq>  f`(?nM)"
      by (rule Int_order_transitive)
    with T1 have 
      "?j\<ra>𝟭 \<ra> (minf(f,𝟬..(M\<rs>𝟭)) \<rs> maxδ(f)) \<lsq>  
      f`(?nM) \<ra> (minf(f,𝟬..(M\<rs>𝟭)) \<rs> maxδ(f))"
      using int_ord_transl_inv by simp
    with IV have "K \<lsq> f`(?nM) \<ra> (minf(f,𝟬..(M\<rs>𝟭)) \<rs> maxδ(f))"
      by (rule Int_order_transitive)
    moreover from A1 I VIII have
      "f`(?nM) \<ra> (minf(f,𝟬..(M\<rs>𝟭))\<rs> maxδ(f)) \<lsq> f`(?nM\<ra>?k)"
      using PositiveSet_def Int_ZF_2_1_L16 by simp
    ultimately have "K \<lsq> f`(?nM\<ra>?k)"
      by (rule Int_order_transitive)
    with VII have "K \<lsq> f`(m)" by simp
    } then have  "m. ?N\<lsq>m  K \<lsq> f`(m)"
      by simp
    with V show ?thesis by auto
qed

textPositive slopes are arbitrarily small for small enough arguments.
  Kind of dual to Int_ZF_2_3_L5›.

lemma (in int1) Int_ZF_2_3_L5A: assumes A1: "f𝒮+" and A2: "K"
  shows "N+. m. N\<lsq>m  f`(\<rm>m) \<lsq> K"
proof -
  from A1 have T1: "abs(f`(𝟬)) \<ra> maxδ(f)  "
    using Int_ZF_2_1_L8 by auto
  with A2 have "abs(f`(𝟬)) \<ra> maxδ(f) \<rs> K  "
    using Int_ZF_1_1_L5 by simp
  with A1 have 
    "N+. m. N\<lsq>m  abs(f`(𝟬)) \<ra> maxδ(f) \<rs> K \<lsq> f`(m)"
    using Int_ZF_2_3_L5 by simp
  then obtain N where I: "N+" and II:
    "m. N\<lsq>m   abs(f`(𝟬)) \<ra> maxδ(f) \<rs> K \<lsq> f`(m)"
    by auto
  { fix m assume A3: "N\<lsq>m"
    with A1 have
      "f`(\<rm>m) \<lsq> abs(f`(𝟬)) \<ra> maxδ(f) \<rs> f`(m)"
      using Int_ZF_2_L1A Int_ZF_2_1_L14 by simp
    moreover
    from II T1 A3 have "abs(f`(𝟬)) \<ra> maxδ(f) \<rs> f`(m) \<lsq> 
      (abs(f`(𝟬)) \<ra> maxδ(f)) \<rs>(abs(f`(𝟬)) \<ra> maxδ(f) \<rs> K)"
      using Int_ZF_2_L10 int_ord_transl_inv by simp
    with A2 T1 have "abs(f`(𝟬)) \<ra> maxδ(f) \<rs> f`(m) \<lsq> K"
      using Int_ZF_1_2_L3 by simp
    ultimately have "f`(\<rm>m) \<lsq> K"
      by (rule Int_order_transitive)
  } then have "m. N\<lsq>m   f`(\<rm>m) \<lsq> K"
    by simp
  with I show ?thesis by auto
qed

(*lemma (in int1) Int_ZF_2_3_L5A: assumes A1: "f∈𝒮+" and A2: "K∈ℤ"
  shows "∃N∈ℤ+. ∀m. m\<lsq>(\<rm>N) ⟶ f`(m) \<lsq> K"
proof -
  from A1 have T1: "abs(f`(𝟬)) \<ra> maxδ(f) ∈ ℤ"
    using Int_ZF_2_1_L8 by auto;
  with A2 have "abs(f`(𝟬)) \<ra> maxδ(f) \<rs> K ∈ ℤ"
    using Int_ZF_1_1_L5 by simp;
  with A1 have 
    "∃N∈ℤ+. ∀m. N\<lsq>m ⟶ abs(f`(𝟬)) \<ra> maxδ(f) \<rs> K \<lsq> f`(m)"
    using Int_ZF_2_3_L5 by simp;
  then obtain N where I: "N∈ℤ+" and II:
    "∀m. N\<lsq>m ⟶  abs(f`(𝟬)) \<ra> maxδ(f) \<rs> K \<lsq> f`(m)"
    by auto;
  { fix m assume A3: "m\<lsq>(\<rm>N)"
    with A1 have T2: "f`(m) ∈ ℤ"
      using Int_ZF_2_L1A Int_ZF_2_1_L2B by simp;
    from A1 I II A3 have
      "abs(f`(𝟬)) \<ra> maxδ(f) \<rs> K \<lsq> f`(\<rm>m)" and
      "f`(\<rm>m) \<lsq> abs(f`(𝟬)) \<ra> maxδ(f) \<rs> f`(m)"
       using PositiveSet_def Int_ZF_2_L10AA Int_ZF_2_L1A Int_ZF_2_1_L14
       by auto;
    then have 
      "abs(f`(𝟬)) \<ra> maxδ(f) \<rs> K \<lsq> abs(f`(𝟬)) \<ra> maxδ(f) \<rs> f`(m)"
      by (rule Int_order_transitive)
    with T1 A2 T2 have "f`(m) \<lsq> K"
      using Int_ZF_2_L10AB by simp; 
  } then have "∀m. m\<lsq>(\<rm>N) ⟶ f`(m) \<lsq> K"
    by simp;
  with I show ?thesis by auto;
qed;*)

textA special case of Int_ZF_2_3_L5› where $K=1$.

corollary (in int1) Int_ZF_2_3_L6: assumes "f𝒮+"
  shows "N+. m. N\<lsq>m  f`(m)  +"
  using assms int_zero_one_are_int Int_ZF_2_3_L5 Int_ZF_1_5_L3
  by simp

textA special case of Int_ZF_2_3_L5› where $m=N$. 

corollary (in int1) Int_ZF_2_3_L6A: assumes "f𝒮+" and "K"
   shows "N+. K \<lsq> f`(N)"
proof -
  from assms have "N+. m. N\<lsq>m  K \<lsq> f`(m)"
    using Int_ZF_2_3_L5 by simp
  then obtain N where I: "N  +"  and II: "m. N\<lsq>m  K \<lsq> f`(m)"
    by auto
  then show ?thesis using PositiveSet_def int_ord_is_refl refl_def
    by auto
qed

textIf values of a slope are not bounded above, 
  then the slope is positive.

lemma (in int1) Int_ZF_2_3_L7: assumes A1: "f𝒮" 
  and A2: "K. n+. K \<lsq> f`(n)"
  shows "f  𝒮+"
proof -
  { fix K assume "K"
    with A2 obtain n where "n+"  "K \<lsq> f`(n)"
      by auto
    moreover from A1 have "+  "  "f:" 
      using PositiveSet_def AlmostHoms_def by auto
    ultimately have "m  f``(+). K \<lsq> m" 
      using func1_1_L15D by auto
  } then have "K. m  f``(+). K \<lsq> m" by simp
  with A1 show "f  𝒮+" using Int_ZF_4_L9 Int_ZF_2_3_L2
    by auto
qed

textFor unbounded slope $f$ either $f\in$𝒮+ of 
  $-f\in$𝒮+.

theorem (in int1) Int_ZF_2_3_L8:
  assumes A1: "f𝒮" and A2: "f  FinRangeFunctions(,)"
  shows "(f  𝒮+) Xor ((\<fm>f)  𝒮+)"
proof -
  have T1: "+  " using PositiveSet_def by auto
  from A1 have T2: "f:"  using AlmostHoms_def by simp
  then have I: "f``(+)  " using func1_1_L6 by auto
  from A1 A2 have "f  𝒮+  (\<fm>f)  𝒮+"
    using Int_ZF_2_3_L2 Int_ZF_2_3_L3 IsBounded_def Int_ZF_2_3_L4A
    by blast
  moreover have "¬(f  𝒮+  (\<fm>f)  𝒮+)"
  proof -
    { assume A3: "f  𝒮+"  and A4: "(\<fm>f)  𝒮+"
      from A3 obtain N1 where 
	I: "N1+" and II: "m. N1\<lsq>m  f`(m)  +"
	using Int_ZF_2_3_L6 by auto
      from A4 obtain N2 where 
	III: "N2+" and IV: "m. N2\<lsq>m  (\<fm>f)`(m)  +"
	using Int_ZF_2_3_L6 by auto
      let ?N = "GreaterOf(IntegerOrder,N1,N2)"
      from I III have "N1 \<lsq> ?N"  "N2 \<lsq> ?N"
	using PositiveSet_def Int_ZF_1_3_L18 by auto
      with A1 II IV have
	"f`(?N)  +"  "(\<fm>f)`(?N)  +"  "(\<fm>f)`(?N) = \<rm>(f`(?N))"
	using Int_ZF_2_L1A PositiveSet_def Int_ZF_2_1_L12A
	by auto
      then have False using Int_ZF_1_5_L8 by simp
    } thus ?thesis by auto
  qed
  ultimately show "(f  𝒮+) Xor ((\<fm>f)  𝒮+)"
    using Xor_def by simp
qed

textThe sum of positive slopes is a positive slope.

theorem (in int1) sum_of_pos_sls_is_pos_sl: 
  assumes A1: "f  𝒮+"  "g  𝒮+"
  shows "f\<fp>g  𝒮+"
proof -
  { fix K assume "K"
    with A1 have "N+. m. N\<lsq>m  K \<lsq> f`(m)"
      using Int_ZF_2_3_L5 by simp
    then obtain N where I: "N+" and II: "m. N\<lsq>m  K \<lsq> f`(m)"
      by auto
    from A1 have "M+. m. M\<lsq>m  𝟬 \<lsq> g`(m)"
      using int_zero_one_are_int Int_ZF_2_3_L5 by simp
    then obtain M where III: "M+" and IV: "m. M\<lsq>m  𝟬 \<lsq> g`(m)"
      by auto
    let ?L = "GreaterOf(IntegerOrder,N,M)"
    from I III have V: "?L  +"  "+  " 
      using GreaterOf_def PositiveSet_def by auto
    moreover from A1 V have "(f\<fp>g)`(?L) = f`(?L) \<ra> g`(?L)"
      using Int_ZF_2_1_L12B by auto
    moreover from I II III IV have "K \<lsq> f`(?L) \<ra> g`(?L)"
      using PositiveSet_def Int_ZF_1_3_L18 Int_ZF_2_L15F
      by simp
    ultimately have "?L  +"  "K \<lsq> (f\<fp>g)`(?L)"
      by auto
    then have "n +. K \<lsq> (f\<fp>g)`(n)"
      by auto
  } with A1 show "f\<fp>g  𝒮+"
    using Int_ZF_2_1_L12C Int_ZF_2_3_L7 by simp
qed

textThe composition of positive slopes is a positive slope.

theorem (in int1) comp_of_pos_sls_is_pos_sl: 
  assumes A1: "f  𝒮+"  "g  𝒮+"
  shows "fg  𝒮+"
proof -
  { fix K assume "K"
    with A1 have "N+. m. N\<lsq>m  K \<lsq> f`(m)"
      using Int_ZF_2_3_L5 by simp
    then obtain N where "N+" and I: "m. N\<lsq>m  K \<lsq> f`(m)"
      by auto
    with A1 have "M+. N \<lsq> g`(M)"
      using PositiveSet_def Int_ZF_2_3_L6A by simp
    then obtain M where "M+"  "N \<lsq> g`(M)"
      by auto
    with A1 I have "M+. K \<lsq> (fg)`(M)"
      using PositiveSet_def Int_ZF_2_1_L10
      by auto
  } with A1 show "fg  𝒮+"
    using Int_ZF_2_1_L11 Int_ZF_2_3_L7
    by simp
qed

textA slope equivalent to a positive one is positive.

lemma (in int1) Int_ZF_2_3_L9: 
  assumes A1: "f  𝒮+" and A2: "f,g  AlEqRel" shows "g  𝒮+"
proof -
  from A2 have T: "g𝒮" and "L. m. abs(f`(m)\<rs>g`(m)) \<lsq> L"
    using Int_ZF_2_1_L9A by auto
   then obtain L where 
     I: "L"  and II: "m. abs(f`(m)\<rs>g`(m)) \<lsq> L"
     by auto
  { fix K assume A3: "K"
    with I have "K\<ra>L  "
      using Int_ZF_1_1_L5 by simp
    with A1 obtain M where III: "M+"  and IV: "K\<ra>L \<lsq> f`(M)"
      using Int_ZF_2_3_L6A by auto
    with A1 A3 I have  "K \<lsq> f`(M)\<rs>L"
      using PositiveSet_def Int_ZF_2_1_L2B Int_ZF_2_L9B
      by simp
    moreover from A1 T II III have  
      "f`(M)\<rs>L \<lsq> g`(M)"
      using PositiveSet_def Int_ZF_2_1_L2B Int_triangle_ineq2
      by simp
    ultimately have "K \<lsq>  g`(M)"
      by (rule Int_order_transitive)
    with III have "n+. K \<lsq> g`(n)"
      by auto
  } with T show "g  𝒮+"
    using Int_ZF_2_3_L7 by simp
qed

textThe set of positive slopes is saturated with respect to the relation of 
  equivalence of slopes.

lemma (in int1) pos_slopes_saturated: shows "IsSaturated(AlEqRel,𝒮+)"
proof -
  have 
    "equiv(𝒮,AlEqRel)" 
    "AlEqRel  𝒮 × 𝒮"
    using Int_ZF_2_1_L9B by auto
  moreover have "𝒮+  𝒮" by auto
  moreover have "f𝒮+. g𝒮. f,g  AlEqRel  g  𝒮+"
    using Int_ZF_2_3_L9 by blast
  ultimately show "IsSaturated(AlEqRel,𝒮+)"
    by (rule EquivClass_3_L3)
qed
  
textA technical lemma involving a projection of the set of positive slopes
  and a logical epression with exclusive or.

lemma (in int1) Int_ZF_2_3_L10:
  assumes A1: "f𝒮"  "g𝒮"
  and A2: "R = {AlEqRel``{s}. s𝒮+}"
  and A3: "(f𝒮+) Xor (g𝒮+)"
  shows "(AlEqRel``{f}  R) Xor (AlEqRel``{g}  R)"
proof -
  from A1 A2 A3 have 
    "equiv(𝒮,AlEqRel)" 
    "IsSaturated(AlEqRel,𝒮+)"
    "𝒮+  𝒮"
    "f𝒮"  "g𝒮"
    "R = {AlEqRel``{s}. s𝒮+}"
    "(f𝒮+) Xor (g𝒮+)"
    using pos_slopes_saturated Int_ZF_2_1_L9B by auto
  then show ?thesis by (rule EquivClass_3_L7)
qed

textIdentity function is a positive slope.

lemma (in int1) Int_ZF_2_3_L11: shows "id()  𝒮+"
proof -
  let ?f = "id()"
  { fix K assume "K" 
    then obtain n where T: "n+" and "K\<lsq>n"
      using Int_ZF_1_5_L9 by auto
    moreover from T have "?f`(n) = n"
      using PositiveSet_def by simp
    ultimately have  "n+" and "K\<lsq>?f`(n)"
      by auto
    then have "n+. K\<lsq>?f`(n)" by auto
  } then show "?f  𝒮+"
    using Int_ZF_2_1_L17 Int_ZF_2_3_L7 by simp
qed
      
textThe identity function is not almost equal to any bounded function.

lemma (in int1) Int_ZF_2_3_L12: assumes A1: "f  FinRangeFunctions(,)"
  shows "¬(id()  f)"
proof -
  { from A1 have "id()  𝒮+"
      using Int_ZF_2_3_L11 by simp
    moreover assume "id(),f  AlEqRel"
    ultimately have  "f  𝒮+"
      by (rule Int_ZF_2_3_L9)
    with A1 have False using Int_ZF_2_3_L1B
      by simp
  } then show "¬(id()  f)" by auto
qed

subsectionInverting slopes

textNot every slope is a 1:1 function. However, we can still invert slopes 
  in the sense that if $f$ is a slope, then we can find a slope $g$ such that
  $f\circ g$  is almost equal to the identity function. 
  The goal of this this section is to establish this fact for positive slopes.


textIf $f$ is a positive slope, then for every positive integer $p$ 
  the set $\{n\in Z_+: p\leq f(n)\}$ is a nonempty subset of positive integers.
  Recall that $f^{-1}(p)$ is the notation for the smallest element of this set.

lemma (in int1) Int_ZF_2_4_L1: 
  assumes A1: "f  𝒮+" and A2: "p+" and A3: "A = {n+. p \<lsq> f`(n)}"
  shows 
  "A  +"  
  "A  0"
  "f¯(p)  A"
  "mA. f¯(p) \<lsq> m"
proof -
  from A3 show I: "A  +" by auto
  from A1 A2 have "n+. p \<lsq> f`(n)"
    using PositiveSet_def Int_ZF_2_3_L6A by simp
  with A3 show II: "A  0" by auto
  from A3 I II show 
    "f¯(p)  A"
    "mA. f¯(p) \<lsq> m"
    using Int_ZF_1_5_L1C by auto
qed

textIf $f$ is a positive slope and $p$ is a positive integer $p$, then 
   $f^{-1}(p)$ (defined as the minimum of the set $\{n\in Z_+: p\leq f(n)\}$ ) 
  is a (well defined) positive integer.

lemma (in int1) Int_ZF_2_4_L2: 
  assumes "f  𝒮+" and "p+"
  shows 
  "f¯(p)  +" 
  "p \<lsq> f`(f¯(p))"
  using assms Int_ZF_2_4_L1 by auto

textIf $f$ is a positive slope and $p$ is a positive integer such 
  that $n\leq f(p)$, then
  $f^{-1}(n) \leq p$.

lemma (in int1) Int_ZF_2_4_L3: 
  assumes "f  𝒮+" and  "m+"  "p+" and "m \<lsq> f`(p)"
  shows "f¯(m) \<lsq> p"
  using assms Int_ZF_2_4_L1 by simp

textAn upper bound $f(f^{-1}(m) -1)$ for positive slopes.

lemma (in int1) Int_ZF_2_4_L4: 
  assumes A1: "f  𝒮+" and A2: "m+" and A3: "f¯(m)\<rs>𝟭  +"
  shows "f`(f¯(m)\<rs>𝟭) \<lsq> m"   "f`(f¯(m)\<rs>𝟭)  m"
proof -
  from A1 A2 have T: "f¯(m)  " using Int_ZF_2_4_L2 PositiveSet_def
    by simp
  from A1 A3 have "f:"  and "f¯(m)\<rs>𝟭  "
    using Int_ZF_2_3_L1 PositiveSet_def by auto
  with A1 A2 have T1: "f`(f¯(m)\<rs>𝟭)  "  "m" 
    using apply_funtype PositiveSet_def by auto
  { assume "m \<lsq> f`(f¯(m)\<rs>𝟭)"
    with A1 A2 A3 have "f¯(m) \<lsq> f¯(m)\<rs>𝟭" 
      by (rule Int_ZF_2_4_L3)
    with T have False using Int_ZF_1_2_L3AA
      by simp
  } then have I: "¬(m \<lsq> f`(f¯(m)\<rs>𝟭))" by auto
  with T1 show "f`(f¯(m)\<rs>𝟭) \<lsq> m"   
    by (rule Int_ZF_2_L19)
  from T1 I show "f`(f¯(m)\<rs>𝟭)  m"
    by (rule Int_ZF_2_L19)
qed

textThe (candidate for) the inverse of a positive slope is nondecreasing.

lemma (in int1) Int_ZF_2_4_L5:
  assumes A1: "f  𝒮+" and A2: "m+" and A3: "m\<lsq>n"
  shows "f¯(m) \<lsq> f¯(n)"
proof -
  from A2 A3 have T: "n  +" using Int_ZF_1_5_L7 by blast
  with A1 have "n \<lsq> f`(f¯(n))" using Int_ZF_2_4_L2
    by simp
  with A3 have "m \<lsq> f`(f¯(n))" by (rule Int_order_transitive)
  with A1 A2 T show "f¯(m) \<lsq> f¯(n)"
    using Int_ZF_2_4_L2 Int_ZF_2_4_L3 by simp
qed
  
textIf $f^{-1}(m)$ is positive and $n$ is a positive integer, then, 
  then $f^{-1}(m+n)-1$ is positive.

lemma (in int1) Int_ZF_2_4_L6: 
  assumes A1: "f  𝒮+" and A2: "m+"  "n+" and 
  A3: "f¯(m)\<rs>𝟭  +"
  shows "f¯(m\<ra>n)\<rs>𝟭  +"
proof -
  from A1 A2 have "f¯(m)\<rs>𝟭 \<lsq>  f¯(m\<ra>n) \<rs> 𝟭"
     using PositiveSet_def Int_ZF_1_5_L7A Int_ZF_2_4_L2 
       Int_ZF_2_4_L5 int_zero_one_are_int Int_ZF_1_1_L4 
       int_ord_transl_inv by simp
  with A3 show "f¯(m\<ra>n)\<rs>𝟭  +" using Int_ZF_1_5_L7
    by blast
qed

textIf $f$ is a slope, then $f(f^{-1}(m+n)-f^{-1}(m) - f^{-1}(n))$ is 
  uniformly bounded above and below. Will it be the messiest IsarMathLib
  proof ever? Only time will tell.

lemma (in int1) Int_ZF_2_4_L7:  assumes A1: "f  𝒮+" and
  A2: "m+. f¯(m)\<rs>𝟭  +"
  shows 
  "U. m+. n+. f`(f¯(m\<ra>n)\<rs>f¯(m)\<rs>f¯(n)) \<lsq> U"
  "N. m+. n+. N \<lsq> f`(f¯(m\<ra>n)\<rs>f¯(m)\<rs>f¯(n))"
proof -
  from A1 have "L. r. f`(r) \<lsq> f`(r\<rs>𝟭) \<ra> L"
    using Int_ZF_2_1_L28 by simp
  then obtain L where 
    I: "L" and II: "r. f`(r) \<lsq> f`(r\<rs>𝟭) \<ra> L"
    by auto
  from A1 have 
    "M. r.p.q. f`(r\<rs>p\<rs>q) \<lsq> f`(r)\<rs>f`(p)\<rs>f`(q)\<ra>M"
    "K. r.p.q. f`(r)\<rs>f`(p)\<rs>f`(q)\<ra>K \<lsq> f`(r\<rs>p\<rs>q)"
    using Int_ZF_2_1_L30 by auto
  then obtain M K where III: "M" and 
    IV: "r.p.q. f`(r\<rs>p\<rs>q) \<lsq> f`(r)\<rs>f`(p)\<rs>f`(q)\<ra>M" 
    and 
    V: "K" and VI: "r.p.q. f`(r)\<rs>f`(p)\<rs>f`(q)\<ra>K \<lsq> f`(r\<rs>p\<rs>q)"
    by auto
  from I III V have 
    "L\<ra>M  "  "(\<rm>L) \<rs> L \<ra> K  " 
    using Int_ZF_1_1_L4 Int_ZF_1_1_L5 by auto
  moreover
    { fix m n
      assume A3: "m+" "n+" 
      have  "f`(f¯(m\<ra>n)\<rs>f¯(m)\<rs>f¯(n)) \<lsq>  L\<ra>M  
	(\<rm>L)\<rs>L\<ra>K \<lsq> f`(f¯(m\<ra>n)\<rs>f¯(m)\<rs>f¯(n))"
      proof -
	let ?r = "f¯(m\<ra>n)"
	let ?p = "f¯(m)"
	let ?q = "f¯(n)"
	from A1 A3 have T1:
	  "?p  +"  "?q  +"  "?r  +"
	  using Int_ZF_2_4_L2 pos_int_closed_add_unfolded by auto
	with A3 have T2:
	  "m  "  "n  "  "?p  "  "?q  "  "?r  " 
	  using PositiveSet_def by auto
	from A2 A3 have T3:
	  "?r\<rs>𝟭  +" "?p\<rs>𝟭  +"  "?q\<rs>𝟭  +"
	  using pos_int_closed_add_unfolded by auto
	from A1 A3 have VII:
	  "m\<ra>n \<lsq> f`(?r)"
	  "m \<lsq> f`(?p)"  
	  "n \<lsq> f`(?q)"  
	  using Int_ZF_2_4_L2 pos_int_closed_add_unfolded by auto
	from A1 A3 T3 have VIII:
	  "f`(?r\<rs>𝟭) \<lsq> m\<ra>n"
	  "f`(?p\<rs>𝟭) \<lsq> m"
	  "f`(?q\<rs>𝟭) \<lsq> n"
	  using pos_int_closed_add_unfolded Int_ZF_2_4_L4 by auto
	have "f`(?r\<rs>?p\<rs>?q) \<lsq> L\<ra>M"
	proof -
	  from IV T2 have "f`(?r\<rs>?p\<rs>?q) \<lsq> f`(?r)\<rs>f`(?p)\<rs>f`(?q)\<ra>M"
	    by simp
	  moreover 
	  from I II T2 VIII have
	    "f`(?r) \<lsq> f`(?r\<rs>𝟭) \<ra> L"
	    "f`(?r\<rs>𝟭) \<ra> L \<lsq> m\<ra>n\<ra>L"
	    using int_ord_transl_inv by auto
	  then have "f`(?r) \<lsq>  m\<ra>n\<ra>L"
	    by (rule Int_order_transitive)
	  with VII have "f`(?r) \<rs> f`(?p) \<lsq> m\<ra>n\<ra>L\<rs>m"
	    using int_ineq_add_sides by simp
	  with I T2 VII have "f`(?r) \<rs> f`(?p) \<rs> f`(?q) \<lsq> n\<ra>L\<rs>n"
	    using Int_ZF_1_2_L9 int_ineq_add_sides by simp
	  with I III T2 have "f`(?r) \<rs> f`(?p) \<rs> f`(?q) \<ra> M \<lsq> L\<ra>M"
	    using Int_ZF_1_2_L3 int_ord_transl_inv by simp
	  ultimately show "f`(?r\<rs>?p\<rs>?q) \<lsq> L\<ra>M"
	    by (rule Int_order_transitive)
	qed
	moreover have "(\<rm>L)\<rs>L \<ra>K \<lsq> f`(?r\<rs>?p\<rs>?q)"
	proof -
	  from I II T2 VIII have
	    "f`(?p) \<lsq> f`(?p\<rs>𝟭) \<ra> L"
	    "f`(?p\<rs>𝟭) \<ra> L \<lsq> m \<ra>L"
	    using int_ord_transl_inv by auto
	  then have "f`(?p) \<lsq>  m \<ra>L"
	    by (rule Int_order_transitive)
	  with VII have "m\<ra>n \<rs>(m\<ra>L) \<lsq> f`(?r) \<rs> f`(?p)"
	    using int_ineq_add_sides by simp
	  with I T2 have "n \<rs> L \<lsq>  f`(?r) \<rs> f`(?p)"
	    using Int_ZF_1_2_L9 by simp
	  moreover
	  from I II T2 VIII have
	    "f`(?q) \<lsq> f`(?q\<rs>𝟭) \<ra> L"
	    "f`(?q\<rs>𝟭) \<ra> L \<lsq> n \<ra>L"
	    using int_ord_transl_inv by auto
	  then have "f`(?q) \<lsq>  n \<ra>L"
	    by (rule Int_order_transitive)
	  ultimately have 
	    "n \<rs> L \<rs> (n\<ra>L) \<lsq>  f`(?r) \<rs> f`(?p) \<rs> f`(?q)"
	    using int_ineq_add_sides by simp
	  with I V T2 have 
	    "(\<rm>L)\<rs>L \<ra>K \<lsq>  f`(?r) \<rs> f`(?p) \<rs> f`(?q) \<ra> K"
	    using Int_ZF_1_2_L3 int_ord_transl_inv by simp
	  moreover from VI T2 have
	    "f`(?r) \<rs> f`(?p) \<rs> f`(?q) \<ra> K \<lsq> f`(?r\<rs>?p\<rs>?q)"
	    by simp
	  ultimately show "(\<rm>L)\<rs>L \<ra>K \<lsq> f`(?r\<rs>?p\<rs>?q)"
	    by (rule Int_order_transitive)
	qed
	ultimately show
	  "f`(?r\<rs>?p\<rs>?q) \<lsq>  L\<ra>M  
	  (\<rm>L)\<rs>L\<ra>K \<lsq> f`(f¯(m\<ra>n)\<rs>f¯(m)\<rs>f¯(n))" 
	  by simp 
      qed 
    }
  ultimately show 
    "U. m+. n+. f`(f¯(m\<ra>n)\<rs>f¯(m)\<rs>f¯(n)) \<lsq> U"
    "N. m+. n+. N \<lsq> f`(f¯(m\<ra>n)\<rs>f¯(m)\<rs>f¯(n))"
    by auto
qed

textThe expression $f^{-1}(m+n)-f^{-1}(m) - f^{-1}(n)$ is uniformly bounded
  for all pairs $\langle m,n \rangle \in$ +×ℤ+. 
  Recall that in the int1›
  context ε(f,x)› is defined so that 
  $\varepsilon(f,\langle m,n \rangle ) = f^{-1}(m+n)-f^{-1}(m) - f^{-1}(n)$.

lemma (in int1) Int_ZF_2_4_L8:  assumes A1: "f  𝒮+" and
  A2: "m+. f¯(m)\<rs>𝟭  +"
  shows "M. x+×+. abs(ε(f,x)) \<lsq> M"
proof -
  from A1 A2 have 
    "U. m+. n+. f`(f¯(m\<ra>n)\<rs>f¯(m)\<rs>f¯(n)) \<lsq> U"
    "N. m+. n+. N \<lsq> f`(f¯(m\<ra>n)\<rs>f¯(m)\<rs>f¯(n))"
    using  Int_ZF_2_4_L7 by auto
  then obtain U N where I:
    "m+. n+. f`(f¯(m\<ra>n)\<rs>f¯(m)\<rs>f¯(n)) \<lsq> U" 
    "m+. n+. N \<lsq> f`(f¯(m\<ra>n)\<rs>f¯(m)\<rs>f¯(n))"
    by auto
  have "+×+  0" using int_one_two_are_pos by auto
  moreover from A1 have "f: "
    using AlmostHoms_def by simp
  moreover from A1 have
    "a.b+.x. b\<lsq>x  a \<lsq> f`(x)"
    using Int_ZF_2_3_L5 by simp
  moreover from A1 have
    "a.b+.y. b\<lsq>y  f`(\<rm>y) \<lsq> a"
    using Int_ZF_2_3_L5A by simp
  moreover have 
    "x+×+. ε(f,x)    f`(ε(f,x)) \<lsq> U  N \<lsq> f`(ε(f,x))"
  proof -
    { fix x assume A3: "x  +×+"
      let ?m = "fst(x)"
      let ?n = "snd(x)"
      from A3 have T: "?m  +"  "?n  +"  "?m\<ra>?n  +"
	using pos_int_closed_add_unfolded by auto
      with A1 have 
	"f¯(?m\<ra>?n)  "  "f¯(?m)  "  "f¯(?n)  "
	using Int_ZF_2_4_L2 PositiveSet_def by auto
      with I T have 
	"ε(f,x)    f`(ε(f,x)) \<lsq> U  N \<lsq> f`(ε(f,x))"
	using Int_ZF_1_1_L5 by auto 
    } thus ?thesis by simp
    qed
  ultimately show "M.x+×+. abs(ε(f,x)) \<lsq> M"
    by (rule Int_ZF_1_6_L4)
qed
  
textThe (candidate for) inverse of a positive slope is a (well defined) 
  function on +.

lemma (in int1) Int_ZF_2_4_L9: 
  assumes A1: "f  𝒮+" and A2: "g = {p,f¯(p). p+}"
  shows 
  "g : ++"
  "g : +"
proof -
  from A1 have 
    "p+. f¯(p)  +" 
    "p+. f¯(p)  " 
    using Int_ZF_2_4_L2 PositiveSet_def by auto
  with A2 show 
    "g : ++"  and  "g : +" 
    using ZF_fun_from_total by auto
qed

textWhat are the values of the (candidate for) the inverse of a positive slope?

lemma (in int1) Int_ZF_2_4_L10: 
  assumes A1: "f  𝒮+" and A2: "g = {p,f¯(p). p+}" and A3: "p+"
  shows "g`(p) = f¯(p)"
proof -
  from A1 A2 have  "g : ++" using Int_ZF_2_4_L9 by simp
  with A2 A3 show "g`(p) = f¯(p)" using ZF_fun_from_tot_val by simp
qed

textThe (candidate for) the inverse of a positive slope is a slope.

lemma (in int1) Int_ZF_2_4_L11: assumes A1: "f  𝒮+" and 
  A2: "m+. f¯(m)\<rs>𝟭  +" and
  A3: "g = {p,f¯(p). p+}"
  shows "OddExtension(,IntegerAddition,IntegerOrder,g)  𝒮"
proof -
  from A1 A2 have "L. x+×+. abs(ε(f,x)) \<lsq> L"
    using Int_ZF_2_4_L8 by simp
  then obtain L where I: "x+×+. abs(ε(f,x)) \<lsq> L"
    by auto
  from A1 A3 have "g : +" using Int_ZF_2_4_L9 
    by simp
  moreover have "m+. n+. abs(δ(g,m,n)) \<lsq> L"
  proof-
    { fix m n
      assume A4: "m+"  "n+"
      then have "m,n  +×+" by simp
      with I have "abs(ε(f,m,n)) \<lsq> L" by simp
      moreover have "ε(f,m,n) = f¯(m\<ra>n) \<rs> f¯(m) \<rs> f¯(n)"
	by simp
      moreover from A1 A3 A4 have
	"f¯(m\<ra>n) = g`(m\<ra>n)"  "f¯(m) = g`(m)"  "f¯(n) = g`(n)"
	using pos_int_closed_add_unfolded Int_ZF_2_4_L10 by auto
      ultimately have "abs(δ(g,m,n)) \<lsq> L" by simp
    } thus "m+. n+. abs(δ(g,m,n)) \<lsq> L" by simp
  qed
  ultimately show ?thesis by (rule Int_ZF_2_1_L24)
qed
  
textEvery positive slope that is at least $2$ on positive integers
  almost has an inverse.


lemma (in int1) Int_ZF_2_4_L12: assumes A1: "f  𝒮+" and 
  A2: "m+. f¯(m)\<rs>𝟭  +"
  shows "h𝒮. fh  id()"
proof -
  let ?g = "{p,f¯(p). p+}"
  let ?h = "OddExtension(,IntegerAddition,IntegerOrder,?g)"
  from A1 have 
    "M. n. f`(n) \<lsq> f`(n\<rs>𝟭) \<ra> M"
    using Int_ZF_2_1_L28 by simp
  then obtain M where 
    I: "M" and II: "n. f`(n) \<lsq> f`(n\<rs>𝟭) \<ra> M"
    by auto
  from A1 A2 have T: "?h  𝒮"
    using Int_ZF_2_4_L11 by simp
  moreover have  "f?h  id()"
  proof -
    from A1 T have "f?h  𝒮" using Int_ZF_2_1_L11 
      by simp
    moreover note I
    moreover
    { fix m assume A3: "m+"
      with A1 have "f¯(m)  "
	using Int_ZF_2_4_L2 PositiveSet_def by simp 
      with II have "f`(f¯(m)) \<lsq> f`(f¯(m)\<rs>𝟭) \<ra> M"
	by simp
      moreover from A1 A2 I A3 have "f`(f¯(m)\<rs>𝟭) \<ra> M \<lsq> m\<ra>M"
	using Int_ZF_2_4_L4 int_ord_transl_inv by simp
      ultimately have "f`(f¯(m)) \<lsq> m\<ra>M"
	by (rule Int_order_transitive)
      moreover from A1 A3 have "m \<lsq> f`(f¯(m))"
	using Int_ZF_2_4_L2 by simp
      moreover from A1 A2 T A3 have "f`(f¯(m)) = (f?h)`(m)"
	using Int_ZF_2_4_L9 Int_ZF_1_5_L11
	  Int_ZF_2_4_L10 PositiveSet_def Int_ZF_2_1_L10
	by simp
      ultimately have "m \<lsq> (f?h)`(m)  (f?h)`(m) \<lsq> m\<ra>M"
	by simp }
    ultimately show "f?h  id()" using Int_ZF_2_1_L32
      by simp
  qed 
  ultimately show "h𝒮. fh  id()"
    by auto
qed

textInt_ZF_2_4_L12› is almost what we need, except that it has an assumption
  that the values of the slope that we get the inverse for are not smaller than $2$ on
  positive integers. The Arthan's proof of Theorem 11 has a mistake where he says "note that
  for all but finitely many $m,n\in N$ $p=g(m)$ and $q=g(n)$ are both positive". Of course
  there may be infinitely many pairs $\langle m,n \rangle$ such that $p,q$ are not both 
  positive. This is however easy to workaround: we just modify the slope by adding a 
  constant so that the slope is large enough on positive integers and then look 
  for the inverse.

theorem (in int1) pos_slope_has_inv: assumes A1: "f  𝒮+"
  shows "g𝒮. fg  (h𝒮. gh  id())"
proof -
  from A1 have "f: "  "𝟭"  "𝟮  "
    using AlmostHoms_def int_zero_one_are_int int_two_three_are_int
    by auto
  moreover from A1 have
     "a.b+.x. b\<lsq>x  a \<lsq> f`(x)"
    using Int_ZF_2_3_L5 by simp
  ultimately have 
    "c. 𝟮 \<lsq> Minimum(IntegerOrder,{n+. 𝟭 \<lsq> f`(n)\<ra>c})"
    by (rule Int_ZF_1_6_L7)
  then obtain c where I: "c" and
    II: "𝟮 \<lsq> Minimum(IntegerOrder,{n+. 𝟭 \<lsq> f`(n)\<ra>c})"
    by auto
  let ?g = "{m,f`(m)\<ra>c. m}"
  from A1 I have III: "?g𝒮" and IV: "f?g" using Int_ZF_2_1_L33 
    by auto
  from IV have "f,?g  AlEqRel" by simp
  with A1 have T: "?g  𝒮+" by (rule Int_ZF_2_3_L9)
  moreover have "m+. ?g¯(m)\<rs>𝟭  +"
  proof
    fix m assume A2: "m+"
    from A1 I II have V: "𝟮 \<lsq> ?g¯(𝟭)"
      using Int_ZF_2_1_L33 PositiveSet_def by simp
    moreover from A2 T have "?g¯(𝟭) \<lsq> ?g¯(m)"
      using Int_ZF_1_5_L3 int_one_two_are_pos Int_ZF_2_4_L5
      by simp
    ultimately have "𝟮 \<lsq> ?g¯(m)"
      by (rule Int_order_transitive)
    then have "𝟮\<rs>𝟭 \<lsq> ?g¯(m)\<rs>𝟭"
      using int_zero_one_are_int Int_ZF_1_1_L4 int_ord_transl_inv
      by simp
    then show  "?g¯(m)\<rs>𝟭  +"
      using int_zero_one_are_int Int_ZF_1_2_L3 Int_ZF_1_5_L3
      by simp
  qed
  ultimately have "h𝒮. ?gh  id()"
    by (rule Int_ZF_2_4_L12)
  with III IV show ?thesis by auto
qed

subsectionCompleteness

textIn this section we consider properties of slopes that are
  needed for the proof of completeness of real numbers constructred
  in Real_ZF_1.thy›. In particular we consider properties
  of embedding of integers into the set of slopes by the mapping
  $m \mapsto m^S$ , where $m^S$ is defined by $m^S(n) = m\cdot n$.

textIf m is an integer, then $m^S$ is a slope whose value
  is $m\cdot n$ for every integer.

lemma (in int1) Int_ZF_2_5_L1: assumes A1: "m  "
  shows 
  "n  . (mS)`(n) = mn"
  "mS  𝒮"
proof -
  from A1 have I: "mS:"
    using Int_ZF_1_1_L5 ZF_fun_from_total by simp
  then show II: "n  . (mS)`(n) = mn" using ZF_fun_from_tot_val
    by simp
  { fix n k
    assume A2: "n"  "k"
    with A1 have T: "mn  "  "mk  "
      using Int_ZF_1_1_L5 by auto
    from A1 A2 II T  have "δ(mS,n,k) = mk \<rs> mk"
      using Int_ZF_1_1_L5 Int_ZF_1_1_L1 Int_ZF_1_2_L3
      by simp
    also from T have " = 𝟬" using Int_ZF_1_1_L4
      by simp
    finally have "δ(mS,n,k) = 𝟬" by simp
    then have "abs(δ(mS,n,k)) \<lsq> 𝟬"
      using Int_ZF_2_L18 int_zero_one_are_int int_ord_is_refl refl_def
      by simp
  } then have "n.k. abs(δ(mS,n,k)) \<lsq> 𝟬"
    by simp
  with I show  "mS  𝒮" by (rule Int_ZF_2_1_L5)
qed

textFor any slope $f$ there is an integer $m$ such that there is some slope $g$ 
  that is almost equal to $m^S$ and dominates $f$ in the sense that $f\leq g$ 
  on positive integers (which implies that either $g$ is almost equal to $f$ or
  $g-f$ is a positive slope. This will be used in Real_ZF_1.thy› to show
  that for any real number there is an integer that (whose real embedding) 
  is greater or equal.

lemma (in int1) Int_ZF_2_5_L2: assumes A1: "f  𝒮"
  shows "m. g𝒮. (mSg  (fg  g\<fp>(\<fm>f)  𝒮+))"
proof -
  from A1 have 
    "m k. m  k  (p. abs(f`(p)) \<lsq> mabs(p)\<ra>k)"
    using Arthan_Lem_8 by simp
  then obtain m k where I: "m" and II: "k" and 
    III: "p. abs(f`(p)) \<lsq> mabs(p)\<ra>k"
    by auto
  let ?g = "{n,mS`(n) \<ra>k. n}"
  from I have IV: "mS  𝒮" using Int_ZF_2_5_L1 by simp
  with II have V: "?g𝒮" and VI: "mS?g" using Int_ZF_2_1_L33 
    by auto
  { fix n assume A2: "n+"
    with A1 have "f`(n)  "
      using Int_ZF_2_1_L2B PositiveSet_def by simp
    then have "f`(n) \<lsq> abs(f`(n))" using Int_ZF_2_L19C 
      by simp
    moreover  
    from III A2 have "abs(f`(n)) \<lsq> mabs(n) \<ra> k"
      using PositiveSet_def by simp
    with A2 have "abs(f`(n)) \<lsq> mn\<ra>k"
      using Int_ZF_1_5_L4A by simp
    ultimately have "f`(n) \<lsq> mn\<ra>k"
      by (rule Int_order_transitive)
    moreover
    from II IV A2 have "?g`(n) = (mS)`(n)\<ra>k"
      using Int_ZF_2_1_L33 PositiveSet_def by simp
    with I A2 have "?g`(n) = mn\<ra>k"
      using Int_ZF_2_5_L1 PositiveSet_def by simp
    ultimately have "f`(n) \<lsq> ?g`(n)"
      by simp
  } then have "n+. f`(n) \<lsq> ?g`(n)"
    by simp
  with A1 V have "f?g  ?g \<fp> (\<fm>f)  𝒮+"
    using Int_ZF_2_3_L4C by simp
  with I V VI show ?thesis by auto
qed

textThe negative of an integer embeds in slopes as a negative of the 
  orgiginal embedding.

lemma (in int1) Int_ZF_2_5_L3: assumes A1:  "m  "
  shows "(\<rm>m)S = \<fm>(mS)"
proof -
  from A1 have "(\<rm>m)S: " and "(\<fm>(mS)): "
    using Int_ZF_1_1_L4 Int_ZF_2_5_L1 AlmostHoms_def Int_ZF_2_1_L12
    by auto
  moreover have "n. ((\<rm>m)S)`(n) = (\<fm>(mS))`(n)"
  proof
    fix n assume A2: "n"
    with A1 have 
      "((\<rm>m)S)`(n) = (\<rm>m)n"
      "(\<fm>(mS))`(n) = \<rm>(mn)"
      using Int_ZF_1_1_L4 Int_ZF_2_5_L1 Int_ZF_2_1_L12A
      by auto
    with A1 A2 show "((\<rm>m)S)`(n) = (\<fm>(mS))`(n)"
      using Int_ZF_1_1_L5 by simp
  qed
  ultimately show "(\<rm>m)S = \<fm>(mS)" using fun_extension_iff
    by simp
qed


textThe sum of embeddings is the embeding of the sum.

lemma (in int1) Int_ZF_2_5_L3A: assumes A1: "m"  "k"
  shows "(mS) \<fp> (kS) = ((m\<ra>k)S)"
proof -
  from A1 have T1: "m\<ra>k  " using Int_ZF_1_1_L5 
    by simp
  with A1 have T2:
    "(mS)  𝒮"  "(kS)  𝒮"
    "(m\<ra>k)S   𝒮"
    "(mS) \<fp> (kS)  𝒮"
    using Int_ZF_2_5_L1 Int_ZF_2_1_L12C by auto
  then have 
    "(mS) \<fp> (kS) : "
    "(m\<ra>k)S : " 
    using AlmostHoms_def by auto
  moreover have "n. ((mS) \<fp> (kS))`(n) = ((m\<ra>k)S)`(n)"
  proof
    fix n assume A2: "n"
    with A1 T1 T2 have  "((mS) \<fp> (kS))`(n) = (m\<ra>k)n"
      using Int_ZF_2_1_L12B Int_ZF_2_5_L1 Int_ZF_1_1_L1
      by simp
    also from T1 A2 have " = ((m\<ra>k)S)`(n)"
      using Int_ZF_2_5_L1 by simp
    finally show "((mS) \<fp> (kS))`(n) = ((m\<ra>k)S)`(n)"
      by simp
  qed
  ultimately show "(mS) \<fp> (kS) = ((m\<ra>k)S)"
    using fun_extension_iff by simp
qed

textThe composition of embeddings is the embeding of the product.

lemma (in int1) Int_ZF_2_5_L3B: assumes A1: "m"  "k"
  shows "(mS)  (kS) = ((mk)S)"
proof -
  from A1 have T1: "mk  " using Int_ZF_1_1_L5 
    by simp
  with A1 have T2:
    "(mS)  𝒮"  "(kS)  𝒮"
    "(mk)S   𝒮"
    "(mS)  (kS)  𝒮"
    using Int_ZF_2_5_L1 Int_ZF_2_1_L11 by auto
  then have 
    "(mS)  (kS) : "
    "(mk)S : " 
    using AlmostHoms_def by auto
  moreover have "n. ((mS)  (kS))`(n) = ((mk)S)`(n)"
  proof
    fix n assume A2: "n"
    with A1 T2 have
      "((mS)  (kS))`(n) = (mS)`(kn)"
       using Int_ZF_2_1_L10 Int_ZF_2_5_L1 by simp
    moreover
    from A1 A2 have "kn  " using Int_ZF_1_1_L5 
      by simp
    with A1 A2 have "(mS)`(kn) = mkn"
      using Int_ZF_2_5_L1 Int_ZF_1_1_L7 by simp
    ultimately have "((mS)  (kS))`(n) = mkn"
      by simp
    also from T1 A2 have "mkn = ((mk)S)`(n)"
      using Int_ZF_2_5_L1 by simp
    finally show "((mS)  (kS))`(n) = ((mk)S)`(n)"
      by simp
  qed
  ultimately show "(mS)  (kS) = ((mk)S)"
    using fun_extension_iff by simp
qed
  

textEmbedding integers in slopes preserves order.

lemma (in int1) Int_ZF_2_5_L4: assumes A1:  "m\<lsq>n"
  shows "(mS)  (nS)  (nS)\<fp>(\<fm>(mS))  𝒮+"
proof -
  from A1 have "mS  𝒮" and "nS  𝒮"
    using Int_ZF_2_L1A Int_ZF_2_5_L1 by auto
  moreover from A1 have "k+. (mS)`(k) \<lsq> (nS)`(k)"
    using Int_ZF_1_3_L13B Int_ZF_2_L1A PositiveSet_def Int_ZF_2_5_L1
    by simp
  ultimately show ?thesis using Int_ZF_2_3_L4C
    by simp
qed

textWe aim at showing that $m\mapsto m^S$ is an injection modulo
  the relation of almost equality. To do that we first show that if
  $m^S$ has finite range, then $m=0$.

lemma (in int1) Int_ZF_2_5_L5: 
  assumes "m" and "mS  FinRangeFunctions(,)"
  shows "m=𝟬"
  using assms FinRangeFunctions_def Int_ZF_2_5_L1 AlmostHoms_def 
    func_imagedef Int_ZF_1_6_L8 by simp

textEmbeddings of two integers are almost equal only if 
  the integers are equal.

lemma (in int1) Int_ZF_2_5_L6: 
  assumes A1: "m"  "k" and A2: "(mS)  (kS)"
  shows "m=k"
proof -
  from A1 have T: "m\<rs>k  " using Int_ZF_1_1_L5 by simp
  from A1 have "(\<fm>(kS)) =  ((\<rm>k)S)"
    using Int_ZF_2_5_L3 by simp
  then have "mS \<fp> (\<fm>(kS)) = (mS) \<fp> ((\<rm>k)S)"
    by simp
  with A1 have "mS \<fp> (\<fm>(kS)) = ((m\<rs>k)S)"
    using Int_ZF_1_1_L4 Int_ZF_2_5_L3A by simp
  moreover from A1 A2 have "mS \<fp> (\<fm>(kS))  FinRangeFunctions(,)"
    using Int_ZF_2_5_L1 Int_ZF_2_1_L9D by simp
  ultimately have "(m\<rs>k)S  FinRangeFunctions(,)"
    by simp
  with T have "m\<rs>k = 𝟬" using Int_ZF_2_5_L5
    by simp
  with A1 show "m=k" by (rule Int_ZF_1_L15)
qed

textEmbedding of $1$ is the identity slope and embedding of zero is a 
  finite range function.

lemma (in int1) Int_ZF_2_5_L7: shows 
  "𝟭S = id()"
  "𝟬S  FinRangeFunctions(,)"
proof -
  have "id() = {x,x. x}"
    using id_def by blast
  then show "𝟭S = id()" using Int_ZF_1_1_L4 by simp
  have "{𝟬S`(n). n} = {𝟬n. n}"
    using int_zero_one_are_int Int_ZF_2_5_L1 by simp
  also have " = {𝟬}" using Int_ZF_1_1_L4 int_not_empty
    by simp
  finally have "{𝟬S`(n). n} = {𝟬}" by simp
  then have "{𝟬S`(n). n}  Fin()"
    using int_zero_one_are_int Finite1_L16 by simp
  moreover have "𝟬S: " 
    using int_zero_one_are_int Int_ZF_2_5_L1 AlmostHoms_def 
    by simp
  ultimately show "𝟬S  FinRangeFunctions(,)"
    using Finite1_L19 by simp  
qed

textA somewhat technical condition for a embedding of an integer 
  to be "less  or equal" (in the sense apriopriate for slopes) than 
  the composition of a slope and another integer (embedding).

lemma (in int1) Int_ZF_2_5_L8: 
  assumes A1: "f  𝒮" and A2: "N  "  "M  " and
  A3: "n+. Mn \<lsq> f`(Nn)"
  shows "MS  f(NS)   (f(NS)) \<fp> (\<fm>(MS))  𝒮+"
proof -
  from A1 A2 have "MS  𝒮"  "f(NS)  𝒮"
    using Int_ZF_2_5_L1 Int_ZF_2_1_L11 by auto
  moreover from A1 A2 A3 have "n+. (MS)`(n) \<lsq> (f(NS))`(n)"
    using Int_ZF_2_5_L1 PositiveSet_def Int_ZF_2_1_L10
    by simp
  ultimately show ?thesis using Int_ZF_2_3_L4C
    by simp
qed

textAnother technical condition for the composition of a slope and 
  an integer (embedding) to be "less  or equal" (in the sense apriopriate 
  for slopes) than embedding of another integer.

lemma (in int1) Int_ZF_2_5_L9: 
  assumes A1: "f  𝒮" and A2: "N  "  "M  " and
  A3: "n+.  f`(Nn) \<lsq> Mn "
  shows "f(NS)  (MS)  (MS) \<fp> (\<fm>(f(NS)))  𝒮+"
proof -
  from A1 A2 have "f(NS)  𝒮"  "MS  𝒮"  
    using Int_ZF_2_5_L1 Int_ZF_2_1_L11 by auto
  moreover from A1 A2 A3 have "n+. (f(NS))`(n) \<lsq> (MS)`(n) "
    using Int_ZF_2_5_L1 PositiveSet_def Int_ZF_2_1_L10
    by simp
  ultimately show ?thesis using Int_ZF_2_3_L4C
    by simp
qed

end