Theory Int_ZF_1

Up to index of Isabelle/ZF/IsarMathLib

theory Int_ZF_1
imports Int_ZF OrderedRing_ZF
begin

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

    Copyright (C) 2005, 2006  Slawomir Kolodynski

    This program is free software; Redistribution and use in source and binary forms, 
    with or without modification, are permitted provided that the following conditions are met:

   1. Redistributions of source code must retain the above copyright notice, 
   this list of conditions and the following disclaimer.
   2. Redistributions in binary form must reproduce the above copyright notice, 
   this list of conditions and the following disclaimer in the documentation and/or 
   other materials provided with the distribution.
   3. The name of the author may not be used to endorse or promote products 
   derived from this software without specific prior written permission.

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.

*)

header {*\isaheader{Int\_ZF\_1.thy}*}

theory Int_ZF_1 imports Int_ZF OrderedRing_ZF 

begin;

text{*This theory file considers the set of integers as an ordered ring.*}

section{*Integers as a ring*}

text{*In this section we show that integers form a commutative ring. *}


text{*The next lemma provides the condition to show that addition is 
  distributive with respect to multiplication.*}

lemma (in int0) Int_ZF_1_1_L1: assumes A1: "a∈\<int>"  "b∈\<int>"  "c∈\<int>" 
  shows 
  "a·(b\<ra>c) = a·b \<ra> a·c" 
  "(b\<ra>c)·a = b·a \<ra> c·a"
  using prems Int_ZF_1_L2 zadd_zmult_distrib zadd_zmult_distrib2
  by auto;

text{*Integers form a commutative ring, hence we can use theorems proven 
  in @{text "ring0"} context (locale).*}

lemma (in int0) Int_ZF_1_1_L2: shows
  "IsAring(\<int>,IntegerAddition,IntegerMultiplication)"
  "IntegerMultiplication {is commutative on} \<int>"
  "ring0(\<int>,IntegerAddition,IntegerMultiplication)"
proof -
  have "∀a∈\<int>.∀b∈\<int>.∀c∈\<int>. 
    a·(b\<ra>c) = a·b \<ra> a·c ∧ (b\<ra>c)·a = b·a \<ra> c·a"
    using Int_ZF_1_1_L1 by simp;
  then have "IsDistributive(\<int>,IntegerAddition,IntegerMultiplication)"
    using IsDistributive_def by simp;
  then show "IsAring(\<int>,IntegerAddition,IntegerMultiplication)"
    "ring0(\<int>,IntegerAddition,IntegerMultiplication)"
    using Int_ZF_1_T1 Int_ZF_1_T2 IsAring_def ring0_def 
    by auto;
  have "∀a∈\<int>.∀b∈\<int>. a·b = b·a" using Int_ZF_1_L4 by simp;
  then show "IntegerMultiplication {is commutative on} \<int>"
    using IsCommutative_def by simp;
qed;

text{*Zero and one are integers.*}

lemma (in int0) int_zero_one_are_int: shows "\<zero>∈\<int>"  "\<one>∈\<int>"
  using Int_ZF_1_1_L2 ring0.Ring_ZF_1_L2 by auto;

text{*Negative of zero is zero.*}

lemma (in int0) int_zero_one_are_intA: shows "(\<rm>\<zero>) = \<zero>"
  using Int_ZF_1_T2 group0.group_inv_of_one by simp;

text{*Properties with one integer.*}

lemma (in int0) Int_ZF_1_1_L4: assumes A1: "a ∈ \<int>"
  shows 
  "a\<ra>\<zero> = a" 
  "\<zero>\<ra>a = a" 
  "a·\<one> = a"   "\<one>·a = a"
  "\<zero>·a = \<zero>"   "a·\<zero> = \<zero>" 
  "(\<rm>a) ∈ \<int>"  "(\<rm>(\<rm>a)) = a"
  "a\<rs>a = \<zero>"   "a\<rs>\<zero> = a"  "\<two>·a = a\<ra>a"
proof -;
  from A1 show 
    "a\<ra>\<zero> = a"   "\<zero>\<ra>a = a"   "a·\<one> = a" 
    "\<one>·a = a"   "a\<rs>a = \<zero>"   "a\<rs>\<zero> = a"  
    "(\<rm>a) ∈ \<int>"  "\<two>·a = a\<ra>a"   "(\<rm>(\<rm>a)) = a"
    using Int_ZF_1_1_L2 ring0.Ring_ZF_1_L3 by auto;
  from A1 show "\<zero>·a = \<zero>"   "a·\<zero> = \<zero>"
    using Int_ZF_1_1_L2 ring0.Ring_ZF_1_L6 by auto;
qed;

text{*Properties that require two integers.*}

lemma (in int0) Int_ZF_1_1_L5: assumes A1: "a∈\<int>"  "b∈\<int>"
  shows 
  "a\<ra>b ∈ \<int>" 
  "a\<rs>b ∈ \<int>" 
  "a·b ∈ \<int>"
  "a\<ra>b = b\<ra>a" 
  "a·b = b·a" 
  "(\<rm>b)\<rs>a = (\<rm>a)\<rs>b" 
  "(\<rm>(a\<ra>b)) = (\<rm>a)\<rs>b"  
  "(\<rm>(a\<rs>b)) = ((\<rm>a)\<ra>b)"
  "(\<rm>a)·b = \<rm>(a·b)" 
  "a·(\<rm>b) = \<rm>(a·b)"
  "(\<rm>a)·(\<rm>b) = a·b"
  using prems Int_ZF_1_1_L2 ring0.Ring_ZF_1_L4 ring0.Ring_ZF_1_L9
    ring0.Ring_ZF_1_L7 ring0.Ring_ZF_1_L7A Int_ZF_1_L4 by auto;

text{*$2$ and $3$ are integers.*}

lemma (in int0) int_two_three_are_int: shows "\<two> ∈ \<int>"  "\<three> ∈ \<int>"
    using int_zero_one_are_int Int_ZF_1_1_L5 by auto;

text{*Another property with two integers.*}

lemma (in int0) Int_ZF_1_1_L5B:
  assumes A1: "a∈\<int>"  "b∈\<int>"
  shows "a\<rs>(\<rm>b) = a\<ra>b"
  using prems Int_ZF_1_1_L2 ring0.Ring_ZF_1_L9
  by simp;

text{*Properties that require three integers.*}

lemma (in int0) Int_ZF_1_1_L6: assumes A1: "a∈\<int>"  "b∈\<int>"  "c∈\<int>"
  shows 
  "a\<rs>(b\<ra>c) = a\<rs>b\<rs>c"
  "a\<rs>(b\<rs>c) = a\<rs>b\<ra>c"
  "a·(b\<rs>c) = a·b \<rs> a·c"
  "(b\<rs>c)·a = b·a \<rs> c·a"
  using prems Int_ZF_1_1_L2 ring0.Ring_ZF_1_L10  ring0.Ring_ZF_1_L8 
  by auto;

text{*One more property with three integers.*}
(* Inclusion of a\<ra>(b\<rs>c) = a\<ra>b\<rs>c causes the previous lemma
  to be unusable by simp in some cases*)
lemma (in int0) Int_ZF_1_1_L6A: assumes A1: "a∈\<int>"  "b∈\<int>"  "c∈\<int>"
  shows "a\<ra>(b\<rs>c) = a\<ra>b\<rs>c"
  using prems Int_ZF_1_1_L2 ring0.Ring_ZF_1_L10A by simp;

text{*Associativity of addition and multiplication.*}

lemma (in int0) Int_ZF_1_1_L7: assumes A1: "a∈\<int>" "b∈\<int>" "c∈\<int>"
  shows 
  "a\<ra>b\<ra>c = a\<ra>(b\<ra>c)"
  "a·b·c = a·(b·c)"
  using prems Int_ZF_1_1_L2 ring0.Ring_ZF_1_L11 by auto;
  
section{*Rearrangement lemmas*}

text{*In this section we collect lemmas about identities related to 
  rearranging the terms in expresssions*}

text{*A formula with a positive integer.*}

lemma (in int0) Int_ZF_1_2_L1: assumes "\<zero>\<lsq>a"
  shows "abs(a)\<ra>\<one> = abs(a\<ra>\<one>)"
  using prems Int_ZF_2_L16 Int_ZF_2_L12A by simp;

text{*A formula with two integers, one positive.*}

lemma (in int0) Int_ZF_1_2_L2: assumes A1: "a∈\<int>" and A2: "\<zero>\<lsq>b"
  shows "a\<ra>(abs(b)\<ra>\<one>)·a = (abs(b\<ra>\<one>)\<ra>\<one>)·a"
proof -
  from A2 have T1: "abs(b\<ra>\<one>) ∈ \<int>"
    using Int_ZF_2_L12A Int_ZF_2_L1A Int_ZF_2_L14 by blast;
  with A1 A2 show ?thesis 
    using Int_ZF_1_2_L1 Int_ZF_1_1_L2 ring0.Ring_ZF_2_L1 
    by simp;
qed;

text{*A couple of formulae about canceling opposite integers.*}

lemma (in int0) Int_ZF_1_2_L3: assumes A1: "a∈\<int>"  "b∈\<int>"
  shows 
  "a\<ra>b\<rs>a = b"
  "a\<ra>(b\<rs>a) = b"
  "a\<ra>b\<rs>b = a" 
  "a\<rs>b\<ra>b = a"
  "(\<rm>a)\<ra>(a\<ra>b) = b"
  "a\<ra>(b\<rs>a) = b"
  "(\<rm>b)\<ra>(a\<ra>b) = a"
  "a\<rs>(b\<ra>a) = \<rm>b"
  "a\<rs>(a\<ra>b) = \<rm>b"
  "a\<rs>(a\<rs>b) = b"
  "a\<rs>b\<rs>a =  \<rm>b"
  "a\<rs>b \<rs> (a\<ra>b) = (\<rm>b)\<rs>b"
  using prems Int_ZF_1_T2 group0.group0_4_L6A group0.group0_2_L16
    group0.group0_2_L16A group0.group0_4_L6AA group0.group0_4_L6AB
    group0.group0_4_L6F group0.group0_4_L6AC by auto;

text{*Subtracting one does not increase integers. This may be moved to a theory
  about ordered rings one day.*}

lemma (in int0) Int_ZF_1_2_L3A: assumes A1: "a\<lsq>b"
  shows "a\<rs>\<one> \<lsq> b" 
proof -
  from A1 have "b\<ra>\<one>\<rs>\<one> = b"
    using Int_ZF_2_L1A int_zero_one_are_int Int_ZF_1_2_L3 by simp;
  moreover from A1 have "a\<rs>\<one> \<lsq> b\<ra>\<one>\<rs>\<one>"
    using Int_ZF_2_L12A int_zero_one_are_int Int_ZF_1_1_L4 int_ord_transl_inv
    by simp;
  ultimately show "a\<rs>\<one> \<lsq> b" by simp;
qed;

text{*Subtracting one does not increase integers, special case.*}

lemma (in int0) Int_ZF_1_2_L3AA: 
  assumes A1: "a∈\<int>" shows 
  "a\<rs>\<one> \<lsq>a"
  "a\<rs>\<one> ≠ a"
  "¬(a\<lsq>a\<rs>\<one>)"
  "¬(a\<ra>\<one> \<lsq>a)"
  "¬(\<one>\<ra>a \<lsq>a)"
proof -
  from A1 have "a\<lsq>a" using int_ord_is_refl refl_def
    by simp;
  then show "a\<rs>\<one> \<lsq>a" using Int_ZF_1_2_L3A
    by simp;
  moreover from A1 show "a\<rs>\<one> ≠ a" using Int_ZF_1_L14 by simp;
  ultimately show I: "¬(a\<lsq>a\<rs>\<one>)" using Int_ZF_2_L19AA
    by blast;
  with A1 show "¬(a\<ra>\<one> \<lsq>a)"
    using int_zero_one_are_int Int_ZF_2_L9B by simp;
  with A1 show "¬(\<one>\<ra>a \<lsq>a)" 
    using int_zero_one_are_int Int_ZF_1_1_L5 by simp;
qed;

text{*A formula with a nonpositive integer.*}

lemma (in int0) Int_ZF_1_2_L4: assumes "a\<lsq>\<zero>"
  shows "abs(a)\<ra>\<one> = abs(a\<rs>\<one>)"
  using prems int_zero_one_are_int Int_ZF_1_2_L3A Int_ZF_2_T1 
      group3.OrderedGroup_ZF_3_L3A Int_ZF_2_L1A 
      int_zero_one_are_int Int_ZF_1_1_L5 by simp;

text{*A formula with two integers, one negative.*}

lemma (in int0) Int_ZF_1_2_L5: assumes A1: "a∈\<int>" and A2: "b\<lsq>\<zero>"
  shows "a\<ra>(abs(b)\<ra>\<one>)·a = (abs(b\<rs>\<one>)\<ra>\<one>)·a"
proof -
  from A2 have "abs(b\<rs>\<one>) ∈ \<int>"
    using int_zero_one_are_int Int_ZF_1_2_L3A Int_ZF_2_L1A Int_ZF_2_L14 
    by blast;
  with A1 A2 show ?thesis 
    using Int_ZF_1_2_L4 Int_ZF_1_1_L2 ring0.Ring_ZF_2_L1
    by simp;
qed;
  
text{*A rearrangement with four integers.*}

lemma (in int0) Int_ZF_1_2_L6: 
  assumes A1: "a∈\<int>"  "b∈\<int>"  "c∈\<int>"  "d∈\<int>"
  shows 
  "a\<rs>(b\<rs>\<one>)·c = (d\<rs>b·c)\<rs>(d\<rs>a\<rs>c)"
proof -;
  from A1 have T1: 
    "(d\<rs>b·c) ∈ \<int>" "d\<rs>a ∈ \<int>" "(\<rm>(b·c)) ∈ \<int>"
    using Int_ZF_1_1_L5 Int_ZF_1_1_L4 by auto;   
  with A1 have 
    "(d\<rs>b·c)\<rs>(d\<rs>a\<rs>c) = (\<rm>(b·c))\<ra>a\<ra>c"
    using Int_ZF_1_1_L6 Int_ZF_1_2_L3 by simp;
  also from A1 T1 have "(\<rm>(b·c))\<ra>a\<ra>c = a\<rs>(b\<rs>\<one>)·c" 
    using int_zero_one_are_int Int_ZF_1_1_L6 Int_ZF_1_1_L4 Int_ZF_1_1_L5
    by simp;
  finally show ?thesis by simp
qed;

text{*Some other rearrangements with two integers.*}

lemma (in int0) Int_ZF_1_2_L7: assumes "a∈\<int>"  "b∈\<int>"
  shows 
  "a·b = (a\<rs>\<one>)·b\<ra>b"
  "a·(b\<ra>\<one>) = a·b\<ra>a"
  "(b\<ra>\<one>)·a = b·a\<ra>a"
  "(b\<ra>\<one>)·a = a\<ra>b·a"
  using prems Int_ZF_1_1_L1 Int_ZF_1_1_L5 int_zero_one_are_int 
    Int_ZF_1_1_L6 Int_ZF_1_1_L4 Int_ZF_1_T2 group0.group0_2_L16 
  by auto;

(*text{*Another rearrangement with two integers and cancelling.*}

lemma (in int0) Int_ZF_1_2_L7A:  assumes A1: "a∈\<int>"  "b∈\<int>"
  shows "a\<rs>b \<ra> (a\<rs>b) = \<rm>(\<two>·b)"
proof -
  from A1 have "a\<rs>b \<rs> (a\<ra>b) = (\<rm>b)\<rs>b"
    using Int_ZF_1_T2 group0.group0_4_L6F by simp;*)
  

text{*Another rearrangement with two integers.*}

lemma (in int0) Int_ZF_1_2_L8: 
  assumes A1: "a∈\<int>" "b∈\<int>"
  shows "a\<ra>\<one>\<ra>(b\<ra>\<one>) = b\<ra>a\<ra>\<two>"
  using prems int_zero_one_are_int Int_ZF_1_T2 group0.group0_4_L8
  by simp;

text{*A couple of rearrangement with three integers.*}

lemma (in int0) Int_ZF_1_2_L9: 
  assumes "a∈\<int>"  "b∈\<int>"  "c∈\<int>"
  shows 
  "(a\<rs>b)\<ra>(b\<rs>c) = a\<rs>c"
  "(a\<rs>b)\<rs>(a\<rs>c) = c\<rs>b"
  "a\<ra>(b\<ra>(c\<rs>a\<rs>b)) = c"
  "(\<rm>a)\<rs>b\<ra>c = c\<rs>a\<rs>b"
  "(\<rm>b)\<rs>a\<ra>c = c\<rs>a\<rs>b"
  "(\<rm>((\<rm>a)\<ra>b\<ra>c)) = a\<rs>b\<rs>c"
  "a\<ra>b\<ra>c\<rs>a = b\<ra>c"
  "a\<ra>b\<rs>(a\<ra>c) = b\<rs>c"
  using prems Int_ZF_1_T2 
    group0.group0_4_L4B group0.group0_4_L6D group0.group0_4_L4D
    group0.group0_4_L6B group0.group0_4_L6E
  by auto;

text{*Another couple of rearrangements with three integers.*}

lemma (in int0) Int_ZF_1_2_L9A: 
  assumes A1: "a∈\<int>"  "b∈\<int>"  "c∈\<int>"
  shows "(\<rm>(a\<rs>b\<rs>c)) = c\<ra>b\<rs>a"
proof -
  from A1 have T: 
    "a\<rs>b ∈ \<int>"  "(\<rm>(a\<rs>b)) ∈ \<int>"  "(\<rm>b) ∈ \<int>" using 
    Int_ZF_1_1_L4 Int_ZF_1_1_L5 by auto;
  with A1 have "(\<rm>(a\<rs>b\<rs>c)) = c \<rs> ((\<rm>b)\<ra>a)"
     using Int_ZF_1_1_L5 by simp;
   also from A1 T have "… = c\<ra>b\<rs>a"
     using Int_ZF_1_1_L6 Int_ZF_1_1_L5B
     by simp;
  finally show "(\<rm>(a\<rs>b\<rs>c)) = c\<ra>b\<rs>a" 
    by simp;
qed;
  
 
text{*Another rearrangement with three integers.*}

lemma (in int0) Int_ZF_1_2_L10: 
  assumes A1: "a∈\<int>" "b∈\<int>" "c∈\<int>"
  shows "(a\<ra>\<one>)·b \<ra> (c\<ra>\<one>)·b = (c\<ra>a\<ra>\<two>)·b"
proof -
  from A1 have "a\<ra>\<one> ∈ \<int>" "c\<ra>\<one> ∈ \<int>" 
    using int_zero_one_are_int Int_ZF_1_1_L5 by auto; 
  with A1 have 
    "(a\<ra>\<one>)·b \<ra> (c\<ra>\<one>)·b = (a\<ra>\<one>\<ra>(c\<ra>\<one>))·b"
    using Int_ZF_1_1_L1 by simp;
  also from A1 have "… = (c\<ra>a\<ra>\<two>)·b"
    using Int_ZF_1_2_L8 by simp;
  finally show ?thesis by simp;
qed;

text{*A technical rearrangement involing inequalities with absolute value.*}

lemma (in int0) Int_ZF_1_2_L10A: 
  assumes A1: "a∈\<int>"  "b∈\<int>"  "c∈\<int>"  "e∈\<int>"
  and A2: "abs(a·b\<rs>c) \<lsq> d"  "abs(b·a\<rs>e) \<lsq> f"
  shows "abs(c\<rs>e) \<lsq>  f\<ra>d"
proof -  
  from A1 A2 have T1:
    "d∈\<int>"  "f∈\<int>"  "a·b ∈ \<int>"  "a·b\<rs>c ∈ \<int>"  "b·a\<rs>e ∈ \<int>"
    using Int_ZF_2_L1A Int_ZF_1_1_L5 by auto;
  with A2 have
    "abs((b·a\<rs>e)\<rs>(a·b\<rs>c)) \<lsq> f \<ra>d"
    using Int_ZF_2_L21 by simp;
  with A1 T1 show "abs(c\<rs>e) \<lsq> f\<ra>d" 
    using Int_ZF_1_1_L5 Int_ZF_1_2_L9 by simp;
qed;

text{*Some arithmetics.*}

lemma (in int0) Int_ZF_1_2_L11: assumes A1: "a∈\<int>"
  shows 
  "a\<ra>\<one>\<ra>\<two> = a\<ra>\<three>"
  "a = \<two>·a \<rs> a"
proof -
  from A1 show "a\<ra>\<one>\<ra>\<two> = a\<ra>\<three>"
    using int_zero_one_are_int int_two_three_are_int Int_ZF_1_T2 group0.group0_4_L4C
    by simp
  from A1 show "a = \<two>·a \<rs> a"
    using int_zero_one_are_int Int_ZF_1_1_L1 Int_ZF_1_1_L4 Int_ZF_1_T2 group0.group0_2_L16
    by simp;
qed;

text{*A simple rearrangement with three integers.*}

lemma (in int0) Int_ZF_1_2_L12: 
  assumes "a∈\<int>"  "b∈\<int>"  "c∈\<int>"
  shows 
  "(b\<rs>c)·a = a·b \<rs> a·c"
  using prems Int_ZF_1_1_L6 Int_ZF_1_1_L5 by simp;

text{*A big rearrangement with five integers.*}

lemma (in int0) Int_ZF_1_2_L13: 
  assumes A1: "a∈\<int>"  "b∈\<int>"  "c∈\<int>" "d∈\<int>"  "x∈\<int>"
  shows "(x\<ra>(a·x\<ra>b)\<ra>c)·d = d·(a\<ra>\<one>)·x \<ra> (b·d\<ra>c·d)"
proof -
  from A1 have T1: 
    "a·x ∈ \<int>"   "(a\<ra>\<one>)·x ∈ \<int>"  
    "(a\<ra>\<one>)·x \<ra> b ∈ \<int>" 
    using Int_ZF_1_1_L5 int_zero_one_are_int by auto
  with A1 have "(x\<ra>(a·x\<ra>b)\<ra>c)·d = ((a\<ra>\<one>)·x \<ra> b)·d \<ra> c·d"
    using Int_ZF_1_1_L7 Int_ZF_1_2_L7 Int_ZF_1_1_L1
    by simp;
  also from A1 T1 have "… = (a\<ra>\<one>)·x·d \<ra> b · d \<ra> c·d"
    using Int_ZF_1_1_L1 by simp;
  finally have "(x\<ra>(a·x\<ra>b)\<ra>c)·d = (a\<ra>\<one>)·x·d \<ra> b·d \<ra> c·d"
    by simp;
  moreover from A1 T1 have "(a\<ra>\<one>)·x·d = d·(a\<ra>\<one>)·x"
    using int_zero_one_are_int Int_ZF_1_1_L5 Int_ZF_1_1_L7 by simp;
  ultimately have "(x\<ra>(a·x\<ra>b)\<ra>c)·d = d·(a\<ra>\<one>)·x \<ra> b·d \<ra> c·d"
    by simp;
  moreover from A1 T1 have 
    "d·(a\<ra>\<one>)·x ∈ \<int>"  "b·d ∈ \<int>"  "c·d ∈ \<int>"
    using int_zero_one_are_int Int_ZF_1_1_L5 by auto;
  ultimately show ?thesis using Int_ZF_1_1_L7 by simp;
