section ‹Integers 1›
theory Int_ZF_1 imports Int_ZF_IML OrderedRing_ZF
begin
text‹This theory file considers the set of integers as an ordered ring.›
subsection‹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∈ℤ" "b∈ℤ" "c∈ℤ"
shows
"a⋅(b\<ra>c) = a⋅b \<ra> a⋅c"
"(b\<ra>c)⋅a = b⋅a \<ra> c⋅a"
using assms 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 ‹ring0› context (locale).›
lemma (in int0) Int_ZF_1_1_L2: shows
"IsAring(ℤ,IntegerAddition,IntegerMultiplication)"
"IntegerMultiplication {is commutative on} ℤ"
"ring0(ℤ,IntegerAddition,IntegerMultiplication)"
proof -
have "∀a∈ℤ.∀b∈ℤ.∀c∈ℤ.
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(ℤ,IntegerAddition,IntegerMultiplication)"
using IsDistributive_def by simp
then show "IsAring(ℤ,IntegerAddition,IntegerMultiplication)"
"ring0(ℤ,IntegerAddition,IntegerMultiplication)"
using Int_ZF_1_T1 Int_ZF_1_T2 IsAring_def ring0_def
by auto
have "∀a∈ℤ.∀b∈ℤ. a⋅b = b⋅a" using Int_ZF_1_L4 by simp
then show "IntegerMultiplication {is commutative on} ℤ"
using IsCommutative_def by simp
qed
text‹Zero and one are integers.›
lemma (in int0) int_zero_one_are_int: shows "𝟬∈ℤ" "𝟭∈ℤ"
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>𝟬) = 𝟬"
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 ∈ ℤ"
shows
"a\<ra>𝟬 = a"
"𝟬\<ra>a = a"
"a⋅𝟭 = a" "𝟭⋅a = a"
"𝟬⋅a = 𝟬" "a⋅𝟬 = 𝟬"
"(\<rm>a) ∈ ℤ" "(\<rm>(\<rm>a)) = a"
"a\<rs>a = 𝟬" "a\<rs>𝟬 = a" "𝟮⋅a = a\<ra>a"
proof -
from A1 show
"a\<ra>𝟬 = a" "𝟬\<ra>a = a" "a⋅𝟭 = a"
"𝟭⋅a = a" "a\<rs>a = 𝟬" "a\<rs>𝟬 = a"
"(\<rm>a) ∈ ℤ" "𝟮⋅a = a\<ra>a" "(\<rm>(\<rm>a)) = a"
using Int_ZF_1_1_L2 ring0.Ring_ZF_1_L3 by auto
from A1 show "𝟬⋅a = 𝟬" "a⋅𝟬 = 𝟬"
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 "a∈ℤ" "b∈ℤ"
shows
"a\<ra>b ∈ ℤ"
"a\<rs>b ∈ ℤ"
"a⋅b ∈ ℤ"
"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 assms 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 "𝟮 ∈ ℤ" "𝟯 ∈ ℤ"
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 "a∈ℤ" "b∈ℤ"
shows "a\<rs>(\<rm>b) = a\<ra>b"
using assms 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 "a∈ℤ" "b∈ℤ" "c∈ℤ"
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 assms Int_ZF_1_1_L2 ring0.Ring_ZF_1_L10 ring0.Ring_ZF_1_L8
by auto
text‹One more property with three integers.›
lemma (in int0) Int_ZF_1_1_L6A: assumes "a∈ℤ" "b∈ℤ" "c∈ℤ"
shows "a\<ra>(b\<rs>c) = a\<ra>b\<rs>c"
using assms 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 "a∈ℤ" "b∈ℤ" "c∈ℤ"
shows
"a\<ra>b\<ra>c = a\<ra>(b\<ra>c)"
"a⋅b⋅c = a⋅(b⋅c)"
using assms Int_ZF_1_1_L2 ring0.Ring_ZF_1_L11 by auto
subsection‹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 "𝟬\<lsq>a"
shows "abs(a)\<ra>𝟭 = abs(a\<ra>𝟭)"
using assms 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∈ℤ" and A2: "𝟬\<lsq>b"
shows "a\<ra>(abs(b)\<ra>𝟭)⋅a = (abs(b\<ra>𝟭)\<ra>𝟭)⋅a"
proof -
from A2 have "abs(b\<ra>𝟭) ∈ ℤ"
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∈ℤ" "b∈ℤ"
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 assms Int_ZF_1_T2 group0.group0_4_L6A group0.inv_cancel_two
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>𝟭 \<lsq> b"
proof -
from A1 have "b\<ra>𝟭\<rs>𝟭 = b"
using Int_ZF_2_L1A int_zero_one_are_int Int_ZF_1_2_L3 by simp
moreover from A1 have "a\<rs>𝟭 \<lsq> b\<ra>𝟭\<rs>𝟭"
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>𝟭 \<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∈ℤ" shows
"a\<rs>𝟭 \<lsq>a"
"a\<rs>𝟭 ≠ a"
"¬(a\<lsq>a\<rs>𝟭)"
"¬(a\<ra>𝟭 \<lsq>a)"
"¬(𝟭\<ra>a \<lsq>a)"
proof -
from A1 have "a\<lsq>a" using int_ord_is_refl refl_def
by simp
then show "a\<rs>𝟭 \<lsq>a" using Int_ZF_1_2_L3A
by simp
moreover from A1 show "a\<rs>𝟭 ≠ a" using Int_ZF_1_L14 by simp
ultimately show I: "¬(a\<lsq>a\<rs>𝟭)" using Int_ZF_2_L19AA
by blast
with A1 show "¬(a\<ra>𝟭 \<lsq>a)"
using int_zero_one_are_int Int_ZF_2_L9B by simp
with A1 show "¬(𝟭\<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>𝟬"
shows "abs(a)\<ra>𝟭 = abs(a\<rs>𝟭)"
using assms 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∈ℤ" and A2: "b\<lsq>𝟬"
shows "a\<ra>(abs(b)\<ra>𝟭)⋅a = (abs(b\<rs>𝟭)\<ra>𝟭)⋅a"
proof -
from A2 have "abs(b\<rs>𝟭) ∈ ℤ"
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∈ℤ" "b∈ℤ" "c∈ℤ" "d∈ℤ"
shows
"a\<rs>(b\<rs>𝟭)⋅c = (d\<rs>b⋅c)\<rs>(d\<rs>a\<rs>c)"
proof -
from A1 have T1:
"(d\<rs>b⋅c) ∈ ℤ" "d\<rs>a ∈ ℤ" "(\<rm>(b⋅c)) ∈ ℤ"
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>𝟭)⋅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∈ℤ" "b∈ℤ"
shows
"a⋅b = (a\<rs>𝟭)⋅b\<ra>b"
"a⋅(b\<ra>𝟭) = a⋅b\<ra>a"
"(b\<ra>𝟭)⋅a = b⋅a\<ra>a"
"(b\<ra>𝟭)⋅a = a\<ra>b⋅a"
using assms 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.inv_cancel_two
by auto
text‹Another rearrangement with two integers.›
lemma (in int0) Int_ZF_1_2_L8:
assumes A1: "a∈ℤ" "b∈ℤ"
shows "a\<ra>𝟭\<ra>(b\<ra>𝟭) = b\<ra>a\<ra>𝟮"
using assms 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∈ℤ" "b∈ℤ" "c∈ℤ"
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 assms 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∈ℤ" "b∈ℤ" "c∈ℤ"
shows "(\<rm>(a\<rs>b\<rs>c)) = c\<ra>b\<rs>a"
proof -
from A1 have T:
"a\<rs>b ∈ ℤ" "(\<rm>(a\<rs>b)) ∈ ℤ" "(\<rm>b) ∈ ℤ" 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∈ℤ" "b∈ℤ" "c∈ℤ"
shows "(a\<ra>𝟭)⋅b \<ra> (c\<ra>𝟭)⋅b = (c\<ra>a\<ra>𝟮)⋅b"
proof -
from A1 have "a\<ra>𝟭 ∈ ℤ" "c\<ra>𝟭 ∈ ℤ"
using int_zero_one_are_int Int_ZF_1_1_L5 by auto
with A1 have
"(a\<ra>𝟭)⋅b \<ra> (c\<ra>𝟭)⋅b = (a\<ra>𝟭\<ra>(c\<ra>𝟭))⋅b"
using Int_ZF_1_1_L1 by simp
also from A1 have "… = (c\<ra>a\<ra>𝟮)⋅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∈ℤ" "b∈ℤ" "c∈ℤ" "e∈ℤ"
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∈ℤ" "f∈ℤ" "a⋅b ∈ ℤ" "a⋅b\<rs>c ∈ ℤ" "b⋅a\<rs>e ∈ ℤ"
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∈ℤ"
shows
"a\<ra>𝟭\<ra>𝟮 = a\<ra>𝟯"
"a = 𝟮⋅a \<rs> a"
proof -
from A1 show "a\<ra>𝟭\<ra>𝟮 = a\<ra>𝟯"
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 = 𝟮⋅a \<rs> a"
using int_zero_one_are_int Int_ZF_1_1_L1 Int_ZF_1_1_L4 Int_ZF_1_T2 group0.inv_cancel_two
by simp
qed
text‹A simple rearrangement with three integers.›
lemma (in int0) Int_ZF_1_2_L12:
assumes "a∈ℤ" "b∈ℤ" "c∈ℤ"
shows
"(b\<rs>c)⋅a = a⋅b \<rs> a⋅c"
using assms 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∈ℤ" "b∈ℤ" "c∈ℤ" "d∈ℤ" "x∈ℤ"
shows "(x\<ra>(a⋅x\<ra>b)\<ra>c)⋅d = d⋅(a\<ra>𝟭)⋅x \<ra> (b⋅d\<ra>c⋅d)"
proof -
from A1 have T1:
"a⋅x ∈ ℤ" "(a\<ra>𝟭)⋅x ∈ ℤ"
"(a\<ra>𝟭)⋅x \<ra> b ∈ ℤ"
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>𝟭)⋅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>𝟭)⋅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>𝟭)⋅x⋅d \<ra> b⋅d \<ra> c⋅d"
by simp
moreover from A1 T1 have "(a\<ra>𝟭)⋅x⋅d = d⋅(a\<ra>𝟭)⋅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>𝟭)⋅x \<ra> b⋅d \<ra> c⋅d"
by simp
moreover from A1 T1 have
"d⋅(a\<ra>𝟭)⋅x ∈ ℤ" "b⋅d ∈ ℤ" "c⋅d ∈ ℤ"
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∈ℤ" "b∈ℤ" "c∈ℤ" "d∈ℤ" "x∈ℤ"
shows "(a⋅x \<ra> b) \<ra> (c⋅x \<ra> d) = (a\<ra>c)⋅x \<ra> (b\<ra>d)"
using assms 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∈ℤ" "b∈ℤ" "c∈ℤ" "d∈ℤ"
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∈ℤ" "b∈ℤ" "c∈ℤ" "d∈ℤ"
shows "a\<ra>(b\<rs>c)\<ra>d = a\<ra>b\<ra>d\<rs>c"
using assms 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∈ℤ" "b∈ℤ" "c∈ℤ"
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∈ℤ" "b∈ℤ" "c∈ℤ"
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
subsection‹Integers as an ordered ring›
text‹We already know from ‹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.›
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: "𝟬\<lsq>a" "𝟬\<lsq>b"
and A3: "𝟬 \<lsq> a⋅b"
shows "𝟬 \<lsq> a⋅(b\<ra>𝟭)"
proof -
from A1 A3 have "𝟬\<ra>𝟬 \<lsq> a⋅b\<ra>a"
using int_ineq_add_sides by simp
with A1 show "𝟬 \<lsq> a⋅(b\<ra>𝟭)"
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: "𝟬\<lsq>a" "𝟬\<lsq>b"
shows "𝟬\<lsq>a⋅b"
proof -
from A1 have "𝟬\<lsq>b" by simp
moreover from A1 have "𝟬 \<lsq> a⋅𝟬" 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. 𝟬\<lsq>m ∧ 𝟬\<lsq>a⋅m ⟶ 𝟬 \<lsq> a⋅(m\<ra>𝟭)"
using Int_ZF_1_3_L1 by simp
ultimately show "𝟬\<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
"ℤ⇧+ {is closed under} IntegerMultiplication"
proof -
{ fix a b assume "a∈ℤ⇧+" "b∈ℤ⇧+"
then have "a⋅b ∈ℤ⇧+"
using Int_ZF_1_3_L2 Int_ZF_2_T1 group3.OrderedGroup_ZF_1_L2
by simp
} then have "∀a∈ℤ⇧+.∀b∈ℤ⇧+.a⋅b ∈ℤ⇧+" by simp
then show ?thesis using IsOpClosed_def by simp
qed
text‹Integers form an ordered ring. All theorems proven in the
‹ring1› context are valid in ‹int0› context.›
theorem (in int0) Int_ZF_1_3_T1: shows
"IsAnOrdRing(ℤ,IntegerAddition,IntegerMultiplication,IntegerOrder)"
"ring1(ℤ,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: "𝟭\<lsq>a" "𝟭\<lsq>b"
and A2: "𝟭 \<lsq> a⋅b"
shows "𝟭 \<lsq> a⋅(b\<ra>𝟭)"
proof -
from A1 A2 have "𝟭\<lsq>𝟮" and "𝟮 \<lsq> a⋅(b\<ra>𝟭)"
using Int_ZF_2_L1A int_ineq_add_sides Int_ZF_2_L16B Int_ZF_1_2_L7
by auto
then show "𝟭 \<lsq> a⋅(b\<ra>𝟭)" 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: "𝟭\<lsq>a" "𝟭\<lsq>b"
shows "𝟭 \<lsq> a⋅b"
proof -
from A1 have "𝟭\<lsq>b" "𝟭\<lsq>a⋅𝟭"
using Int_ZF_2_L1A Int_ZF_1_1_L4 by auto
moreover from A1 have
"∀m. 𝟭\<lsq>m ∧ 𝟭 \<lsq> a⋅m ⟶ 𝟭 \<lsq> a⋅(m\<ra>𝟭)"
using Int_ZF_1_3_L3_indstep by simp
ultimately show "𝟭 \<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∈ℤ" "b∈ℤ"
shows
"abs((\<rm>a)⋅b) = abs(a⋅b)"
"abs(a⋅(\<rm>b)) = abs(a⋅b)"
"abs((\<rm>a)⋅(\<rm>b)) = abs(a⋅b)"
using assms 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∈ℤ" "b∈ℤ"
shows "abs(a⋅b) = abs(a)⋅abs(b)"
using assms 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 "𝟬\<lsq>a"
shows "𝟬\<lsq>𝟮⋅a"
using assms 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∈ℤ" "b∈ℤ"
shows "¬(b\<lsq>a) ⟷ a\<ra>𝟭 \<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>𝟭 \<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>𝟭 \<lsq> b" by simp
next assume A4: "a\<ra>𝟭 \<lsq> b"
{ assume "b\<lsq>a"
with A4 have "a\<ra>𝟭 \<lsq> a" by (rule Int_order_transitive)
moreover from A1 have "a \<lsq> a\<ra>𝟭"
using Int_ZF_2_L12B by simp
ultimately have "a\<ra>𝟭 = 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∈ℤ" "b∈ℤ"
shows "b\<lsq>a ∨ a\<ra>𝟭 \<lsq> b"
using A1 Int_ZF_1_3_L6 by auto
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∈ℤ" "b∈ℤ" and A2: "¬(b\<lsq>a)"
shows "a \<lsq> b\<rs>𝟭"
proof -
from A1 A2 have "a\<ra>𝟭 \<rs> 𝟭 \<lsq> b \<rs> 𝟭"
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>𝟭"
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>𝟭 \<lsq> b"
"a \<lsq> b\<rs>𝟭"
proof -
from A1 have T: "a∈ℤ" "b∈ℤ" 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>𝟭 \<lsq> b"
"a \<lsq> b\<rs>𝟭"
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∈ℤ" "b∈ℤ"
shows "a=b ∨ (a \<lsq> b\<rs>𝟭) ∨ (b\<ra>𝟭 \<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 ‹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∈ℤ"
shows "a=𝟬 ∨ (a \<lsq> \<rm>𝟭) ∨ (𝟭\<lsq>a)"
proof -
from A1 have "a=𝟬 ∨ (a \<lsq> 𝟬 \<rs>𝟭) ∨ (𝟬 \<ra>𝟭 \<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∈ℤ"
shows "¬(a\<lsq>𝟬) ⟷ 𝟭 \<lsq> a"
using assms 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∈ℤ" "b∈ℤ"
and "¬(a\<lsq>𝟬)" "¬(b\<lsq>𝟬)"
shows "¬((a⋅b) \<lsq> 𝟬)"
using assms 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∈ℤ" "b∈ℤ"
and A2: "¬(b\<lsq>𝟬)" and A3: "a⋅b \<lsq> 𝟬"
shows "a\<lsq>𝟬"
proof -
{ assume "¬(a\<lsq>𝟬)"
with A1 A2 have "¬((a⋅b) \<lsq> 𝟬)" using Int_ZF_1_3_L8
by simp
} with A3 show "a\<lsq>𝟬" 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∈ℤ" "b∈ℤ"
shows "a\<lsq>b ⟷ a\<rs>b \<lsq> 𝟬"
using assms 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 "𝟬 \<lsq> b\<rs>a"
using assms 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∈ℤ" "b∈ℤ" "c∈ℤ"
and A2: "a⋅c \<lsq> b⋅c" and A4: "¬(c\<lsq>𝟬)"
shows "a \<lsq> b"
proof -
from A1 A4 have "a\<rs>b ∈ ℤ" "c∈ℤ" "¬(c\<lsq>𝟬)"
using Int_ZF_1_1_L5 by auto
moreover from A1 A2 have "(a\<rs>b)⋅c \<lsq> 𝟬"
using Int_ZF_1_1_L5 Int_ZF_1_3_L10 Int_ZF_1_1_L6
by simp
ultimately have "a\<rs>b \<lsq> 𝟬" 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∈ℤ" "b∈ℤ"
and A2: "¬(abs(a) \<lsq> abs(b))"
shows "¬(abs(a) \<lsq> 𝟬)"
proof -
{ assume "abs(a) \<lsq> 𝟬"
moreover from A1 have "𝟬 \<lsq> abs(a)" using int_abs_nonneg
by simp
ultimately have "abs(a) = 𝟬" by (rule Int_ZF_2_L3)
with A1 A2 have False using int_abs_nonneg by simp
} then show "¬(abs(a) \<lsq> 𝟬)" 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>𝟬" and "𝟬\<lsq>b"
shows "a⋅b \<lsq> 𝟬"
using assms 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: "𝟬\<lsq>c"
shows
"a⋅c \<lsq> b⋅c"
"c⋅a \<lsq> c⋅b"
using assms 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 "𝟭\<lsq>a" and "b\<lsq>c" and "(a\<ra>𝟭)⋅c \<lsq> d"
shows "(a\<ra>𝟭)⋅b \<lsq> d"
proof -
from assms have
"(a\<ra>𝟭)⋅b \<lsq> (a\<ra>𝟭)⋅c"
"(a\<ra>𝟭)⋅c \<lsq> d"
using Int_ZF_2_L16C Int_ZF_1_3_L13 by auto
then show "(a\<ra>𝟭)⋅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∈ℤ⇩+"
shows
"a⋅c \<lsq> b⋅c"
"c⋅a \<lsq> c⋅b"
proof -
let ?R = "ℤ"
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∈ℤ" "b∈ℤ" "c∈ℤ" "d∈ℤ"
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) ∈ ℤ" "abs(b) ∈ ℤ"
"abs(a)⋅abs(b) ∈ ℤ"
"abs(a)⋅d ∈ ℤ"
"c⋅d ∈ ℤ"
"abs(b)\<ra>d ∈ ℤ"
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∈ℤ" "n∈ℤ"
and A2: "¬(abs(m) \<lsq> abs(n))"
shows "n \<lsq> abs(m)" "m≠𝟬"
proof -
from A1 have T1: "n \<lsq> abs(n)"
using Int_ZF_2_L19C by simp
from A1 have "abs(n) ∈ ℤ" "abs(m) ∈ ℤ"
using Int_ZF_2_L14 by auto
moreover note A2
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≠𝟬" 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: "𝟬 \<lsq> m"
shows "(\<rm>m) \<lsq> 𝟬"
proof -
from A1 have "(\<rm>m) \<lsq> (\<rm>𝟬)"
using Int_ZF_2_L10 by simp
then show "(\<rm>m) \<lsq> 𝟬" 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∈ℤ"
shows
"(\<rm>abs(m)) \<lsq> abs(m)"
"(\<rm>abs(m))..abs(m) ≠ 0"
proof -
from A1 have "(\<rm>abs(m)) \<lsq> 𝟬" "𝟬 \<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∈ℤ" "n∈ℤ"
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 assms 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∈ℤ" and A2: "abs(m) \<lsq> n"
shows
"(\<rm>n) \<lsq> m" "m \<lsq> n"
"m ∈ (\<rm>n)..n"
"𝟬 \<lsq> n"
using assms 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∈ℤ" and A2: "abs(m) \<lsq> n" and A3: "𝟬\<lsq>k"
shows "(\<rm>(n\<ra>k)) \<lsq> m"
using assms 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) ∈ ℤ ∧ abs(b(x)) \<lsq> L"
shows "IsBounded({b(x). x∈X},IntegerOrder)"
proof -
let ?G = "ℤ"
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 assms 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∈ℤ} ∈ Fin(ℤ)"
shows "∃L∈ℤ. ∀x∈ℤ. abs(b(x)) \<lsq> L"
using assms 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→ℤ" and "A⊆X" and "∀x∈A. abs(f`(x)) \<lsq> L"
shows "IsBounded(f``(A),IntegerOrder)"
proof -
let ?G = "ℤ"
let ?P = "IntegerAddition"
let ?r = "IntegerOrder"
from assms 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:ℤ→ℤ" and "∀m∈ℤ. abs(f`(m)) \<lsq> L"
shows "f``(ℤ) ∈ Fin(ℤ)"
proof -
from assms have "f:ℤ→ℤ" "ℤ ⊆ ℤ" "∀m∈ℤ. abs(f`(m)) \<lsq> L"
by auto
then have "IsBounded(f``(ℤ),IntegerOrder)"
by (rule Int_ZF_1_3_L20B)
then show "f``(ℤ) ∈ Fin(ℤ)" 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∈ℤ" "b∈ℤ" "c∈ℤ"
shows "abs(a\<rs>b\<rs>c) \<lsq> abs(a) \<ra> abs(b) \<ra> abs(c)"
proof -
from A1 have T: "a\<rs>b ∈ ℤ" "abs(c) ∈ ℤ"
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> 𝟮⋅c"
using assms 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∈ℤ" and "b\<lsq> c\<ra>a"
shows "abs(b\<rs>a) \<lsq> c"
using assms 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∈ℤ" "b∈ℤ" "c∈ℤ" "d∈ℤ"
shows "abs(a\<rs>c) \<lsq> abs(a\<ra>b) \<ra> abs(c\<ra>d) \<ra> abs(b\<rs>d)"
using assms 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 ‹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∈ℤ" and A3: "b\<lsq> a\<ra>c"
shows "abs(b\<rs>a) \<lsq> c"
proof -
from A1 have "a ∈ ℤ"
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
subsection‹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(ℤ)" 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) ∈ ℤ"
"Minimum(IntegerOrder,A) ∈ ℤ"
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⊆ℤ" using FinD by simp
ultimately show
"Maximum(IntegerOrder,A) ∈ ℤ"
"Minimum(IntegerOrder,A) ∈ ℤ"
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) ∈ ℤ"
"Minimum(IntegerOrder,A) ∈ ℤ"
using assms 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"
"∀x∈A. Minimum(IntegerOrder,A) \<lsq> x"
proof -
from A1 A2 have
"IntegerOrder {is total on} ℤ"
"trans(IntegerOrder)"
"IntegerOrder ⊆ ℤ×ℤ"
"∀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)
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) ∈ ℤ"
"∀x∈A. x \<lsq> Maximum(IntegerOrder,A)"
proof -
from A1 A2 have
"IntegerOrder {is total on} ℤ"
"trans(IntegerOrder)" and
I: "IntegerOrder ⊆ ℤ×ℤ" 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 ⊆ ℤ" by (rule Order_ZF_3_L1A)
with II show "Maximum(IntegerOrder,A) ∈ ℤ" 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∈ℤ. F(q) ∈ ℤ"
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) ∈ ℤ"
"Minimum(IntegerOrder,K) ∈ ℤ"
"∀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(ℤ)" using Int_bounded_iff_fin
by simp
with A3 have "{F(q). q ∈ A} ∈ Fin(ℤ)"
by (rule fin_image_fin)
with A2 A4 have T1: "K ∈ Fin(ℤ)" "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) ∈ ℤ"
"Minimum(IntegerOrder,K) ∈ ℤ"
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∈ℤ" "b∈ℤ" "c∈ℤ"
shows
"Maximum(IntegerOrder,{a,b,c}) ∈ ℤ"
"a \<lsq> Maximum(IntegerOrder,{a,b,c})"
"b \<lsq> Maximum(IntegerOrder,{a,b,c})"
"c \<lsq> Maximum(IntegerOrder,{a,b,c})"
using assms 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:ℤ→ℤ" and A2: "a\<lsq>b"
shows
"maxf(f,a..b) ∈ ℤ"
"∀c ∈ a..b. f`(c) \<lsq> maxf(f,a..b)"
"∃c ∈ a..b. f`(c) = maxf(f,a..b)"
"minf(f,a..b) ∈ ℤ"
"∀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∈ℤ" "b∈ℤ" "a..b ⊆ ℤ"
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)) ∈ ℤ"
"Minimum(IntegerOrder,f``(a..b)) ∈ f``(a..b)"
"∀x∈f``(a..b). Minimum(IntegerOrder,f``(a..b)) \<lsq> x"
"Minimum(IntegerOrder,f``(a..b)) ∈ ℤ"
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) ∈ ℤ"
"∀c ∈ a..b. f`(c) \<lsq> maxf(f,a..b)"
"∃c ∈ a..b. f`(c) = maxf(f,a..b)"
"minf(f,a..b) ∈ ℤ"
"∀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
subsection‹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 rephrase 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 "ℤ⇩+ {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∈ℤ⇩+" "b∈ℤ⇩+" shows "a\<ra>b ∈ ℤ⇩+"
using assms pos_int_closed_add IsOpClosed_def
by simp
text‹‹ℤ⇧+› is bounded below.›
lemma (in int0) Int_ZF_1_5_L1: shows
"IsBoundedBelow(ℤ⇧+,IntegerOrder)"
"IsBoundedBelow(ℤ⇩+,IntegerOrder)"
using Nonnegative_def PositiveSet_def IsBoundedBelow_def by auto
text‹Subsets of ‹ℤ⇧+› are bounded below.›
lemma (in int0) Int_ZF_1_5_L1A: assumes "A ⊆ ℤ⇧+"
shows "IsBoundedBelow(A,IntegerOrder)"
using assms Int_ZF_1_5_L1 Order_ZF_3_L12 by blast
text‹Subsets of ‹ℤ⇩+› are bounded below.›
lemma (in int0) Int_ZF_1_5_L1B: assumes A1: "A ⊆ ℤ⇩+"
shows "IsBoundedBelow(A,IntegerOrder)"
using A1 Int_ZF_1_5_L1 Order_ZF_3_L12 by blast
text‹Every nonempty subset of positive integers has a mimimum.›
lemma (in int0) Int_ZF_1_5_L1C: assumes "A ⊆ ℤ⇩+" and "A ≠ 0"
shows
"HasAminimum(IntegerOrder,A)"
"Minimum(IntegerOrder,A) ∈ A"
"∀x∈A. Minimum(IntegerOrder,A) \<lsq> x"
using assms 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 ⊆ ℤ⇧+" and A2: "A ∉ Fin(ℤ)" and A3: "D∈ℤ"
shows "∃n∈A. D\<lsq>n"
proof -
{ assume "∀n∈A. ¬(D\<lsq>n)"
moreover from A1 A3 have "D∈ℤ" "∀n∈A. n∈ℤ"
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 have "IsBounded(A,IntegerOrder)"
using Int_ZF_1_5_L1A IsBounded_def by simp
with A2 have False using 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 ‹Int_ZF_1_5_L2›, except we have ‹ℤ⇩+›
instead of ‹ℤ⇧+› here.›
lemma (in int0) Int_ZF_1_5_L2A:
assumes A1: "A ⊆ ℤ⇩+" and A2: "A ∉ Fin(ℤ)" and A3: "D∈ℤ"
shows "∃n∈A. D\<lsq>n"
proof -
{ assume "∀n∈A. ¬(D\<lsq>n)"
moreover from A1 A3 have "D∈ℤ" "∀n∈A. n∈ℤ"
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 have "IsBounded(A,IntegerOrder)"
using Int_ZF_1_5_L1B IsBounded_def by simp
with A2 have False using 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∈ℤ"
shows "Exactly_1_of_3_holds (m=𝟬,m∈ℤ⇩+,(\<rm>m)∈ℤ⇩+)"
using assms 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∈ℤ"
shows "m=𝟬 ∨ m∈ℤ⇩+ ∨ (\<rm>m) ∈ ℤ⇩+"
using assms 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∈ℤ⇩+ ⟷ 𝟭\<lsq>m"
proof
assume "m∈ℤ⇩+" then have "𝟬\<lsq>m" "m≠𝟬"
using PositiveSet_def by auto
then have "𝟬\<ra>𝟭 \<lsq> m"
using Int_ZF_4_L1B by auto
then show "𝟭\<lsq>m"
using int_zero_one_are_int Int_ZF_1_T2 group0.group0_2_L2
by simp
next assume "𝟭\<lsq>m"
then have "m∈ℤ" "𝟬\<lsq>m" "m≠𝟬"
using Int_ZF_2_L1A Int_ZF_2_L16C by auto
then show "m∈ℤ⇩+" 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∈ℤ⇩+" "b∈ℤ⇩+"
shows "a⋅b ∈ ℤ⇩+"
using assms 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
"ℤ⇩+ {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(ℤ,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 "ℤ⇧+ = ℤ⇩+ ∪ {𝟬}"
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:ℤ→ℤ" and A2: "K∈ℤ" "N∈ℤ"
shows "∃C∈ℤ. ∀n∈ℤ⇩+. K \<lsq> f`(n) \<ra> C ⟶ N\<lsq>n"
proof -
from A2 have "N\<lsq>𝟭 ∨ 𝟮\<lsq>N"
using int_zero_one_are_int no_int_between
by simp
moreover
{ assume A3: "N\<lsq>𝟭"
let ?C = "𝟬"
have "?C ∈ ℤ" using int_zero_one_are_int
by simp
moreover
{ fix n assume "n∈ℤ⇩+"
then have "𝟭 \<lsq> n" using Int_ZF_1_5_L3
by simp
with A3 have "N\<lsq>n" by (rule Int_order_transitive)
} then have "∀n∈ℤ⇩+. K \<lsq> f`(n) \<ra> ?C ⟶ N\<lsq>n"
by auto
ultimately have "∃C∈ℤ. ∀n∈ℤ⇩+. K \<lsq> f`(n) \<ra> C ⟶ N\<lsq>n"
by auto }
moreover
{ let ?C = "K \<rs> 𝟭 \<rs> maxf(f,𝟭..(N\<rs>𝟭))"
assume "𝟮\<lsq>N"
then have "𝟮\<rs>𝟭 \<lsq> N\<rs>𝟭"
using int_zero_one_are_int Int_ZF_1_1_L4 int_ord_transl_inv
by simp
then have I: "𝟭 \<lsq> N\<rs>𝟭"
using int_zero_one_are_int Int_ZF_1_2_L3 by simp
with A1 A2 have T:
"maxf(f,𝟭..(N\<rs>𝟭)) ∈ ℤ" "K\<rs>𝟭 ∈ ℤ" "?C ∈ ℤ"
using Int_ZF_1_4_L2 Int_ZF_1_1_L5 int_zero_one_are_int
by auto
moreover
{ fix n assume A4: "n∈ℤ⇩+"
{ assume A5: "K \<lsq> f`(n) \<ra> ?C" and "¬(N\<lsq>n)"
with A2 A4 have "n \<lsq> N\<rs>𝟭"
using PositiveSet_def Int_ZF_1_3_L6A by simp
with A4 have "n ∈ 𝟭..(N\<rs>𝟭)"
using Int_ZF_1_5_L3 Interval_def by auto
with A1 I T have "f`(n)\<ra>?C \<lsq> maxf(f,𝟭..(N\<rs>𝟭)) \<ra> ?C"
using Int_ZF_1_4_L2 int_ord_transl_inv by simp
with T have "f`(n)\<ra>?C \<lsq> K\<rs>𝟭"
using Int_ZF_1_2_L3 by simp
with A5 have "K \<lsq> K\<rs>𝟭"
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∈ℤ⇩+. K \<lsq> f`(n) \<ra> ?C ⟶ N\<lsq>n"
by simp
ultimately have "∃C∈ℤ. ∀n∈ℤ⇩+. 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∈ℤ⇩+" shows "abs(a) = a"
using assms Int_ZF_2_T1 group3.OrderedGroup_ZF_3_L2B
by simp
text‹One and two are in ‹ℤ⇩+›.›
lemma (in int0) int_one_two_are_pos: shows "𝟭 ∈ ℤ⇩+" "𝟮 ∈ ℤ⇩+"
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 ‹ℤ⇩+› by a function defined on integers
is not empty.›
lemma (in int0) Int_ZF_1_5_L5: assumes A1: "f : ℤ→X"
shows "f``(ℤ⇩+) ≠ 0"
proof -
have "ℤ⇩+ ⊆ ℤ" using PositiveSet_def by auto
with A1 show "f``(ℤ⇩+) ≠ 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 ∈ ℤ⇩+"
shows
"𝟬 \<lsq> n\<rs>𝟭"
"𝟬 ∈ 𝟬..(n\<rs>𝟭)"
"𝟬..(n\<rs>𝟭) ⊆ ℤ"
proof -
from A1 have "𝟭 \<lsq> n" "(\<rm>𝟭) ∈ ℤ"
using Int_ZF_1_5_L3 int_zero_one_are_int Int_ZF_1_1_L4
by auto
then have "𝟭\<rs>𝟭 \<lsq> n\<rs>𝟭"
using int_ord_transl_inv by simp
then show "𝟬 \<lsq> n\<rs>𝟭"
using int_zero_one_are_int Int_ZF_1_1_L4 by simp
then show "𝟬 ∈ 𝟬..(n\<rs>𝟭)"
using int_zero_one_are_int int_ord_is_refl refl_def Order_ZF_2_L1B
by simp
show "𝟬..(n\<rs>𝟭) ⊆ ℤ"
using Int_ZF_2_L1B Order_ZF_2_L6 by simp
qed
text‹Intgers greater than one in ‹ℤ⇩+› belong to ‹ℤ⇩+›.
This is a property of ordered groups and follows from
‹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 ∈ ℤ⇩+" and "a\<lsq>b"
shows "b ∈ ℤ⇩+"
proof-
from assms have "𝟭\<lsq>a" "a\<lsq>b"
using Int_ZF_1_5_L3 by auto
then have "𝟭\<lsq>b" by (rule Int_order_transitive)
then show "b ∈ ℤ⇩+" 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∈ℤ" "b ∈ ℤ⇩+"
shows "a \<lsq> a\<ra>b" "a ≠ a\<ra>b" "a\<ra>b ∈ ℤ"
using assms 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∈ℤ"
shows
"a \<lsq> GreaterOf(IntegerOrder,𝟭,a)"
"GreaterOf(IntegerOrder,𝟭,a) ∈ ℤ⇩+"
"GreaterOf(IntegerOrder,𝟭,a) \<ra> 𝟭 ∈ ℤ⇩+"
"a \<lsq> GreaterOf(IntegerOrder,𝟭,a) \<ra> 𝟭"
"a ≠ GreaterOf(IntegerOrder,𝟭,a) \<ra> 𝟭"
using assms int_zero_not_one Int_ZF_1_3_T1 ring1.OrdRing_ZF_3_L12
by auto
text‹The opposite of an element of ‹ℤ⇩+› cannot belong to
‹ℤ⇩+›.›
lemma (in int0) Int_ZF_1_5_L8: assumes "a ∈ ℤ⇩+"
shows "(\<rm>a) ∉ ℤ⇩+"
using assms Int_ZF_2_T1 group3.OrderedGroup_ZF_1_L20
by simp
text‹For every integer there is one in ‹ℤ⇩+› that is greater or
equal.›
lemma (in int0) Int_ZF_1_5_L9: assumes "a∈ℤ"
shows "∃b∈ℤ⇩+. a\<lsq>b"
using assms int_not_trivial Int_ZF_2_T1 group3.OrderedGroup_ZF_1_L23
by simp
text‹A theorem about odd extensions. Recall from ‹OrdereGroup_ZF.thy›
that the odd extension of an integer function $f$ defined on ‹ℤ⇩+›
is the odd function on ‹ℤ› equal to $f$ on ‹ℤ⇩+›.
First we show that the odd extension is defined on ‹ℤ›.›
lemma (in int0) Int_ZF_1_5_L10: assumes "f : ℤ⇩+→ℤ"
shows "OddExtension(ℤ,IntegerAddition,IntegerOrder,f) : ℤ→ℤ"
using assms Int_ZF_2_T1 group3.odd_ext_props by simp
text‹On ‹ℤ⇩+›, the odd extension of $f$ is the same as $f$.›
lemma (in int0) Int_ZF_1_5_L11: assumes "f : ℤ⇩+→ℤ" and "a ∈ ℤ⇩+" and
"g = OddExtension(ℤ,IntegerAddition,IntegerOrder,f)"
shows "g`(a) = f`(a)"
using assms Int_ZF_2_T1 group3.odd_ext_props by simp
text‹On ‹\<sm>ℤ⇩+›, the value of the odd extension of $f$
is the negative of $f(-a)$.›
lemma (in int0) Int_ZF_1_5_L12:
assumes "f : ℤ⇩+→ℤ" and "a ∈ (\<sm>ℤ⇩+)" and
"g = OddExtension(ℤ,IntegerAddition,IntegerOrder,f)"
shows "g`(a) = \<rm>(f`(\<rm>a))"
using assms Int_ZF_2_T1 group3.odd_ext_props by simp
text‹Odd extensions are odd on ‹ℤ›.›
lemma (in int0) int_oddext_is_odd:
assumes "f : ℤ⇩+→ℤ" and "a∈ℤ" and
"g = OddExtension(ℤ,IntegerAddition,IntegerOrder,f)"
shows "g`(\<rm>a) = \<rm>(g`(a))"
using assms 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: ℤ→ℤ" shows
"(∀a∈ℤ. f`(\<rm>a) = (\<rm>f`(a))) ⟷ (∀a∈ℤ. (\<rm>(f`(\<rm>a))) = f`(a))"
using assms 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 : ℤ⇩+→ℤ" and "a∈ℤ" and
"g = OddExtension(ℤ,IntegerAddition,IntegerOrder,f)"
shows "(\<rm>g`(\<rm>a)) = g`(a)"
using assms Int_ZF_2_T1 group3.oddext_is_odd_alt by simp
subsection‹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 ‹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: ℤ→ℤ" and
"∀a∈ℤ.∃b∈ℤ⇩+.∀x. b\<lsq>x ⟶ a \<lsq> f`(x)" and "A ⊆ ℤ" and
"IsBoundedAbove(f``(A),IntegerOrder)"
shows "IsBoundedAbove(A,IntegerOrder)"
using assms 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: ℤ→ℤ" and
A3: "∀a∈ℤ.∃b∈ℤ⇩+.∀x. b\<lsq>x ⟶ a \<lsq> f`(x)" and
A4: "∀x∈X. b(x) ∈ ℤ ∧ f`(b(x)) \<lsq> U"
shows "∃u.∀x∈X. b(x) \<lsq> u"
proof -
let ?G = "ℤ"
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 ‹ Int_ZF_1_6_L2›.›
lemma (in int0) Int_ZF_1_6_L3: assumes A1: "X≠0" and A2: "f: ℤ→ℤ" and
A3: "∀a∈ℤ.∃b∈ℤ⇩+.∀y. b\<lsq>y ⟶ f`(\<rm>y) \<lsq> a" and
A4: "∀x∈X. b(x) ∈ ℤ ∧ L \<lsq> f`(b(x))"
shows "∃l.∀x∈X. l \<lsq> b(x)"
proof -
let ?G = "ℤ"
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 ‹Int_ZF_1_6_L2› and
‹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
‹OrderedGroup_ZF›.›
lemma (in int0) Int_ZF_1_6_L4:
assumes A1: "X≠0" and A2: "f: ℤ→ℤ" and
A3: "∀a∈ℤ.∃b∈ℤ⇩+.∀x. b\<lsq>x ⟶ a \<lsq> f`(x)" and
A4: "∀a∈ℤ.∃b∈ℤ⇩+.∀y. b\<lsq>y ⟶ f`(\<rm>y) \<lsq> a" and
A5: "∀x∈X. b(x) ∈ ℤ ∧ f`(b(x)) \<lsq> U ∧ L \<lsq> f`(b(x))"
shows "∃M.∀x∈X. abs(b(x)) \<lsq> M"
proof -
let ?G = "ℤ"
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: ℤ→ℤ" and A2: "N∈ℤ" 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(ℤ)"
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 ⊆ ℤ" 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∈ℤ" and
A2: "∀m. N\<lsq>m ⟶ L \<lsq> f`(m)" and
A3: "f: ℤ→ℤ" and A4: "K∈ℤ"
shows "∃c∈ℤ. ∀n∈ℤ⇩+. K \<lsq> f`(n)\<ra>c"
proof -
have "IsBoundedBelow(ℤ⇩+,IntegerOrder)"
using Int_ZF_1_5_L1 by simp
with A3 A1 A2 have "IsBoundedBelow(f``(ℤ⇩+),IntegerOrder)"
by (rule Int_ZF_1_6_L5)
with A1 obtain l where I: "∀y∈f``(ℤ⇩+). l \<lsq> y"
using Int_ZF_1_5_L5 IsBoundedBelow_def by auto
let ?c = "K\<rs>l"
from A3 have "f``(ℤ⇩+) ≠ 0" using Int_ZF_1_5_L5
by simp
then have "∃y. y ∈ f``(ℤ⇩+)" by (rule nonempty_has_element)
then obtain y where "y ∈ f``(ℤ⇩+)" by auto
with A4 I have T: "l ∈ ℤ" "?c ∈ ℤ"
using Int_ZF_2_L1A Int_ZF_1_1_L5 by auto
{ fix n assume A5: "n∈ℤ⇩+"
have "ℤ⇩+ ⊆ ℤ" 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∈ℤ⇩+. 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: ℤ→ℤ" and A2: "K∈ℤ" "N∈ℤ" and
A3: "∀a∈ℤ.∃b∈ℤ⇩+.∀x. b\<lsq>x ⟶ a \<lsq> f`(x)"
shows "∃C∈ℤ. N \<lsq> Minimum(IntegerOrder,{n∈ℤ⇩+. K \<lsq> f`(n)\<ra>C})"
proof -
from A1 A2 have "∃C∈ℤ. ∀n∈ℤ⇩+. K \<lsq> f`(n) \<ra> C ⟶ N\<lsq>n"
using Int_ZF_1_5_L4 by simp
then obtain C where I: "C∈ℤ" and
II: "∀n∈ℤ⇩+. 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∈ℤ⇩+. K \<lsq> f`(n)\<ra>C})"
proof -
from A2 A3 I have "∃n∈ℤ⇩+.∀x. n\<lsq>x ⟶ K\<rs>C \<lsq> f`(x)"
using Int_ZF_1_1_L5 by simp
then obtain n where
"n∈ℤ⇩+" and "∀x. n\<lsq>x ⟶ K\<rs>C \<lsq> f`(x)"
by auto
with A2 I have
"{n∈ℤ⇩+. K \<lsq> f`(n)\<ra>C} ≠ 0"
"{n∈ℤ⇩+. K \<lsq> f`(n)\<ra>C} ⊆ ℤ⇩+"
using int_ord_is_refl refl_def PositiveSet_def Int_ZF_2_L9C
by auto
then show "HasAminimum(IntegerOrder,{n∈ℤ⇩+. K \<lsq> f`(n)\<ra>C})"
using Int_ZF_1_5_L1C by simp
qed
moreover from II have
"∀n ∈ {n∈ℤ⇩+. K \<lsq> f`(n)\<ra>C}. ⟨N,n⟩ ∈ IntegerOrder"
by auto
ultimately have
"⟨N,Minimum(IntegerOrder,{n∈ℤ⇩+. 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∈ℤ" and A2: "{a⋅x. x∈ℤ} ∈ Fin(ℤ)"
shows "a = 𝟬"
proof -
from A1 have "a=𝟬 ∨ (a \<lsq> \<rm>𝟭) ∨ (𝟭\<lsq>a)"
using Int_ZF_1_3_L6C by simp
moreover
{ assume "a \<lsq> \<rm>𝟭"
then have "{a⋅x. x∈ℤ} ∉ Fin(ℤ)"
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 "𝟭\<lsq>a"
then have "{a⋅x. x∈ℤ} ∉ Fin(ℤ)"
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 = 𝟬" by auto
qed
subsection‹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∈ℤ. F(q) ∈ ℤ" and
A2: "∀q∈ℤ. F(q)⋅abs(q) \<lsq> A⋅abs(q) \<ra> B" and
A3: "A∈ℤ" "B∈ℤ"
shows "∃L. ∀p∈ℤ. F(p) \<lsq> L"
proof -
let ?I = "(\<rm>abs(B))..abs(B)"
let ?K = "{F(q). q ∈ ?I}"
let ?M = "Maximum(IntegerOrder,?K)"
let ?L = "GreaterOf(IntegerOrder,?M,A\<ra>𝟭)"
from A3 A1 have C1:
"IsBounded(?I,IntegerOrder)"
"?I ≠ 0"
"∀q∈ℤ. F(q) ∈ ℤ"
"?K = {F(q). q ∈ ?I}"
using Order_ZF_3_L11 Int_ZF_1_3_L17 by auto
then have "?M ∈ ℤ" by (rule Int_ZF_1_4_L1)
with A3 have T1: "?M \<lsq> ?L" "A\<ra>𝟭 \<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∈ℤ" have "F(p) \<lsq> ?L"
proof -
{ 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 have "F(p) \<lsq> ?L" by (rule Int_order_transitive) }
moreover
{ assume A5: "¬(abs(p) \<lsq> abs(B))"
from A3 A2 A4 have
"A⋅abs(p) ∈ ℤ" "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>𝟭)⋅abs(p)"
using Int_ZF_2_L14 Int_ZF_1_2_L7 by simp
moreover from A3 A1 A4 A5 have
"F(p) ∈ ℤ" "A\<ra>𝟭 ∈ ℤ" "abs(p) ∈ ℤ"
"¬(abs(p) \<lsq> 𝟬)"
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>𝟭"
using Int_ineq_simpl_positive by simp
moreover from T1 have "A\<ra>𝟭 \<lsq> ?L" by simp
ultimately have "F(p) \<lsq> ?L" by (rule Int_order_transitive) }
ultimately show ?thesis by blast
qed
} then have "∀p∈ℤ. F(p) \<lsq> ?L" by simp
thus ?thesis by auto
qed
text‹A lemma about splitting (not really, there is some overlap)
the ‹ℤ×ℤ› 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∈ℤ" "b∈ℤ"
shows
"𝟬\<lsq>a ∧ 𝟬\<lsq>b ∨ a\<lsq>𝟬 ∧ b\<lsq>𝟬 ∨
a\<lsq>𝟬 ∧ 𝟬\<lsq>b ∧ 𝟬 \<lsq> a\<ra>b ∨ a\<lsq>𝟬 ∧ 𝟬\<lsq>b ∧ a\<ra>b \<lsq> 𝟬 ∨
𝟬\<lsq>a ∧ b\<lsq>𝟬 ∧ 𝟬 \<lsq> a\<ra>b ∨ 𝟬\<lsq>a ∧ b\<lsq>𝟬 ∧ a\<ra>b \<lsq> 𝟬"
using assms Int_ZF_2_T1 group3.OrdGroup_6cases by simp
end