Mostly show sqrt 2 is irrational (#53)

This commit is contained in:
Patrick Stevens
2019-10-22 20:56:58 +01:00
committed by GitHub
parent 6eaa181104
commit f5f4cf1376
8 changed files with 244 additions and 37 deletions

View File

@@ -53,6 +53,7 @@ open import Monoids.Definition
open import Semirings.Definition
open import Semirings.Solver
open import CauchyCompletion
open import Fields.CauchyCompletion.Definition
open import Fields.CauchyCompletion.Addition
module Everything.Safe where

View File

@@ -5,6 +5,7 @@
open import Numbers.Primes.PrimeNumbers
open import Numbers.Primes.IntegerFactorisation
open import Numbers.Rationals
open import Numbers.RationalsLemmas
open import Logic.PropositionalLogic
open import Logic.PropositionalLogicExamples

View File

@@ -15,7 +15,7 @@ open import Functions
open import LogicalFormulae
open import Numbers.Naturals.Naturals
module CauchyCompletion {m n o : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A A A} {_*_ : A A A} {_<_ : Rel {m} {o} A} {pOrder : SetoidPartialOrder S _<_} (tOrder : SetoidTotalOrder {_<_ = _<_} pOrder) {R : Ring S _+_ _*_} (order : OrderedRing R tOrder) (F : Field R) (charNot2 : Setoid.__ S ((Ring.1R R) + (Ring.1R R)) (Ring.0R R) False) where
module Fields.CauchyCompletion.Addition {m n o : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A A A} {_*_ : A A A} {_<_ : Rel {m} {o} A} {pOrder : SetoidPartialOrder S _<_} {tOrder : SetoidTotalOrder {_<_ = _<_} pOrder} {R : Ring S _+_ _*_} (order : OrderedRing R tOrder) (F : Field R) (charNot2 : Setoid.__ S ((Ring.1R R) + (Ring.1R R)) (Ring.0R R) False) where
open Setoid S
open SetoidTotalOrder tOrder
@@ -26,23 +26,9 @@ open Ring R
open Group additiveGroup
open Field F
open import Fields.CauchyCompletion.Definition order F
open import Rings.Orders.Lemmas(order)
cauchy : Sequence A Set (m o)
cauchy s = (ε : A) (0R < ε) Sg (λ N {m n : } (N <N m) (N <N n) abs ((index s m) -R (index s n)) < ε)
record CauchyCompletion : Set (m o) where
field
elts : Sequence A
converges : cauchy elts
injection : A CauchyCompletion
CauchyCompletion.elts (injection a) = constSequence a
CauchyCompletion.converges (injection a) = λ ε 0<e 0 , λ {m} {n} _ _ <WellDefined (symmetric (identityOfIndiscernablesRight __ (absWellDefined (index (constSequence a) m + inverse (index (constSequence a) n)) 0R (t m n)) (absZero order))) reflexive 0<e
where
t : (m n : ) index (constSequence a) m + inverse (index (constSequence a) n) 0R
t m n = identityOfIndiscernablesLeft __ (identityOfIndiscernablesLeft __ invRight (equalityCommutative (applyEquality (λ i a + inverse i) (indexAndConst a n)))) (applyEquality (_+ inverse (index (constSequence a) n)) (equalityCommutative (indexAndConst a m)))
halve : (a : A) Sg A (λ i i + i a)
-- TODO: we need to know the characteristic already
halve a with allInvertible (1R + 1R) charNot2

View File

@@ -0,0 +1,44 @@
{-# OPTIONS --safe --warning=error --without-K --guardedness #-}
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Setoids.Setoids
open import Rings.Definition
open import Rings.Lemmas
open import Rings.Order
open import Groups.Definition
open import Groups.Groups
open import Fields.Fields
open import Sets.EquivalenceRelations
open import Sequences
open import Setoids.Orders
open import Functions
open import LogicalFormulae
open import Numbers.Naturals.Naturals
module Fields.CauchyCompletion.Definition {m n o : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A A A} {_*_ : A A A} {_<_ : Rel {m} {o} A} {pOrder : SetoidPartialOrder S _<_} {tOrder : SetoidTotalOrder {_<_ = _<_} pOrder} {R : Ring S _+_ _*_} (order : OrderedRing R tOrder) (F : Field R) where
open Setoid S
open SetoidTotalOrder tOrder
open SetoidPartialOrder pOrder
open Equivalence eq
open OrderedRing order
open Ring R
open Group additiveGroup
open Field F
open import Rings.Orders.Lemmas(order)
cauchy : Sequence A Set (m o)
cauchy s = (ε : A) (0R < ε) Sg (λ N {m n : } (N <N m) (N <N n) abs ((index s m) -R (index s n)) < ε)
record CauchyCompletion : Set (m o) where
field
elts : Sequence A
converges : cauchy elts
injection : A CauchyCompletion
CauchyCompletion.elts (injection a) = constSequence a
CauchyCompletion.converges (injection a) = λ ε 0<e 0 , λ {m} {n} _ _ <WellDefined (symmetric (identityOfIndiscernablesRight __ (absWellDefined (index (constSequence a) m + inverse (index (constSequence a) n)) 0R (t m n)) (absZero order))) reflexive 0<e
where
t : (m n : ) index (constSequence a) m + inverse (index (constSequence a) n) 0R
t m n = identityOfIndiscernablesLeft __ (identityOfIndiscernablesLeft __ invRight (equalityCommutative (applyEquality (λ i a + inverse i) (indexAndConst a n)))) (applyEquality (_+ inverse (index (constSequence a) n)) (equalityCommutative (indexAndConst a m)))

View File

@@ -33,6 +33,9 @@ a *Q b = fieldOfFractionsTimes IntDom a b
0Q :
0Q = Ring.0R Ring
injectionQ :
injectionQ z = z ,, (nonneg 1 , λ ())
Field : Field Ring
Field = fieldOfFractions IntDom

View File

@@ -0,0 +1,114 @@
{-# OPTIONS --safe --warning=error #-}
open import WellFoundedInduction
open import LogicalFormulae
open import Numbers.Naturals.Naturals
open import Numbers.Integers.Integers
open import Numbers.Integers.Order
open import Groups.Groups
open import Groups.Definition
open import Rings.Definition
open import Rings.Order
open import Fields.Fields
open import Numbers.Primes.PrimeNumbers
open import Setoids.Setoids
open import Setoids.Orders
open import Functions
open import Fields.FieldOfFractions
open import Fields.FieldOfFractionsOrder
open import Sets.EquivalenceRelations
open import Numbers.Rationals
open import Semirings.Definition
open import Orders
module Numbers.RationalsLemmas where
open OrderedRing OrderedRing
open import Rings.Orders.Lemmas OrderedRing
open SetoidTotalOrder (totalOrderToSetoidTotalOrder Order)
evenOrOdd : (a : ) (Sg (λ i i *N 2 a)) || (Sg (λ i succ (i *N 2) a))
evenOrOdd zero = inl (zero , refl)
evenOrOdd (succ zero) = inr (zero , refl)
evenOrOdd (succ (succ a)) with evenOrOdd a
evenOrOdd (succ (succ a)) | inl (a/2 , even) = inl (succ a/2 , applyEquality (λ i succ (succ i)) even)
evenOrOdd (succ (succ a)) | inr (a/2-1 , odd) = inr (succ a/2-1 , applyEquality (λ i succ (succ i)) odd)
parity : (a b : ) succ (2 *N a) 2 *N b False
parity zero (succ b) pr rewrite Semiring.commutative Semiring b (succ (b +N 0)) = bad pr
where
bad : (1 succ (succ ((b +N 0) +N b))) False
bad ()
parity (succ a) (succ b) pr rewrite Semiring.commutative Semiring b (succ (b +N 0)) | Semiring.commutative Semiring a (succ (a +N 0)) | Semiring.commutative Semiring (a +N 0) a | Semiring.commutative Semiring (b +N 0) b = parity a b (succInjective (succInjective pr))
sqrt0 : (p : ) p *N p 0 p 0
sqrt0 zero pr = refl
sqrt0 (succ p) ()
-- So as to give us something easy to induct down, introduce a silly extra variable
evil' : (k : ) ((a b : ) (0 <N a) (pr : k a +N b) a *N a (b *N b) *N 2 False)
evil' = rec <NWellfounded (λ z (x x₁ : ) (pr' : 0 <N x) (x₂ : z x +N x₁) (x₃ : x *N x (x₁ *N x₁) *N 2) False) go
where
go : (k : ) (indHyp : (k' : ) (k'<k : k' <N k) (r s : ) (0<r : 0 <N r) (r+s : k' r +N s) (x₇ : r *N r (s *N s) *N 2) False) (a b : ) (0<a : 0 <N a) (a+b : k a +N b) (pr : a *N a (b *N b) *N 2) False
go k indHyp a b 0<a a+b pr = contr
where
open import Semirings.Solver Semiring multiplicationNIsCommutative
aEven : Sg (λ i i *N 2 a)
aEven with evenOrOdd a
aEven | inl x = x
aEven | inr (a/2-1 , odd) rewrite equalityCommutative odd =
-- Derive a pretty mechanical contradiction using the automatic solver.
-- This line looks hellish, but it was almost completely mechanical.
exFalso (parity (a/2-1 +N (a/2-1 *N succ (a/2-1 *N 2))) (b *N b) (transitivity (
-- truly mechanical bit starts here
from (succ (plus (plus (const a/2-1) (times (const a/2-1) (succ (times (const a/2-1) (succ (succ zero)))))) (plus (plus (const a/2-1) (times (const a/2-1) (succ (times (const a/2-1) (succ (succ zero)))))) (times zero (plus (const a/2-1) (times (const a/2-1) (succ (times (const a/2-1) (succ (succ zero))))))))))
to succ (plus (times (const a/2-1) (succ (succ zero))) (times (times (const a/2-1) (succ (succ zero))) (succ (times (const a/2-1) (succ (succ zero))))))
-- truly mechanical bit ends here
by applyEquality (λ i succ (a/2-1 +N i)) (
-- Grinding out some manipulations
transitivity (equalityCommutative (Semiring.+Associative Semiring a/2-1 _ _)) (applyEquality (a/2-1 +N_) (transitivity (Semiring.commutative Semiring (a/2-1 *N (a/2-1 +N a/2-1)) _) (transitivity (equalityCommutative (Semiring.+Associative Semiring a/2-1 _ _)) (applyEquality (a/2-1 +N_) (transitivity (equalityCommutative (Semiring.+Associative Semiring a/2-1 _ _)) (applyEquality (a/2-1 +N_) (transitivity (equalityCommutative (Semiring.+DistributesOver* Semiring a/2-1 _ _)) (applyEquality (a/2-1 *N_) (equalityCommutative (Semiring.+Associative Semiring a/2-1 _ _))))))))))
)) (transitivity pr (multiplicationNIsCommutative (b *N b) 2))))
next : (underlying aEven *N 2) *N (underlying aEven *N 2) (b *N b) *N 2
next with aEven
... | a/2 , even rewrite even = pr
next2 : (underlying aEven *N 2) *N underlying aEven b *N b
next2 = productCancelsRight 2 _ _ (le 1 refl) (transitivity (equalityCommutative (Semiring.*Associative Semiring (underlying aEven *N 2) _ _)) next)
next3 : b *N b (underlying aEven *N underlying aEven) *N 2
next3 = equalityCommutative (transitivity (transitivity (equalityCommutative (Semiring.*Associative Semiring (underlying aEven) _ _)) (multiplicationNIsCommutative (underlying aEven) _)) next2)
halveDecreased : underlying aEven <N a
halveDecreased with aEven
halveDecreased | zero , even rewrite equalityCommutative even = exFalso (PartialOrder.irreflexive (TotalOrder.order TotalOrder) 0<a)
halveDecreased | succ a/2 , even = le a/2 (transitivity (applyEquality succ (transitivity (Semiring.commutative Semiring a/2 _) (applyEquality succ (transitivity (doubleIsAddTwo a/2) (multiplicationNIsCommutative 2 a/2))))) even)
reduced : b +N underlying aEven <N k
reduced with lessRespectsAddition b halveDecreased
... | bl rewrite a+b = identityOfIndiscernablesLeft _<N_ bl (Semiring.commutative Semiring _ b)
0<b : 0 <N b
0<b with TotalOrder.totality TotalOrder 0 b
0<b | inl (inl 0<b) = 0<b
0<b | inl (inr ())
0<b | inr 0=b rewrite equalityCommutative 0=b = exFalso (PartialOrder.irreflexive (TotalOrder.order TotalOrder) {0} (identityOfIndiscernablesRight _<N_ 0<a (sqrt0 a pr)))
contr : False
contr = indHyp (b +N underlying aEven) reduced b (underlying aEven) 0<b refl next3
evil : (a b : ) (0 <N a) a *N a (b *N b) *N 2 False
evil a b 0<a = evil' (a +N b) a b 0<a refl
--sqrt2Irrational : (a : ) (a *Q a) =Q (injectionQ (nonneg 2)) False
--sqrt2Irrational (numerator ,, (denominator , denom!=0)) pr = {!!}
-- where
-- -- We have in hand `pr`, which is the following (almost by definition):
-- pr' : (numerator *Z numerator) ≡ (denominator *Z denominator) *Z nonneg 2
-- pr' = transitivity (equalityCommutative (transitivity (Ring.*Commutative Ring) (Ring.identIsIdent Ring))) pr
--
-- -- Move into the naturals so that we can do nice divisibility things.
-- lemma1 : abs ((denominator *Z denominator) *Z nonneg 2) ≡ (abs denominator *Z abs denominator) *Z nonneg 2
-- lemma1 = transitivity (absRespectsTimes (denominator *Z denominator) (nonneg 2)) (applyEquality (_*Z nonneg 2) (absRespectsTimes denominator denominator))
-- amenableToNaturals : (abs numerator) *Z (abs numerator) ≡ ((abs denominator) *Z (abs denominator)) *Z nonneg 2
-- amenableToNaturals = transitivity (equalityCommutative (absRespectsTimes numerator numerator)) (transitivity (applyEquality abs pr') lemma1)

View File

@@ -10,28 +10,33 @@ open import Setoids.Setoids
open import Setoids.Orders
open import Sets.EquivalenceRelations
module Rings.Lemmas where
module Rings.Lemmas {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} {_*_ : A A A} (R : Ring S _+_ _*_) where
ringMinusExtracts : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} {_*_ : A A A} (R : Ring S _+_ _*_) {x y : A} Setoid.__ S (x * Group.inverse (Ring.additiveGroup R) y) (Group.inverse (Ring.additiveGroup R) (x * y))
ringMinusExtracts {S = S} {_+_ = _+_} {_*_ = _*_} R {x = x} {y} = transferToRight' additiveGroup (transitive (symmetric *DistributesOver+) (transitive (*WellDefined reflexive invLeft) (Ring.timesZero R)))
open Setoid S
open Ring R
open Group additiveGroup
ringMinusExtracts : {x y : A} Setoid.__ S (x * Group.inverse (Ring.additiveGroup R) y) (Group.inverse (Ring.additiveGroup R) (x * y))
ringMinusExtracts {x = x} {y} = transferToRight' additiveGroup (transitive (symmetric *DistributesOver+) (transitive (*WellDefined reflexive invLeft) (Ring.timesZero R)))
where
open Equivalence eq
ringMinusExtracts' : {x y : A} Setoid.__ S ((Group.inverse (Ring.additiveGroup R) x) * y) (Group.inverse (Ring.additiveGroup R) (x * y))
ringMinusExtracts' {x = x} {y} = transitive *Commutative (transitive ringMinusExtracts (inverseWellDefined additiveGroup *Commutative))
where
open Equivalence eq
twoNegativesTimes : {a b : A} (inverse a) * (inverse b) a * b
twoNegativesTimes {a} {b} = transitive (ringMinusExtracts) (transitive (inverseWellDefined additiveGroup ringMinusExtracts') (invTwice additiveGroup (a * b)))
where
open Setoid S
open Equivalence eq
open Ring R
open Group additiveGroup
groupLemmaMove0G : {a b : _} {A : Set a} {_·_ : A A A} {S : Setoid {a} {b} A} (G : Group S _·_) {x : A} (Setoid.__ S (Group.0G G) (Group.inverse G x)) Setoid.__ S x (Group.0G G)
groupLemmaMove0G {S = S} G {x} pr = transitive (symmetric (invInv G)) (transitive (symmetric p) (invIdent G))
where
open Group G
open Setoid S
open Equivalence eq
p : inverse 0G inverse (inverse x)
open Equivalence (Setoid.eq S)
p : Setoid.__ S (Group.inverse G (Group.0G G)) (Group.inverse G (Group.inverse G x))
p = inverseWellDefined G pr
groupLemmaMove0G' : {a b : _} {A : Set a} {_·_ : A A A} {S : Setoid {a} {b} A} (G : Group S _·_) {x : A} Setoid.__ S x (Group.0G G) (Setoid.__ S (Group.0G G) (Group.inverse G x))
groupLemmaMove0G' {S = S} G {x} pr = transferToRight' G (transitive identLeft pr)
where
open Group G
open Setoid S
open Equivalence eq
groupLemmaMove0G' {S = S} G {x} pr = transferToRight' G (Equivalence.transitive (Setoid.eq S) (Group.identLeft G) pr)

View File

@@ -10,7 +10,6 @@ open import Setoids.Setoids
open import Functions
open import Sets.EquivalenceRelations
open import Rings.Definition
open import Rings.Lemmas
open import Rings.Order
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
@@ -24,6 +23,9 @@ open SetoidTotalOrder tOrder
open Ring R
open Group additiveGroup
open import Rings.Lemmas R
ringAddInequalities : {w x y z : A} w < x y < z (w + y) < (x + z)
ringAddInequalities {w = w} {x} {y} {z} w<x y<z = transitive (orderRespectsAddition w<x y) (<WellDefined groupIsAbelian groupIsAbelian (orderRespectsAddition y<z x))
@@ -131,7 +133,7 @@ ringCanMultiplyByPositive {x} {y} {c} 0<c x<y = SetoidPartialOrder.<WellDefined
p1 : 0R < ((y * c) + ((Group.inverse additiveGroup x) * c))
p1 = SetoidPartialOrder.<WellDefined pOrder reflexive (Equivalence.transitive eq *Commutative (Equivalence.transitive eq *DistributesOver+ ((Group.+WellDefined additiveGroup) *Commutative *Commutative))) (OrderedRing.orderRespectsMultiplication order have 0<c)
p' : 0R < ((y * c) + (Group.inverse additiveGroup (x * c)))
p' = SetoidPartialOrder.<WellDefined pOrder reflexive (Group.+WellDefined additiveGroup reflexive (Equivalence.transitive eq (Equivalence.transitive eq *Commutative (ringMinusExtracts R)) (inverseWellDefined additiveGroup *Commutative))) p1
p' = SetoidPartialOrder.<WellDefined pOrder reflexive (Group.+WellDefined additiveGroup reflexive (Equivalence.transitive eq (Equivalence.transitive eq *Commutative ringMinusExtracts) (inverseWellDefined additiveGroup *Commutative))) p1
q : (0R + (x * c)) < (((y * c) + (Group.inverse additiveGroup (x * c))) + (x * c))
q = OrderedRing.orderRespectsAddition order p' (x * c)
q' : (x * c) < ((y * c) + 0R)
@@ -144,7 +146,7 @@ ringCanCancelPositive {x} {y} {c} 0<c xc<yc = SetoidPartialOrder.<WellDefined pO
have : 0R < ((y * c) + (Group.inverse additiveGroup (x * c)))
have = SetoidPartialOrder.<WellDefined pOrder (Group.invRight additiveGroup) reflexive (OrderedRing.orderRespectsAddition order xc<yc (Group.inverse additiveGroup _))
p1 : 0R < ((y * c) + ((Group.inverse additiveGroup x) * c))
p1 = SetoidPartialOrder.<WellDefined pOrder reflexive (Group.+WellDefined additiveGroup reflexive (symmetric (Equivalence.transitive eq (*Commutative) (Equivalence.transitive eq (ringMinusExtracts R) (inverseWellDefined additiveGroup *Commutative))))) have
p1 = SetoidPartialOrder.<WellDefined pOrder reflexive (Group.+WellDefined additiveGroup reflexive (symmetric (Equivalence.transitive eq (*Commutative) (Equivalence.transitive eq ringMinusExtracts (inverseWellDefined additiveGroup *Commutative))))) have
q : 0R < ((y + Group.inverse additiveGroup x) * c)
q = SetoidPartialOrder.<WellDefined pOrder reflexive (Equivalence.transitive eq (Equivalence.transitive eq (Group.+WellDefined additiveGroup *Commutative *Commutative) (symmetric *DistributesOver+)) *Commutative) p1
q' : 0R < (y + Group.inverse additiveGroup x)
@@ -197,7 +199,7 @@ ringCanMultiplyByNegative {x} {y} {c} c<0 x<y = ringSwapNegatives u
t : (x * Group.inverse additiveGroup c) < (y * Group.inverse additiveGroup c)
t = ringCanMultiplyByPositive 0<-c x<y
u : (Group.inverse additiveGroup (x * c)) < Group.inverse additiveGroup (y * c)
u = SetoidPartialOrder.<WellDefined pOrder (ringMinusExtracts R) (ringMinusExtracts R) t
u = SetoidPartialOrder.<WellDefined pOrder ringMinusExtracts ringMinusExtracts t
ringCanCancelNegative : {x y c : A} c < (Ring.0R R) (x * c) < (y * c) y < x
ringCanCancelNegative {x} {y} {c} c<0 xc<yc = r
@@ -206,7 +208,7 @@ ringCanCancelNegative {x} {y} {c} c<0 xc<yc = r
p0 : 0R < ((y * c) + inverse (x * c))
p0 = SetoidPartialOrder.<WellDefined pOrder invRight reflexive (OrderedRing.orderRespectsAddition order xc<yc (inverse (x * c)))
p1 : 0R < ((y * c) + ((inverse x) * c))
p1 = SetoidPartialOrder.<WellDefined pOrder reflexive (Group.+WellDefined additiveGroup reflexive (Equivalence.transitive eq (inverseWellDefined additiveGroup *Commutative) (Equivalence.transitive eq (symmetric (ringMinusExtracts R)) *Commutative))) p0
p1 = SetoidPartialOrder.<WellDefined pOrder reflexive (Group.+WellDefined additiveGroup reflexive (Equivalence.transitive eq (inverseWellDefined additiveGroup *Commutative) (Equivalence.transitive eq (symmetric ringMinusExtracts) *Commutative))) p0
p2 : 0R < ((y + inverse x) * c)
p2 = SetoidPartialOrder.<WellDefined pOrder reflexive (Equivalence.transitive eq (Group.+WellDefined additiveGroup *Commutative *Commutative) (Equivalence.transitive eq (symmetric *DistributesOver+) *Commutative)) p1
q : (y + inverse x) < 0R
@@ -243,3 +245,54 @@ absNegation a | inr 0=a with totality 0G (inverse a)
absNegation a | inr 0=a | inl (inl 0<-a) = exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (inverseWellDefined additiveGroup (Equivalence.symmetric eq 0=a)) (invIdent additiveGroup)) 0<-a))
absNegation a | inr 0=a | inl (inr -a<0) = exFalso (irreflexive {0G} (<WellDefined (Equivalence.transitive eq (inverseWellDefined additiveGroup (Equivalence.symmetric eq 0=a)) (invIdent additiveGroup)) (Equivalence.reflexive eq) -a<0))
absNegation a | inr 0=a | inr 0=-a = Equivalence.transitive eq (Equivalence.symmetric eq 0=a) 0=-a
lemm4 : (a b : A) (0G < a) (b < 0G) (a * b) < 0G
lemm4 a b 0<a b<0 with orderRespectsMultiplication 0<a (lemm2 _ b<0)
... | bl = <WellDefined (invTwice additiveGroup _) (Equivalence.reflexive eq) (lemm2' _ (<WellDefined (Equivalence.reflexive eq) ringMinusExtracts bl))
lemm5 : (a b : A) (a < 0G) (b < 0G) 0G < (a * b)
lemm5 a b a<0 b<0 with orderRespectsMultiplication (lemm2 _ a<0) (lemm2 _ b<0)
... | bl = <WellDefined (Equivalence.reflexive eq) twoNegativesTimes bl
absRespectsTimes : (a b : A) abs (a * b) (abs a) * (abs b)
absRespectsTimes a b with totality 0R a
absRespectsTimes a b | inl (inl 0<a) with totality 0R b
absRespectsTimes a b | inl (inl 0<a) | inl (inl 0<b) with totality 0R (a * b)
absRespectsTimes a b | inl (inl 0<a) | inl (inl 0<b) | inl (inl 0<ab) = Equivalence.reflexive eq
absRespectsTimes a b | inl (inl 0<a) | inl (inl 0<b) | inl (inr ab<0) = exFalso (irreflexive {0G} (SetoidPartialOrder.transitive pOrder (orderRespectsMultiplication 0<a 0<b) ab<0))
absRespectsTimes a b | inl (inl 0<a) | inl (inl 0<b) | inr 0=ab = exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=ab) (orderRespectsMultiplication 0<a 0<b)))
absRespectsTimes a b | inl (inl 0<a) | inl (inr b<0) with totality 0R (a * b)
absRespectsTimes a b | inl (inl 0<a) | inl (inr b<0) | inl (inl 0<ab) with <WellDefined (Equivalence.reflexive eq) ringMinusExtracts (orderRespectsMultiplication 0<a (lemm2 b b<0))
... | bl = exFalso (irreflexive {0G} (SetoidPartialOrder.transitive pOrder 0<ab (<WellDefined (invTwice additiveGroup _) (Equivalence.reflexive eq) (lemm2' _ bl))))
absRespectsTimes a b | inl (inl 0<a) | inl (inr b<0) | inl (inr ab<0) = Equivalence.symmetric eq ringMinusExtracts
absRespectsTimes a b | inl (inl 0<a) | inl (inr b<0) | inr 0=ab = exFalso (irreflexive {0G} (<WellDefined (Equivalence.symmetric eq 0=ab) (Equivalence.reflexive eq) (lemm4 a b 0<a b<0)))
absRespectsTimes a b | inl (inl 0<a) | inr 0=b with totality 0R (a * b)
absRespectsTimes a b | inl (inl 0<a) | inr 0=b | inl (inl 0<ab) = exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) (timesZero {a})) 0<ab))
absRespectsTimes a b | inl (inl 0<a) | inr 0=b | inl (inr ab<0) = exFalso ((irreflexive {0G} (<WellDefined (Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) (timesZero {a})) (Equivalence.reflexive eq) ab<0)))
absRespectsTimes a b | inl (inl 0<a) | inr 0=b | inr 0=ab = Equivalence.reflexive eq
absRespectsTimes a b | inl (inr a<0) with totality 0R b
absRespectsTimes a b | inl (inr a<0) | inl (inl 0<b) with totality 0R (a * b)
absRespectsTimes a b | inl (inr a<0) | inl (inl 0<b) | inl (inl 0<ab) = exFalso (irreflexive {0G} (SetoidPartialOrder.transitive pOrder 0<ab (<WellDefined *Commutative (Equivalence.reflexive eq) (lemm4 b a 0<b a<0))))
absRespectsTimes a b | inl (inr a<0) | inl (inl 0<b) | inl (inr ab<0) = Equivalence.symmetric eq ringMinusExtracts'
absRespectsTimes a b | inl (inr a<0) | inl (inl 0<b) | inr 0=ab = exFalso (irreflexive {0G} (<WellDefined (Equivalence.symmetric eq (Equivalence.transitive eq 0=ab *Commutative)) (Equivalence.reflexive eq) (lemm4 b a 0<b a<0)))
absRespectsTimes a b | inl (inr a<0) | inl (inr b<0) with totality 0R (a * b)
absRespectsTimes a b | inl (inr a<0) | inl (inr b<0) | inl (inl 0<ab) = Equivalence.symmetric eq twoNegativesTimes
absRespectsTimes a b | inl (inr a<0) | inl (inr b<0) | inl (inr ab<0) = exFalso (irreflexive {0G} (SetoidPartialOrder.transitive pOrder (lemm5 a b a<0 b<0) ab<0))
absRespectsTimes a b | inl (inr a<0) | inl (inr b<0) | inr 0=ab = exFalso (exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=ab) (lemm5 a b a<0 b<0))))
absRespectsTimes a b | inl (inr a<0) | inr 0=b with totality 0R (a * b)
absRespectsTimes a b | inl (inr a<0) | inr 0=b | inl (inl 0<ab) = exFalso (irreflexive {0R} (<WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) (timesZero {a})) 0<ab))
absRespectsTimes a b | inl (inr a<0) | inr 0=b | inl (inr ab<0) = exFalso (irreflexive {0R} (<WellDefined (Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) timesZero) (Equivalence.reflexive eq) ab<0))
absRespectsTimes a b | inl (inr a<0) | inr 0=b | inr 0=ab = Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) (Equivalence.transitive eq (Equivalence.transitive eq timesZero (Equivalence.symmetric eq timesZero)) (*WellDefined (Equivalence.reflexive eq) 0=b))
absRespectsTimes a b | inr 0=a with totality 0R b
absRespectsTimes a b | inr 0=a | inl (inl 0<b) with totality 0R (a * b)
absRespectsTimes a b | inr 0=a | inl (inl 0<b) | inl (inl 0<ab) = Equivalence.reflexive eq
absRespectsTimes a b | inr 0=a | inl (inl 0<b) | inl (inr ab<0) = Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.transitive eq (inverseWellDefined additiveGroup (Equivalence.transitive eq (*WellDefined (Equivalence.symmetric eq 0=a) (Equivalence.reflexive eq)) (Equivalence.transitive eq *Commutative timesZero))) (invIdent additiveGroup)) (Equivalence.transitive eq (Equivalence.symmetric eq timesZero) *Commutative)) (*WellDefined 0=a (Equivalence.reflexive eq))
absRespectsTimes a b | inr 0=a | inl (inl 0<b) | inr 0=ab = Equivalence.reflexive eq
absRespectsTimes a b | inr 0=a | inl (inr b<0) with totality 0R (a * b)
absRespectsTimes a b | inr 0=a | inl (inr b<0) | inl (inl 0<ab) = Equivalence.transitive eq (Equivalence.transitive eq (*WellDefined (Equivalence.symmetric eq 0=a) (Equivalence.reflexive eq)) *Commutative) (Equivalence.transitive eq timesZero (Equivalence.transitive eq (Equivalence.symmetric eq (Equivalence.transitive eq *Commutative timesZero)) (*WellDefined 0=a (Equivalence.reflexive eq))))
absRespectsTimes a b | inr 0=a | inl (inr b<0) | inl (inr ab<0) = Equivalence.symmetric eq ringMinusExtracts
absRespectsTimes a b | inr 0=a | inl (inr b<0) | inr 0=ab = Equivalence.transitive eq (Equivalence.transitive eq (*WellDefined (Equivalence.symmetric eq 0=a) (Equivalence.reflexive eq)) *Commutative) (Equivalence.transitive eq timesZero (Equivalence.transitive eq (Equivalence.symmetric eq (Equivalence.transitive eq *Commutative timesZero)) (*WellDefined 0=a (Equivalence.reflexive eq))))
absRespectsTimes a b | inr 0=a | inr 0=b with totality 0R (a * b)
absRespectsTimes a b | inr 0=a | inr 0=b | inl (inl 0<ab) = exFalso (irreflexive {0R} (<WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) timesZero) 0<ab))
absRespectsTimes a b | inr 0=a | inr 0=b | inl (inr ab<0) = exFalso (irreflexive {0R} (<WellDefined (Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) timesZero) (Equivalence.reflexive eq) ab<0))
absRespectsTimes a b | inr 0=a | inr 0=b | inr 0=ab = Equivalence.reflexive eq