qed;

text{*Rerrangement about adding linear functions.*}

lemma (in int0) Int_ZF_1_2_L14:
  assumes "a∈\<int>"  "b∈\<int>"  "c∈\<int>" "d∈\<int>"  "x∈\<int>"
  shows "(a·x \<ra> b) \<ra> (c·x \<ra> d) = (a\<ra>c)·x \<ra> (b\<ra>d)"
  using prems Int_ZF_1_1_L2 ring0.Ring_ZF_2_L3 by simp;

text{*A rearrangement with four integers. 
  Again we have to use the generic set notation to use a theorem proven in 
  different context.*}

lemma (in int0) Int_ZF_1_2_L15: assumes A1: "a∈\<int>"  "b∈\<int>"  "c∈\<int>" "d∈\<int>"
  and A2: "a = b\<rs>c\<rs>d"
  shows 
  "d = b\<rs>a\<rs>c"
  "d = (\<rm>a)\<ra>b\<rs>c"
  "b = a\<ra>d\<ra>c"
proof -
  let ?G = "int"
  let ?f = "IntegerAddition"
  from A1 A2 have I:
    "group0(?G, ?f)"   "?f {is commutative on} ?G"
    "a ∈ ?G"  "b ∈ ?G" "c ∈ ?G"  "d ∈ ?G"
    "a = ?f`⟨?f`⟨b,GroupInv(?G, ?f)`(c)⟩,GroupInv(?G, ?f)`(d)⟩"
    using Int_ZF_1_T2 by auto;
  then have  
    "d = ?f`⟨?f`⟨b,GroupInv(?G, ?f)`(a)⟩,GroupInv(?G,?f)`(c)⟩"
    by (rule group0.group0_4_L9);
  then show "d = b\<rs>a\<rs>c" by simp;
  from I have "d = ?f`⟨?f`⟨GroupInv(?G, ?f)`(a),b⟩, GroupInv(?G, ?f)`(c)⟩"
    by (rule group0.group0_4_L9);
  thus "d = (\<rm>a)\<ra>b\<rs>c"
    by simp;
  from I have "b = ?f`⟨?f`⟨a, d⟩,c⟩"
    by (rule group0.group0_4_L9);
  thus "b = a\<ra>d\<ra>c" by simp
qed;

text{*A rearrangement with four integers. Property of groups.*}

lemma (in int0) Int_ZF_1_2_L16:
  assumes "a∈\<int>"  "b∈\<int>"  "c∈\<int>" "d∈\<int>"
  shows "a\<ra>(b\<rs>c)\<ra>d = a\<ra>b\<ra>d\<rs>c"
  using prems Int_ZF_1_T2 group0.group0_4_L8 by simp

text{*Some rearrangements with three integers. Properties of groups.*}

lemma (in int0) Int_ZF_1_2_L17:
  assumes A1: "a∈\<int>"  "b∈\<int>"  "c∈\<int>"
  shows 
  "a\<ra>b\<rs>c\<ra>(c\<rs>b) = a"
  "a\<ra>(b\<ra>c)\<rs>c = a\<ra>b"
proof -
  let ?G = "int"
  let ?f = "IntegerAddition"
  from A1 have I:
    "group0(?G, ?f)"
    "a ∈ ?G"  "b ∈ ?G" "c ∈ ?G"
    using Int_ZF_1_T2 by auto;
  then have 
    "?f`⟨?f`⟨?f`⟨a,b⟩,GroupInv(?G, ?f)`(c)⟩,?f`⟨c,GroupInv(?G, ?f)`(b)⟩⟩ = a"
    by (rule group0.group0_2_L14A);
  thus "a\<ra>b\<rs>c\<ra>(c\<rs>b) = a" by simp;
  from I have
    "?f`⟨?f`⟨a,?f`⟨b,c⟩⟩,GroupInv(?G, ?f)`(c)⟩ = ?f`⟨a,b⟩"
    by (rule group0.group0_2_L14A)
  thus "a\<ra>(b\<ra>c)\<rs>c = a\<ra>b" by simp
qed;

text{*Another rearrangement with three integers. Property of abelian groups.*}

lemma (in int0) Int_ZF_1_2_L18:
  assumes A1: "a∈\<int>"  "b∈\<int>"  "c∈\<int>"
  shows "a\<ra>b\<rs>c\<ra>(c\<rs>a) = b"
proof -
  let ?G = "int"
  let ?f = "IntegerAddition"
  from A1 have
    "group0(?G, ?f)"   "?f {is commutative on} ?G"
    "a ∈ ?G"  "b ∈ ?G" "c ∈ ?G"
    using Int_ZF_1_T2 by auto;
  then have
    "?f`⟨?f`⟨?f`⟨a,b⟩,GroupInv(?G, ?f)`(c)⟩,?f`⟨c,GroupInv(?G, ?f)`(a)⟩⟩ = b"
    by (rule group0.group0_4_L6D);
  thus "a\<ra>b\<rs>c\<ra>(c\<rs>a) = b" by simp;
qed;

section{*Integers as an ordered ring*}

text{*We already know from @{text "Int_ZF"} that integers with addition 
  form a linearly ordered group. To show that integers form 
  an ordered ring we need the fact that the set of nonnegative integers
  is closed under multiplication. Since we don't have the theory
  of oredered rings we temporarily put some facts about integers as an ordered
  ring in this section.*}

text{* We start with the property that a product of 
  nonnegative integers is nonnegative. The proof is by induction and the next
  lemma is the induction step.*}

lemma (in int0) Int_ZF_1_3_L1: assumes A1: "\<zero>\<lsq>a"  "\<zero>\<lsq>b"
  and A3: "\<zero> \<lsq> a·b"
  shows "\<zero> \<lsq> a·(b\<ra>\<one>)"
proof -
  from A1 A3 have "\<zero>\<ra>\<zero> \<lsq> a·b\<ra>a"
    using int_ineq_add_sides by simp;
  with A1 show "\<zero> \<lsq> a·(b\<ra>\<one>)"
    using int_zero_one_are_int Int_ZF_1_1_L4 Int_ZF_2_L1A Int_ZF_1_2_L7 
    by simp;
qed;

text{* Product of nonnegative integers is nonnegative.*}
  
lemma (in int0) Int_ZF_1_3_L2: assumes A1: "\<zero>\<lsq>a"  "\<zero>\<lsq>b" 
  shows "\<zero>\<lsq>a·b"
proof -
  from A1 have "\<zero>\<lsq>b" by simp
  moreover from A1 have "\<zero> \<lsq> a·\<zero>" using
    Int_ZF_2_L1A Int_ZF_1_1_L4 int_zero_one_are_int int_ord_is_refl refl_def
    by simp;
  moreover from A1 have 
    "∀m. \<zero>\<lsq>m ∧ \<zero>\<lsq>a·m --> \<zero> \<lsq> a·(m\<ra>\<one>)"
    using Int_ZF_1_3_L1 by simp;
  ultimately show "\<zero>\<lsq>a·b" by (rule Induction_on_int);
qed;

text{*The set of nonnegative integers is closed under  multiplication.*}

lemma (in int0) Int_ZF_1_3_L2A: shows 
  "\<int>+ {is closed under} IntegerMultiplication"
proof -
  { fix a b assume "a∈\<int>+"  "b∈\<int>+"
    then have "a·b ∈\<int>+"
      using Int_ZF_1_3_L2 Int_ZF_2_T1 group3.OrderedGroup_ZF_1_L2 
      by simp;
  } then have "∀a∈\<int>+.∀b∈\<int>+.a·b ∈\<int>+" by simp; 
  then show ?thesis using IsOpClosed_def by simp
qed;

text{*Integers form an ordered ring. All theorems proven in the 
  @{text "ring1"} context are valid in @{text "int0"} context.*}

theorem (in int0) Int_ZF_1_3_T1: shows 
  "IsAnOrdRing(\<int>,IntegerAddition,IntegerMultiplication,IntegerOrder)"
  "ring1(\<int>,IntegerAddition,IntegerMultiplication,IntegerOrder)"
  using Int_ZF_1_1_L2 Int_ZF_2_L1B Int_ZF_1_3_L2A Int_ZF_2_T1
    OrdRing_ZF_1_L6 OrdRing_ZF_1_L2 by auto;
  
text{*Product of integers that are greater that one is greater than one. 
  The proof is by induction and 
  the next step is the induction step.*}

lemma (in int0) Int_ZF_1_3_L3_indstep:  
  assumes A1: "\<one>\<lsq>a"  "\<one>\<lsq>b"
  and A2: "\<one> \<lsq> a·b"
  shows "\<one> \<lsq> a·(b\<ra>\<one>)"
proof -
   from A1 A2 have "\<one>\<lsq>\<two>" and "\<two> \<lsq> a·(b\<ra>\<one>)"
    using Int_ZF_2_L1A int_ineq_add_sides Int_ZF_2_L16B Int_ZF_1_2_L7 
    by auto;
  then show "\<one> \<lsq> a·(b\<ra>\<one>)" by (rule Int_order_transitive);
qed;

text{*Product of integers that are greater that one is greater than one. *}

lemma (in int0) Int_ZF_1_3_L3: 
  assumes A1: "\<one>\<lsq>a" "\<one>\<lsq>b"
  shows "\<one> \<lsq> a·b"
proof -
  from A1 have "\<one>\<lsq>b"  "\<one>\<lsq>a·\<one>"
    using Int_ZF_2_L1A Int_ZF_1_1_L4 by auto;
  moreover from A1 have 
    "∀m. \<one>\<lsq>m ∧ \<one> \<lsq> a·m --> \<one> \<lsq> a·(m\<ra>\<one>)"
    using Int_ZF_1_3_L3_indstep by simp;
  ultimately show "\<one> \<lsq> a·b" by (rule Induction_on_int);
qed;

text{*$|a\cdot (-b)| = |(-a)\cdot b| = |(-a)\cdot (-b)| = |a\cdot b|$
  This is a property of ordered rings..*}

lemma (in int0) Int_ZF_1_3_L4: assumes "a∈\<int>"  "b∈\<int>"
  shows 
  "abs((\<rm>a)·b) = abs(a·b)"
  "abs(a·(\<rm>b)) = abs(a·b)"
  "abs((\<rm>a)·(\<rm>b)) = abs(a·b)"
  using prems Int_ZF_1_1_L5 Int_ZF_2_L17 by auto;

text{*Absolute value of a product is the product of absolute values.
  Property of ordered rings.*}

lemma (in int0) Int_ZF_1_3_L5:
  assumes A1: "a∈\<int>"  "b∈\<int>"
  shows "abs(a·b) = abs(a)·abs(b)"
  using prems Int_ZF_1_3_T1 ring1.OrdRing_ZF_2_L5 by simp;

text{*Double nonnegative is nonnegative. Property of ordered rings.*}

lemma (in int0) Int_ZF_1_3_L5A: assumes "\<zero>\<lsq>a"
  shows "\<zero>\<lsq>\<two>·a"
  using prems Int_ZF_1_3_T1 ring1.OrdRing_ZF_1_L5A by simp;


text{*The next lemma shows what happens when one integer is not
  greater or equal than another.*}

lemma (in int0) Int_ZF_1_3_L6: 
  assumes A1: "a∈\<int>"  "b∈\<int>" 
  shows "¬(b\<lsq>a) <-> a\<ra>\<one> \<lsq> b"
proof;
  assume A3: "¬(b\<lsq>a)"
  with A1 have  "a\<lsq>b" by (rule Int_ZF_2_L19);
  then have "a = b   ∨  a\<ra>\<one> \<lsq> b"
    using Int_ZF_4_L1B by simp; 
  moreover from A1 A3 have "a≠b" by (rule Int_ZF_2_L19);
  ultimately show "a\<ra>\<one> \<lsq> b" by simp;
next assume A4: "a\<ra>\<one> \<lsq> b"
  { assume "b\<lsq>a" 
    with A4 have "a\<ra>\<one> \<lsq> a" by (rule Int_order_transitive);
    moreover from A1 have "a \<lsq> a\<ra>\<one>"
      using Int_ZF_2_L12B by simp;
    ultimately have "a\<ra>\<one> = a"
      by (rule Int_ZF_2_L3); 
    with A1 have False using Int_ZF_1_L14 by simp; 
  } then show "¬(b\<lsq>a)" by auto
qed;

text{*Another form of stating that there are no integers
  between integers $m$ and $m+1$.*}

corollary (in int0) no_int_between: assumes A1: "a∈\<int>"  "b∈\<int>"
  shows "b\<lsq>a ∨ a\<ra>\<one> \<lsq> b"
  using A1 Int_ZF_1_3_L6 by auto; (* A1 has to be here *)    


text{*Another way of saying what it means that one integer is not
  greater or equal than another.*}

corollary (in int0) Int_ZF_1_3_L6A:
  assumes A1: "a∈\<int>"  "b∈\<int>" and A2: "¬(b\<lsq>a)"
  shows "a \<lsq> b\<rs>\<one>"
proof -
  from A1 A2 have "a\<ra>\<one> \<rs> \<one> \<lsq> b \<rs> \<one>"
    using Int_ZF_1_3_L6 int_zero_one_are_int Int_ZF_1_1_L4 
      int_ord_transl_inv by simp;
  with A1 show "a \<lsq> b\<rs>\<one>"
    using int_zero_one_are_int Int_ZF_1_2_L3
    by simp;
qed;

text{*Yet another form of stating that there are nointegers between
  $m$ and $m+1$.*}

lemma (in int0) no_int_between1: 
  assumes A1: "a\<lsq>b"  and A2: "a≠b"
  shows 
  "a\<ra>\<one> \<lsq> b"
  "a \<lsq> b\<rs>\<one>"
proof -
  from A1 have T: "a∈\<int>"  "b∈\<int>" using Int_ZF_2_L1A
    by auto;
  { assume "b\<lsq>a"
    with A1 have "a=b" by (rule Int_ZF_2_L3);
    with A2 have False by simp }
  then have "¬(b\<lsq>a)" by auto;
  with T show 
    "a\<ra>\<one> \<lsq> b" 
    "a \<lsq> b\<rs>\<one>"
    using no_int_between Int_ZF_1_3_L6A by auto;
qed;

text{*We can decompose proofs into three cases: $a=b$, $a\leq b-1b$ or 
  $a\geq b+1b$. *}

lemma (in int0) Int_ZF_1_3_L6B: assumes A1: "a∈\<int>"  "b∈\<int>"
  shows "a=b ∨ (a \<lsq> b\<rs>\<one>) ∨ (b\<ra>\<one> \<lsq>a)"
proof -
  from A1 have "a=b ∨ (a\<lsq>b ∧ a≠b) ∨ (b\<lsq>a ∧ b≠a)"
    using Int_ZF_2_T1 group3.OrderedGroup_ZF_1_L31
    by simp;
  then show ?thesis using no_int_between1
    by auto;
qed;

text{*A special case of @{text "Int_ZF_1_3_L6B"} when $b=0$. This
  allows to split the proofs in cases $a\leq -1$, $a=0$ and $a\geq 1$.*}

corollary (in int0) Int_ZF_1_3_L6C: assumes A1: "a∈\<int>"
  shows "a=\<zero> ∨ (a \<lsq> \<rm>\<one>) ∨ (\<one>\<lsq>a)"  
proof -
  from A1 have "a=\<zero> ∨ (a \<lsq> \<zero> \<rs>\<one>) ∨ (\<zero> \<ra>\<one> \<lsq>a)"
    using int_zero_one_are_int Int_ZF_1_3_L6B by simp;
  then show ?thesis using Int_ZF_1_1_L4 int_zero_one_are_int
    by simp;
qed;
  

text{*An integer is not less or equal zero iff it is greater or equal one.*}

lemma (in int0) Int_ZF_1_3_L7: assumes "a∈\<int>" 
  shows "¬(a\<lsq>\<zero>) <-> \<one> \<lsq> a"
  using prems int_zero_one_are_int Int_ZF_1_3_L6 Int_ZF_1_1_L4
  by simp;

text{*Product of positive integers is positive.*}

lemma (in int0) Int_ZF_1_3_L8: 
  assumes "a∈\<int>"  "b∈\<int>" 
  and "¬(a\<lsq>\<zero>)"  "¬(b\<lsq>\<zero>)"
  shows "¬((a·b) \<lsq> \<zero>)"
  using prems Int_ZF_1_3_L7 Int_ZF_1_3_L3 Int_ZF_1_1_L5 Int_ZF_1_3_L7
  by simp;

text{*If $a\cdot b$ is nonnegative and $b$ is positive, then $a$ is 
  nonnegative. Proof by contradiction.*}

lemma (in int0) Int_ZF_1_3_L9:
  assumes A1: "a∈\<int>"  "b∈\<int>"
  and A2:  "¬(b\<lsq>\<zero>)" and A3: "a·b \<lsq> \<zero>" 
  shows "a\<lsq>\<zero>"
proof -
  { assume "¬(a\<lsq>\<zero>)"
    with A1 A2 have "¬((a·b) \<lsq> \<zero>)" using Int_ZF_1_3_L8
      by simp;
  } with A3 show "a\<lsq>\<zero>" by auto;
qed;

text{*One integer is less or equal another iff the difference is nonpositive.*}

lemma (in int0) Int_ZF_1_3_L10:
  assumes "a∈\<int>"  "b∈\<int>"
  shows "a\<lsq>b <-> a\<rs>b \<lsq> \<zero>"
  using prems Int_ZF_2_T1 group3.OrderedGroup_ZF_1_L9
  by simp;

text{*Some conclusions from the fact that one integer
  is less or equal than another.*}

lemma (in int0) Int_ZF_1_3_L10A: assumes "a\<lsq>b"
  shows "\<zero> \<lsq> b\<rs>a"
  using prems Int_ZF_2_T1 group3.OrderedGroup_ZF_1_L12A
  by simp;

text{*We can simplify out a positive element on both sides of an 
  inequality.*}

lemma (in int0) Int_ineq_simpl_positive: 
  assumes A1: "a∈\<int>"  "b∈\<int>"  "c∈\<int>" 
  and A2: "a·c \<lsq> b·c" and A4: "¬(c\<lsq>\<zero>)"
  shows "a \<lsq> b"
proof -
  from A1 A4 have "a\<rs>b ∈  \<int>"  "c∈\<int>"  "¬(c\<lsq>\<zero>)"
    using Int_ZF_1_1_L5 by auto;
  moreover from A1 A2 have "(a\<rs>b)·c \<lsq> \<zero>"
    using Int_ZF_1_1_L5 Int_ZF_1_3_L10 Int_ZF_1_1_L6
    by simp;
  ultimately have "a\<rs>b \<lsq> \<zero>" by (rule Int_ZF_1_3_L9);
  with A1 show "a \<lsq> b" using Int_ZF_1_3_L10 by simp;
qed;

text{*A technical lemma about conclusion from an inequality between
  absolute values. This is a property of ordered rings.*}

lemma (in int0) Int_ZF_1_3_L11:
  assumes A1: "a∈\<int>"  "b∈\<int>"
  and A2: "¬(abs(a) \<lsq> abs(b))"
  shows "¬(abs(a) \<lsq> \<zero>)"
proof -
  { assume "abs(a) \<lsq> \<zero>"
    moreover from A1 have "\<zero> \<lsq> abs(a)" using int_abs_nonneg
      by simp;
    ultimately have "abs(a) = \<zero>" by (rule Int_ZF_2_L3);
    with A1 A2 have False using int_abs_nonneg by simp;
  } then show  "¬(abs(a) \<lsq> \<zero>)" by auto;
qed;

text{*Negative times positive is negative. This a property of ordered rings.*}

lemma (in int0) Int_ZF_1_3_L12:
  assumes "a\<lsq>\<zero>"  and "\<zero>\<lsq>b"
  shows "a·b \<lsq> \<zero>"
  using prems Int_ZF_1_3_T1 ring1.OrdRing_ZF_1_L8 
  by simp;

text{*We can multiply an inequality by a nonnegative number. 
  This is a property of ordered rings.*}

lemma (in int0) Int_ZF_1_3_L13:
  assumes A1: "a\<lsq>b" and A2: "\<zero>\<lsq>c"
  shows 
  "a·c \<lsq> b·c"
  "c·a \<lsq> c·b"
  using prems Int_ZF_1_3_T1 ring1.OrdRing_ZF_1_L9 by auto;

text{*A technical lemma about decreasing a factor in an inequality.*}

lemma (in int0) Int_ZF_1_3_L13A:
  assumes "\<one>\<lsq>a" and "b\<lsq>c" and "(a\<ra>\<one>)·c \<lsq> d"
  shows "(a\<ra>\<one>)·b \<lsq> d"
proof -
  from prems have 
    "(a\<ra>\<one>)·b \<lsq> (a\<ra>\<one>)·c"
    "(a\<ra>\<one>)·c \<lsq> d"
    using Int_ZF_2_L16C Int_ZF_1_3_L13 by auto;
  then show "(a\<ra>\<one>)·b \<lsq> d" by (rule Int_order_transitive);
qed;

text{*We can multiply an inequality by a positive number. 
  This is a property of ordered rings.*}

lemma (in int0) Int_ZF_1_3_L13B:  
  assumes A1: "a\<lsq>b" and A2: "c∈\<int>+"
  shows  
  "a·c \<lsq> b·c"  
  "c·a \<lsq> c·b"
proof -
  let ?R = "\<int>"
  let ?A = "IntegerAddition"
  let ?M = "IntegerMultiplication"
  let ?r = "IntegerOrder"
  from A1 A2 have 
    "ring1(?R, ?A, ?M, ?r)" 
    "⟨a,b⟩ ∈ ?r"
    "c ∈ PositiveSet(?R, ?A, ?r)"
    using Int_ZF_1_3_T1 by auto;
  then show 
    "a·c \<lsq> b·c"  
    "c·a \<lsq> c·b"
    using ring1.OrdRing_ZF_1_L9A by auto;
qed;

text{*A rearrangement with four integers and absolute value.*}

lemma (in int0) Int_ZF_1_3_L14:
  assumes  A1: "a∈\<int>"  "b∈\<int>"  "c∈\<int>"  "d∈\<int>" 
  shows "abs(a·b)\<ra>(abs(a)\<ra>c)·d = (d\<ra>abs(b))·abs(a)\<ra>c·d"
proof -;
  from A1 have T1: 
    "abs(a) ∈ \<int>"  "abs(b) ∈ \<int>" 
    "abs(a)·abs(b) ∈ \<int>" 
    "abs(a)·d ∈ \<int>" 
    "c·d ∈ \<int>"
    "abs(b)\<ra>d ∈ \<int>" 
    using Int_ZF_2_L14 Int_ZF_1_1_L5 by auto;
  with A1 have "abs(a·b)\<ra>(abs(a)\<ra>c)·d = abs(a)·(abs(b)\<ra>d)\<ra>c·d"
    using Int_ZF_1_3_L5 Int_ZF_1_1_L1 Int_ZF_1_1_L7 by simp;
  with A1 T1 show ?thesis using Int_ZF_1_1_L5 by simp;
qed;

text{*A technical lemma about what happens when one absolute value is
  not greater or equal than another.*}

lemma (in int0) Int_ZF_1_3_L15: assumes A1: "m∈\<int>" "n∈\<int>"
  and A2: "¬(abs(m) \<lsq> abs(n))"
  shows "n \<lsq> abs(m)"  "m≠\<zero>" 
proof -
  from A1 have T1: "n \<lsq> abs(n)" 
    using Int_ZF_2_L19C by simp;
  from A1 have "abs(n) ∈ \<int>"  "abs(m) ∈ \<int>"
    using Int_ZF_2_L14 by auto;
  moreover from A2 have "¬(abs(m) \<lsq> abs(n))" .
  ultimately have "abs(n) \<lsq> abs(m)"
    by (rule Int_ZF_2_L19);
  with T1 show  "n \<lsq> abs(m)" by (rule Int_order_transitive);
  from A1 A2 show "m≠\<zero>" using Int_ZF_2_L18 int_abs_nonneg by auto;
qed;

text{*Negative of a nonnegative is nonpositive.*}

lemma (in int0) Int_ZF_1_3_L16: assumes A1: "\<zero> \<lsq> m"
  shows "(\<rm>m) \<lsq> \<zero>"
proof -
  from A1 have "(\<rm>m) \<lsq> (\<rm>\<zero>)"
    using Int_ZF_2_L10 by simp;
  then show "(\<rm>m) \<lsq> \<zero>" using Int_ZF_1_L11
    by simp;
qed;

text{*Some statements about intervals centered at $0$.*}

lemma (in int0) Int_ZF_1_3_L17: assumes A1: "m∈\<int>"
  shows 
  "(\<rm>abs(m)) \<lsq> abs(m)"
  "(\<rm>abs(m))..abs(m) ≠ 0"
proof -
  from A1 have "(\<rm>abs(m)) \<lsq> \<zero>"  "\<zero> \<lsq> abs(m)" 
    using int_abs_nonneg Int_ZF_1_3_L16 by auto;
  then show "(\<rm>abs(m)) \<lsq> abs(m)" by (rule Int_order_transitive);
  then have "abs(m) ∈ (\<rm>abs(m))..abs(m)"
    using int_ord_is_refl Int_ZF_2_L1A Order_ZF_2_L2 by simp;
  thus "(\<rm>abs(m))..abs(m) ≠ 0" by auto;
qed

text{*The greater of two integers is indeed greater than both, 
  and the smaller one is smaller that both.*}

lemma (in int0) Int_ZF_1_3_L18: assumes A1: "m∈\<int>"  "n∈\<int>"
  shows 
  "m \<lsq> GreaterOf(IntegerOrder,m,n)"
  "n \<lsq> GreaterOf(IntegerOrder,m,n)"
  "SmallerOf(IntegerOrder,m,n) \<lsq> m"
  "SmallerOf(IntegerOrder,m,n) \<lsq> n"
  using prems Int_ZF_2_T1 Order_ZF_3_L2 by auto;

text{*If $|m| \leq n$, then $m \in -n..n$.*}

lemma (in int0) Int_ZF_1_3_L19: 
  assumes A1: "m∈\<int>" and A2: "abs(m) \<lsq> n"
  shows 
  "(\<rm>n) \<lsq> m"  "m \<lsq> n"
  "m ∈ (\<rm>n)..n"
  "\<zero> \<lsq> n"
  using prems Int_ZF_2_T1 group3.OrderedGroup_ZF_3_L8  
    group3.OrderedGroup_ZF_3_L8A Order_ZF_2_L1 
  by auto;

text{* A slight generalization of the above lemma. *}

lemma (in int0) Int_ZF_1_3_L19A: 
  assumes A1: "m∈\<int>" and A2: "abs(m) \<lsq> n" and A3: "\<zero>\<lsq>k"
  shows "(\<rm>(n\<ra>k)) \<lsq> m"
  using prems Int_ZF_2_T1 group3.OrderedGroup_ZF_3_L8B
  by simp;

text{*Sets of integers that have absolute value bounded are bounded.*}

lemma (in int0) Int_ZF_1_3_L20:
  assumes A1: "∀x∈X. b(x) ∈ \<int> ∧ abs(b(x)) \<lsq> L"
  shows "IsBounded({b(x). x∈X},IntegerOrder)"
proof -;
  let ?G = "\<int>"
  let ?P = "IntegerAddition"
  let ?r = "IntegerOrder"
  from A1 have
    "group3(?G, ?P, ?r)"
    "?r {is total on} ?G"
    "∀x∈X. b(x) ∈ ?G ∧ ⟨AbsoluteValue(?G, ?P, ?r) ` b(x), L⟩ ∈ ?r"
    using Int_ZF_2_T1 by auto;
  then show "IsBounded({b(x). x∈X},IntegerOrder)"
    by (rule group3.OrderedGroup_ZF_3_L9A);
qed;

text{*If a set is bounded, then the absolute values of the elements of that
  set are bounded.*}

lemma (in int0) Int_ZF_1_3_L20A: assumes "IsBounded(A,IntegerOrder)"
  shows "∃L. ∀a∈A. abs(a) \<lsq> L"
  using prems Int_ZF_2_T1 group3.OrderedGroup_ZF_3_L10A 
  by simp;

text{*Absolute vaues of integers from a finite image of integers are bounded
  by an integer.*}

lemma (in int0) Int_ZF_1_3_L20AA: 
  assumes A1: "{b(x). x∈\<int>} ∈ Fin(\<int>)"
  shows "∃L∈\<int>. ∀x∈\<int>. abs(b(x)) \<lsq> L"
  using prems int_not_empty Int_ZF_2_T1 group3.OrderedGroup_ZF_3_L11A
  by simp;

text{*If absolute values of values of some integer function are bounded,
  then the image a set from the domain is a bounded set.*}

lemma (in int0) Int_ZF_1_3_L20B:
  assumes "f:X->\<int>" and "A⊆X" and "∀x∈A. abs(f`(x)) \<lsq> L"
  shows  "IsBounded(f``(A),IntegerOrder)"
proof -
  let ?G = "\<int>"
  let ?P = "IntegerAddition"
  let ?r = "IntegerOrder"
  from prems have 
    "group3(?G, ?P, ?r)"
    "?r {is total on} ?G"
    "f:X->?G"
    "A⊆X"
    "∀x∈A. ⟨AbsoluteValue(?G, ?P, ?r)`(f`(x)), L⟩ ∈ ?r"
    using Int_ZF_2_T1 by auto;
  then show "IsBounded(f``(A), ?r)" 
    by (rule group3.OrderedGroup_ZF_3_L9B);
qed;

text{*A special case of the previous lemma for a function from integers to 
  integers.*}

corollary (in int0) Int_ZF_1_3_L20C:
  assumes "f:\<int>->\<int>" and "∀m∈\<int>. abs(f`(m)) \<lsq> L"
  shows "f``(\<int>) ∈ Fin(\<int>)"
proof -
  from prems have "f:\<int>->\<int>" "\<int> ⊆ \<int>"  "∀m∈\<int>. abs(f`(m)) \<lsq> L"
    by auto;
  then have "IsBounded(f``(\<int>),IntegerOrder)"
    by (rule Int_ZF_1_3_L20B);
  then show "f``(\<int>) ∈ Fin(\<int>)" using Int_bounded_iff_fin
    by simp;
qed;

text{*A triangle inequality with three integers. Property of 
  linearly ordered abelian groups.*}

lemma (in int0) int_triangle_ineq3: 
  assumes A1: "a∈\<int>"  "b∈\<int>"  "c∈\<int>"
  shows "abs(a\<rs>b\<rs>c) \<lsq> abs(a) \<ra> abs(b) \<ra> abs(c)"
proof -
  from A1 have T: "a\<rs>b ∈ \<int>"  "abs(c) ∈ \<int>"
    using Int_ZF_1_1_L5 Int_ZF_2_L14 by auto;
  with A1 have "abs(a\<rs>b\<rs>c) \<lsq> abs(a\<rs>b) \<ra> abs(c)"
    using Int_triangle_ineq1 by simp;
  moreover from A1 T have 
    "abs(a\<rs>b) \<ra> abs(c) \<lsq>  abs(a) \<ra> abs(b) \<ra> abs(c)"
    using Int_triangle_ineq1 int_ord_transl_inv by simp;
  ultimately show ?thesis by (rule Int_order_transitive);
qed;

text{*If $a\leq c$ and $b\leq c$, then $a+b\leq 2\cdot c$. 
  Property of ordered rings.*}

lemma (in int0) Int_ZF_1_3_L21:
  assumes A1: "a\<lsq>c"  "b\<lsq>c" shows "a\<ra>b \<lsq> \<two>·c"
  using prems Int_ZF_1_3_T1 ring1.OrdRing_ZF_2_L6 by simp;

text{*If an integer $a$ is between $b$ and $b+c$, then
  $|b-a| \leq c$. Property of ordered groups.*}

lemma (in int0) Int_ZF_1_3_L22: 
  assumes "a\<lsq>b" and "c∈\<int>" and "b\<lsq> c\<ra>a"
  shows "abs(b\<rs>a) \<lsq> c"
  using prems Int_ZF_2_T1 group3.OrderedGroup_ZF_3_L8C
  by simp;

text{*An application of the triangle inequality with four
  integers. Property of linearly ordered abelian groups.*}

lemma (in int0) Int_ZF_1_3_L22A: 
  assumes  "a∈\<int>"  "b∈\<int>"  "c∈\<int>"  "d∈\<int>"
  shows "abs(a\<rs>c) \<lsq> abs(a\<ra>b) \<ra> abs(c\<ra>d) \<ra> abs(b\<rs>d)"
  using prems Int_ZF_1_T2 Int_ZF_2_T1 group3.OrderedGroup_ZF_3_L7F
  by simp;

text{*If an integer $a$ is between $b$ and $b+c$, then
  $|b-a| \leq c$. Property of ordered groups.
  A version of @{text "Int_ZF_1_3_L22"} with sligtly different
  assumptions.*}

lemma (in int0) Int_ZF_1_3_L23: 
  assumes A1: "a\<lsq>b" and A2: "c∈\<int>" and A3: "b\<lsq> a\<ra>c"
  shows "abs(b\<rs>a) \<lsq> c"
proof -
  from A1 have "a ∈ \<int>"
    using Int_ZF_2_L1A by simp;
  with A2 A3 have "b\<lsq> c\<ra>a"
    using Int_ZF_1_1_L5 by simp;
  with A1 A2 show "abs(b\<rs>a) \<lsq> c"
    using Int_ZF_1_3_L22 by simp;
qed;

section{*Maximum and minimum of a set of integers*}

text{*In this section we provide some sufficient conditions for
  integer subsets to have extrema (maxima and minima).*}

text{*Finite nonempty subsets of integers attain maxima and minima.*}

theorem (in int0) Int_fin_have_max_min:
  assumes A1: "A ∈ Fin(\<int>)" and A2: "A≠0"
  shows 
  "HasAmaximum(IntegerOrder,A)"
  "HasAminimum(IntegerOrder,A)"  
  "Maximum(IntegerOrder,A) ∈ A"
  "Minimum(IntegerOrder,A) ∈ A"
  "∀x∈A. x \<lsq> Maximum(IntegerOrder,A)"
  "∀x∈A. Minimum(IntegerOrder,A) \<lsq> x"
  "Maximum(IntegerOrder,A) ∈ \<int>"
  "Minimum(IntegerOrder,A) ∈ \<int>"
proof -
  from A1 have 
    "A=0 ∨ HasAmaximum(IntegerOrder,A)" and
    "A=0 ∨ HasAminimum(IntegerOrder,A)"
    using Int_ZF_2_T1 Int_ZF_2_L6 Finite_ZF_1_1_T1A Finite_ZF_1_1_T1B
    by auto;
  with A2 show 
    "HasAmaximum(IntegerOrder,A)"
    "HasAminimum(IntegerOrder,A)"  
    by auto;
  from A1 A2 show 
    "Maximum(IntegerOrder,A) ∈ A"
    "Minimum(IntegerOrder,A) ∈ A"
    "∀x∈A. x \<lsq> Maximum(IntegerOrder,A)"
    "∀x∈A. Minimum(IntegerOrder,A) \<lsq> x"
    using Int_ZF_2_T1 Finite_ZF_1_T2 by auto;
  moreover from A1 have "A⊆\<int>" using FinD by simp;
  ultimately show 
    "Maximum(IntegerOrder,A) ∈ \<int>"
    "Minimum(IntegerOrder,A) ∈ \<int>"
    by auto;
qed;

text{*Bounded nonempty integer subsets attain maximum and minimum.*}

theorem (in int0) Int_bounded_have_max_min:
  assumes "IsBounded(A,IntegerOrder)" and "A≠0"
  shows 
  "HasAmaximum(IntegerOrder,A)"
  "HasAminimum(IntegerOrder,A)"  
  "Maximum(IntegerOrder,A) ∈ A"
  "Minimum(IntegerOrder,A) ∈ A"
  "∀x∈A. x \<lsq> Maximum(IntegerOrder,A)"
  "∀x∈A. Minimum(IntegerOrder,A) \<lsq> x"
  "Maximum(IntegerOrder,A) ∈ \<int>"
  "Minimum(IntegerOrder,A) ∈ \<int>"
  using prems Int_fin_have_max_min Int_bounded_iff_fin
  by auto;

text{*Nonempty set of integers that is bounded below attains its minimum.*}

theorem (in int0) int_bounded_below_has_min:
  assumes A1: "IsBoundedBelow(A,IntegerOrder)" and A2: "A≠0"
  shows "
  HasAminimum(IntegerOrder,A)"
  "Minimum(IntegerOrder,A) ∈ A"
  (*"Minimum(IntegerOrder,A) ∈ \<int>"*)
  "∀x∈A. Minimum(IntegerOrder,A) \<lsq> x"
proof -
  from A1 A2 have
    "IntegerOrder {is total on} \<int>"
    "trans(IntegerOrder)"
    "IntegerOrder ⊆ \<int>×\<int>"
    "∀A. IsBounded(A,IntegerOrder) ∧ A≠0 --> HasAminimum(IntegerOrder,A)"
    "A≠0"  "IsBoundedBelow(A,IntegerOrder)"
    using Int_ZF_2_T1 Int_ZF_2_L6 Int_ZF_2_L1B Int_bounded_have_max_min
    by auto;
  then show "HasAminimum(IntegerOrder,A)"
    by (rule Order_ZF_4_L11); (*blast works too*)
  then show 
    "Minimum(IntegerOrder,A) ∈ A"
    "∀x∈A. Minimum(IntegerOrder,A) \<lsq> x"
    using Int_ZF_2_L4 Order_ZF_4_L4 by auto;
qed;

text{*Nonempty set of integers that is bounded above attains its
  maximum.*}

theorem (in int0) int_bounded_above_has_max:
  assumes A1: "IsBoundedAbove(A,IntegerOrder)" and A2: "A≠0"
  shows 
  "HasAmaximum(IntegerOrder,A)"
  "Maximum(IntegerOrder,A) ∈ A"
  "Maximum(IntegerOrder,A) ∈ \<int>"
  "∀x∈A. x \<lsq> Maximum(IntegerOrder,A)"
proof -
  from A1 A2 have
    "IntegerOrder {is total on} \<int>"
    "trans(IntegerOrder)" and
    I: "IntegerOrder ⊆ \<int>×\<int>" and
    "∀A. IsBounded(A,IntegerOrder) ∧ A≠0 --> HasAmaximum(IntegerOrder,A)"
    "A≠0"  "IsBoundedAbove(A,IntegerOrder)"
    using Int_ZF_2_T1 Int_ZF_2_L6 Int_ZF_2_L1B Int_bounded_have_max_min
    by auto;
  then show "HasAmaximum(IntegerOrder,A)"
    by (rule Order_ZF_4_L11A);
  then show 
    II: "Maximum(IntegerOrder,A) ∈ A" and
    "∀x∈A. x \<lsq> Maximum(IntegerOrder,A)"
    using Int_ZF_2_L4 Order_ZF_4_L3 by auto;
  from I A1 have "A ⊆ \<int>" by (rule Order_ZF_3_L1A)
  with II show "Maximum(IntegerOrder,A) ∈ \<int>" by auto
qed;
 
text{*A set defined by separation over a bounded set attains its maximum
  and minimum.*}

lemma (in int0) Int_ZF_1_4_L1:
  assumes A1: "IsBounded(A,IntegerOrder)" and A2: "A≠0"
  and A3: "∀q∈\<int>. F(q) ∈ \<int>"
  and A4: "K = {F(q). q ∈ A}"
  shows
  "HasAmaximum(IntegerOrder,K)"
  "HasAminimum(IntegerOrder,K)"  
  "Maximum(IntegerOrder,K) ∈ K"
  "Minimum(IntegerOrder,K) ∈ K"
  "Maximum(IntegerOrder,K) ∈ \<int>"
  "Minimum(IntegerOrder,K) ∈ \<int>"
  "∀q∈A. F(q) \<lsq> Maximum(IntegerOrder,K)"
  "∀q∈A. Minimum(IntegerOrder,K) \<lsq> F(q)"
  "IsBounded(K,IntegerOrder)"
proof -
  from A1 have "A ∈ Fin(\<int>)" using Int_bounded_iff_fin
    by simp;
  with A3 have "{F(q). q ∈ A} ∈ Fin(\<int>)"
    by (rule Finite1_L6);
  with A2 A4 have T1: "K ∈ Fin(\<int>)"  "K≠0" by auto;
  then show T2: 
    "HasAmaximum(IntegerOrder,K)"
    "HasAminimum(IntegerOrder,K)"  
    and "Maximum(IntegerOrder,K) ∈ K"
    "Minimum(IntegerOrder,K) ∈ K"
    "Maximum(IntegerOrder,K) ∈ \<int>"
    "Minimum(IntegerOrder,K) ∈ \<int>"
    using Int_fin_have_max_min by auto;
  { fix q assume "q∈A" 
    with A4 have "F(q) ∈ K" by auto;
    with T1 have 
      "F(q) \<lsq> Maximum(IntegerOrder,K)"
      "Minimum(IntegerOrder,K) \<lsq> F(q)"
      using Int_fin_have_max_min by auto;
  } then show 
      "∀q∈A. F(q) \<lsq> Maximum(IntegerOrder,K)"
      "∀q∈A. Minimum(IntegerOrder,K) \<lsq> F(q)"
    by auto;
  from T2 show "IsBounded(K,IntegerOrder)"
    using Order_ZF_4_L7 Order_ZF_4_L8A IsBounded_def
    by simp;
qed;

text{*A three element set has a maximume and minimum.*}

lemma (in int0) Int_ZF_1_4_L1A: assumes A1: "a∈\<int>"  "b∈\<int>"  "c∈\<int>"
  shows 
  "Maximum(IntegerOrder,{a,b,c}) ∈  \<int>"
  "a \<lsq> Maximum(IntegerOrder,{a,b,c})"
  "b \<lsq> Maximum(IntegerOrder,{a,b,c})"
  "c \<lsq> Maximum(IntegerOrder,{a,b,c})"
  using prems Int_ZF_2_T1 Finite_ZF_1_L2A by auto;

text{*Integer functions attain maxima and minima over intervals.*}

lemma (in int0) Int_ZF_1_4_L2:
  assumes A1: "f:\<int>->\<int>" and A2: "a\<lsq>b"
  shows
  "maxf(f,a..b) ∈ \<int>"
  "∀c ∈ a..b. f`(c) \<lsq> maxf(f,a..b)"
  "∃c ∈ a..b. f`(c) = maxf(f,a..b)"
  "minf(f,a..b) ∈ \<int>"
  "∀c ∈ a..b. minf(f,a..b) \<lsq> f`(c)"
  "∃c ∈ a..b. f`(c) = minf(f,a..b)"
proof -
  from A2 have T: "a∈\<int>"  "b∈\<int>"  "a..b ⊆ \<int>"
    using Int_ZF_2_L1A Int_ZF_2_L1B Order_ZF_2_L6 
    by auto;
  with A1 A2 have
    "Maximum(IntegerOrder,f``(a..b)) ∈ f``(a..b)" 
    "∀x∈f``(a..b). x \<lsq> Maximum(IntegerOrder,f``(a..b))"
    "Maximum(IntegerOrder,f``(a..b)) ∈ \<int>"
    "Minimum(IntegerOrder,f``(a..b)) ∈ f``(a..b)" 
    "∀x∈f``(a..b). Minimum(IntegerOrder,f``(a..b)) \<lsq> x"
    "Minimum(IntegerOrder,f``(a..b)) ∈ \<int>"
    using Int_ZF_4_L8 Int_ZF_2_T1 group3.OrderedGroup_ZF_2_L6
      Int_fin_have_max_min by auto;
  with A1 T show
    "maxf(f,a..b) ∈ \<int>"
    "∀c ∈ a..b. f`(c) \<lsq> maxf(f,a..b)"
    "∃c ∈ a..b. f`(c) = maxf(f,a..b)"
    "minf(f,a..b) ∈ \<int>"
    "∀c ∈ a..b. minf(f,a..b) \<lsq> f`(c)"
    "∃c ∈ a..b. f`(c) = minf(f,a..b)"
    using func_imagedef by auto;
qed;

section{*The set of nonnegative integers*}

text{*The set of nonnegative integers looks like the set of natural numbers.
  We explore that in this section. We also rephrasse some lemmas about the set
  of positive integers known from the theory of oredered grups.*}

text{*The set of positive integers is closed under addition.*}

lemma (in int0) pos_int_closed_add: 
  shows "\<int>+ {is closed under} IntegerAddition"
  using Int_ZF_2_T1 group3.OrderedGroup_ZF_1_L13 by simp;

text{*Text expended version of the fact that the set of positive integers 
  is closed under addition*}

lemma (in int0) pos_int_closed_add_unfolded: 
  assumes "a∈\<int>+"  "b∈\<int>+"  shows "a\<ra>b ∈ \<int>+"
  using prems pos_int_closed_add IsOpClosed_def
  by simp;

text{*@{text "\<int>+"} is bounded below.*}

lemma (in int0) Int_ZF_1_5_L1: shows 
  "IsBoundedBelow(\<int>+,IntegerOrder)"
  "IsBoundedBelow(\<int>+,IntegerOrder)"
  using Nonnegative_def PositiveSet_def IsBoundedBelow_def by auto;

text{*Subsets of @{text "\<int>+"} are bounded below.*}

lemma (in int0) Int_ZF_1_5_L1A: assumes A1: "A ⊆ \<int>+" 
  shows "IsBoundedBelow(A,IntegerOrder)"
  using A1 Int_ZF_1_5_L1 Order_ZF_3_L12 by blast;

text{*Subsets of @{text "\<int>+"} are bounded below.*}

lemma (in int0) Int_ZF_1_5_L1B: assumes A1: "A ⊆ \<int>+" 
  shows "IsBoundedBelow(A,IntegerOrder)"
  using A1 Int_ZF_1_5_L1 Order_ZF_3_L12 by blast; (*A1 label has to be here*)

text{*Every nonempty subset of positive integers has a mimimum. *}

lemma (in int0) Int_ZF_1_5_L1C: assumes "A ⊆ \<int>+" and "A ≠ 0"
  shows 
  "HasAminimum(IntegerOrder,A)"
  "Minimum(IntegerOrder,A) ∈ A"
  "∀x∈A. Minimum(IntegerOrder,A) \<lsq> x"
  using prems Int_ZF_1_5_L1B int_bounded_below_has_min by auto;

text{*Infinite subsets of $Z^+$ do not have a maximum - If $A\subseteq Z^+$
  then for every integer we can find one in the set that is not smaller.*}

lemma (in int0) Int_ZF_1_5_L2:
  assumes A1: "A ⊆ \<int>+"  and A2: "A ∉ Fin(\<int>)" and A3: "D∈\<int>"
  shows "∃n∈A. D\<lsq>n"
proof -
  { assume "∀n∈A. ¬(D\<lsq>n)" 
    moreover from A1 A3 have "D∈\<int>"  "∀n∈A. n∈\<int>" 
      using Nonnegative_def by auto;
    ultimately have "∀n∈A. n\<lsq>D"
      using Int_ZF_2_L19 by blast;
    hence "∀n∈A. ⟨n,D⟩ ∈ IntegerOrder" by simp;
    then have "IsBoundedAbove(A,IntegerOrder)"
      by (rule Order_ZF_3_L10);
    with A1 A2 have False using Int_ZF_1_5_L1A IsBounded_def 
      Int_bounded_iff_fin by auto;
  } thus ?thesis by auto;
qed;

text{*Infinite subsets of $Z_+$ do not have a maximum - If $A\subseteq Z_+$
  then for every integer we can find one in the set that is not smaller.  
  This is very similar to @{text "Int_ZF_1_5_L2"}, except we have @{text "\<int>+"}
  instead of @{text "\<int>+"} here.*}

lemma (in int0) Int_ZF_1_5_L2A:
  assumes A1: "A ⊆ \<int>+"  and A2: "A ∉ Fin(\<int>)" and A3: "D∈\<int>"
  shows "∃n∈A. D\<lsq>n"
proof -
{ assume "∀n∈A. ¬(D\<lsq>n)" 
    moreover from A1 A3 have "D∈\<int>"  "∀n∈A. n∈\<int>" 
      using PositiveSet_def by auto;
    ultimately have "∀n∈A. n\<lsq>D"
      using Int_ZF_2_L19 by blast;
    hence "∀n∈A. ⟨n,D⟩ ∈ IntegerOrder" by simp;
    then have "IsBoundedAbove(A,IntegerOrder)"
      by (rule Order_ZF_3_L10);
    with A1 A2 have False using Int_ZF_1_5_L1B IsBounded_def 
      Int_bounded_iff_fin by auto;
  } thus ?thesis by auto;
qed;

text{*An integer is either positive, zero, or its opposite is postitive.*}

lemma (in int0) Int_decomp: assumes "m∈\<int>"
  shows "Exactly_1_of_3_holds (m=\<zero>,m∈\<int>+,(\<rm>m)∈\<int>+)"
  using prems Int_ZF_2_T1 group3.OrdGroup_decomp 
  by simp;

text{*An integer is zero, positive, or it's inverse is positive.*}

lemma (in int0) int_decomp_cases: assumes "m∈\<int>"
  shows "m=\<zero> ∨ m∈\<int>+ ∨ (\<rm>m) ∈ \<int>+"
  using prems Int_ZF_2_T1 group3.OrderedGroup_ZF_1_L14
  by simp;

text{*An integer is in the positive set iff it is greater or equal one.*}

lemma (in int0) Int_ZF_1_5_L3: shows "m∈\<int>+ <-> \<one>\<lsq>m"
proof
  assume "m∈\<int>+" then have "\<zero>\<lsq>m"  "m≠\<zero>"
    using PositiveSet_def by auto;
  then have "\<zero>\<ra>\<one> \<lsq> m" 
    using Int_ZF_4_L1B by auto;
  then show "\<one>\<lsq>m" 
    using int_zero_one_are_int Int_ZF_1_T2 group0.group0_2_L2
    by simp;
next assume "\<one>\<lsq>m"
  then have "m∈\<int>"  "\<zero>\<lsq>m"  "m≠\<zero>"
    using Int_ZF_2_L1A Int_ZF_2_L16C by auto;
  then show "m∈\<int>+" using PositiveSet_def by auto;
qed;

text{*The set of positive integers is closed under multiplication.
  The unfolded form.*}

lemma (in int0) pos_int_closed_mul_unfold: 
  assumes "a∈\<int>+"  "b∈\<int>+"
  shows "a·b ∈ \<int>+"
  using prems Int_ZF_1_5_L3 Int_ZF_1_3_L3 by simp;

text{*The set of positive integers is closed under multiplication.*}

lemma (in int0) pos_int_closed_mul: shows
  "\<int>+ {is closed under} IntegerMultiplication"
  using pos_int_closed_mul_unfold IsOpClosed_def
  by simp;

text{*It is an overkill to prove that the ring of integers has no zero
  divisors this way, but why not?*}

lemma (in int0) int_has_no_zero_divs: 
  shows "HasNoZeroDivs(\<int>,IntegerAddition,IntegerMultiplication)"
  using pos_int_closed_mul Int_ZF_1_3_T1 ring1.OrdRing_ZF_3_L3
  by simp;

text{*Nonnegative integers are positive ones plus zero.*}

lemma (in int0) Int_ZF_1_5_L3A: shows "\<int>+ = \<int>+ ∪ {\<zero>}"
  using Int_ZF_2_T1 group3.OrderedGroup_ZF_1_L24 by simp;

text{*We can make a function smaller than any constant on a given interval
  of positive integers by adding another constant.*}


lemma (in int0) Int_ZF_1_5_L4: 
  assumes A1: "f:\<int>->\<int>" and A2: "K∈\<int>" "N∈\<int>"
  shows "∃C∈\<int>. ∀n∈\<int>+. K \<lsq> f`(n) \<ra> C --> N\<lsq>n"
proof -
  from A2 have "N\<lsq>\<one> ∨ \<two>\<lsq>N"
    using int_zero_one_are_int no_int_between
    by simp;
  moreover
  { assume A3: "N\<lsq>\<one>"
    let ?C = "\<zero>"
    have "?C ∈ \<int>" using int_zero_one_are_int
      by simp
    moreover
    { fix n assume "n∈\<int>+"
      then have "\<one> \<lsq> n" using Int_ZF_1_5_L3
        by simp 
      with A3 have "N\<lsq>n" by (rule Int_order_transitive)
    } then have  "∀n∈\<int>+. K \<lsq> f`(n) \<ra> ?C --> N\<lsq>n"
      by auto
    ultimately have "∃C∈\<int>. ∀n∈\<int>+. K \<lsq> f`(n) \<ra> C --> N\<lsq>n"
      by auto; }
  moreover
  { let ?C = "K \<rs> \<one> \<rs> maxf(f,\<one>..(N\<rs>\<one>))"
    assume "\<two>\<lsq>N"
    then have "\<two>\<rs>\<one> \<lsq> N\<rs>\<one>"
      using int_zero_one_are_int Int_ZF_1_1_L4 int_ord_transl_inv
      by simp;
    then have I: "\<one> \<lsq> N\<rs>\<one>"
      using int_zero_one_are_int Int_ZF_1_2_L3 by simp;
    with A1 A2 have T: 
      "maxf(f,\<one>..(N\<rs>\<one>)) ∈ \<int>"  "K\<rs>\<one> ∈ \<int>"  "?C ∈ \<int>"
      using Int_ZF_1_4_L2 Int_ZF_1_1_L5 int_zero_one_are_int
      by auto;
    moreover
    { fix n assume A4: "n∈\<int>+" 
      { assume A5: "K \<lsq> f`(n) \<ra> ?C" and "¬(N\<lsq>n)"
        with A2 A4 have "n \<lsq> N\<rs>\<one>"  
          using PositiveSet_def Int_ZF_1_3_L6A by simp;
        with A4 have "n ∈ \<one>..(N\<rs>\<one>)"
          using Int_ZF_1_5_L3 Interval_def by auto;
        with A1 I T have "f`(n)\<ra>?C \<lsq> maxf(f,\<one>..(N\<rs>\<one>)) \<ra> ?C"
          using Int_ZF_1_4_L2 int_ord_transl_inv by simp;
        with T have "f`(n)\<ra>?C \<lsq> K\<rs>\<one>"
          using Int_ZF_1_2_L3 by simp;
        with A5 have "K \<lsq>  K\<rs>\<one>"
          by (rule Int_order_transitive)
        with A2 have False using Int_ZF_1_2_L3AA by simp;
      } then have "K \<lsq> f`(n) \<ra> ?C --> N\<lsq>n"
        by auto 
    } then have "∀n∈\<int>+. K \<lsq> f`(n) \<ra> ?C --> N\<lsq>n"
      by simp;
    ultimately have "∃C∈\<int>. ∀n∈\<int>+. K \<lsq> f`(n) \<ra> C --> N\<lsq>n" 
      by auto } 
  ultimately show ?thesis by auto;
qed;

text{*Absolute value is identity on positive integers.*}

lemma (in int0) Int_ZF_1_5_L4A: 
  assumes "a∈\<int>+" shows "abs(a) = a"
  using prems Int_ZF_2_T1 group3.OrderedGroup_ZF_3_L2B
  by simp;

text{*One and two are in @{text "\<int>+"}.*}

lemma (in int0) int_one_two_are_pos: shows "\<one> ∈ \<int>+"  "\<two> ∈ \<int>+"
  using int_zero_one_are_int int_ord_is_refl refl_def Int_ZF_1_5_L3
  Int_ZF_2_L16B by auto;

text{*The image of @{text "\<int>+"} by a function defined on integers 
  is not empty.*}

lemma (in int0) Int_ZF_1_5_L5: assumes A1: "f : \<int>->X"
  shows "f``(\<int>+) ≠ 0"
proof -
  have "\<int>+ ⊆ \<int>" using PositiveSet_def by auto;
  with A1 show "f``(\<int>+) ≠ 0"
    using int_one_two_are_pos func_imagedef by auto;
qed;

text{*If $n$ is positive, then $n-1$ is nonnegative.*}

lemma (in int0) Int_ZF_1_5_L6: assumes A1: "n ∈ \<int>+"
  shows 
  "\<zero> \<lsq> n\<rs>\<one>"
  "\<zero> ∈ \<zero>..(n\<rs>\<one>)"
  "\<zero>..(n\<rs>\<one>) ⊆ \<int>"
proof -
  from A1 have "\<one> \<lsq> n"  "(\<rm>\<one>) ∈ \<int>"
    using Int_ZF_1_5_L3 int_zero_one_are_int Int_ZF_1_1_L4 
    by auto;
  then have "\<one>\<rs>\<one> \<lsq> n\<rs>\<one>"
    using int_ord_transl_inv by simp;
  then show "\<zero> \<lsq> n\<rs>\<one>"
    using int_zero_one_are_int Int_ZF_1_1_L4 by simp;
  then show "\<zero> ∈ \<zero>..(n\<rs>\<one>)"
    using int_zero_one_are_int int_ord_is_refl refl_def Order_ZF_2_L1B
    by simp;
  show "\<zero>..(n\<rs>\<one>) ⊆ \<int>"
    using Int_ZF_2_L1B Order_ZF_2_L6 by simp;
qed;

text{*Intgers greater than one in @{text "\<int>+"} belong to @{text "\<int>+"}.
  This is a property of ordered groups and follows from 
  @{text "OrderedGroup_ZF_1_L19"}, but Isabelle's simplifier has problems 
  using that result directly, so we reprove it specifically for integers.*}

lemma (in int0) Int_ZF_1_5_L7:  assumes "a ∈ \<int>+" and "a\<lsq>b"
  shows "b ∈ \<int>+"
proof-
  from prems have "\<one>\<lsq>a"  "a\<lsq>b"
    using Int_ZF_1_5_L3 by auto;
  then have "\<one>\<lsq>b" by (rule Int_order_transitive);
  then show "b ∈ \<int>+" using Int_ZF_1_5_L3 by simp;
qed;

text{*Adding a positive integer increases integers.*}

lemma (in int0) Int_ZF_1_5_L7A: assumes "a∈\<int>"  "b ∈ \<int>+"
  shows "a \<lsq> a\<ra>b"  "a ≠ a\<ra>b"  "a\<ra>b ∈ \<int>"
  using prems Int_ZF_2_T1 group3.OrderedGroup_ZF_1_L22
  by auto;

text{*For any integer $m$ the greater of $m$ and $1$ is a positive
  integer that is greater or equal than $m$. If we add $1$ to it we
  get a positive integer that is strictly greater than $m$.*}

lemma (in int0) Int_ZF_1_5_L7B: assumes "a∈\<int>"
  shows 
  "a \<lsq> GreaterOf(IntegerOrder,\<one>,a)"
  "GreaterOf(IntegerOrder,\<one>,a) ∈ \<int>+"
  "GreaterOf(IntegerOrder,\<one>,a) \<ra> \<one> ∈ \<int>+"
  "a \<lsq> GreaterOf(IntegerOrder,\<one>,a) \<ra> \<one>"  
  "a ≠ GreaterOf(IntegerOrder,\<one>,a) \<ra> \<one>"
  using prems int_zero_not_one Int_ZF_1_3_T1 ring1.OrdRing_ZF_3_L12
  by auto;


(*text{*Adding one increases integers - one more version useful
  for some proofs by contradiction.*}

lemma (in int0) Int_ZF_1_5_L7B: assumes A1: "a∈\<int>"
  shows 
  "¬(a\<ra>\<one> \<lsq>a)"
  "¬(\<one>\<ra>a \<lsq>a)"
proof -
  { assume A2: "a\<ra>\<one> \<lsq>a"
    moreover from A1 have "a \<lsq> a\<ra>\<one>"
      using Int_ZF_2_L12B by simp
    ultimately have "a\<ra>\<one> = a" by (rule Int_ZF_2_L3) done somewhere else*)  
text{*The opposite of an element of @{text "\<int>+"} cannot belong to
  @{text "\<int>+"}.*}

lemma (in int0) Int_ZF_1_5_L8: assumes "a ∈ \<int>+"
  shows "(\<rm>a) ∉ \<int>+"
  using prems Int_ZF_2_T1 group3.OrderedGroup_ZF_1_L20
  by simp;

text{*For every integer there is one in @{text "\<int>+"} that is greater or 
  equal.*}

lemma (in int0) Int_ZF_1_5_L9: assumes "a∈\<int>"
  shows "∃b∈\<int>+. a\<lsq>b"
  using prems int_not_trivial Int_ZF_2_T1 group3.OrderedGroup_ZF_1_L23
  by simp;

text{*A theorem about odd extensions. Recall from @{text "OrdereGroup_ZF.thy"}
  that the odd extension of an integer function $f$ defined on @{text "\<int>+"} 
  is the odd function on @{text "\<int>"} equal to $f$ on @{text "\<int>+"}. 
  First we show that the odd extension is defined on @{text "\<int>"}.*}

lemma (in int0) Int_ZF_1_5_L10: assumes "f : \<int>+->\<int>"
  shows "OddExtension(\<int>,IntegerAddition,IntegerOrder,f) : \<int>->\<int>"
  using prems Int_ZF_2_T1 group3.odd_ext_props by simp;

text{*On @{text "\<int>+"}, the odd extension of $f$ is the same as $f$.*}

lemma (in int0) Int_ZF_1_5_L11: assumes "f : \<int>+->\<int>" and "a ∈ \<int>+" and
  "g = OddExtension(\<int>,IntegerAddition,IntegerOrder,f)"
  shows "g`(a) = f`(a)"
  using prems Int_ZF_2_T1 group3.odd_ext_props by simp;

text{*On @{text "\<sm>\<int>+"}, the value of the odd extension of $f$ 
  is the negative of $f(-a)$. *}

lemma (in int0) Int_ZF_1_5_L12:  
  assumes "f : \<int>+->\<int>" and "a ∈ (\<sm>\<int>+)" and
  "g = OddExtension(\<int>,IntegerAddition,IntegerOrder,f)"
  shows "g`(a) = \<rm>(f`(\<rm>a))"
  using prems Int_ZF_2_T1 group3.odd_ext_props by simp;

text{*Odd extensions are odd on @{text "\<int>"}.*}

lemma (in int0) int_oddext_is_odd:
  assumes "f : \<int>+->\<int>" and "a∈\<int>" and
  "g = OddExtension(\<int>,IntegerAddition,IntegerOrder,f)"
  shows "g`(\<rm>a) = \<rm>(g`(a))"
  using prems Int_ZF_2_T1 group3.oddext_is_odd by simp;


text{*Alternative definition of an odd function.*}

lemma (in int0) Int_ZF_1_5_L13:  assumes A1: "f: \<int>->\<int>" shows 
  "(∀a∈\<int>. f`(\<rm>a) = (\<rm>f`(a))) <-> (∀a∈\<int>. (\<rm>(f`(\<rm>a))) = f`(a))"
  using prems Int_ZF_1_T2 group0.group0_6_L2 by simp;

text{*Another way of expressing the fact that odd extensions are odd.*}

lemma (in int0) int_oddext_is_odd_alt:
  assumes "f : \<int>+->\<int>" and "a∈\<int>" and
  "g = OddExtension(\<int>,IntegerAddition,IntegerOrder,f)"
  shows "(\<rm>g`(\<rm>a)) = g`(a)"
  using prems Int_ZF_2_T1 group3.oddext_is_odd_alt by simp;

section{*Functions with infinite limits*}

text{*In this section we consider functions (integer sequences) that have
  infinite limits. An integer function has infinite positive limit if it 
  is arbitrarily
  large for large enough arguments. Similarly, a function has infinite negative limit 
  if it is arbitrarily small for small enough arguments. 
  The material in this 
  come mostly from the section in @{text "OrderedGroup_ZF.thy"} 
  with he same title.
  Here we rewrite the theorems from that section in the notation 
  we use for integers and
  add some results specific for the ordered group of integers. *}

text{*If an image of a set by a function with infinite positive limit 
  is bounded above, then the set itself is bounded above.*}

lemma (in int0) Int_ZF_1_6_L1: assumes "f: \<int>->\<int>" and
  "∀a∈\<int>.∃b∈\<int>+.∀x. b\<lsq>x --> a \<lsq> f`(x)" and "A ⊆ \<int>" and
  "IsBoundedAbove(f``(A),IntegerOrder)"
  shows "IsBoundedAbove(A,IntegerOrder)"
  using prems int_not_trivial Int_ZF_2_T1 group3.OrderedGroup_ZF_7_L1
  by simp;

text{*If an image of a set defined by separation 
  by a function with infinite positive limit 
  is bounded above, then the set itself is bounded above.*}

lemma (in int0) Int_ZF_1_6_L2: assumes A1: "X≠0" and A2: "f: \<int>->\<int>" and 
  A3: "∀a∈\<int>.∃b∈\<int>+.∀x. b\<lsq>x --> a \<lsq> f`(x)" and
  A4: "∀x∈X. b(x) ∈ \<int>  ∧ f`(b(x)) \<lsq> U"
  shows "∃u.∀x∈X. b(x) \<lsq> u"
proof -
  let ?G = "\<int>"
  let ?P = "IntegerAddition"
  let ?r = "IntegerOrder"
  from A1 A2 A3 A4 have 
    "group3(?G, ?P, ?r)" 
    "?r {is total on} ?G" 
    "?G ≠ {TheNeutralElement(?G, ?P)}"
    "X≠0"  "f: ?G->?G"
    "∀a∈?G. ∃b∈PositiveSet(?G, ?P, ?r). ∀y. ⟨b, y⟩ ∈ ?r --> ⟨a, f`(y)⟩ ∈ ?r"
    "∀x∈X. b(x) ∈ ?G ∧ ⟨f`(b(x)), U⟩ ∈ ?r"
    using int_not_trivial Int_ZF_2_T1 by auto;
  then have "∃u. ∀x∈X. ⟨b(x), u⟩ ∈ ?r" by (rule group3.OrderedGroup_ZF_7_L2);
  thus ?thesis by simp;
qed;

text{*If an image of a set defined by separation 
  by a integer function with infinite negative limit 
  is bounded below, then the set itself is bounded above.
  This is dual to @{text " Int_ZF_1_6_L2"}.*}

lemma (in int0) Int_ZF_1_6_L3: assumes A1: "X≠0" and A2: "f: \<int>->\<int>" and 
  A3: "∀a∈\<int>.∃b∈\<int>+.∀y. b\<lsq>y --> f`(\<rm>y) \<lsq> a" and
  A4: "∀x∈X. b(x) ∈ \<int>  ∧ L \<lsq> f`(b(x))"
  shows "∃l.∀x∈X. l \<lsq> b(x)"
proof -
  let ?G = "\<int>"
  let ?P = "IntegerAddition"
  let ?r = "IntegerOrder"
  from A1 A2 A3 A4 have 
    "group3(?G, ?P, ?r)" 
    "?r {is total on} ?G" 
    "?G ≠ {TheNeutralElement(?G, ?P)}"
    "X≠0"  "f: ?G->?G"
    "∀a∈?G. ∃b∈PositiveSet(?G, ?P, ?r). ∀y. 
    ⟨b, y⟩ ∈ ?r --> ⟨f`(GroupInv(?G, ?P)`(y)),a⟩ ∈ ?r"
    "∀x∈X. b(x) ∈ ?G ∧ ⟨L,f`(b(x))⟩ ∈ ?r"
    using int_not_trivial Int_ZF_2_T1 by auto;
  then have "∃l. ∀x∈X. ⟨l, b(x)⟩ ∈ ?r" by (rule group3.OrderedGroup_ZF_7_L3);
  thus ?thesis by simp;
qed;

text{*The next lemma combines @{text "Int_ZF_1_6_L2"} and 
  @{text "Int_ZF_1_6_L3"} to show that if the image of a set 
  defined by separation by a function with infinite limits is bounded,
  then the set itself is bounded. The proof again uses directly a fact from
  @{text "OrderedGroup_ZF.thy "}.*}

lemma (in int0) Int_ZF_1_6_L4:
  assumes A1: "X≠0" and A2: "f: \<int>->\<int>" and 
  A3: "∀a∈\<int>.∃b∈\<int>+.∀x. b\<lsq>x --> a \<lsq> f`(x)" and
  A4: "∀a∈\<int>.∃b∈\<int>+.∀y. b\<lsq>y --> f`(\<rm>y) \<lsq> a" and
  A5: "∀x∈X. b(x) ∈ \<int> ∧ f`(b(x)) \<lsq> U ∧ L \<lsq> f`(b(x))"
  shows "∃M.∀x∈X. abs(b(x)) \<lsq> M"
proof -
  let ?G = "\<int>"
  let ?P = "IntegerAddition"
  let ?r = "IntegerOrder"
  from A1 A2 A3 A4 A5 have
    "group3(?G, ?P, ?r)" 
    "?r {is total on} ?G" 
    "?G ≠ {TheNeutralElement(?G, ?P)}"
    "X≠0"  "f: ?G->?G"
    "∀a∈?G. ∃b∈PositiveSet(?G, ?P, ?r). ∀y. ⟨b, y⟩ ∈ ?r --> ⟨a, f`(y)⟩ ∈ ?r"
    "∀a∈?G. ∃b∈PositiveSet(?G, ?P, ?r). ∀y. 
    ⟨b, y⟩ ∈ ?r --> ⟨f`(GroupInv(?G, ?P)`(y)),a⟩ ∈ ?r"
    "∀x∈X. b(x) ∈ ?G ∧ ⟨L,f`(b(x))⟩ ∈ ?r ∧ ⟨f`(b(x)), U⟩ ∈ ?r"
    using int_not_trivial Int_ZF_2_T1 by auto;
  then have "∃M. ∀x∈X. ⟨AbsoluteValue(?G, ?P, ?r) ` b(x), M⟩ ∈ ?r"
    by (rule group3.OrderedGroup_ZF_7_L4);
  thus ?thesis by simp;
qed;

text{*If a function is larger than some constant for arguments large
  enough, then the image of a set that is bounded below is bounded below.
  This is not true for ordered groups in general, but only for those for which
  bounded sets are finite.
  This does not require the function to have infinite limit, but such 
  functions do have this property. *}

lemma (in int0) Int_ZF_1_6_L5:
  assumes A1: "f: \<int>->\<int>" and A2: "N∈\<int>" and
  A3: "∀m. N\<lsq>m --> L \<lsq> f`(m)" and 
  A4: "IsBoundedBelow(A,IntegerOrder)"
  shows "IsBoundedBelow(f``(A),IntegerOrder)"
proof -
  from A2 A4 have "A = {x∈A. x\<lsq>N} ∪ {x∈A. N\<lsq>x}"
    using Int_ZF_2_T1 Int_ZF_2_L1C Order_ZF_1_L5 
    by simp;
  moreover have 
    "f``({x∈A. x\<lsq>N} ∪ {x∈A. N\<lsq>x}) =
    f``{x∈A. x\<lsq>N} ∪ f``{x∈A. N\<lsq>x}"
    by (rule image_Un);
  ultimately have "f``(A) = f``{x∈A. x\<lsq>N} ∪ f``{x∈A. N\<lsq>x}"
    by simp;
  moreover have "IsBoundedBelow(f``{x∈A. x\<lsq>N},IntegerOrder)"
  proof -
    let ?B = "{x∈A. x\<lsq>N}"
    from A4 have "?B ∈ Fin(\<int>)"
      using Order_ZF_3_L16 Int_bounded_iff_fin by auto;
    with A1 have  "IsBounded(f``(?B),IntegerOrder)"
      using Finite1_L6A Int_bounded_iff_fin by simp;
    then show "IsBoundedBelow(f``(?B),IntegerOrder)"
      using IsBounded_def by simp;
  qed;
  moreover have "IsBoundedBelow(f``{x∈A. N\<lsq>x},IntegerOrder)"
  proof -
    let ?C = "{x∈A. N\<lsq>x}"
    from A4 have "?C ⊆ \<int>" using Int_ZF_2_L1C by auto;
    with A1 A3 have "∀y ∈ f``(?C). ⟨L,y⟩ ∈ IntegerOrder"
      using func_imagedef by simp;
    then show "IsBoundedBelow(f``(?C),IntegerOrder)"
      by (rule Order_ZF_3_L9);
  qed;
  ultimately show "IsBoundedBelow(f``(A),IntegerOrder)"
    using Int_ZF_2_T1 Int_ZF_2_L6 Int_ZF_2_L1B Order_ZF_3_L6
    by simp;
qed;

text{*A function that has an infinite limit can be made arbitrarily large
  on positive integers by adding a constant. This does not actually require
  the function to have infinite limit, just to be larger than a constant
  for arguments large enough.*}

lemma (in int0) Int_ZF_1_6_L6: assumes A1: "N∈\<int>" and
  A2: "∀m. N\<lsq>m --> L \<lsq> f`(m)" and
  A3: "f: \<int>->\<int>" and A4: "K∈\<int>"
  shows "∃c∈\<int>. ∀n∈\<int>+. K \<lsq> f`(n)\<ra>c"
proof -
  have "IsBoundedBelow(\<int>+,IntegerOrder)"
    using Int_ZF_1_5_L1 by simp;
  with A3 A1 A2 have "IsBoundedBelow(f``(\<int>+),IntegerOrder)"
    by (rule Int_ZF_1_6_L5);
  with A1 obtain l where I: "∀y∈f``(\<int>+). l \<lsq> y"
    using Int_ZF_1_5_L5 IsBoundedBelow_def by auto;
  let ?c = "K\<rs>l"
  from A3 have "f``(\<int>+) ≠ 0" using Int_ZF_1_5_L5
    by simp;
  then have "∃y. y ∈ f``(\<int>+)" by (rule nonempty_has_element);
  then obtain y where "y ∈ f``(\<int>+)" by auto;
  with A4 I have T: "l ∈ \<int>"  "?c ∈ \<int>"
    using Int_ZF_2_L1A Int_ZF_1_1_L5 by auto;
  { fix n assume A5: "n∈\<int>+"
    have "\<int>+ ⊆ \<int>" using PositiveSet_def by auto;
    with A3 I T A5 have "l \<ra> ?c \<lsq> f`(n) \<ra> ?c"
      using func_imagedef int_ord_transl_inv by auto;
     with I T have "l \<ra> ?c \<lsq> f`(n) \<ra> ?c"
      using int_ord_transl_inv by simp;
    with A4 T have "K \<lsq>  f`(n) \<ra> ?c"
      using Int_ZF_1_2_L3 by simp;
  } then have "∀n∈\<int>+. K \<lsq>  f`(n) \<ra> ?c" by simp;
  with T show ?thesis by auto;
qed;

text{*If a function has infinite limit, then we can add such constant
  such that minimum of those arguments for which the function (plus the constant) 
  is larger than another given constant is greater than a third constant.
  It is not as complicated as it sounds.*}

lemma (in int0) Int_ZF_1_6_L7: 
  assumes A1: "f: \<int>->\<int>" and A2: "K∈\<int>"  "N∈\<int>" and
  A3: "∀a∈\<int>.∃b∈\<int>+.∀x. b\<lsq>x --> a \<lsq> f`(x)"
  shows "∃C∈\<int>. N \<lsq> Minimum(IntegerOrder,{n∈\<int>+. K \<lsq> f`(n)\<ra>C})"
proof -
  from A1 A2 have "∃C∈\<int>. ∀n∈\<int>+. K \<lsq> f`(n) \<ra> C --> N\<lsq>n"
    using Int_ZF_1_5_L4 by simp
  then obtain C where I: "C∈\<int>" and
    II: "∀n∈\<int>+. K \<lsq> f`(n) \<ra> C --> N\<lsq>n"
    by auto
  have "antisym(IntegerOrder)" using Int_ZF_2_L4 by simp;
  moreover have "HasAminimum(IntegerOrder,{n∈\<int>+. K \<lsq> f`(n)\<ra>C})"
  proof -
    from A2 A3 I have "∃n∈\<int>+.∀x. n\<lsq>x --> K\<rs>C \<lsq> f`(x)"
      using Int_ZF_1_1_L5 by simp;
    then obtain n where 
      "n∈\<int>+" and "∀x. n\<lsq>x --> K\<rs>C \<lsq>  f`(x)"
      by auto;
    with A2 I have 
      "{n∈\<int>+. K \<lsq> f`(n)\<ra>C} ≠ 0"
      "{n∈\<int>+. K \<lsq> f`(n)\<ra>C} ⊆ \<int>+"
      using int_ord_is_refl refl_def PositiveSet_def Int_ZF_2_L9C
      by auto;
    then show "HasAminimum(IntegerOrder,{n∈\<int>+. K \<lsq> f`(n)\<ra>C})"
      using Int_ZF_1_5_L1C by simp;
  qed
  moreover from II have 
    "∀n ∈ {n∈\<int>+. K \<lsq> f`(n)\<ra>C}. ⟨N,n⟩ ∈ IntegerOrder" 
    by auto; (* we can not put ?A here *)
  ultimately have 
    "⟨N,Minimum(IntegerOrder,{n∈\<int>+. K \<lsq> f`(n)\<ra>C})⟩ ∈ IntegerOrder"
    by (rule Order_ZF_4_L12);
  with I show ?thesis by auto;
qed;

text{*For any integer $m$ the function $k\mapsto m\cdot k$ has an infinite limit
  (or negative of that). This is why we put some properties of these functions 
  here, even though they properly belong to a (yet nonexistent) section on 
  homomorphisms. The next lemma shows that the set $\{a\cdot x: x\in Z\}$
  can finite only if $a=0$.*}

lemma (in int0) Int_ZF_1_6_L8: 
  assumes A1: "a∈\<int>" and A2: "{a·x. x∈\<int>} ∈ Fin(\<int>)"
  shows "a = \<zero>"
proof -
  from A1 have "a=\<zero> ∨ (a \<lsq> \<rm>\<one>) ∨ (\<one>\<lsq>a)"
    using Int_ZF_1_3_L6C by simp;
  moreover
  { assume "a \<lsq> \<rm>\<one>"
    then have "{a·x. x∈\<int>} ∉ Fin(\<int>)"
      using int_zero_not_one Int_ZF_1_3_T1 ring1.OrdRing_ZF_3_L6
      by simp;
    with A2 have False by simp }
  moreover
  { assume "\<one>\<lsq>a" 
    then have "{a·x. x∈\<int>} ∉ Fin(\<int>)"
      using int_zero_not_one Int_ZF_1_3_T1 ring1.OrdRing_ZF_3_L5
    by simp; 
  with A2 have False by simp }
  ultimately show  "a = \<zero>" by auto;
qed;
    
section{*Miscelaneous*}

text{*In this section we put some technical lemmas needed in various other 
  places that are hard to classify.*}

text{*Suppose we have an integer expression (a meta-function)$F$ such that 
  $F(p)|p|$ is bounded by a linear function of $|p|$, that is for some integers
  $A,B$ we have $F(p)|p|\leq A|p|+B.$ We show that $F$ is then bounded. 
  The proof is easy, we just divide both sides by $|p|$ 
  and take the limit (just kidding). *}

lemma (in int0) Int_ZF_1_7_L1:
  assumes A1: "∀q∈\<int>. F(q) ∈ \<int>" and 
  A2:  "∀q∈\<int>. F(q)·abs(q) \<lsq> A·abs(q) \<ra> B" and 
  A3: "A∈\<int>"  "B∈\<int>"
  shows "∃L. ∀p∈\<int>. F(p) \<lsq> L"
proof -
  let ?I = "(\<rm>abs(B))..abs(B)"
  def DefK: K == "{F(q). q ∈ ?I}"
  let ?M = "Maximum(IntegerOrder,K)"
  let ?L = "GreaterOf(IntegerOrder,?M,A\<ra>\<one>)"
  from A3 A1 DefK have C1:
    "IsBounded(?I,IntegerOrder)"   
    "?I ≠ 0"
    "∀q∈\<int>. F(q) ∈ \<int>"
    "K = {F(q). q ∈ ?I}"
    using Order_ZF_3_L11 Int_ZF_1_3_L17 by auto;
  then have "?M ∈ \<int>" by (rule Int_ZF_1_4_L1);
  with A3 have T1: "?M \<lsq> ?L"  "A\<ra>\<one> \<lsq> ?L"
    using int_zero_one_are_int Int_ZF_1_1_L5 Int_ZF_1_3_L18
    by auto;
  from C1 have T2: "∀q∈?I. F(q) \<lsq> ?M"
    by (rule Int_ZF_1_4_L1);
  { fix p assume A4: "p∈\<int>" have "F(p) \<lsq> ?L" 
    proof (cases "abs(p) \<lsq> abs(B)")
      assume "abs(p) \<lsq> abs(B)"
      with A4 T1 T2 have "F(p) \<lsq> ?M"  "?M \<lsq> ?L"
        using Int_ZF_1_3_L19 by auto;
      then show "F(p) \<lsq> ?L" by (rule Int_order_transitive);
    next assume A5: "¬(abs(p) \<lsq> abs(B))"
      from A3 A2 A4 have 
        "A·abs(p) ∈ \<int>"  "F(p)·abs(p) \<lsq> A·abs(p) \<ra> B"
        using Int_ZF_2_L14 Int_ZF_1_1_L5 by auto;
      moreover from A3 A4 A5 have "B \<lsq> abs(p)"
        using Int_ZF_1_3_L15 by simp;
      ultimately have
        "F(p)·abs(p) \<lsq> A·abs(p) \<ra> abs(p)"
        using Int_ZF_2_L15A by blast;
      with A3 A4 have "F(p)·abs(p) \<lsq> (A\<ra>\<one>)·abs(p)"
        using Int_ZF_2_L14 Int_ZF_1_2_L7 by simp;
      moreover from A3 A1 A4 A5 have 
        "F(p) ∈ \<int>"  "A\<ra>\<one> ∈ \<int>"  "abs(p) ∈ \<int>"
         "¬(abs(p) \<lsq> \<zero>)"
        using int_zero_one_are_int Int_ZF_1_1_L5 Int_ZF_2_L14 Int_ZF_1_3_L11
        by auto;
      ultimately have "F(p) \<lsq> A\<ra>\<one>"
        using Int_ineq_simpl_positive by simp;
      moreover from T1 have  "A\<ra>\<one> \<lsq> ?L" by simp;
      ultimately show "F(p) \<lsq> ?L" by (rule Int_order_transitive);
    qed
 } then have "∀p∈\<int>. F(p) \<lsq> ?L" by simp
 thus ?thesis by auto;
qed;

text{*A lemma about splitting (not really, there is some overlap) 
  the @{text "\<int>×\<int>"} into six subsets (cases). The subsets are as follows:
  first and third qaudrant, and second and fourth quadrant farther split
  by the $b =-a$ line. *}

lemma (in int0) int_plane_split_in6: assumes "a∈\<int>"  "b∈\<int>"
  shows
  "\<zero>\<lsq>a ∧ \<zero>\<lsq>b  ∨  a\<lsq>\<zero> ∧ b\<lsq>\<zero>  ∨  
  a\<lsq>\<zero> ∧ \<zero>\<lsq>b ∧ \<zero> \<lsq> a\<ra>b  ∨ a\<lsq>\<zero> ∧ \<zero>\<lsq>b ∧ a\<ra>b \<lsq> \<zero>  ∨  
  \<zero>\<lsq>a ∧ b\<lsq>\<zero> ∧ \<zero> \<lsq> a\<ra>b  ∨  \<zero>\<lsq>a ∧ b\<lsq>\<zero> ∧ a\<ra>b \<lsq> \<zero>"
  using prems Int_ZF_2_T1 group3.OrdGroup_6cases by simp;

end

Integers as a ring

lemma Int_ZF_1_1_L1:

  [| a ∈ int; b ∈ int; c ∈ int |]
  ==> IntegerMultiplication ` ⟨a, IntegerAddition ` ⟨b, c⟩⟩ =
      IntegerAddition `
      ⟨IntegerMultiplication ` ⟨a, b⟩, IntegerMultiplication ` ⟨a, c⟩⟩
  [| a ∈ int; b ∈ int; c ∈ int |]
  ==> IntegerMultiplication ` ⟨IntegerAddition ` ⟨b, c⟩, a⟩ =
      IntegerAddition `
      ⟨IntegerMultiplication ` ⟨b, a⟩, IntegerMultiplication ` ⟨c, a⟩⟩

lemma Int_ZF_1_1_L2:

  IsAring(int, IntegerAddition, IntegerMultiplication)
  IntegerMultiplication {is commutative on} int
  ring0(int, IntegerAddition, IntegerMultiplication)

lemma int_zero_one_are_int:

  TheNeutralElement(int, IntegerAddition) ∈ int
  TheNeutralElement(int, IntegerMultiplication) ∈ int

lemma int_zero_one_are_intA:

  GroupInv(int, IntegerAddition) ` TheNeutralElement(int, IntegerAddition) =
  TheNeutralElement(int, IntegerAddition)

lemma Int_ZF_1_1_L4:

  a ∈ int ==> IntegerAddition ` ⟨a, TheNeutralElement(int, IntegerAddition)⟩ = a
  a ∈ int ==> IntegerAddition ` ⟨TheNeutralElement(int, IntegerAddition), a⟩ = a
  a ∈ int
  ==> IntegerMultiplication ` ⟨a, TheNeutralElement(int, IntegerMultiplication)⟩ =
      a
  a ∈ int
  ==> IntegerMultiplication ` ⟨TheNeutralElement(int, IntegerMultiplication), a⟩ =
      a
  a ∈ int
  ==> IntegerMultiplication ` ⟨TheNeutralElement(int, IntegerAddition), a⟩ =
      TheNeutralElement(int, IntegerAddition)
  a ∈ int
  ==> IntegerMultiplication ` ⟨a, TheNeutralElement(int, IntegerAddition)⟩ =
      TheNeutralElement(int, IntegerAddition)
  a ∈ int ==> GroupInv(int, IntegerAddition) ` a ∈ int
  a ∈ int
  ==> GroupInv(int, IntegerAddition) ` (GroupInv(int, IntegerAddition) ` a) = a
  a ∈ int
  ==> IntegerAddition ` ⟨a, GroupInv(int, IntegerAddition) ` a⟩ =
      TheNeutralElement(int, IntegerAddition)
  a ∈ int
  ==> IntegerAddition `
      ⟨a, GroupInv(int, IntegerAddition) `
          TheNeutralElement(int, IntegerAddition)⟩ =
      a
  a ∈ int
  ==> IntegerMultiplication `
      ⟨IntegerAddition `
       ⟨TheNeutralElement(int, IntegerMultiplication),
        TheNeutralElement(int, IntegerMultiplication)⟩,
       a⟩ =
      IntegerAddition ` ⟨a, a

lemma Int_ZF_1_1_L5:

  [| a ∈ int; b ∈ int |] ==> IntegerAddition ` ⟨a, b⟩ ∈ int
  [| a ∈ int; b ∈ int |]
  ==> IntegerAddition ` ⟨a, GroupInv(int, IntegerAddition) ` b⟩ ∈ int
  [| a ∈ int; b ∈ int |] ==> IntegerMultiplication ` ⟨a, b⟩ ∈ int
  [| a ∈ int; b ∈ int |] ==> IntegerAddition ` ⟨a, b⟩ = IntegerAddition ` ⟨b, a
  [| a ∈ int; b ∈ int |]
  ==> IntegerMultiplication ` ⟨a, b⟩ = IntegerMultiplication ` ⟨b, a
  [| a ∈ int; b ∈ int |]
  ==> IntegerAddition `
      ⟨GroupInv(int, IntegerAddition) ` b, GroupInv(int, IntegerAddition) ` a⟩ =
      IntegerAddition `
      ⟨GroupInv(int, IntegerAddition) ` a, GroupInv(int, IntegerAddition) ` b
  [| a ∈ int; b ∈ int |]
  ==> GroupInv(int, IntegerAddition) ` (IntegerAddition ` ⟨a, b⟩) =
      IntegerAddition `
      ⟨GroupInv(int, IntegerAddition) ` a, GroupInv(int, IntegerAddition) ` b
  [| a ∈ int; b ∈ int |]
  ==> GroupInv(int, IntegerAddition) `
      (IntegerAddition ` ⟨a, GroupInv(int, IntegerAddition) ` b⟩) =
      IntegerAddition ` ⟨GroupInv(int, IntegerAddition) ` a, b
  [| a ∈ int; b ∈ int |]
  ==> IntegerMultiplication ` ⟨GroupInv(int, IntegerAddition) ` a, b⟩ =
      GroupInv(int, IntegerAddition) ` (IntegerMultiplication ` ⟨a, b⟩)
  [| a ∈ int; b ∈ int |]
  ==> IntegerMultiplication ` ⟨a, GroupInv(int, IntegerAddition) ` b⟩ =
      GroupInv(int, IntegerAddition) ` (IntegerMultiplication ` ⟨a, b⟩)
  [| a ∈ int; b ∈ int |]
  ==> IntegerMultiplication `
      ⟨GroupInv(int, IntegerAddition) ` a, GroupInv(int, IntegerAddition) ` b⟩ =
      IntegerMultiplication ` ⟨a, b

lemma int_two_three_are_int:

  IntegerAddition `
  ⟨TheNeutralElement(int, IntegerMultiplication),
   TheNeutralElement(int, IntegerMultiplication)⟩ ∈
  int
  IntegerAddition `
  ⟨IntegerAddition `
   ⟨TheNeutralElement(int, IntegerMultiplication),
    TheNeutralElement(int, IntegerMultiplication)⟩,
   TheNeutralElement(int, IntegerMultiplication)⟩ ∈
  int

lemma Int_ZF_1_1_L5B:

  [| a ∈ int; b ∈ int |]
  ==> IntegerAddition `
      ⟨a, GroupInv(int, IntegerAddition) ` (GroupInv(int, IntegerAddition) ` b)⟩ =
      IntegerAddition ` ⟨a, b

lemma Int_ZF_1_1_L6:

  [| a ∈ int; b ∈ int; c ∈ int |]
  ==> IntegerAddition `
      ⟨a, GroupInv(int, IntegerAddition) ` (IntegerAddition ` ⟨b, c⟩)⟩ =
      IntegerAddition `
      ⟨IntegerAddition ` ⟨a, GroupInv(int, IntegerAddition) ` b⟩,
       GroupInv(int, IntegerAddition) ` c
  [| a ∈ int; b ∈ int; c ∈ int |]
  ==> IntegerAddition `
      ⟨a, GroupInv(int, IntegerAddition) `
          (IntegerAddition ` ⟨b, GroupInv(int, IntegerAddition) ` c⟩)⟩ =
      IntegerAddition `
      ⟨IntegerAddition ` ⟨a, GroupInv(int, IntegerAddition) ` b⟩, c
  [| a ∈ int; b ∈ int; c ∈ int |]
  ==> IntegerMultiplication `
      ⟨a, IntegerAddition ` ⟨b, GroupInv(int, IntegerAddition) ` c⟩⟩ =
      IntegerAddition `
      ⟨IntegerMultiplication ` ⟨a, b⟩,
       GroupInv(int, IntegerAddition) ` (IntegerMultiplication ` ⟨a, c⟩)⟩
  [| a ∈ int; b ∈ int; c ∈ int |]
  ==> IntegerMultiplication `
      ⟨IntegerAddition ` ⟨b, GroupInv(int, IntegerAddition) ` c⟩, a⟩ =
      IntegerAddition `
      ⟨IntegerMultiplication ` ⟨b, a⟩,
       GroupInv(int, IntegerAddition) ` (IntegerMultiplication ` ⟨c, a⟩)⟩

lemma Int_ZF_1_1_L6A:

  [| a ∈ int; b ∈ int; c ∈ int |]
  ==> IntegerAddition `
      ⟨a, IntegerAddition ` ⟨b, GroupInv(int, IntegerAddition) ` c⟩⟩ =
      IntegerAddition `
      ⟨IntegerAddition ` ⟨a, b⟩, GroupInv(int, IntegerAddition) ` c

lemma Int_ZF_1_1_L7:

  [| a ∈ int; b ∈ int; c ∈ int |]
  ==> IntegerAddition ` ⟨IntegerAddition ` ⟨a, b⟩, c⟩ =
      IntegerAddition ` ⟨a, IntegerAddition ` ⟨b, c⟩⟩
  [| a ∈ int; b ∈ int; c ∈ int |]
  ==> IntegerMultiplication ` ⟨IntegerMultiplication ` ⟨a, b⟩, c⟩ =
      IntegerMultiplication ` ⟨a, IntegerMultiplication ` ⟨b, c⟩⟩

Rearrangement lemmas

lemma Int_ZF_1_2_L1:

  ⟨TheNeutralElement(int, IntegerAddition), a⟩ ∈ IntegerOrder
  ==> IntegerAddition `
      ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) ` a,
       TheNeutralElement(int, IntegerMultiplication)⟩ =
      AbsoluteValue(int, IntegerAddition, IntegerOrder) `
      (IntegerAddition ` ⟨a, TheNeutralElement(int, IntegerMultiplication)⟩)

lemma Int_ZF_1_2_L2:

  [| a ∈ int; ⟨TheNeutralElement(int, IntegerAddition), b⟩ ∈ IntegerOrder |]
  ==> IntegerAddition `
      ⟨a, IntegerMultiplication `
          ⟨IntegerAddition `
           ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) ` b,
            TheNeutralElement(int, IntegerMultiplication)⟩,
           a⟩⟩ =
      IntegerMultiplication `
      ⟨IntegerAddition `
       ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) `
        (IntegerAddition ` ⟨b, TheNeutralElement(int, IntegerMultiplication)⟩),
        TheNeutralElement(int, IntegerMultiplication)⟩,
       a

lemma Int_ZF_1_2_L3:

  [| a ∈ int; b ∈ int |]
  ==> IntegerAddition `
      ⟨IntegerAddition ` ⟨a, b⟩, GroupInv(int, IntegerAddition) ` a⟩ =
      b
  [| a ∈ int; b ∈ int |]
  ==> IntegerAddition `
      ⟨a, IntegerAddition ` ⟨b, GroupInv(int, IntegerAddition) ` a⟩⟩ =
      b
  [| a ∈ int; b ∈ int |]
  ==> IntegerAddition `
      ⟨IntegerAddition ` ⟨a, b⟩, GroupInv(int, IntegerAddition) ` b⟩ =
      a
  [| a ∈ int; b ∈ int |]
  ==> IntegerAddition `
      ⟨IntegerAddition ` ⟨a, GroupInv(int, IntegerAddition) ` b⟩, b⟩ =
      a
  [| a ∈ int; b ∈ int |]
  ==> IntegerAddition `
      ⟨GroupInv(int, IntegerAddition) ` a, IntegerAddition ` ⟨a, b⟩⟩ =
      b
  [| a ∈ int; b ∈ int |]
  ==> IntegerAddition `
      ⟨a, IntegerAddition ` ⟨b, GroupInv(int, IntegerAddition) ` a⟩⟩ =
      b
  [| a ∈ int; b ∈ int |]
  ==> IntegerAddition `
      ⟨GroupInv(int, IntegerAddition) ` b, IntegerAddition ` ⟨a, b⟩⟩ =
      a
  [| a ∈ int; b ∈ int |]
  ==> IntegerAddition `
      ⟨a, GroupInv(int, IntegerAddition) ` (IntegerAddition ` ⟨b, a⟩)⟩ =
      GroupInv(int, IntegerAddition) ` b
  [| a ∈ int; b ∈ int |]
  ==> IntegerAddition `
      ⟨a, GroupInv(int, IntegerAddition) ` (IntegerAddition ` ⟨a, b⟩)⟩ =
      GroupInv(int, IntegerAddition) ` b
  [| a ∈ int; b ∈ int |]
  ==> IntegerAddition `
      ⟨a, GroupInv(int, IntegerAddition) `
          (IntegerAddition ` ⟨a, GroupInv(int, IntegerAddition) ` b⟩)⟩ =
      b
  [| a ∈ int; b ∈ int |]
  ==> IntegerAddition `
      ⟨IntegerAddition ` ⟨a, GroupInv(int, IntegerAddition) ` b⟩,
       GroupInv(int, IntegerAddition) ` a⟩ =
      GroupInv(int, IntegerAddition) ` b
  [| a ∈ int; b ∈ int |]
  ==> IntegerAddition `
      ⟨IntegerAddition ` ⟨a, GroupInv(int, IntegerAddition) ` b⟩,
       GroupInv(int, IntegerAddition) ` (IntegerAddition ` ⟨a, b⟩)⟩ =
      IntegerAddition `
      ⟨GroupInv(int, IntegerAddition) ` b, GroupInv(int, IntegerAddition) ` b

lemma Int_ZF_1_2_L3A:

a, b⟩ ∈ IntegerOrder
  ==> ⟨IntegerAddition `
       ⟨a, GroupInv(int, IntegerAddition) `
           TheNeutralElement(int, IntegerMultiplication)⟩,
       b⟩ ∈
      IntegerOrder

lemma Int_ZF_1_2_L3AA:

  a ∈ int
  ==> ⟨IntegerAddition `
       ⟨a, GroupInv(int, IntegerAddition) `
           TheNeutralElement(int, IntegerMultiplication)⟩,
       a⟩ ∈
      IntegerOrder
  a ∈ int
  ==> IntegerAddition `
      ⟨a, GroupInv(int, IntegerAddition) `
          TheNeutralElement(int, IntegerMultiplication)⟩ ≠
      a
  a ∈ int
  ==> ⟨a, IntegerAddition `
          ⟨a, GroupInv(int, IntegerAddition) `
              TheNeutralElement(int, IntegerMultiplication)⟩⟩ ∉
      IntegerOrder
  a ∈ int
  ==> ⟨IntegerAddition ` ⟨a, TheNeutralElement(int, IntegerMultiplication)⟩, a⟩ ∉
      IntegerOrder
  a ∈ int
  ==> ⟨IntegerAddition ` ⟨TheNeutralElement(int, IntegerMultiplication), a⟩, a⟩ ∉
      IntegerOrder

lemma Int_ZF_1_2_L4:

a, TheNeutralElement(int, IntegerAddition)⟩ ∈ IntegerOrder
  ==> IntegerAddition `
      ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) ` a,
       TheNeutralElement(int, IntegerMultiplication)⟩ =
      AbsoluteValue(int, IntegerAddition, IntegerOrder) `
      (IntegerAddition `
       ⟨a, GroupInv(int, IntegerAddition) `
           TheNeutralElement(int, IntegerMultiplication)⟩)

lemma Int_ZF_1_2_L5:

  [| a ∈ int; ⟨b, TheNeutralElement(int, IntegerAddition)⟩ ∈ IntegerOrder |]
  ==> IntegerAddition `
      ⟨a, IntegerMultiplication `
          ⟨IntegerAddition `
           ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) ` b,
            TheNeutralElement(int, IntegerMultiplication)⟩,
           a⟩⟩ =
      IntegerMultiplication `
      ⟨IntegerAddition `
       ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) `
        (IntegerAddition `
         ⟨b, GroupInv(int, IntegerAddition) `
             TheNeutralElement(int, IntegerMultiplication)⟩),
        TheNeutralElement(int, IntegerMultiplication)⟩,
       a

lemma Int_ZF_1_2_L6:

  [| a ∈ int; b ∈ int; c ∈ int; d ∈ int |]
  ==> IntegerAddition `
      ⟨a, GroupInv(int, IntegerAddition) `
          (IntegerMultiplication `
           ⟨IntegerAddition `
            ⟨b, GroupInv(int, IntegerAddition) `
                TheNeutralElement(int, IntegerMultiplication)⟩,
            c⟩)⟩ =
      IntegerAddition `
      ⟨IntegerAddition `
       ⟨d, GroupInv(int, IntegerAddition) ` (IntegerMultiplication ` ⟨b, c⟩)⟩,
       GroupInv(int, IntegerAddition) `
       (IntegerAddition `
        ⟨IntegerAddition ` ⟨d, GroupInv(int, IntegerAddition) ` a⟩,
         GroupInv(int, IntegerAddition) ` c⟩)⟩

lemma Int_ZF_1_2_L7:

  [| a ∈ int; b ∈ int |]
  ==> IntegerMultiplication ` ⟨a, b⟩ =
      IntegerAddition `
      ⟨IntegerMultiplication `
       ⟨IntegerAddition `
        ⟨a, GroupInv(int, IntegerAddition) `
            TheNeutralElement(int, IntegerMultiplication)⟩,
        b⟩,
       b
  [| a ∈ int; b ∈ int |]
  ==> IntegerMultiplication `
      ⟨a, IntegerAddition ` ⟨b, TheNeutralElement(int, IntegerMultiplication)⟩⟩ =
      IntegerAddition ` ⟨IntegerMultiplication ` ⟨a, b⟩, a
  [| a ∈ int; b ∈ int |]
  ==> IntegerMultiplication `
      ⟨IntegerAddition ` ⟨b, TheNeutralElement(int, IntegerMultiplication)⟩, a⟩ =
      IntegerAddition ` ⟨IntegerMultiplication ` ⟨b, a⟩, a
  [| a ∈ int; b ∈ int |]
  ==> IntegerMultiplication `
      ⟨IntegerAddition ` ⟨b, TheNeutralElement(int, IntegerMultiplication)⟩, a⟩ =
      IntegerAddition ` ⟨a, IntegerMultiplication ` ⟨b, a⟩⟩

lemma Int_ZF_1_2_L8:

  [| a ∈ int; b ∈ int |]
  ==> IntegerAddition `
      ⟨IntegerAddition ` ⟨a, TheNeutralElement(int, IntegerMultiplication)⟩,
       IntegerAddition ` ⟨b, TheNeutralElement(int, IntegerMultiplication)⟩⟩ =
      IntegerAddition `
      ⟨IntegerAddition ` ⟨b, a⟩,
       IntegerAddition `
       ⟨TheNeutralElement(int, IntegerMultiplication),
        TheNeutralElement(int, IntegerMultiplication)⟩⟩

lemma Int_ZF_1_2_L9:

  [| a ∈ int; b ∈ int; c ∈ int |]
  ==> IntegerAddition `
      ⟨IntegerAddition ` ⟨a, GroupInv(int, IntegerAddition) ` b⟩,
       IntegerAddition ` ⟨b, GroupInv(int, IntegerAddition) ` c⟩⟩ =
      IntegerAddition ` ⟨a, GroupInv(int, IntegerAddition) ` c
  [| a ∈ int; b ∈ int; c ∈ int |]
  ==> IntegerAddition `
      ⟨IntegerAddition ` ⟨a, GroupInv(int, IntegerAddition) ` b⟩,
       GroupInv(int, IntegerAddition) `
       (IntegerAddition ` ⟨a, GroupInv(int, IntegerAddition) ` c⟩)⟩ =
      IntegerAddition ` ⟨c, GroupInv(int, IntegerAddition) ` b
  [| a ∈ int; b ∈ int; c ∈ int |]
  ==> IntegerAddition `
      ⟨a, IntegerAddition `
          ⟨b, IntegerAddition `
              ⟨IntegerAddition ` ⟨c, GroupInv(int, IntegerAddition) ` a⟩,
               GroupInv(int, IntegerAddition) ` b⟩⟩⟩ =
      c
  [| a ∈ int; b ∈ int; c ∈ int |]
  ==> IntegerAddition `
      ⟨IntegerAddition `
       ⟨GroupInv(int, IntegerAddition) ` a, GroupInv(int, IntegerAddition) ` b⟩,
       c⟩ =
      IntegerAddition `
      ⟨IntegerAddition ` ⟨c, GroupInv(int, IntegerAddition) ` a⟩,
       GroupInv(int, IntegerAddition) ` b
  [| a ∈ int; b ∈ int; c ∈ int |]
  ==> IntegerAddition `
      ⟨IntegerAddition `
       ⟨GroupInv(int, IntegerAddition) ` b, GroupInv(int, IntegerAddition) ` a⟩,
       c⟩ =
      IntegerAddition `
      ⟨IntegerAddition ` ⟨c, GroupInv(int, IntegerAddition) ` a⟩,
       GroupInv(int, IntegerAddition) ` b
  [| a ∈ int; b ∈ int; c ∈ int |]
  ==> GroupInv(int, IntegerAddition) `
      (IntegerAddition `
       ⟨IntegerAddition ` ⟨GroupInv(int, IntegerAddition) ` a, b⟩, c⟩) =
      IntegerAddition `
      ⟨IntegerAddition ` ⟨a, GroupInv(int, IntegerAddition) ` b⟩,
       GroupInv(int, IntegerAddition) ` c
  [| a ∈ int; b ∈ int; c ∈ int |]
  ==> IntegerAddition `
      ⟨IntegerAddition ` ⟨IntegerAddition ` ⟨a, b⟩, c⟩,
       GroupInv(int, IntegerAddition) ` a⟩ =
      IntegerAddition ` ⟨b, c
  [| a ∈ int; b ∈ int; c ∈ int |]
  ==> IntegerAddition `
      ⟨IntegerAddition ` ⟨a, b⟩,
       GroupInv(int, IntegerAddition) ` (IntegerAddition ` ⟨a, c⟩)⟩ =
      IntegerAddition ` ⟨b, GroupInv(int, IntegerAddition) ` c

lemma Int_ZF_1_2_L9A:

  [| a ∈ int; b ∈ int; c ∈ int |]
  ==> GroupInv(int, IntegerAddition) `
      (IntegerAddition `
       ⟨IntegerAddition ` ⟨a, GroupInv(int, IntegerAddition) ` b⟩,
        GroupInv(int, IntegerAddition) ` c⟩) =
      IntegerAddition `
      ⟨IntegerAddition ` ⟨c, b⟩, GroupInv(int, IntegerAddition) ` a

lemma Int_ZF_1_2_L10:

  [| a ∈ int; b ∈ int; c ∈ int |]
  ==> IntegerAddition `
      ⟨IntegerMultiplication `
       ⟨IntegerAddition ` ⟨a, TheNeutralElement(int, IntegerMultiplication)⟩, b⟩,
       IntegerMultiplication `
       ⟨IntegerAddition ` ⟨c, TheNeutralElement(int, IntegerMultiplication)⟩,
        b⟩⟩ =
      IntegerMultiplication `
      ⟨IntegerAddition `
       ⟨IntegerAddition ` ⟨c, a⟩,
        IntegerAddition `
        ⟨TheNeutralElement(int, IntegerMultiplication),
         TheNeutralElement(int, IntegerMultiplication)⟩⟩,
       b

lemma Int_ZF_1_2_L10A:

  [| a ∈ int; b ∈ int; c ∈ int; e ∈ int;
     ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) `
      (IntegerAddition `
       ⟨IntegerMultiplication ` ⟨a, b⟩, GroupInv(int, IntegerAddition) ` c⟩),
      d⟩ ∈
     IntegerOrder;
     ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) `
      (IntegerAddition `
       ⟨IntegerMultiplication ` ⟨b, a⟩, GroupInv(int, IntegerAddition) ` e⟩),
      f⟩ ∈
     IntegerOrder |]
  ==> ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) `
       (IntegerAddition ` ⟨c, GroupInv(int, IntegerAddition) ` e⟩),
       IntegerAddition ` ⟨f, d⟩⟩ ∈
      IntegerOrder

lemma Int_ZF_1_2_L11:

  a ∈ int
  ==> IntegerAddition `
      ⟨IntegerAddition ` ⟨a, TheNeutralElement(int, IntegerMultiplication)⟩,
       IntegerAddition `
       ⟨TheNeutralElement(int, IntegerMultiplication),
        TheNeutralElement(int, IntegerMultiplication)⟩⟩ =
      IntegerAddition `
      ⟨a, IntegerAddition `
          ⟨IntegerAddition `
           ⟨TheNeutralElement(int, IntegerMultiplication),
            TheNeutralElement(int, IntegerMultiplication)⟩,
           TheNeutralElement(int, IntegerMultiplication)⟩⟩
  a ∈ int
  ==> a = IntegerAddition `
          ⟨IntegerMultiplication `
           ⟨IntegerAddition `
            ⟨TheNeutralElement(int, IntegerMultiplication),
             TheNeutralElement(int, IntegerMultiplication)⟩,
            a⟩,
           GroupInv(int, IntegerAddition) ` a

lemma Int_ZF_1_2_L12:

  [| a ∈ int; b ∈ int; c ∈ int |]
  ==> IntegerMultiplication `
      ⟨IntegerAddition ` ⟨b, GroupInv(int, IntegerAddition) ` c⟩, a⟩ =
      IntegerAddition `
      ⟨IntegerMultiplication ` ⟨a, b⟩,
       GroupInv(int, IntegerAddition) ` (IntegerMultiplication ` ⟨a, c⟩)⟩

lemma Int_ZF_1_2_L13:

  [| a ∈ int; b ∈ int; c ∈ int; d ∈ int; x ∈ int |]
  ==> IntegerMultiplication `
      ⟨IntegerAddition `
       ⟨IntegerAddition `
        ⟨x, IntegerAddition ` ⟨IntegerMultiplication ` ⟨a, x⟩, b⟩⟩,
        c⟩,
       d⟩ =
      IntegerAddition `
      ⟨IntegerMultiplication `
       ⟨IntegerMultiplication `
        ⟨d, IntegerAddition ` ⟨a, TheNeutralElement(int, IntegerMultiplication)⟩⟩,
        x⟩,
       IntegerAddition `
       ⟨IntegerMultiplication ` ⟨b, d⟩, IntegerMultiplication ` ⟨c, d⟩⟩⟩

lemma Int_ZF_1_2_L14:

  [| a ∈ int; b ∈ int; c ∈ int; d ∈ int; x ∈ int |]
  ==> IntegerAddition `
      ⟨IntegerAddition ` ⟨IntegerMultiplication ` ⟨a, x⟩, b⟩,
       IntegerAddition ` ⟨IntegerMultiplication ` ⟨c, x⟩, d⟩⟩ =
      IntegerAddition `
      ⟨IntegerMultiplication ` ⟨IntegerAddition ` ⟨a, c⟩, x⟩,
       IntegerAddition ` ⟨b, d⟩⟩

lemma Int_ZF_1_2_L15:

  [| a ∈ int; b ∈ int; c ∈ int; d ∈ int;
     a = IntegerAddition `
         ⟨IntegerAddition ` ⟨b, GroupInv(int, IntegerAddition) ` c⟩,
          GroupInv(int, IntegerAddition) ` d⟩ |]
  ==> d = IntegerAddition `
          ⟨IntegerAddition ` ⟨b, GroupInv(int, IntegerAddition) ` a⟩,
           GroupInv(int, IntegerAddition) ` c
  [| a ∈ int; b ∈ int; c ∈ int; d ∈ int;
     a = IntegerAddition `
         ⟨IntegerAddition ` ⟨b, GroupInv(int, IntegerAddition) ` c⟩,
          GroupInv(int, IntegerAddition) ` d⟩ |]
  ==> d = IntegerAddition `
          ⟨IntegerAddition ` ⟨GroupInv(int, IntegerAddition) ` a, b⟩,
           GroupInv(int, IntegerAddition) ` c
  [| a ∈ int; b ∈ int; c ∈ int; d ∈ int;
     a = IntegerAddition `
         ⟨IntegerAddition ` ⟨b, GroupInv(int, IntegerAddition) ` c⟩,
          GroupInv(int, IntegerAddition) ` d⟩ |]
  ==> b = IntegerAddition ` ⟨IntegerAddition ` ⟨a, d⟩, c

lemma Int_ZF_1_2_L16:

  [| a ∈ int; b ∈ int; c ∈ int; d ∈ int |]
  ==> IntegerAddition `
      ⟨IntegerAddition `
       ⟨a, IntegerAddition ` ⟨b, GroupInv(int, IntegerAddition) ` c⟩⟩,
       d⟩ =
      IntegerAddition `
      ⟨IntegerAddition ` ⟨IntegerAddition ` ⟨a, b⟩, d⟩,
       GroupInv(int, IntegerAddition) ` c

lemma Int_ZF_1_2_L17:

  [| a ∈ int; b ∈ int; c ∈ int |]
  ==> IntegerAddition `
      ⟨IntegerAddition `
       ⟨IntegerAddition ` ⟨a, b⟩, GroupInv(int, IntegerAddition) ` c⟩,
       IntegerAddition ` ⟨c, GroupInv(int, IntegerAddition) ` b⟩⟩ =
      a
  [| a ∈ int; b ∈ int; c ∈ int |]
  ==> IntegerAddition `
      ⟨IntegerAddition ` ⟨a, IntegerAddition ` ⟨b, c⟩⟩,
       GroupInv(int, IntegerAddition) ` c⟩ =
      IntegerAddition ` ⟨a, b

lemma Int_ZF_1_2_L18:

  [| a ∈ int; b ∈ int; c ∈ int |]
  ==> IntegerAddition `
      ⟨IntegerAddition `
       ⟨IntegerAddition ` ⟨a, b⟩, GroupInv(int, IntegerAddition) ` c⟩,
       IntegerAddition ` ⟨c, GroupInv(int, IntegerAddition) ` a⟩⟩ =
      b

Integers as an ordered ring

lemma Int_ZF_1_3_L1:

  [| ⟨TheNeutralElement(int, IntegerAddition), a⟩ ∈ IntegerOrder;
     ⟨TheNeutralElement(int, IntegerAddition), b⟩ ∈ IntegerOrder;
     ⟨TheNeutralElement(int, IntegerAddition), IntegerMultiplication ` ⟨a, b⟩⟩ ∈
     IntegerOrder |]
  ==> ⟨TheNeutralElement(int, IntegerAddition),
       IntegerMultiplication `
       ⟨a, IntegerAddition `
           ⟨b, TheNeutralElement(int, IntegerMultiplication)⟩⟩⟩ ∈
      IntegerOrder

lemma Int_ZF_1_3_L2:

  [| ⟨TheNeutralElement(int, IntegerAddition), a⟩ ∈ IntegerOrder;
     ⟨TheNeutralElement(int, IntegerAddition), b⟩ ∈ IntegerOrder |]
  ==> ⟨TheNeutralElement(int, IntegerAddition), IntegerMultiplication ` ⟨a, b⟩⟩ ∈
      IntegerOrder

lemma Int_ZF_1_3_L2A:

  Nonnegative(int, IntegerAddition, IntegerOrder) {is closed under}
  IntegerMultiplication

theorem Int_ZF_1_3_T1:

  IsAnOrdRing(int, IntegerAddition, IntegerMultiplication, IntegerOrder)
  ring1(int, IntegerAddition, IntegerMultiplication, IntegerOrder)

lemma Int_ZF_1_3_L3_indstep:

  [| ⟨TheNeutralElement(int, IntegerMultiplication), a⟩ ∈ IntegerOrder;
     ⟨TheNeutralElement(int, IntegerMultiplication), b⟩ ∈ IntegerOrder;
     ⟨TheNeutralElement(int, IntegerMultiplication),
      IntegerMultiplication ` ⟨a, b⟩⟩ ∈
     IntegerOrder |]
  ==> ⟨TheNeutralElement(int, IntegerMultiplication),
       IntegerMultiplication `
       ⟨a, IntegerAddition `
           ⟨b, TheNeutralElement(int, IntegerMultiplication)⟩⟩⟩ ∈
      IntegerOrder

lemma Int_ZF_1_3_L3:

  [| ⟨TheNeutralElement(int, IntegerMultiplication), a⟩ ∈ IntegerOrder;
     ⟨TheNeutralElement(int, IntegerMultiplication), b⟩ ∈ IntegerOrder |]
  ==> ⟨TheNeutralElement(int, IntegerMultiplication),
       IntegerMultiplication ` ⟨a, b⟩⟩ ∈
      IntegerOrder

lemma Int_ZF_1_3_L4:

  [| a ∈ int; b ∈ int |]
  ==> AbsoluteValue(int, IntegerAddition, IntegerOrder) `
      (IntegerMultiplication ` ⟨GroupInv(int, IntegerAddition) ` a, b⟩) =
      AbsoluteValue(int, IntegerAddition, IntegerOrder) `
      (IntegerMultiplication ` ⟨a, b⟩)
  [| a ∈ int; b ∈ int |]
  ==> AbsoluteValue(int, IntegerAddition, IntegerOrder) `
      (IntegerMultiplication ` ⟨a, GroupInv(int, IntegerAddition) ` b⟩) =
      AbsoluteValue(int, IntegerAddition, IntegerOrder) `
      (IntegerMultiplication ` ⟨a, b⟩)
  [| a ∈ int; b ∈ int |]
  ==> AbsoluteValue(int, IntegerAddition, IntegerOrder) `
      (IntegerMultiplication `
       ⟨GroupInv(int, IntegerAddition) ` a, GroupInv(int, IntegerAddition) ` b⟩) =
      AbsoluteValue(int, IntegerAddition, IntegerOrder) `
      (IntegerMultiplication ` ⟨a, b⟩)

lemma Int_ZF_1_3_L5:

  [| a ∈ int; b ∈ int |]
  ==> AbsoluteValue(int, IntegerAddition, IntegerOrder) `
      (IntegerMultiplication ` ⟨a, b⟩) =
      IntegerMultiplication `
      ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) ` a,
       AbsoluteValue(int, IntegerAddition, IntegerOrder) ` b

lemma Int_ZF_1_3_L5A:

  ⟨TheNeutralElement(int, IntegerAddition), a⟩ ∈ IntegerOrder
  ==> ⟨TheNeutralElement(int, IntegerAddition),
       IntegerMultiplication `
       ⟨IntegerAddition `
        ⟨TheNeutralElement(int, IntegerMultiplication),
         TheNeutralElement(int, IntegerMultiplication)⟩,
        a⟩⟩ ∈
      IntegerOrder

lemma Int_ZF_1_3_L6:

  [| a ∈ int; b ∈ int |]
  ==> ⟨b, a⟩ ∉ IntegerOrder <->
      ⟨IntegerAddition ` ⟨a, TheNeutralElement(int, IntegerMultiplication)⟩, b⟩ ∈
      IntegerOrder

corollary no_int_between:

  [| a ∈ int; b ∈ int |]
  ==> ⟨b, a⟩ ∈ IntegerOrder ∨
      ⟨IntegerAddition ` ⟨a, TheNeutralElement(int, IntegerMultiplication)⟩, b⟩ ∈
      IntegerOrder

corollary Int_ZF_1_3_L6A:

  [| a ∈ int; b ∈ int; ⟨b, a⟩ ∉ IntegerOrder |]
  ==> ⟨a, IntegerAddition `
          ⟨b, GroupInv(int, IntegerAddition) `
              TheNeutralElement(int, IntegerMultiplication)⟩⟩ ∈
      IntegerOrder

lemma no_int_between1:

  [| ⟨a, b⟩ ∈ IntegerOrder; ab |]
  ==> ⟨IntegerAddition ` ⟨a, TheNeutralElement(int, IntegerMultiplication)⟩, b⟩ ∈
      IntegerOrder
  [| ⟨a, b⟩ ∈ IntegerOrder; ab |]
  ==> ⟨a, IntegerAddition `
          ⟨b, GroupInv(int, IntegerAddition) `
              TheNeutralElement(int, IntegerMultiplication)⟩⟩ ∈
      IntegerOrder

lemma Int_ZF_1_3_L6B:

  [| a ∈ int; b ∈ int |]
  ==> a = b ∨
      ⟨a, IntegerAddition `
          ⟨b, GroupInv(int, IntegerAddition) `
              TheNeutralElement(int, IntegerMultiplication)⟩⟩ ∈
      IntegerOrder ∨
      ⟨IntegerAddition ` ⟨b, TheNeutralElement(int, IntegerMultiplication)⟩, a⟩ ∈
      IntegerOrder

corollary Int_ZF_1_3_L6C:

  a ∈ int
  ==> a = TheNeutralElement(int, IntegerAddition) ∨
      ⟨a, GroupInv(int, IntegerAddition) `
          TheNeutralElement(int, IntegerMultiplication)⟩ ∈
      IntegerOrder ∨
      ⟨TheNeutralElement(int, IntegerMultiplication), a⟩ ∈ IntegerOrder

lemma Int_ZF_1_3_L7:

  a ∈ int
  ==> ⟨a, TheNeutralElement(int, IntegerAddition)⟩ ∉ IntegerOrder <->
      ⟨TheNeutralElement(int, IntegerMultiplication), a⟩ ∈ IntegerOrder

lemma Int_ZF_1_3_L8:

  [| a ∈ int; b ∈ int;
     ⟨a, TheNeutralElement(int, IntegerAddition)⟩ ∉ IntegerOrder;
     ⟨b, TheNeutralElement(int, IntegerAddition)⟩ ∉ IntegerOrder |]
  ==> ⟨IntegerMultiplication ` ⟨a, b⟩, TheNeutralElement(int, IntegerAddition)⟩ ∉
      IntegerOrder

lemma Int_ZF_1_3_L9:

  [| a ∈ int; b ∈ int;
     ⟨b, TheNeutralElement(int, IntegerAddition)⟩ ∉ IntegerOrder;
     ⟨IntegerMultiplication ` ⟨a, b⟩, TheNeutralElement(int, IntegerAddition)⟩ ∈
     IntegerOrder |]
  ==> ⟨a, TheNeutralElement(int, IntegerAddition)⟩ ∈ IntegerOrder

lemma Int_ZF_1_3_L10:

  [| a ∈ int; b ∈ int |]
  ==> ⟨a, b⟩ ∈ IntegerOrder <->
      ⟨IntegerAddition ` ⟨a, GroupInv(int, IntegerAddition) ` b⟩,
       TheNeutralElement(int, IntegerAddition)⟩ ∈
      IntegerOrder

lemma Int_ZF_1_3_L10A:

a, b⟩ ∈ IntegerOrder
  ==> ⟨TheNeutralElement(int, IntegerAddition),
       IntegerAddition ` ⟨b, GroupInv(int, IntegerAddition) ` a⟩⟩ ∈
      IntegerOrder

lemma Int_ineq_simpl_positive:

  [| a ∈ int; b ∈ int; c ∈ int;
     ⟨IntegerMultiplication ` ⟨a, c⟩, IntegerMultiplication ` ⟨b, c⟩⟩ ∈
     IntegerOrder;
     ⟨c, TheNeutralElement(int, IntegerAddition)⟩ ∉ IntegerOrder |]
  ==> ⟨a, b⟩ ∈ IntegerOrder

lemma Int_ZF_1_3_L11:

  [| a ∈ int; b ∈ int;
     ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) ` a,
      AbsoluteValue(int, IntegerAddition, IntegerOrder) ` b⟩ ∉
     IntegerOrder |]
  ==> ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) ` a,
       TheNeutralElement(int, IntegerAddition)⟩ ∉
      IntegerOrder

lemma Int_ZF_1_3_L12:

  [| ⟨a, TheNeutralElement(int, IntegerAddition)⟩ ∈ IntegerOrder;
     ⟨TheNeutralElement(int, IntegerAddition), b⟩ ∈ IntegerOrder |]
  ==> ⟨IntegerMultiplication ` ⟨a, b⟩, TheNeutralElement(int, IntegerAddition)⟩ ∈
      IntegerOrder

lemma Int_ZF_1_3_L13:

  [| ⟨a, b⟩ ∈ IntegerOrder;
     ⟨TheNeutralElement(int, IntegerAddition), c⟩ ∈ IntegerOrder |]
  ==> ⟨IntegerMultiplication ` ⟨a, c⟩, IntegerMultiplication ` ⟨b, c⟩⟩ ∈
      IntegerOrder
  [| ⟨a, b⟩ ∈ IntegerOrder;
     ⟨TheNeutralElement(int, IntegerAddition), c⟩ ∈ IntegerOrder |]
  ==> ⟨IntegerMultiplication ` ⟨c, a⟩, IntegerMultiplication ` ⟨c, b⟩⟩ ∈
      IntegerOrder

lemma Int_ZF_1_3_L13A:

  [| ⟨TheNeutralElement(int, IntegerMultiplication), a⟩ ∈ IntegerOrder;
     ⟨b, c⟩ ∈ IntegerOrder;
     ⟨IntegerMultiplication `
      ⟨IntegerAddition ` ⟨a, TheNeutralElement(int, IntegerMultiplication)⟩, c⟩,
      d⟩ ∈
     IntegerOrder |]
  ==> ⟨IntegerMultiplication `
       ⟨IntegerAddition ` ⟨a, TheNeutralElement(int, IntegerMultiplication)⟩, b⟩,
       d⟩ ∈
      IntegerOrder

lemma Int_ZF_1_3_L13B:

  [| ⟨a, b⟩ ∈ IntegerOrder; c ∈ PositiveSet(int, IntegerAddition, IntegerOrder) |]
  ==> ⟨IntegerMultiplication ` ⟨a, c⟩, IntegerMultiplication ` ⟨b, c⟩⟩ ∈
      IntegerOrder
  [| ⟨a, b⟩ ∈ IntegerOrder; c ∈ PositiveSet(int, IntegerAddition, IntegerOrder) |]
  ==> ⟨IntegerMultiplication ` ⟨c, a⟩, IntegerMultiplication ` ⟨c, b⟩⟩ ∈
      IntegerOrder

lemma Int_ZF_1_3_L14:

  [| a ∈ int; b ∈ int; c ∈ int; d ∈ int |]
  ==> IntegerAddition `
      ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) `
       (IntegerMultiplication ` ⟨a, b⟩),
       IntegerMultiplication `
       ⟨IntegerAddition `
        ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) ` a, c⟩,
        d⟩⟩ =
      IntegerAddition `
      ⟨IntegerMultiplication `
       ⟨IntegerAddition `
        ⟨d, AbsoluteValue(int, IntegerAddition, IntegerOrder) ` b⟩,
        AbsoluteValue(int, IntegerAddition, IntegerOrder) ` a⟩,
       IntegerMultiplication ` ⟨c, d⟩⟩

lemma Int_ZF_1_3_L15:

  [| m ∈ int; n ∈ int;
     ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) ` m,
      AbsoluteValue(int, IntegerAddition, IntegerOrder) ` n⟩ ∉
     IntegerOrder |]
  ==> ⟨n, AbsoluteValue(int, IntegerAddition, IntegerOrder) ` m⟩ ∈ IntegerOrder
  [| m ∈ int; n ∈ int;
     ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) ` m,
      AbsoluteValue(int, IntegerAddition, IntegerOrder) ` n⟩ ∉
     IntegerOrder |]
  ==> m ≠ TheNeutralElement(int, IntegerAddition)

lemma Int_ZF_1_3_L16:

  ⟨TheNeutralElement(int, IntegerAddition), m⟩ ∈ IntegerOrder
  ==> ⟨GroupInv(int, IntegerAddition) ` m,
       TheNeutralElement(int, IntegerAddition)⟩ ∈
      IntegerOrder

lemma Int_ZF_1_3_L17:

  m ∈ int
  ==> ⟨GroupInv(int, IntegerAddition) `
       (AbsoluteValue(int, IntegerAddition, IntegerOrder) ` m),
       AbsoluteValue(int, IntegerAddition, IntegerOrder) ` m⟩ ∈
      IntegerOrder
  m ∈ int
  ==> Interval
       (IntegerOrder,
        GroupInv(int, IntegerAddition) `
        (AbsoluteValue(int, IntegerAddition, IntegerOrder) ` m),
        AbsoluteValue(int, IntegerAddition, IntegerOrder) ` m) ≠
      0

lemma Int_ZF_1_3_L18:

  [| m ∈ int; n ∈ int |] ==> ⟨m, GreaterOf(IntegerOrder, m, n)⟩ ∈ IntegerOrder
  [| m ∈ int; n ∈ int |] ==> ⟨n, GreaterOf(IntegerOrder, m, n)⟩ ∈ IntegerOrder
  [| m ∈ int; n ∈ int |] ==> ⟨SmallerOf(IntegerOrder, m, n), m⟩ ∈ IntegerOrder
  [| m ∈ int; n ∈ int |] ==> ⟨SmallerOf(IntegerOrder, m, n), n⟩ ∈ IntegerOrder

lemma Int_ZF_1_3_L19:

  [| m ∈ int;
     ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) ` m, n⟩ ∈ IntegerOrder |]
  ==> ⟨GroupInv(int, IntegerAddition) ` n, m⟩ ∈ IntegerOrder
  [| m ∈ int;
     ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) ` m, n⟩ ∈ IntegerOrder |]
  ==> ⟨m, n⟩ ∈ IntegerOrder
  [| m ∈ int;
     ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) ` m, n⟩ ∈ IntegerOrder |]
  ==> m ∈ Interval(IntegerOrder, GroupInv(int, IntegerAddition) ` n, n)
  [| m ∈ int;
     ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) ` m, n⟩ ∈ IntegerOrder |]
  ==> ⟨TheNeutralElement(int, IntegerAddition), n⟩ ∈ IntegerOrder

lemma Int_ZF_1_3_L19A:

  [| m ∈ int;
     ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) ` m, n⟩ ∈ IntegerOrder;
     ⟨TheNeutralElement(int, IntegerAddition), k⟩ ∈ IntegerOrder |]
  ==> ⟨GroupInv(int, IntegerAddition) ` (IntegerAddition ` ⟨n, k⟩), m⟩ ∈
      IntegerOrder

lemma Int_ZF_1_3_L20:

xX. b(x) ∈ int ∧
        ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) ` b(x), L⟩ ∈
        IntegerOrder
  ==> IsBounded({b(x) . xX}, IntegerOrder)

lemma Int_ZF_1_3_L20A:

  IsBounded(A, IntegerOrder)
  ==> ∃L. ∀aA. ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) ` a, L⟩ ∈
                IntegerOrder

lemma Int_ZF_1_3_L20AA:

  {b(x) . x ∈ int} ∈ Fin(int)
  ==> ∃L∈int.
         ∀x∈int.
            ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) ` b(x), L⟩ ∈
            IntegerOrder

lemma Int_ZF_1_3_L20B:

  [| fX -> int; AX;
     ∀xA. ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) ` (f ` x), L⟩ ∈
           IntegerOrder |]
  ==> IsBounded(f `` A, IntegerOrder)

corollary Int_ZF_1_3_L20C:

  [| f ∈ int -> int;
     ∀m∈int.
        ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) ` (f ` m), L⟩ ∈
        IntegerOrder |]
  ==> f `` int ∈ Fin(int)

lemma int_triangle_ineq3:

  [| a ∈ int; b ∈ int; c ∈ int |]
  ==> ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) `
       (IntegerAddition `
        ⟨IntegerAddition ` ⟨a, GroupInv(int, IntegerAddition) ` b⟩,
         GroupInv(int, IntegerAddition) ` c⟩),
       IntegerAddition `
       ⟨IntegerAddition `
        ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) ` a,
         AbsoluteValue(int, IntegerAddition, IntegerOrder) ` b⟩,
        AbsoluteValue(int, IntegerAddition, IntegerOrder) ` c⟩⟩ ∈
      IntegerOrder

lemma Int_ZF_1_3_L21:

  [| ⟨a, c⟩ ∈ IntegerOrder; ⟨b, c⟩ ∈ IntegerOrder |]
  ==> ⟨IntegerAddition ` ⟨a, b⟩,
       IntegerMultiplication `
       ⟨IntegerAddition `
        ⟨TheNeutralElement(int, IntegerMultiplication),
         TheNeutralElement(int, IntegerMultiplication)⟩,
        c⟩⟩ ∈
      IntegerOrder

lemma Int_ZF_1_3_L22:

  [| ⟨a, b⟩ ∈ IntegerOrder; c ∈ int;
     ⟨b, IntegerAddition ` ⟨c, a⟩⟩ ∈ IntegerOrder |]
  ==> ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) `
       (IntegerAddition ` ⟨b, GroupInv(int, IntegerAddition) ` a⟩),
       c⟩ ∈
      IntegerOrder

lemma Int_ZF_1_3_L22A:

  [| a ∈ int; b ∈ int; c ∈ int; d ∈ int |]
  ==> ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) `
       (IntegerAddition ` ⟨a, GroupInv(int, IntegerAddition) ` c⟩),
       IntegerAddition `
       ⟨IntegerAddition `
        ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) `
         (IntegerAddition ` ⟨a, b⟩),
         AbsoluteValue(int, IntegerAddition, IntegerOrder) `
         (IntegerAddition ` ⟨c, d⟩)⟩,
        AbsoluteValue(int, IntegerAddition, IntegerOrder) `
        (IntegerAddition ` ⟨b, GroupInv(int, IntegerAddition) ` d⟩)⟩⟩ ∈
      IntegerOrder

lemma Int_ZF_1_3_L23:

  [| ⟨a, b⟩ ∈ IntegerOrder; c ∈ int;
     ⟨b, IntegerAddition ` ⟨a, c⟩⟩ ∈ IntegerOrder |]
  ==> ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) `
       (IntegerAddition ` ⟨b, GroupInv(int, IntegerAddition) ` a⟩),
       c⟩ ∈
      IntegerOrder

Maximum and minimum of a set of integers

theorem Int_fin_have_max_min:

  [| A ∈ Fin(int); A ≠ 0 |] ==> HasAmaximum(IntegerOrder, A)
  [| A ∈ Fin(int); A ≠ 0 |] ==> HasAminimum(IntegerOrder, A)
  [| A ∈ Fin(int); A ≠ 0 |] ==> Maximum(IntegerOrder, A) ∈ A
  [| A ∈ Fin(int); A ≠ 0 |] ==> Minimum(IntegerOrder, A) ∈ A
  [| A ∈ Fin(int); A ≠ 0 |] ==> ∀xA. ⟨x, Maximum(IntegerOrder, A)⟩ ∈ IntegerOrder
  [| A ∈ Fin(int); A ≠ 0 |] ==> ∀xA. ⟨Minimum(IntegerOrder, A), x⟩ ∈ IntegerOrder
  [| A ∈ Fin(int); A ≠ 0 |] ==> Maximum(IntegerOrder, A) ∈ int
  [| A ∈ Fin(int); A ≠ 0 |] ==> Minimum(IntegerOrder, A) ∈ int

theorem Int_bounded_have_max_min:

  [| IsBounded(A, IntegerOrder); A ≠ 0 |] ==> HasAmaximum(IntegerOrder, A)
  [| IsBounded(A, IntegerOrder); A ≠ 0 |] ==> HasAminimum(IntegerOrder, A)
  [| IsBounded(A, IntegerOrder); A ≠ 0 |] ==> Maximum(IntegerOrder, A) ∈ A
  [| IsBounded(A, IntegerOrder); A ≠ 0 |] ==> Minimum(IntegerOrder, A) ∈ A
  [| IsBounded(A, IntegerOrder); A ≠ 0 |]
  ==> ∀xA. ⟨x, Maximum(IntegerOrder, A)⟩ ∈ IntegerOrder
  [| IsBounded(A, IntegerOrder); A ≠ 0 |]
  ==> ∀xA. ⟨Minimum(IntegerOrder, A), x⟩ ∈ IntegerOrder
  [| IsBounded(A, IntegerOrder); A ≠ 0 |] ==> Maximum(IntegerOrder, A) ∈ int
  [| IsBounded(A, IntegerOrder); A ≠ 0 |] ==> Minimum(IntegerOrder, A) ∈ int

theorem int_bounded_below_has_min:

  [| IsBoundedBelow(A, IntegerOrder); A ≠ 0 |] ==> HasAminimum(IntegerOrder, A)
  [| IsBoundedBelow(A, IntegerOrder); A ≠ 0 |] ==> Minimum(IntegerOrder, A) ∈ A
  [| IsBoundedBelow(A, IntegerOrder); A ≠ 0 |]
  ==> ∀xA. ⟨Minimum(IntegerOrder, A), x⟩ ∈ IntegerOrder

theorem int_bounded_above_has_max:

  [| IsBoundedAbove(A, IntegerOrder); A ≠ 0 |] ==> HasAmaximum(IntegerOrder, A)
  [| IsBoundedAbove(A, IntegerOrder); A ≠ 0 |] ==> Maximum(IntegerOrder, A) ∈ A
  [| IsBoundedAbove(A, IntegerOrder); A ≠ 0 |] ==> Maximum(IntegerOrder, A) ∈ int
  [| IsBoundedAbove(A, IntegerOrder); A ≠ 0 |]
  ==> ∀xA. ⟨x, Maximum(IntegerOrder, A)⟩ ∈ IntegerOrder

lemma Int_ZF_1_4_L1:

  [| IsBounded(A, IntegerOrder); A ≠ 0; ∀q∈int. F(q) ∈ int; K = {F(q) . qA} |]
  ==> HasAmaximum(IntegerOrder, K)
  [| IsBounded(A, IntegerOrder); A ≠ 0; ∀q∈int. F(q) ∈ int; K = {F(q) . qA} |]
  ==> HasAminimum(IntegerOrder, K)
  [| IsBounded(A, IntegerOrder); A ≠ 0; ∀q∈int. F(q) ∈ int; K = {F(q) . qA} |]
  ==> Maximum(IntegerOrder, K) ∈ K
  [| IsBounded(A, IntegerOrder); A ≠ 0; ∀q∈int. F(q) ∈ int; K = {F(q) . qA} |]
  ==> Minimum(IntegerOrder, K) ∈ K
  [| IsBounded(A, IntegerOrder); A ≠ 0; ∀q∈int. F(q) ∈ int; K = {F(q) . qA} |]
  ==> Maximum(IntegerOrder, K) ∈ int
  [| IsBounded(A, IntegerOrder); A ≠ 0; ∀q∈int. F(q) ∈ int; K = {F(q) . qA} |]
  ==> Minimum(IntegerOrder, K) ∈ int
  [| IsBounded(A, IntegerOrder); A ≠ 0; ∀q∈int. F(q) ∈ int; K = {F(q) . qA} |]
  ==> ∀qA. ⟨F(q), Maximum(IntegerOrder, K)⟩ ∈ IntegerOrder
  [| IsBounded(A, IntegerOrder); A ≠ 0; ∀q∈int. F(q) ∈ int; K = {F(q) . qA} |]
  ==> ∀qA. ⟨Minimum(IntegerOrder, K), F(q)⟩ ∈ IntegerOrder
  [| IsBounded(A, IntegerOrder); A ≠ 0; ∀q∈int. F(q) ∈ int; K = {F(q) . qA} |]
  ==> IsBounded(K, IntegerOrder)

lemma Int_ZF_1_4_L1A:

  [| a ∈ int; b ∈ int; c ∈ int |] ==> Maximum(IntegerOrder, {a, b, c}) ∈ int
  [| a ∈ int; b ∈ int; c ∈ int |]
  ==> ⟨a, Maximum(IntegerOrder, {a, b, c})⟩ ∈ IntegerOrder
  [| a ∈ int; b ∈ int; c ∈ int |]
  ==> ⟨b, Maximum(IntegerOrder, {a, b, c})⟩ ∈ IntegerOrder
  [| a ∈ int; b ∈ int; c ∈ int |]
  ==> ⟨c, Maximum(IntegerOrder, {a, b, c})⟩ ∈ IntegerOrder

lemma Int_ZF_1_4_L2:

  [| f ∈ int -> int; ⟨a, b⟩ ∈ IntegerOrder |]
  ==> Maximum(IntegerOrder, f `` Interval(IntegerOrder, a, b)) ∈ int
  [| f ∈ int -> int; ⟨a, b⟩ ∈ IntegerOrder |]
  ==> ∀c∈Interval(IntegerOrder, a, b).
         ⟨f ` c, Maximum(IntegerOrder, f `` Interval(IntegerOrder, a, b))⟩ ∈
         IntegerOrder
  [| f ∈ int -> int; ⟨a, b⟩ ∈ IntegerOrder |]
  ==> ∃c∈Interval(IntegerOrder, a, b).
         f ` c = Maximum(IntegerOrder, f `` Interval(IntegerOrder, a, b))
  [| f ∈ int -> int; ⟨a, b⟩ ∈ IntegerOrder |]
  ==> Minimum(IntegerOrder, f `` Interval(IntegerOrder, a, b)) ∈ int
  [| f ∈ int -> int; ⟨a, b⟩ ∈ IntegerOrder |]
  ==> ∀c∈Interval(IntegerOrder, a, b).
         ⟨Minimum(IntegerOrder, f `` Interval(IntegerOrder, a, b)), f ` c⟩ ∈
         IntegerOrder
  [| f ∈ int -> int; ⟨a, b⟩ ∈ IntegerOrder |]
  ==> ∃c∈Interval(IntegerOrder, a, b).
         f ` c = Minimum(IntegerOrder, f `` Interval(IntegerOrder, a, b))

The set of nonnegative integers

lemma pos_int_closed_add:

  PositiveSet(int, IntegerAddition, IntegerOrder) {is closed under}
  IntegerAddition

lemma pos_int_closed_add_unfolded:

  [| a ∈ PositiveSet(int, IntegerAddition, IntegerOrder);
     b ∈ PositiveSet(int, IntegerAddition, IntegerOrder) |]
  ==> IntegerAddition ` ⟨a, b⟩ ∈ PositiveSet(int, IntegerAddition, IntegerOrder)

lemma Int_ZF_1_5_L1:

  IsBoundedBelow(Nonnegative(int, IntegerAddition, IntegerOrder), IntegerOrder)
  IsBoundedBelow(PositiveSet(int, IntegerAddition, IntegerOrder), IntegerOrder)

lemma Int_ZF_1_5_L1A:

  A ⊆ Nonnegative(int, IntegerAddition, IntegerOrder)
  ==> IsBoundedBelow(A, IntegerOrder)

lemma Int_ZF_1_5_L1B:

  A ⊆ PositiveSet(int, IntegerAddition, IntegerOrder)
  ==> IsBoundedBelow(A, IntegerOrder)

lemma Int_ZF_1_5_L1C:

  [| A ⊆ PositiveSet(int, IntegerAddition, IntegerOrder); A ≠ 0 |]
  ==> HasAminimum(IntegerOrder, A)
  [| A ⊆ PositiveSet(int, IntegerAddition, IntegerOrder); A ≠ 0 |]
  ==> Minimum(IntegerOrder, A) ∈ A
  [| A ⊆ PositiveSet(int, IntegerAddition, IntegerOrder); A ≠ 0 |]
  ==> ∀xA. ⟨Minimum(IntegerOrder, A), x⟩ ∈ IntegerOrder

lemma Int_ZF_1_5_L2:

  [| A ⊆ Nonnegative(int, IntegerAddition, IntegerOrder); A ∉ Fin(int); D ∈ int |]
  ==> ∃nA. ⟨D, n⟩ ∈ IntegerOrder

lemma Int_ZF_1_5_L2A:

  [| A ⊆ PositiveSet(int, IntegerAddition, IntegerOrder); A ∉ Fin(int); D ∈ int |]
  ==> ∃nA. ⟨D, n⟩ ∈ IntegerOrder

lemma Int_decomp:

  m ∈ int
  ==> Exactly_1_of_3_holds
       (m = TheNeutralElement(int, IntegerAddition),
        m ∈ PositiveSet(int, IntegerAddition, IntegerOrder),
        GroupInv(int, IntegerAddition) ` m ∈
        PositiveSet(int, IntegerAddition, IntegerOrder))

lemma int_decomp_cases:

  m ∈ int
  ==> m = TheNeutralElement(int, IntegerAddition) ∨
      m ∈ PositiveSet(int, IntegerAddition, IntegerOrder) ∨
      GroupInv(int, IntegerAddition) ` m ∈
      PositiveSet(int, IntegerAddition, IntegerOrder)

lemma Int_ZF_1_5_L3:

  m ∈ PositiveSet(int, IntegerAddition, IntegerOrder) <->
  ⟨TheNeutralElement(int, IntegerMultiplication), m⟩ ∈ IntegerOrder

lemma pos_int_closed_mul_unfold:

  [| a ∈ PositiveSet(int, IntegerAddition, IntegerOrder);
     b ∈ PositiveSet(int, IntegerAddition, IntegerOrder) |]
  ==> IntegerMultiplication ` ⟨a, b⟩ ∈
      PositiveSet(int, IntegerAddition, IntegerOrder)

lemma pos_int_closed_mul:

  PositiveSet(int, IntegerAddition, IntegerOrder) {is closed under}
  IntegerMultiplication

lemma int_has_no_zero_divs:

  HasNoZeroDivs(int, IntegerAddition, IntegerMultiplication)

lemma Int_ZF_1_5_L3A:

  Nonnegative(int, IntegerAddition, IntegerOrder) =
  PositiveSet(int, IntegerAddition, IntegerOrder) ∪
  {TheNeutralElement(int, IntegerAddition)}

lemma Int_ZF_1_5_L4:

  [| f ∈ int -> int; K ∈ int; N ∈ int |]
  ==> ∃C∈int.
         ∀n∈PositiveSet(int, IntegerAddition, IntegerOrder).
            ⟨K, IntegerAddition ` ⟨f ` n, C⟩⟩ ∈ IntegerOrder -->
            ⟨N, n⟩ ∈ IntegerOrder

lemma Int_ZF_1_5_L4A:

  a ∈ PositiveSet(int, IntegerAddition, IntegerOrder)
  ==> AbsoluteValue(int, IntegerAddition, IntegerOrder) ` a = a

lemma int_one_two_are_pos:

  TheNeutralElement(int, IntegerMultiplication) ∈
  PositiveSet(int, IntegerAddition, IntegerOrder)
  IntegerAddition `
  ⟨TheNeutralElement(int, IntegerMultiplication),
   TheNeutralElement(int, IntegerMultiplication)⟩ ∈
  PositiveSet(int, IntegerAddition, IntegerOrder)

lemma Int_ZF_1_5_L5:

  f ∈ int -> X ==> f `` PositiveSet(int, IntegerAddition, IntegerOrder) ≠ 0

lemma Int_ZF_1_5_L6:

  n ∈ PositiveSet(int, IntegerAddition, IntegerOrder)
  ==> ⟨TheNeutralElement(int, IntegerAddition),
       IntegerAddition `
       ⟨n, GroupInv(int, IntegerAddition) `
           TheNeutralElement(int, IntegerMultiplication)⟩⟩ ∈
      IntegerOrder
  n ∈ PositiveSet(int, IntegerAddition, IntegerOrder)
  ==> TheNeutralElement(int, IntegerAddition) ∈
      Interval
       (IntegerOrder, TheNeutralElement(int, IntegerAddition),
        IntegerAddition `
        ⟨n, GroupInv(int, IntegerAddition) `
            TheNeutralElement(int, IntegerMultiplication)⟩)
  n ∈ PositiveSet(int, IntegerAddition, IntegerOrder)
  ==> Interval
       (IntegerOrder, TheNeutralElement(int, IntegerAddition),
        IntegerAddition `
        ⟨n, GroupInv(int, IntegerAddition) `
            TheNeutralElement(int, IntegerMultiplication)⟩) ⊆
      int

lemma Int_ZF_1_5_L7:

  [| a ∈ PositiveSet(int, IntegerAddition, IntegerOrder); ⟨a, b⟩ ∈ IntegerOrder |]
  ==> b ∈ PositiveSet(int, IntegerAddition, IntegerOrder)

lemma Int_ZF_1_5_L7A:

  [| a ∈ int; b ∈ PositiveSet(int, IntegerAddition, IntegerOrder) |]
  ==> ⟨a, IntegerAddition ` ⟨a, b⟩⟩ ∈ IntegerOrder
  [| a ∈ int; b ∈ PositiveSet(int, IntegerAddition, IntegerOrder) |]
  ==> a ≠ IntegerAddition ` ⟨a, b
  [| a ∈ int; b ∈ PositiveSet(int, IntegerAddition, IntegerOrder) |]
  ==> IntegerAddition ` ⟨a, b⟩ ∈ int

lemma Int_ZF_1_5_L7B:

  a ∈ int
  ==> ⟨a, GreaterOf
           (IntegerOrder, TheNeutralElement(int, IntegerMultiplication), a)⟩ ∈
      IntegerOrder
  a ∈ int
  ==> GreaterOf(IntegerOrder, TheNeutralElement(int, IntegerMultiplication), a) ∈
      PositiveSet(int, IntegerAddition, IntegerOrder)
  a ∈ int
  ==> IntegerAddition `
      ⟨GreaterOf(IntegerOrder, TheNeutralElement(int, IntegerMultiplication), a),
       TheNeutralElement(int, IntegerMultiplication)⟩ ∈
      PositiveSet(int, IntegerAddition, IntegerOrder)
  a ∈ int
  ==> ⟨a, IntegerAddition `
          ⟨GreaterOf
            (IntegerOrder, TheNeutralElement(int, IntegerMultiplication), a),
           TheNeutralElement(int, IntegerMultiplication)⟩⟩ ∈
      IntegerOrder
  a ∈ int
  ==> a ≠ IntegerAddition `
          ⟨GreaterOf
            (IntegerOrder, TheNeutralElement(int, IntegerMultiplication), a),
           TheNeutralElement(int, IntegerMultiplication)⟩

lemma Int_ZF_1_5_L8:

  a ∈ PositiveSet(int, IntegerAddition, IntegerOrder)
  ==> GroupInv(int, IntegerAddition) ` a ∉
      PositiveSet(int, IntegerAddition, IntegerOrder)

lemma Int_ZF_1_5_L9:

  a ∈ int
  ==> ∃b∈PositiveSet(int, IntegerAddition, IntegerOrder). ⟨a, b⟩ ∈ IntegerOrder

lemma Int_ZF_1_5_L10:

  f ∈ PositiveSet(int, IntegerAddition, IntegerOrder) -> int
  ==> OddExtension(int, IntegerAddition, IntegerOrder, f) ∈ int -> int

lemma Int_ZF_1_5_L11:

  [| f ∈ PositiveSet(int, IntegerAddition, IntegerOrder) -> int;
     a ∈ PositiveSet(int, IntegerAddition, IntegerOrder);
     g = OddExtension(int, IntegerAddition, IntegerOrder, f) |]
  ==> g ` a = f ` a

lemma Int_ZF_1_5_L12:

  [| f ∈ PositiveSet(int, IntegerAddition, IntegerOrder) -> int;
     a ∈ GroupInv(int, IntegerAddition) ``
         PositiveSet(int, IntegerAddition, IntegerOrder);
     g = OddExtension(int, IntegerAddition, IntegerOrder, f) |]
  ==> g ` a =
      GroupInv(int, IntegerAddition) ` (f ` (GroupInv(int, IntegerAddition) ` a))

lemma int_oddext_is_odd:

  [| f ∈ PositiveSet(int, IntegerAddition, IntegerOrder) -> int; a ∈ int;
     g = OddExtension(int, IntegerAddition, IntegerOrder, f) |]
  ==> g ` (GroupInv(int, IntegerAddition) ` a) =
      GroupInv(int, IntegerAddition) ` (g ` a)

lemma Int_ZF_1_5_L13:

  f ∈ int -> int
  ==> (∀a∈int.
          f ` (GroupInv(int, IntegerAddition) ` a) =
          GroupInv(int, IntegerAddition) ` (f ` a)) <->
      (∀a∈int.
          GroupInv(int, IntegerAddition) `
          (f ` (GroupInv(int, IntegerAddition) ` a)) =
          f ` a)

lemma int_oddext_is_odd_alt:

  [| f ∈ PositiveSet(int, IntegerAddition, IntegerOrder) -> int; a ∈ int;
     g = OddExtension(int, IntegerAddition, IntegerOrder, f) |]
  ==> GroupInv(int, IntegerAddition) `
      (g ` (GroupInv(int, IntegerAddition) ` a)) =
      g ` a

Functions with infinite limits

lemma Int_ZF_1_6_L1:

  [| f ∈ int -> int;
     ∀a∈int.
        ∃b∈PositiveSet(int, IntegerAddition, IntegerOrder).
           ∀x. ⟨b, x⟩ ∈ IntegerOrder --> ⟨a, f ` x⟩ ∈ IntegerOrder;
     A ⊆ int; IsBoundedAbove(f `` A, IntegerOrder) |]
  ==> IsBoundedAbove(A, IntegerOrder)

lemma Int_ZF_1_6_L2:

  [| X ≠ 0; f ∈ int -> int;
     ∀a∈int.
        ∃b∈PositiveSet(int, IntegerAddition, IntegerOrder).
           ∀x. ⟨b, x⟩ ∈ IntegerOrder --> ⟨a, f ` x⟩ ∈ IntegerOrder;
     ∀xX. b(x) ∈ int ∧ ⟨f ` b(x), U⟩ ∈ IntegerOrder |]
  ==> ∃u. ∀xX. ⟨b(x), u⟩ ∈ IntegerOrder

lemma Int_ZF_1_6_L3:

  [| X ≠ 0; f ∈ int -> int;
     ∀a∈int.
        ∃b∈PositiveSet(int, IntegerAddition, IntegerOrder).
           ∀y. ⟨b, y⟩ ∈ IntegerOrder -->
               ⟨f ` (GroupInv(int, IntegerAddition) ` y), a⟩ ∈ IntegerOrder;
     ∀xX. b(x) ∈ int ∧ ⟨L, f ` b(x)⟩ ∈ IntegerOrder |]
  ==> ∃l. ∀xX. ⟨l, b(x)⟩ ∈ IntegerOrder

lemma Int_ZF_1_6_L4:

  [| X ≠ 0; f ∈ int -> int;
     ∀a∈int.
        ∃b∈PositiveSet(int, IntegerAddition, IntegerOrder).
           ∀x. ⟨b, x⟩ ∈ IntegerOrder --> ⟨a, f ` x⟩ ∈ IntegerOrder;
     ∀a∈int.
        ∃b∈PositiveSet(int, IntegerAddition, IntegerOrder).
           ∀y. ⟨b, y⟩ ∈ IntegerOrder -->
               ⟨f ` (GroupInv(int, IntegerAddition) ` y), a⟩ ∈ IntegerOrder;
     ∀xX. b(x) ∈ int ∧
           ⟨f ` b(x), U⟩ ∈ IntegerOrder ∧ ⟨L, f ` b(x)⟩ ∈ IntegerOrder |]
  ==> ∃M. ∀xX. ⟨AbsoluteValue(int, IntegerAddition, IntegerOrder) ` b(x), M⟩ ∈
                IntegerOrder

lemma Int_ZF_1_6_L5:

  [| f ∈ int -> int; N ∈ int;
     ∀m. ⟨N, m⟩ ∈ IntegerOrder --> ⟨L, f ` m⟩ ∈ IntegerOrder;
     IsBoundedBelow(A, IntegerOrder) |]
  ==> IsBoundedBelow(f `` A, IntegerOrder)

lemma Int_ZF_1_6_L6:

  [| N ∈ int; ∀m. ⟨N, m⟩ ∈ IntegerOrder --> ⟨L, f ` m⟩ ∈ IntegerOrder;
     f ∈ int -> int; K ∈ int |]
  ==> ∃c∈int.
         ∀n∈PositiveSet(int, IntegerAddition, IntegerOrder).
            ⟨K, IntegerAddition ` ⟨f ` n, c⟩⟩ ∈ IntegerOrder

lemma Int_ZF_1_6_L7:

  [| f ∈ int -> int; K ∈ int; N ∈ int;
     ∀a∈int.
        ∃b∈PositiveSet(int, IntegerAddition, IntegerOrder).
           ∀x. ⟨b, x⟩ ∈ IntegerOrder --> ⟨a, f ` x⟩ ∈ IntegerOrder |]
  ==> ∃C∈int.
         ⟨N, Minimum
              (IntegerOrder,
               {n ∈ PositiveSet(int, IntegerAddition, IntegerOrder) .
                ⟨K, IntegerAddition ` ⟨f ` n, C⟩⟩ ∈ IntegerOrder})⟩ ∈
         IntegerOrder

lemma Int_ZF_1_6_L8:

  [| a ∈ int; {IntegerMultiplication ` ⟨a, x⟩ . x ∈ int} ∈ Fin(int) |]
  ==> a = TheNeutralElement(int, IntegerAddition)

Miscelaneous

lemma Int_ZF_1_7_L1:

  [| ∀q∈int. F(q) ∈ int;
     ∀q∈int.
        ⟨IntegerMultiplication `
         ⟨F(q), AbsoluteValue(int, IntegerAddition, IntegerOrder) ` q⟩,
         IntegerAddition `
         ⟨IntegerMultiplication `
          ⟨A, AbsoluteValue(int, IntegerAddition, IntegerOrder) ` q⟩,
          B⟩⟩ ∈
        IntegerOrder;
     A ∈ int; B ∈ int |]
  ==> ∃L. ∀p∈int. ⟨F(p), L⟩ ∈ IntegerOrder

lemma int_plane_split_in6:

  [| a ∈ int; b ∈ int |]
  ==> ⟨TheNeutralElement(int, IntegerAddition), a⟩ ∈ IntegerOrder ∧
      ⟨TheNeutralElement(int, IntegerAddition), b⟩ ∈ IntegerOrder ∨
      ⟨a, TheNeutralElement(int, IntegerAddition)⟩ ∈ IntegerOrder ∧
      ⟨b, TheNeutralElement(int, IntegerAddition)⟩ ∈ IntegerOrder ∨
      ⟨a, TheNeutralElement(int, IntegerAddition)⟩ ∈ IntegerOrder ∧
      ⟨TheNeutralElement(int, IntegerAddition), b⟩ ∈ IntegerOrder ∧
      ⟨TheNeutralElement(int, IntegerAddition), IntegerAddition ` ⟨a, b⟩⟩ ∈
      IntegerOrder ∨
      ⟨a, TheNeutralElement(int, IntegerAddition)⟩ ∈ IntegerOrder ∧
      ⟨TheNeutralElement(int, IntegerAddition), b⟩ ∈ IntegerOrder ∧
      ⟨IntegerAddition ` ⟨a, b⟩, TheNeutralElement(int, IntegerAddition)⟩ ∈
      IntegerOrder ∨
      ⟨TheNeutralElement(int, IntegerAddition), a⟩ ∈ IntegerOrder ∧
      ⟨b, TheNeutralElement(int, IntegerAddition)⟩ ∈ IntegerOrder ∧
      ⟨TheNeutralElement(int, IntegerAddition), IntegerAddition ` ⟨a, b⟩⟩ ∈
      IntegerOrder ∨
      ⟨TheNeutralElement(int, IntegerAddition), a⟩ ∈ IntegerOrder ∧
      ⟨b, TheNeutralElement(int, IntegerAddition)⟩ ∈ IntegerOrder ∧
      ⟨IntegerAddition ` ⟨a, b⟩, TheNeutralElement(int, IntegerAddition)⟩ ∈
      IntegerOrder