Split out much more structure (#37)

This commit is contained in:
Patrick Stevens
2019-08-18 14:57:41 +01:00
committed by GitHub
parent 69389bb889
commit 1f26064502
52 changed files with 2137 additions and 2014 deletions

View File

@@ -2,14 +2,14 @@
-- This file contains everything that can be compiled in --safe mode.
open import Numbers.Naturals
open import Numbers.Naturals.Naturals
open import Numbers.BinaryNaturals.Definition
open import Lists.Lists
open import Groups.Groups
open import Groups.FinitePermutations
open import Groups.GroupsLemmas
open import Groups.Lemmas
open import Fields.Fields
open import Fields.FieldOfFractions
@@ -34,4 +34,8 @@ open import WellFoundedInduction
open import ClassicalLogic.ClassicalFive
open import Monoids.Definition
open import Semirings.Definition
module Everything.Safe where

View File

@@ -27,5 +27,7 @@ open import Rings.Examples.Proofs
open import Groups.FreeGroups
open import Groups.Groups2
open import Groups.Examples.ExampleSheet1
open import Groups.LectureNotes.Lecture1
module Everything.WithK where

View File

@@ -2,8 +2,8 @@
open import LogicalFormulae
open import Groups.Groups
open import Groups.GroupDefinition
open import Groups.GroupsLemmas
open import Groups.Definition
open import Groups.Lemmas
open import Rings.Definition
open import Rings.Lemmas
open import Rings.IntegralDomains

View File

@@ -2,8 +2,8 @@
open import LogicalFormulae
open import Groups.Groups
open import Groups.GroupDefinition
open import Groups.GroupsLemmas
open import Groups.Definition
open import Groups.Lemmas
open import Rings.Definition
open import Rings.Lemmas
open import Rings.IntegralDomains

View File

@@ -2,7 +2,7 @@
open import LogicalFormulae
open import Groups.Groups
open import Groups.GroupDefinition
open import Groups.Definition
open import Rings.Definition
open import Rings.Lemmas
open import Setoids.Setoids

View File

@@ -8,7 +8,7 @@ open import Numbers.Naturals
open import Numbers.Integers
open import Sets.FinSet
open import Groups.Groups
open import Groups.GroupDefinition
open import Groups.Definition
module Groups.CyclicGroups where

View File

@@ -4,10 +4,10 @@ open import LogicalFormulae
open import Setoids.Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals
open import Numbers.Naturals.Naturals
open import Sets.FinSet
module Groups.GroupDefinition where
module Groups.Definition where
record Group {lvl1 lvl2} {A : Set lvl1} (S : Setoid {lvl1} {lvl2} A) (_·_ : A A A) : Set (lsuc lvl1 lvl2) where
open Setoid S

View File

@@ -1,17 +1,19 @@
{-# OPTIONS --safe --warning=error #-}
open import LogicalFormulae
open import Setoids
open import Setoids.Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Naturals
open import Integers
open import FinSet
open import Groups
open import Rings
open import Fields
open import Numbers.Naturals.Naturals
open import Numbers.Integers
open import Sets.FinSet
open import Groups.Definition
open import Groups.Groups
open import Rings.Definition
open import Rings.Lemmas
open import Fields.Fields
module GroupsExampleSheet1 where
module Groups.Examples.ExampleSheet1 where
{-
Question 1: e is the unique solution of x^2 = x

View File

@@ -1,19 +1,20 @@
{-# OPTIONS --safe --warning=error #-}
open import Groups
open import Groups.Definition
open import Orders
open import Integers
open import Setoids
open import Numbers.Integers
open import Setoids.Setoids
open import LogicalFormulae
open import FinSet
open import Sets.FinSet
open import Functions
open import Naturals
open import Numbers.Naturals
open import IntegersModN
open import RingExamples
open import Rings.Examples.Examples
open import PrimeNumbers
open import Groups2
open import Groups.Groups2
open import Groups.CyclicGroups
module GroupsExamples where
module Groups.Examples.Examples where
elementPowZ : (n : ) (elementPower Group (nonneg 1) n) nonneg n
elementPowZ zero = refl
elementPowZ (succ n) rewrite elementPowZ n = refl

View File

@@ -1,7 +1,7 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Numbers.Naturals -- for length
open import Numbers.Naturals.Naturals -- for length
open import Lists.Lists
open import Functions
open import Groups.SymmetryGroups

View File

@@ -5,10 +5,10 @@ open import WithK
open import Setoids.Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals
open import Numbers.NaturalsWithK
open import Numbers.Naturals.Naturals
open import Numbers.Naturals.WithK
open import Sets.FinSet
open import Groups.GroupDefinition
open import Groups.Definition
open import Groups.SymmetryGroups
open import DecidableSet

View File

@@ -4,9 +4,9 @@ open import LogicalFormulae
open import Setoids.Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals
open import Numbers.Naturals.Naturals
open import Sets.FinSet
open import Groups.GroupDefinition
open import Groups.Definition
module Groups.Groups where

View File

@@ -1,14 +1,14 @@
{-# OPTIONS --safe --warning=error #-}
open import Groups.Groups
open import Groups.GroupDefinition
open import Groups.Definition
open import Orders
open import Numbers.Integers
open import Setoids.Setoids
open import LogicalFormulae
open import Sets.FinSet
open import Functions
open import Numbers.Naturals
open import Numbers.Naturals.Naturals
open import IntegersModN
open import Rings.Examples.Examples
open import PrimeNumbers

View File

@@ -4,13 +4,13 @@ open import LogicalFormulae
open import Setoids.Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals
open import Numbers.Naturals.Naturals
open import Numbers.Integers
open import Numbers.Rationals
open import Sets.FinSet
open import Groups.GroupDefinition
open import Groups.Definition
open import Groups.Groups
open import Rings.RingDefinition
open import Rings.Definition
open import IntegersModN
module Groups.LectureNotes.Lecture1 where

View File

@@ -4,11 +4,11 @@ open import LogicalFormulae
open import Setoids.Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals
open import Numbers.Naturals.Naturals
open import Groups.Groups
open import Groups.GroupDefinition
open import Groups.Definition
module Groups.GroupsLemmas where
module Groups.Lemmas where
invInv : {a b : _} {A : Set a} {_·_ : A A A} {S : Setoid {a} {b} A} (G : Group S _·_) {x : A} Setoid.__ S (Group.inverse G (Group.inverse G x)) x
invInv {S = S} G {x} = symmetric (transferToRight' G invRight)
where

View File

@@ -4,10 +4,10 @@ open import LogicalFormulae
open import Setoids.Setoids
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals
open import Numbers.Naturals.Naturals
open import Sets.FinSet
open import Groups.Groups
open import Groups.GroupDefinition
open import Groups.Definition
module Groups.SymmetryGroups where
data SymmetryGroupElements {a b : _} {A : Set a} (S : Setoid {a} {b} A) : Set (a b) where

View File

@@ -1,15 +1,18 @@
{-# OPTIONS --safe --warning=error #-}
open import LogicalFormulae
open import Groups.GroupDefinition
open import Groups.Definition
open import Groups.Groups
open import Numbers.Naturals
open import Numbers.Naturals.Naturals
open import Numbers.Naturals.Addition -- TODO remove this dependency
open import PrimeNumbers
open import Setoids.Setoids
open import Sets.FinSet
open import Sets.FinSetWithK
open import Functions
open import Numbers.NaturalsWithK
open import Numbers.Naturals.WithK
open import Semirings.Definition
open import Orders
module IntegersModN where
record n (n : ) (pr : 0 <N n) : Set where
@@ -24,7 +27,7 @@ module IntegersModN where
cancelSumFromInequality {a} {b} {c} (le x proof) = le x help
where
help : succ x +N b c
help rewrite equalityCommutative (additionNIsAssociative (succ x) a b) | additionNIsCommutative x a | additionNIsAssociative (succ a) x b | additionNIsCommutative a (x +N b) | additionNIsCommutative (succ (x +N b)) a = canSubtractFromEqualityLeft {a} proof
help rewrite Semiring.+Associative Semiring (succ x) a b | Semiring.commutative Semiring x a | equalityCommutative (Semiring.+Associative Semiring (succ a) x b) | Semiring.commutative Semiring a (x +N b) | Semiring.commutative Semiring (succ (x +N b)) a = canSubtractFromEqualityLeft {a} proof
_+n_ : {n : } {pr : 0 <N n} n n pr n n pr n n pr
_+n_ {0} {le x ()} a b
@@ -45,15 +48,15 @@ module IntegersModN where
plusZnIdentityRight : {n : } {pr : 0 <N n} (a : n n pr) (a +n record { x = 0 ; xLess = pr }) a
plusZnIdentityRight {zero} {()} a
plusZnIdentityRight {succ x} {_} record { x = a ; xLess = aLess } with orderIsTotal (a +N 0) (succ x)
plusZnIdentityRight {succ x} {_} record { x = a ; xLess = aLess } | inl (inl a<sx) = equalityZn _ _ (additionNIsCommutative a 0)
plusZnIdentityRight {succ x} {_} record { x = a ; xLess = aLess } | inl (inl a<sx) = equalityZn _ _ (Semiring.commutative Semiring a 0)
plusZnIdentityRight {succ x} {_} record { x = a ; xLess = aLess } | inl (inr sx<a) = exFalso (f aLess sx<a)
where
f : (aL : a <N succ x) (sx<a : succ x <N a +N 0) False
f aL sx<a rewrite additionNIsCommutative a 0 = orderIsIrreflexive aL sx<a
f aL sx<a rewrite Semiring.commutative Semiring a 0 = orderIsIrreflexive aL sx<a
plusZnIdentityRight {succ x} {_} record { x = a ; xLess = aLess } | inr a=sx = exFalso (f aLess a=sx)
where
f : (aL : a <N succ x) (a=sx : a +N 0 succ x) False
f aL a=sx rewrite additionNIsCommutative a 0 | a=sx = lessIrreflexive aL
f aL a=sx rewrite Semiring.commutative Semiring a 0 | a=sx = TotalOrder.irreflexive TotalOrder aL
plusZnIdentityLeft : {n : } {pr : 0 <N n} (a : n n pr) (record { x = 0 ; xLess = pr }) +n a a
@@ -61,7 +64,7 @@ module IntegersModN where
plusZnIdentityLeft {succ n} {pr} record { x = x ; xLess = xLess } with orderIsTotal x (succ n)
plusZnIdentityLeft {succ n} {pr} record { x = x ; xLess = xLess } | inl (inl x<succn) rewrite <NRefl x<succn xLess = refl
plusZnIdentityLeft {succ n} {pr} record { x = x ; xLess = xLess } | inl (inr succn<x) = exFalso (cannotBeLeqAndG {x} {succ n} (inl xLess) succn<x)
plusZnIdentityLeft {succ n} {pr} record { x = x ; xLess = xLess } | inr x=succn rewrite x=succn = exFalso (lessIrreflexive xLess)
plusZnIdentityLeft {succ n} {pr} record { x = x ; xLess = xLess } | inr x=succn rewrite x=succn = exFalso (TotalOrder.irreflexive TotalOrder xLess)
subLemma : {a b c : } (pr1 : a <N b) (pr2 : a <N c) (pr : b c) (subtractionNResult.result (-N (inl pr1))) (subtractionNResult.result (-N (inl pr2)))
subLemma {a} {b} {c} a<b a<c b=c rewrite b=c | <NRefl a<b a<c = refl
@@ -70,47 +73,47 @@ module IntegersModN where
plusZnCommutative {zero} {()} a b
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } with orderIsTotal (a +N b) (succ n)
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inl a+b<sn) with orderIsTotal (b +N a) (succ n)
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inl a+b<sn) | inl (inl b+a<sn) = equalityZn _ _ (additionNIsCommutative a b)
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inl a+b<sn) | inl (inl b+a<sn) = equalityZn _ _ (Semiring.commutative Semiring a b)
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inl a+b<sn) | inl (inr sn<b+a) = exFalso (f a+b<sn sn<b+a)
where
f : (a +N b) <N succ n succ n <N b +N a False
f pr1 pr2 rewrite additionNIsCommutative b a = lessIrreflexive (orderIsTransitive pr2 pr1)
f pr1 pr2 rewrite Semiring.commutative Semiring b a = TotalOrder.irreflexive TotalOrder (orderIsTransitive pr2 pr1)
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inl a+b<sn) | inr b+a=sn = exFalso (f a+b<sn b+a=sn)
where
f : (a +N b) <N succ n b +N a succ n False
f pr1 eq rewrite additionNIsCommutative b a | eq = lessIrreflexive pr1
f pr1 eq rewrite Semiring.commutative Semiring b a | eq = TotalOrder.irreflexive TotalOrder pr1
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inr sn<a+b) with orderIsTotal (b +N a) (succ n)
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inr sn<a+b) | inl (inl b+a<sn) = exFalso (f sn<a+b b+a<sn)
where
f : succ n <N a +N b b +N a <N succ n False
f pr1 pr2 rewrite additionNIsCommutative b a = lessIrreflexive (orderIsTransitive sn<a+b b+a<sn)
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inr sn<a+b) | inl (inr sn<b+a) = equalityZn _ _ (subLemma sn<a+b sn<b+a (additionNIsCommutative a b))
f pr1 pr2 rewrite Semiring.commutative Semiring b a = TotalOrder.irreflexive TotalOrder (orderIsTransitive sn<a+b b+a<sn)
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inr sn<a+b) | inl (inr sn<b+a) = equalityZn _ _ (subLemma sn<a+b sn<b+a (Semiring.commutative Semiring a b))
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inr sn<a+b) | inr sn=b+a = exFalso (f sn<a+b sn=b+a)
where
f : succ n <N a +N b b +N a succ n False
f pr eq rewrite additionNIsCommutative b a | eq = lessIrreflexive pr
f pr eq rewrite Semiring.commutative Semiring b a | eq = TotalOrder.irreflexive TotalOrder pr
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inr sn=a+b with orderIsTotal (b +N a) (succ n)
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inr sn=a+b | inl (inl b+a<sn) = exFalso f
where
f : False
f rewrite additionNIsCommutative b a | sn=a+b = lessIrreflexive b+a<sn
f rewrite Semiring.commutative Semiring b a | sn=a+b = TotalOrder.irreflexive TotalOrder b+a<sn
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inr sn=a+b | inl (inr sn<b+a) = exFalso f
where
f : False
f rewrite additionNIsCommutative b a | sn=a+b = lessIrreflexive sn<b+a
f rewrite Semiring.commutative Semiring b a | sn=a+b = TotalOrder.irreflexive TotalOrder sn<b+a
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inr sn=a+b | inr sn=b+a = equalityZn _ _ refl
help30 : {a b c n : } (c <N n) (a +N b n) (n<b+c : n <N b +N c) (n <N a +N subtractionNResult.result (-N (inl n<b+c))) False
help30 {a} {b} {c} {n} c<n a+b=n n<b+c x = lessIrreflexive (orderIsTransitive pr5 c<n)
help30 {a} {b} {c} {n} c<n a+b=n n<b+c x = TotalOrder.irreflexive TotalOrder (orderIsTransitive pr5 c<n)
where
pr : n +N n <N a +N (subtractionNResult.result (-N (inl n<b+c)) +N n)
pr = identityOfIndiscernablesRight _ _ _ _<N_ (additionPreservesInequality n x) (additionNIsAssociative a _ n)
pr = identityOfIndiscernablesRight _ _ _ _<N_ (additionPreservesInequality n x) (equalityCommutative (Semiring.+Associative Semiring a _ n))
pr2 : n +N n <N a +N (b +N c)
pr2 = identityOfIndiscernablesRight _ _ _ _<N_ pr (applyEquality (a +N_) (addMinus (inl n<b+c)))
pr3 : n +N n <N (a +N b) +N c
pr3 rewrite additionNIsAssociative a b c = pr2
pr3 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) = pr2
pr4 : n +N n <N c +N n
pr4 rewrite additionNIsCommutative c n = identityOfIndiscernablesRight _ _ _ _<N_ pr3 (applyEquality (_+N c) a+b=n)
pr4 rewrite Semiring.commutative Semiring c n = identityOfIndiscernablesRight _ _ _ _<N_ pr3 (applyEquality (_+N c) a+b=n)
pr5 : n <N c
pr5 = subtractionPreservesInequality n pr4
@@ -118,34 +121,34 @@ module IntegersModN where
help31 {a} {b} {c} {n} a+b=n n<b+c = canSubtractFromEqualityLeft pr2
where
pr1 : a +N (n +N subtractionNResult.result (-N (inl n<b+c))) n +N c
pr1 rewrite addMinus' (inl n<b+c) | equalityCommutative (additionNIsAssociative a b c) | a+b=n = refl
pr1 rewrite addMinus' (inl n<b+c) | Semiring.+Associative Semiring a b c | a+b=n = refl
pr2 : n +N (a +N subtractionNResult.result (-N (inl n<b+c))) n +N c
pr2 = identityOfIndiscernablesLeft _ _ _ _≡_ pr1 (lemm a n (subtractionNResult.result (-N (inl n<b+c))))
where
lemm : (a b c : ) a +N (b +N c) b +N (a +N c)
lemm a b c rewrite equalityCommutative (additionNIsAssociative a b c) | additionNIsCommutative a b | additionNIsAssociative b a c = refl
lemm a b c rewrite Semiring.+Associative Semiring a b c | Semiring.commutative Semiring a b | equalityCommutative (Semiring.+Associative Semiring b a c) = refl
help7 : {a b c n : } b +N c n a <N n (n<a+b : n <N a +N b) (subtractionNResult.result (-N (inl n<a+b)) +N c n) False
help7 {a} {b} {c} {n} b+c=n a<n n<a+b x = lessIrreflexive (identityOfIndiscernablesLeft _ _ _ _<N_ a<n pr5)
help7 {a} {b} {c} {n} b+c=n a<n n<a+b x = TotalOrder.irreflexive TotalOrder (identityOfIndiscernablesLeft _ _ _ _<N_ a<n pr5)
where
pr : (n +N subtractionNResult.result (-N (inl n<a+b))) +N c n +N n
pr = identityOfIndiscernablesLeft _ _ _ _≡_ (applyEquality (n +N_) x) (equalityCommutative (additionNIsAssociative n _ c))
pr = identityOfIndiscernablesLeft _ _ _ _≡_ (applyEquality (n +N_) x) (Semiring.+Associative Semiring n _ c)
pr2 : (a +N b) +N c n +N n
pr2 = identityOfIndiscernablesLeft _ _ _ _≡_ pr (applyEquality (_+N c) (addMinus' (inl n<a+b)))
pr3 : a +N (b +N c) n +N n
pr3 rewrite equalityCommutative (additionNIsAssociative a b c) = pr2
pr3 rewrite Semiring.+Associative Semiring a b c = pr2
pr4 : a +N n n +N n
pr4 = identityOfIndiscernablesLeft _ _ _ _≡_ pr3 (applyEquality (a +N_) b+c=n)
pr5 : a n
pr5 = canSubtractFromEqualityRight pr4
help29 : {a b c n : } (c <N n) (n<b+c : n <N b +N c) (a +N subtractionNResult.result (-N (inl n<b+c))) n a +N b n False
help29 {a} {b} {c} {n} c<n n<b+c pr a+b=n = lessIrreflexive (identityOfIndiscernablesLeft _ _ _ _<N_ c<n p4)
help29 {a} {b} {c} {n} c<n n<b+c pr a+b=n = TotalOrder.irreflexive TotalOrder (identityOfIndiscernablesLeft _ _ _ _<N_ c<n p4)
where
p1 : a +N (subtractionNResult.result (-N (inl n<b+c)) +N n) n +N n
p1 = identityOfIndiscernablesLeft _ _ _ _≡_ (applyEquality (_+N n) pr) (additionNIsAssociative a _ n)
p1 = identityOfIndiscernablesLeft _ _ _ _≡_ (applyEquality (_+N n) pr) (equalityCommutative (Semiring.+Associative Semiring a _ n))
p2 : (a +N b) +N c n +N n
p2 rewrite additionNIsAssociative a b c = identityOfIndiscernablesLeft _ _ _ _≡_ p1 (applyEquality (a +N_) (addMinus (inl n<b+c)))
p2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) = identityOfIndiscernablesLeft _ _ _ _≡_ p1 (applyEquality (a +N_) (addMinus (inl n<b+c)))
p3 : n +N c n +N n
p3 = identityOfIndiscernablesLeft _ _ _ _≡_ p2 (applyEquality (_+N c) a+b=n)
p4 : c n
@@ -155,288 +158,288 @@ module IntegersModN where
help28 {a} {b} {c} {n} pr a+b=n = canSubtractFromEqualityLeft p2
where
p1 : a +N (b +N c) n +N c
p1 rewrite equalityCommutative (additionNIsAssociative a b c) | a+b=n = refl
p1 rewrite Semiring.+Associative Semiring a b c | a+b=n = refl
p2 : n +N subtractionNResult.result (-N (inl pr)) n +N c
p2 = identityOfIndiscernablesLeft _ _ _ _≡_ p1 (equalityCommutative (addMinus' (inl pr)))
help27 : {a b c n : } (a +N b succ n) (a +N (b +N c) <N succ n) a +N (b +N c) c
help27 {a} {b} {zero} {n} a+b=sn a+b+c<sn rewrite additionNIsCommutative b 0 | a+b=sn = exFalso (lessIrreflexive a+b+c<sn)
help27 {a} {b} {succ c} {n} a+b=sn a+b+c<sn rewrite equalityCommutative (additionNIsAssociative a b (succ c)) | a+b=sn = exFalso (cannotAddAndEnlarge' a+b+c<sn)
help27 {a} {b} {zero} {n} a+b=sn a+b+c<sn rewrite Semiring.commutative Semiring b 0 | a+b=sn = exFalso (TotalOrder.irreflexive TotalOrder a+b+c<sn)
help27 {a} {b} {succ c} {n} a+b=sn a+b+c<sn rewrite Semiring.+Associative Semiring a b (succ c) | a+b=sn = exFalso (cannotAddAndEnlarge' a+b+c<sn)
help26 : {a b c n : } (a +N b n) (a +N (b +N c) n) 0 c
help26 {a} {b} {c} {n} a+b=n a+b+c=n rewrite (equalityCommutative (additionNIsAssociative a b c)) | a+b=n | additionNIsCommutative n c = equalityCommutative (canSubtractFromEqualityRight a+b+c=n)
help26 {a} {b} {c} {n} a+b=n a+b+c=n rewrite Semiring.+Associative Semiring a b c | a+b=n | Semiring.commutative Semiring n c = equalityCommutative (canSubtractFromEqualityRight a+b+c=n)
help25 : {a b c n : } (a +N b n) (b +N c n) (a +N 0 c)
help25 {a} {b} {c} {n} a+b=n b+c=n rewrite additionNIsCommutative a 0 | additionNIsCommutative a b | equalityCommutative a+b=n = equalityCommutative (canSubtractFromEqualityLeft b+c=n)
help25 {a} {b} {c} {n} a+b=n b+c=n rewrite Semiring.commutative Semiring a 0 | Semiring.commutative Semiring a b | equalityCommutative a+b=n = equalityCommutative (canSubtractFromEqualityLeft b+c=n)
help24 : {a n : } (a <N n) (n <N a +N 0) False
help24 {a} {n} a<n n<a+0 rewrite additionNIsCommutative a 0 = lessIrreflexive (orderIsTransitive a<n n<a+0)
help24 {a} {n} a<n n<a+0 rewrite Semiring.commutative Semiring a 0 = TotalOrder.irreflexive TotalOrder (orderIsTransitive a<n n<a+0)
help23 : {a n : } (a <N n) (a +N 0 n) False
help23 {a} {n} a<n a+0=n rewrite additionNIsCommutative a 0 | a+0=n = lessIrreflexive a<n
help23 {a} {n} a<n a+0=n rewrite Semiring.commutative Semiring a 0 | a+0=n = TotalOrder.irreflexive TotalOrder a<n
help22 : {a b c n : } (a +N b n) (c n) (n<b+c : n <N b +N c) (n <N a +N subtractionNResult.result (-N (inl n<b+c))) False
help22 {a} {b} {c} {.c} a+b=n refl n<b+c pr = lessIrreflexive (identityOfIndiscernablesRight _ _ _ _<N_ pr4 a+b=n)
help22 {a} {b} {c} {.c} a+b=n refl n<b+c pr = TotalOrder.irreflexive TotalOrder (identityOfIndiscernablesRight _ _ _ _<N_ pr4 a+b=n)
where
pr1 : c +N c <N a +N (subtractionNResult.result (-N (inl n<b+c)) +N c)
pr1 = identityOfIndiscernablesRight _ _ _ _<N_ (additionPreservesInequality c pr) (additionNIsAssociative a _ c)
pr1 = identityOfIndiscernablesRight _ _ _ _<N_ (additionPreservesInequality c pr) (equalityCommutative (Semiring.+Associative Semiring a _ c))
pr2 : c +N c <N a +N (b +N c)
pr2 = identityOfIndiscernablesRight _ _ _ _<N_ pr1 (applyEquality (a +N_) (addMinus (inl n<b+c)))
pr3 : c +N c <N (a +N b) +N c
pr3 = identityOfIndiscernablesRight _ _ _ _<N_ pr2 (equalityCommutative (additionNIsAssociative a b c))
pr3 = identityOfIndiscernablesRight _ _ _ _<N_ pr2 (Semiring.+Associative Semiring a b c)
pr4 : c <N a +N b
pr4 = subtractionPreservesInequality c pr3
help21 : {a b c n : } (a +N b n) (b +N c n) (c n) (a <N n) False
help21 {a} {b} {c} {.c} a+b=n b+c=n refl a<n = lessIrreflexive (identityOfIndiscernablesLeft _ _ _ _<N_ a<n a=c)
help21 {a} {b} {c} {.c} a+b=n b+c=n refl a<n = TotalOrder.irreflexive TotalOrder (identityOfIndiscernablesLeft _ _ _ _<N_ a<n a=c)
where
b=0 : b 0
a=c : a c
a=c = identityOfIndiscernablesLeft _ _ _ _≡_ a+b=n lemm
where
lemm : a +N b a
lemm rewrite b=0 | additionNIsCommutative a 0 = refl
lemm rewrite b=0 | Semiring.commutative Semiring a 0 = refl
b=0' : (b c : ) (b +N c c) b 0
b=0' zero c b+c=c = refl
b=0' (succ b) zero b+c=c = naughtE (equalityCommutative b+c=c)
b=0' (succ b) (succ c) b+c=c = b=0' (succ b) c bl2
where
bl2 : succ b +N c c
bl2 rewrite additionNIsCommutative b c | additionNIsCommutative (succ c) b = succInjective b+c=c
bl2 rewrite Semiring.commutative Semiring b c | Semiring.commutative Semiring (succ c) b = succInjective b+c=c
b=0 = b=0' b c b+c=n
help20 : {a b c n : } (c n) (a +N b n) (n<b+c : n <N b +N c) (a +N subtractionNResult.result (-N (inl n<b+c)) <N n) False
help20 {a} {b} {c} {n} c=n a+b=n n<b+c pr = lessIrreflexive (identityOfIndiscernablesLeft _ _ _ _<N_ pr5 c=n)
help20 {a} {b} {c} {n} c=n a+b=n n<b+c pr = TotalOrder.irreflexive TotalOrder (identityOfIndiscernablesLeft _ _ _ _<N_ pr5 c=n)
where
pr1 : a +N (subtractionNResult.result (-N (inl n<b+c)) +N n) <N n +N n
pr1 = identityOfIndiscernablesLeft _ _ _ _<N_ (additionPreservesInequality n pr) (additionNIsAssociative a _ n)
pr1 = identityOfIndiscernablesLeft _ _ _ _<N_ (additionPreservesInequality n pr) (equalityCommutative (Semiring.+Associative Semiring a _ n))
pr2 : a +N (b +N c) <N n +N n
pr2 = identityOfIndiscernablesLeft _ _ _ _<N_ pr1 (applyEquality (a +N_) (addMinus (inl n<b+c)))
pr3 : (a +N b) +N c <N n +N n
pr3 rewrite additionNIsAssociative a b c = pr2
pr3 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) = pr2
pr4 : c +N n <N n +N n
pr4 = identityOfIndiscernablesLeft _ _ _ _<N_ pr3 (transitivity (applyEquality (_+N c) a+b=n) (additionNIsCommutative n c))
pr4 = identityOfIndiscernablesLeft _ _ _ _<N_ pr3 (transitivity (applyEquality (_+N c) a+b=n) (Semiring.commutative Semiring n c))
pr5 : c <N n
pr5 = subtractionPreservesInequality n pr4
help19 : {a b c n : } (b+c<n : b +N c <N n) (n<a+b : n <N a +N b) (a <N n) (subtractionNResult.result (-N (inl n<a+b)) +N c n) False
help19 {a} {b} {c} {n} b+c<n n<a+b a<n pr = lessIrreflexive (identityOfIndiscernablesLeft _ _ _ _<N_ r q')
help19 {a} {b} {c} {n} b+c<n n<a+b a<n pr = TotalOrder.irreflexive TotalOrder (identityOfIndiscernablesLeft _ _ _ _<N_ r q')
where
p : (n +N subtractionNResult.result (-N (inl n<a+b))) +N c n +N n
p = identityOfIndiscernablesLeft _ _ _ _≡_ (applyEquality (n +N_ ) pr) (equalityCommutative (additionNIsAssociative n _ c))
p = identityOfIndiscernablesLeft _ _ _ _≡_ (applyEquality (n +N_ ) pr) (Semiring.+Associative Semiring n _ c)
q : (a +N b) +N c n +N n
q = identityOfIndiscernablesLeft _ _ _ _≡_ p (applyEquality (_+N c) (addMinus' (inl n<a+b)))
q' : a +N (b +N c) n +N n
q' = identityOfIndiscernablesLeft _ _ _ _≡_ q (additionNIsAssociative a b c)
q' = identityOfIndiscernablesLeft _ _ _ _≡_ q (equalityCommutative (Semiring.+Associative Semiring a b c))
r : a +N (b +N c) <N n +N n
r = addStrongInequalities a<n b+c<n
help18 : {a b c n : } (b+c<n : b +N c <N n) (n<a+b : n <N a +N b) (a <N n) (n <N subtractionNResult.result (-N (inl n<a+b)) +N c) False
help18 {a} {b} {c} {n} b+c<n n<a+b a<n pr = lessIrreflexive (orderIsTransitive p4 a<n)
help18 {a} {b} {c} {n} b+c<n n<a+b a<n pr = TotalOrder.irreflexive TotalOrder (orderIsTransitive p4 a<n)
where
p : n +N n <N (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
p = identityOfIndiscernablesRight _ _ _ _<N_ (additionPreservesInequalityOnLeft n pr) (equalityCommutative (additionNIsAssociative n _ c))
p = identityOfIndiscernablesRight _ _ _ _<N_ (additionPreservesInequalityOnLeft n pr) (Semiring.+Associative Semiring n _ c)
p' : n +N n <N (a +N b) +N c
p' = identityOfIndiscernablesRight _ _ _ _<N_ p (applyEquality (_+N c) (addMinus' (inl n<a+b)))
p2 : n +N n <N a +N (b +N c)
p2 = identityOfIndiscernablesRight _ _ _ _<N_ p' (additionNIsAssociative a b c)
p2 = identityOfIndiscernablesRight _ _ _ _<N_ p' (equalityCommutative (Semiring.+Associative Semiring a b c))
p3 : n +N n <N a +N n
p3 = orderIsTransitive p2 (additionPreservesInequalityOnLeft a b+c<n)
p4 : n <N a
p4 = subtractionPreservesInequality n p3
help17 : {a b c n : } (n<b+c : n <N b +N c) (n<a+b : n <N a +N b) (a +N subtractionNResult.result (-N (inl n<b+c)) <N n) (subtractionNResult.result (-N (inl n<a+b)) +N c) n False
help17 {a} {b} {c} {n} n<b+c n<a+b p1 p2 = lessIrreflexive (identityOfIndiscernablesLeft _ _ _ _<N_ pr1'' pr3)
help17 {a} {b} {c} {n} n<b+c n<a+b p1 p2 = TotalOrder.irreflexive TotalOrder (identityOfIndiscernablesLeft _ _ _ _<N_ pr1'' pr3)
where
pr1' : a +N (subtractionNResult.result (-N (inl n<b+c)) +N n) <N n +N n
pr1' = identityOfIndiscernablesLeft _ _ _ _<N_ (additionPreservesInequality n p1) (additionNIsAssociative a _ n)
pr1' = identityOfIndiscernablesLeft _ _ _ _<N_ (additionPreservesInequality n p1) (equalityCommutative (Semiring.+Associative Semiring a _ n))
pr1'' : a +N (b +N c) <N n +N n
pr1'' = identityOfIndiscernablesLeft _ _ _ _<N_ pr1' (applyEquality (a +N_) (addMinus (inl n<b+c)))
pr2' : (n +N subtractionNResult.result (-N (inl n<a+b))) +N c n +N n
pr2' = identityOfIndiscernablesLeft _ _ _ _≡_ (applyEquality (n +N_) p2) (equalityCommutative (additionNIsAssociative n _ c))
pr2' = identityOfIndiscernablesLeft _ _ _ _≡_ (applyEquality (n +N_) p2) (Semiring.+Associative Semiring n _ c)
pr2'' : (a +N b) +N c n +N n
pr2'' = identityOfIndiscernablesLeft _ _ _ _≡_ pr2' (applyEquality (_+N c) (addMinus' (inl n<a+b)))
pr3 : a +N (b +N c) n +N n
pr3 = identityOfIndiscernablesLeft _ _ _ _≡_ pr2'' (additionNIsAssociative a b c)
pr3 = identityOfIndiscernablesLeft _ _ _ _≡_ pr2'' (equalityCommutative (Semiring.+Associative Semiring a b c))
help16 : {a b c n : } (n<b+c : n <N b +N c) (n<a+b : n <N a +N b) (a +N subtractionNResult.result (-N (inl n<b+c))) <N n (pr : n <N subtractionNResult.result (-N (inl n<a+b)) +N c) a +N subtractionNResult.result (-N (inl n<b+c)) subtractionNResult.result (-N (inl pr))
help16 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = exFalso (lessIrreflexive (orderIsTransitive pr3 pr1''))
help16 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = exFalso (TotalOrder.irreflexive TotalOrder (orderIsTransitive pr3 pr1''))
where
pr1' : a +N (subtractionNResult.result (-N (inl n<b+c)) +N n) <N n +N n
pr1' = identityOfIndiscernablesLeft _ _ _ _<N_ (additionPreservesInequality n pr1) (additionNIsAssociative a _ n)
pr1' = identityOfIndiscernablesLeft _ _ _ _<N_ (additionPreservesInequality n pr1) (equalityCommutative (Semiring.+Associative Semiring a _ n))
pr1'' : a +N (b +N c) <N n +N n
pr1'' = identityOfIndiscernablesLeft _ _ _ _<N_ pr1' (applyEquality (a +N_) (addMinus (inl n<b+c)))
pr2' : n +N n <N (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
pr2' = identityOfIndiscernablesRight _ _ _ _<N_ (additionPreservesInequalityOnLeft n pr2) (equalityCommutative (additionNIsAssociative n _ c))
pr2' = identityOfIndiscernablesRight _ _ _ _<N_ (additionPreservesInequalityOnLeft n pr2) (Semiring.+Associative Semiring n _ c)
pr2'' : n +N n <N (a +N b) +N c
pr2'' = identityOfIndiscernablesRight _ _ _ _<N_ pr2' (applyEquality (_+N c) (addMinus' (inl n<a+b)))
pr3 : n +N n <N a +N (b +N c)
pr3 = identityOfIndiscernablesRight _ _ _ _<N_ pr2'' (additionNIsAssociative a b c)
pr3 = identityOfIndiscernablesRight _ _ _ _<N_ pr2'' (equalityCommutative (Semiring.+Associative Semiring a b c))
help15 : {a b c n : } (n<b+c : n <N b +N c) (n<a+b : n <N a +N b) (n <N a +N subtractionNResult.result (-N (inl n<b+c))) (subtractionNResult.result (-N (inl n<a+b)) +N c) <N n False
help15 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = lessIrreflexive (orderIsTransitive p2'' p1')
help15 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = TotalOrder.irreflexive TotalOrder (orderIsTransitive p2'' p1')
where
p1 : (n +N subtractionNResult.result (-N (inl n<a+b))) +N c <N n +N n
p1 = identityOfIndiscernablesLeft _ _ _ _<N_ (additionPreservesInequalityOnLeft n pr2) (equalityCommutative (additionNIsAssociative n _ c))
p1 = identityOfIndiscernablesLeft _ _ _ _<N_ (additionPreservesInequalityOnLeft n pr2) (Semiring.+Associative Semiring n _ c)
p1' : (a +N b) +N c <N n +N n
p1' = identityOfIndiscernablesLeft _ _ _ _<N_ p1 (applyEquality (_+N c) (addMinus' (inl n<a+b)))
p2 : n +N n <N a +N (subtractionNResult.result (-N (inl n<b+c)) +N n)
p2 = identityOfIndiscernablesRight _ _ _ _<N_ (additionPreservesInequality n pr1) (additionNIsAssociative a _ n)
p2 = identityOfIndiscernablesRight _ _ _ _<N_ (additionPreservesInequality n pr1) (equalityCommutative (Semiring.+Associative Semiring a _ n))
p2' : n +N n <N a +N (b +N c)
p2' = identityOfIndiscernablesRight _ _ _ _<N_ p2 (applyEquality (a +N_) (addMinus (inl n<b+c)))
p2'' : n +N n <N (a +N b) +N c
p2'' = identityOfIndiscernablesRight _ _ _ _<N_ p2' (equalityCommutative (additionNIsAssociative a b c))
p2'' = identityOfIndiscernablesRight _ _ _ _<N_ p2' (Semiring.+Associative Semiring a b c)
help14 : {a b c n : } (n<b+c : n <N b +N c) (n<a+b : n <N a +N b) (pr1 : n <N a +N subtractionNResult.result (-N (inl n<b+c))) (pr2 : n <N subtractionNResult.result (-N (inl n<a+b)) +N c) subtractionNResult.result (-N (inl pr1)) subtractionNResult.result (-N (inl pr2))
help14 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = equivalentSubtraction _ _ _ _ pr1 pr2 (transitivity (equalityCommutative (additionNIsAssociative n _ c)) (transitivity (applyEquality (_+N c) (addMinus' (inl n<a+b))) (transitivity (additionNIsAssociative a b c) (equalityCommutative p2))))
help14 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = equivalentSubtraction _ _ _ _ pr1 pr2 (transitivity (Semiring.+Associative Semiring n _ c) (transitivity (applyEquality (_+N c) (addMinus' (inl n<a+b))) (transitivity (equalityCommutative (Semiring.+Associative Semiring a b c)) (equalityCommutative p2))))
where
p1 : (a +N subtractionNResult.result (-N (inl n<b+c))) +N n a +N (subtractionNResult.result (-N (inl n<b+c)) +N n)
p1 = additionNIsAssociative a _ n
p1 = equalityCommutative (Semiring.+Associative Semiring a _ n)
p2 : (a +N subtractionNResult.result (-N (inl n<b+c))) +N n a +N (b +N c)
p2 = identityOfIndiscernablesRight _ _ _ _≡_ p1 (applyEquality (a +N_) (addMinus (inl n<b+c)))
help13 : {a b c n : } (n<b+c : n <N b +N c) (n<a+b : n <N a +N b) (n <N a +N subtractionNResult.result (-N (inl n<b+c))) (subtractionNResult.result (-N (inl n<a+b)) +N c n) False
help13 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = lessIrreflexive (identityOfIndiscernablesRight _ _ _ _<N_ lemm1' lemm3)
help13 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = TotalOrder.irreflexive TotalOrder (identityOfIndiscernablesRight _ _ _ _<N_ lemm1' lemm3)
where
lemm1 : n +N n <N a +N (subtractionNResult.result (-N (inl n<b+c)) +N n)
lemm1 = identityOfIndiscernablesRight _ _ _ _<N_ (additionPreservesInequality n pr1) (additionNIsAssociative a _ n)
lemm1 = identityOfIndiscernablesRight _ _ _ _<N_ (additionPreservesInequality n pr1) (equalityCommutative (Semiring.+Associative Semiring a _ n))
lemm1' : n +N n <N a +N (b +N c)
lemm1' = identityOfIndiscernablesRight _ _ _ _<N_ lemm1 (applyEquality (a +N_) (addMinus (inl n<b+c)))
lemm2 : (n +N subtractionNResult.result (-N (inl n<a+b))) +N c n +N n
lemm2 = identityOfIndiscernablesLeft _ _ _ _≡_ (applyEquality (n +N_) pr2) (equalityCommutative (additionNIsAssociative n _ c))
lemm2 = identityOfIndiscernablesLeft _ _ _ _≡_ (applyEquality (n +N_) pr2) (Semiring.+Associative Semiring n _ c)
lemm2' : (a +N b) +N c n +N n
lemm2' = identityOfIndiscernablesLeft _ _ _ _≡_ lemm2 (applyEquality (_+N c) (addMinus' (inl n<a+b)))
lemm3 : a +N (b +N c) n +N n
lemm3 rewrite equalityCommutative (additionNIsAssociative a b c) = lemm2'
lemm3 rewrite Semiring.+Associative Semiring a b c = lemm2'
help12 : {a b c n : } (n<b+c : n <N b +N c) (n<a+b : n <N a +N b) (a +N subtractionNResult.result (-N (inl n<b+c))) n subtractionNResult.result (-N (inl n<a+b)) +N c <N n False
help12 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = lessIrreflexive lemm4
help12 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = TotalOrder.irreflexive TotalOrder lemm4
where
pr : {a b c : } a +N (b +N c) b +N (a +N c)
pr {a} {b} {c} rewrite equalityCommutative (additionNIsAssociative a b c) | additionNIsCommutative a b | additionNIsAssociative b a c = refl
pr {a} {b} {c} rewrite Semiring.+Associative Semiring a b c | Semiring.commutative Semiring a b | equalityCommutative (Semiring.+Associative Semiring b a c) = refl
lemm1 : (n +N subtractionNResult.result (-N (inl n<a+b))) +N c <N n +N n
lemm1 = identityOfIndiscernablesLeft _ _ _ _<N_ (additionPreservesInequalityOnLeft n pr2) (equalityCommutative (additionNIsAssociative n _ c ))
lemm1 = identityOfIndiscernablesLeft _ _ _ _<N_ (additionPreservesInequalityOnLeft n pr2) (Semiring.+Associative Semiring n _ c)
lemm2 : (a +N b) +N c <N n +N n
lemm2 = identityOfIndiscernablesLeft _ _ _ _<N_ lemm1 (applyEquality (_+N c) (addMinus' (inl n<a+b)))
lemm1' : a +N (subtractionNResult.result (-N (inl n<b+c)) +N n) n +N n
lemm1' = identityOfIndiscernablesLeft _ _ _ _≡_ (applyEquality (_+N n) pr1) (additionNIsAssociative a _ n)
lemm1' = identityOfIndiscernablesLeft _ _ _ _≡_ (applyEquality (_+N n) pr1) (equalityCommutative (Semiring.+Associative Semiring a _ n))
lemm2' : a +N (b +N c) n +N n
lemm2' = identityOfIndiscernablesLeft _ _ _ _≡_ lemm1' (applyEquality (a +N_) (addMinus (inl n<b+c)))
lemm3 : (a +N b) +N c n +N n
lemm3 rewrite additionNIsAssociative a b c = lemm2'
lemm3 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) = lemm2'
lemm4 : (a +N b) +N c <N (a +N b) +N c
lemm4 = identityOfIndiscernablesRight _ _ _ _<N_ lemm2 (equalityCommutative lemm3)
help11 : {a b c n : } (a <N n) (b +N c n) (n<a+b : n <N a +N b) (n <N subtractionNResult.result (-N (inl n<a+b)) +N c) False
help11 {a} {b} {c} {n} a<n b+c=n n<a+b pr1 = lessIrreflexive (orderIsTransitive a<n lemm5)
help11 {a} {b} {c} {n} a<n b+c=n n<a+b pr1 = TotalOrder.irreflexive TotalOrder (orderIsTransitive a<n lemm5)
where
pr : {a b c : } a +N (b +N c) b +N (a +N c)
pr {a} {b} {c} rewrite equalityCommutative (additionNIsAssociative a b c) | additionNIsCommutative a b | additionNIsAssociative b a c = refl
pr {a} {b} {c} rewrite Semiring.+Associative Semiring a b c | Semiring.commutative Semiring a b | equalityCommutative (Semiring.+Associative Semiring b a c) = refl
lemm : n +N n <N (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
lemm = identityOfIndiscernablesRight _ _ _ _<N_ (additionPreservesInequalityOnLeft n pr1) (equalityCommutative (additionNIsAssociative n _ c))
lemm = identityOfIndiscernablesRight _ _ _ _<N_ (additionPreservesInequalityOnLeft n pr1) (Semiring.+Associative Semiring n _ c)
lemm2 : n +N n <N (a +N b) +N c
lemm2 = identityOfIndiscernablesRight _ _ _ _<N_ lemm (applyEquality (_+N c) (addMinus' (inl n<a+b)))
lemm3 : n +N n <N a +N (b +N c)
lemm3 = identityOfIndiscernablesRight _ _ _ _<N_ lemm2 (additionNIsAssociative a b c)
lemm3 = identityOfIndiscernablesRight _ _ _ _<N_ lemm2 (equalityCommutative (Semiring.+Associative Semiring a b c))
lemm4 : n +N n <N a +N n
lemm4 = identityOfIndiscernablesRight _ _ _ _<N_ lemm3 (applyEquality (a +N_) b+c=n)
lemm5 : n <N a
lemm5 = subtractionPreservesInequality n lemm4
help10 : {a b c n : } (n<b+c : n <N b +N c) (n<a+b : n <N a +N b) (a +N subtractionNResult.result (-N (inl n<b+c)) n) (n <N subtractionNResult.result (-N (inl n<a+b)) +N c) False
help10 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = lessIrreflexive lemm6
help10 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = TotalOrder.irreflexive TotalOrder lemm6
where
pr : {a b c : } a +N (b +N c) b +N (a +N c)
pr {a} {b} {c} rewrite equalityCommutative (additionNIsAssociative a b c) | additionNIsCommutative a b | additionNIsAssociative b a c = refl
pr {a} {b} {c} rewrite Semiring.+Associative Semiring a b c | Semiring.commutative Semiring a b | equalityCommutative (Semiring.+Associative Semiring b a c) = refl
lemm : a +N (n +N subtractionNResult.result (-N (inl n<b+c))) n +N n
lemm = identityOfIndiscernablesLeft _ _ _ _≡_ (applyEquality (n +N_) pr1) (pr {n} {a})
lemm2 : a +N (b +N c) n +N n
lemm2 = identityOfIndiscernablesLeft _ _ _ _≡_ lemm (applyEquality (a +N_) (addMinus' (inl n<b+c)))
lemm3 : n +N n <N (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
lemm3 = identityOfIndiscernablesRight _ _ _ _<N_ (additionPreservesInequalityOnLeft n pr2) (equalityCommutative (additionNIsAssociative n _ c))
lemm3 = identityOfIndiscernablesRight _ _ _ _<N_ (additionPreservesInequalityOnLeft n pr2) (Semiring.+Associative Semiring n _ c)
lemm4 : n +N n <N (a +N b) +N c
lemm4 = identityOfIndiscernablesRight _ _ _ _<N_ lemm3 (applyEquality (_+N c) (addMinus' (inl n<a+b)))
lemm5 : n +N n <N a +N (b +N c)
lemm5 = identityOfIndiscernablesRight _ _ _ _<N_ lemm4 (additionNIsAssociative a b c)
lemm5 = identityOfIndiscernablesRight _ _ _ _<N_ lemm4 (equalityCommutative (Semiring.+Associative Semiring a b c))
lemm6 : a +N (b +N c) <N a +N (b +N c)
lemm6 = identityOfIndiscernablesLeft _ _ _ _<N_ lemm5 (equalityCommutative lemm2)
help9 : {a n : } (a +N 0 n) (a <N n) False
help9 {a} {n} n=a+0 a<n rewrite additionNIsCommutative a 0 | n=a+0 = lessIrreflexive a<n
help9 {a} {n} n=a+0 a<n rewrite Semiring.commutative Semiring a 0 | n=a+0 = TotalOrder.irreflexive TotalOrder a<n
help8 : {a n : } (n <N a +N 0) (a <N n) False
help8 {a} {n} n<a+0 a<n rewrite additionNIsCommutative a 0 = lessIrreflexive (orderIsTransitive a<n n<a+0)
help8 {a} {n} n<a+0 a<n rewrite Semiring.commutative Semiring a 0 = TotalOrder.irreflexive TotalOrder (orderIsTransitive a<n n<a+0)
help6 : {a b c n : } (b +N c n) (n<a+b : n <N a +N b) (a +N 0 subtractionNResult.result (-N (inl n<a+b)) +N c)
help6 {a} {b} {c} {n} b+c=n n<a+b rewrite additionNIsCommutative a 0 = canSubtractFromEqualityLeft {n} lem'
help6 {a} {b} {c} {n} b+c=n n<a+b rewrite Semiring.commutative Semiring a 0 = canSubtractFromEqualityLeft {n} lem'
where
lem : n +N a (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
lem rewrite addMinus' (inl n<a+b) | additionNIsAssociative a b c | b+c=n = additionNIsCommutative n a
lem rewrite addMinus' (inl n<a+b) | equalityCommutative (Semiring.+Associative Semiring a b c) | b+c=n = Semiring.commutative Semiring n a
lem' : n +N a n +N (subtractionNResult.result (-N (inl n<a+b)) +N c)
lem' = identityOfIndiscernablesRight _ _ _ _≡_ lem (additionNIsAssociative n _ c)
lem' = identityOfIndiscernablesRight _ _ _ _≡_ lem (equalityCommutative (Semiring.+Associative Semiring n _ c))
help5 : {a b c n : } (n<b+c : n <N b +N c) (n<a+b : n <N a +N b) a +N subtractionNResult.result (-N (inl n<b+c)) subtractionNResult.result (-N (inl n<a+b)) +N c
help5 {a} {b} {c} {n} n<b+c n<a+b = canSubtractFromEqualityLeft {n} lemma''
where
lemma : a +N (n +N subtractionNResult.result (-N (inl n<b+c))) (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
lemma rewrite addMinus' (inl n<b+c) | addMinus' (inl n<a+b) = equalityCommutative (additionNIsAssociative a b c)
lemma rewrite addMinus' (inl n<b+c) | addMinus' (inl n<a+b) = Semiring.+Associative Semiring a b c
lemma' : a +N (n +N subtractionNResult.result (-N (inl n<b+c))) n +N (subtractionNResult.result (-N (inl n<a+b)) +N c)
lemma' = identityOfIndiscernablesRight _ _ _ _≡_ lemma (additionNIsAssociative n _ c)
lemma' = identityOfIndiscernablesRight _ _ _ _≡_ lemma (equalityCommutative (Semiring.+Associative Semiring n _ c))
lemma'' : n +N (a +N subtractionNResult.result (-N (inl n<b+c))) n +N (subtractionNResult.result (-N (inl n<a+b)) +N c)
lemma'' = identityOfIndiscernablesLeft _ _ _ _≡_ lemma' (pr {a} {n} {subtractionNResult.result (-N (inl n<b+c))})
where
pr : {a b c : } a +N (b +N c) b +N (a +N c)
pr {a} {b} {c} rewrite equalityCommutative (additionNIsAssociative a b c) | additionNIsCommutative a b | additionNIsAssociative b a c = refl
pr {a} {b} {c} rewrite Semiring.+Associative Semiring a b c | Semiring.commutative Semiring a b | equalityCommutative (Semiring.+Associative Semiring b a c) = refl
help4 : {a b c n : } (n<a+'b+c : n <N a +N (b +N c)) (n<a+b : n <N a +N b) (subtractionNResult.result (-N (inl n<a+'b+c)) subtractionNResult.result (-N (inl n<a+b)) +N c)
help4 {a} {b} {c} {n} n<a+'b+c n<a+b = canSubtractFromEqualityLeft lemma'
where
lemma : (n +N subtractionNResult.result (-N (inl n<a+'b+c))) (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
lemma rewrite addMinus' (inl n<a+'b+c) | addMinus' (inl n<a+b) = equalityCommutative (additionNIsAssociative a b c)
lemma rewrite addMinus' (inl n<a+'b+c) | addMinus' (inl n<a+b) = Semiring.+Associative Semiring a b c
lemma' : n +N subtractionNResult.result (-N (inl n<a+'b+c)) n +N (subtractionNResult.result (-N (inl n<a+b)) +N c)
lemma' = identityOfIndiscernablesRight _ _ _ _≡_ lemma (additionNIsAssociative n (subtractionNResult.result (-N (inl n<a+b))) c)
lemma' = identityOfIndiscernablesRight _ _ _ _≡_ lemma (equalityCommutative (Semiring.+Associative Semiring n (subtractionNResult.result (-N (inl n<a+b))) c))
help3 : {a b c n : } (a <N n) (b <N n) (c <N n) (a +N b <N n) (pr : n <N b +N c) a +N subtractionNResult.result (-N (inl pr)) n False
help3 {a} {b} {c} {n} a<n b<n c<n a+b<n n<b+c pr = lessIrreflexive (orderIsTransitive (inter4 inter3) c<n)
help3 {a} {b} {c} {n} a<n b<n c<n a+b<n n<b+c pr = TotalOrder.irreflexive TotalOrder (orderIsTransitive (inter4 inter3) c<n)
where
inter : a +N (n +N subtractionNResult.result (-N (inl n<b+c))) n +N n
inter = identityOfIndiscernablesLeft _ _ _ _≡_ (applyEquality (n +N_) pr) (lemma n a (subtractionNResult.result (-N (inl n<b+c))))
where
lemma : (a b c : ) a +N (b +N c) b +N (a +N c)
lemma a b c rewrite equalityCommutative (additionNIsAssociative a b c) | equalityCommutative (additionNIsAssociative b a c) = applyEquality (_+N c) (additionNIsCommutative a b)
lemma a b c rewrite Semiring.+Associative Semiring a b c | Semiring.+Associative Semiring b a c = applyEquality (_+N c) (Semiring.commutative Semiring a b)
inter2 : n +N n a +N (b +N c)
inter2 = equalityCommutative (identityOfIndiscernablesLeft _ _ _ _≡_ inter (applyEquality (a +N_) (addMinus' (inl n<b+c))))
inter3 : n +N n <N n +N c
inter3 rewrite inter2 | equalityCommutative (additionNIsAssociative a b c) = additionPreservesInequality c a+b<n
inter3 rewrite inter2 | Semiring.+Associative Semiring a b c = additionPreservesInequality c a+b<n
inter4 : (n +N n <N n +N c) n <N c
inter4 pr rewrite additionNIsCommutative n c = subtractionPreservesInequality n pr
inter4 pr rewrite Semiring.commutative Semiring n c = subtractionPreservesInequality n pr
help2 : {a b c n : } (sn<b+c : succ n <N b +N c) (sn<a+b+c : succ n <N (a +N b) +N c) a +N subtractionNResult.result (-N (inl sn<b+c)) subtractionNResult.result (-N (inl sn<a+b+c))
help2 {a} {b} {c} {n} sn<b+c sn<a+b+c = res inter
where
inter : a +N (subtractionNResult.result (-N (inl sn<b+c)) +N succ n) subtractionNResult.result (-N (inl sn<a+b+c)) +N succ n
inter rewrite addMinus (inl sn<b+c) | addMinus (inl sn<a+b+c) = equalityCommutative (additionNIsAssociative a b c)
inter rewrite addMinus (inl sn<b+c) | addMinus (inl sn<a+b+c) = Semiring.+Associative Semiring a b c
res : (a +N (subtractionNResult.result (-N (inl sn<b+c)) +N succ n) subtractionNResult.result (-N (inl sn<a+b+c)) +N succ n) a +N subtractionNResult.result (-N (inl sn<b+c)) subtractionNResult.result (-N (inl sn<a+b+c))
res pr = canSubtractFromEqualityRight {_} {succ n} (identityOfIndiscernablesLeft _ _ _ _≡_ pr (equalityCommutative (additionNIsAssociative a (subtractionNResult.result (-N (inl sn<b+c))) (succ n))))
res pr = canSubtractFromEqualityRight {_} {succ n} (identityOfIndiscernablesLeft _ _ _ _≡_ pr (Semiring.+Associative Semiring a (subtractionNResult.result (-N (inl sn<b+c))) (succ n)))
help1 : {a b c n : } (sn<b+c : succ n <N b +N c) (pr1 : succ n <N a +N subtractionNResult.result (-N (inl sn<b+c))) (a +N b <N succ n) (a <N succ n) (b <N succ n) (c <N succ n) False
help1 {a} {b} {c} {n} sn<b+c pr1 a+b<sn a<sn b<sn c<sn with -N (inl sn<b+c)
help1 {a} {b} {c} {n} sn<b+c pr1 a+b<sn a<sn b<sn c<sn | record { result = b+c-sn ; pr = Prb+c-sn } = ans
where
b+c-nNonzero : b+c-sn 0 False
b+c-nNonzero pr rewrite (equalityCommutative Prb+c-sn) | pr | additionNIsCommutative n 0 = lessIrreflexive sn<b+c
b+c-nNonzero pr rewrite (equalityCommutative Prb+c-sn) | pr | Semiring.commutative Semiring n 0 = TotalOrder.irreflexive TotalOrder sn<b+c
2sn<a+b+c' : succ n +N succ n <N succ n +N (a +N b+c-sn)
2sn<a+b+c' = additionPreservesInequalityOnLeft (succ n) pr1
2sn<a+b+c'' : succ n +N succ n <N a +N (succ n +N b+c-sn)
2sn<a+b+c'' rewrite equalityCommutative (additionNIsAssociative a (succ n) b+c-sn) | additionNIsCommutative a (succ n) | additionNIsAssociative (succ n) a b+c-sn = 2sn<a+b+c'
2sn<a+b+c'' rewrite Semiring.+Associative Semiring a (succ n) b+c-sn | Semiring.commutative Semiring a (succ n) | equalityCommutative (Semiring.+Associative Semiring (succ n) a b+c-sn) = 2sn<a+b+c'
eep : succ n +N succ n <N a +N (b +N c)
eep rewrite equalityCommutative Prb+c-sn = 2sn<a+b+c''
eep2 : a +N (b +N c) <N succ n +N c
eep2 rewrite equalityCommutative (additionNIsAssociative a b c) = additionPreservesInequality c a+b<sn
eep2 rewrite Semiring.+Associative Semiring a b c = additionPreservesInequality c a+b<sn
eep2' : a +N (b +N c) <N succ n +N succ n
eep2' = orderIsTransitive eep2 (additionPreservesInequalityOnLeft (succ n) c<sn)
ans : False
ans = lessIrreflexive (orderIsTransitive eep eep2')
ans = TotalOrder.irreflexive TotalOrder (orderIsTransitive eep eep2')
plusZnAssociative : {n : } {pr : 0 <N n} (a b c : n n pr) a +n (b +n c) ((a +n b) +n c)
plusZnAssociative {zero} {()}
@@ -444,29 +447,29 @@ module IntegersModN where
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) with orderIsTotal ((a +N b) +N c) (succ n)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) with orderIsTotal (b +N c) (succ n)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inl b+c<sn) with orderIsTotal (a +N (b +N c)) (succ n)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inl b+c<sn) | inl (inl a+'b+c<sn) = equalityZn _ _ (equalityCommutative (additionNIsAssociative a b c))
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inl b+c<sn) | inl (inl a+'b+c<sn) = equalityZn _ _ (Semiring.+Associative Semiring a b c)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inl b+c<sn) | inl (inr sn<a+'b+c) = exFalso (false {succ n} a+b+c<sn sn<a+'b+c)
where
false : {x : } (a +N b) +N c <N succ n succ n <N a +N (b +N c) False
false pr1 pr2 rewrite additionNIsAssociative a b c = lessIrreflexive (orderIsTransitive pr1 pr2)
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) = TotalOrder.irreflexive TotalOrder (orderIsTransitive pr1 pr2)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inl b+c<sn) | inr sn=a+b+c = exFalso (false a+b+c<sn sn=a+b+c)
where
false : {x : } (a +N b) +N c <N x (a +N (b +N c)) x False
false p1 p2 rewrite additionNIsAssociative a b c | p2 = lessIrreflexive p1
false p1 p2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | p2 = TotalOrder.irreflexive TotalOrder p1
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) with orderIsTotal (a +N (b +N c)) (succ n)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inl _) with orderIsTotal (a +N subtractionNResult.result (-N (inl sn<b+c))) (succ n)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inl _) | inl (inl x) with -N (inl sn<b+c)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inl a+'b+c<sn) | inl (inl x) | record { result = result ; pr = pr } = exFalso (false a+'b+c<sn pr)
where
false : a +N (b +N c) <N succ n succ n +N result b +N c False
false pr1 pr2 rewrite equalityCommutative pr2 | equalityCommutative (additionNIsAssociative a (succ n) result) | additionNIsCommutative a (succ n) | additionNIsAssociative (succ n) a result = cannotAddAndEnlarge' pr1
false pr1 pr2 rewrite equalityCommutative pr2 | Semiring.+Associative Semiring a (succ n) result | Semiring.commutative Semiring a (succ n) | equalityCommutative (Semiring.+Associative Semiring (succ n) a result) = cannotAddAndEnlarge' pr1
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inl _) | inl (inr x) with -N (inl sn<b+c)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inl a+'b+c<sn) | inl (inr x) | record { result = result ; pr = pr } = exFalso false
where
lemma : a +N (succ n +N result) a +N (b +N c)
lemma' : a +N (succ n +N result) <N succ n
lemma'' : succ n +N (a +N result) <N succ n
lemma'' = identityOfIndiscernablesLeft _ _ _ _<N_ lemma' (transitivity (equalityCommutative (additionNIsAssociative a (succ n) result)) (transitivity (applyEquality (λ t t +N result) (additionNIsCommutative a (succ n))) (additionNIsAssociative (succ n) a result)))
lemma'' = identityOfIndiscernablesLeft _ _ _ _<N_ lemma' (transitivity (Semiring.+Associative Semiring a (succ n) result) (transitivity (applyEquality (λ t t +N result) (Semiring.commutative Semiring a (succ n))) (equalityCommutative (Semiring.+Associative Semiring (succ n) a result))))
lemma = applyEquality (λ i a +N i) pr
lemma' rewrite lemma = a+'b+c<sn
false : False
@@ -477,42 +480,42 @@ module IntegersModN where
lemma : a +N (succ n +N result) <N succ n
lemma rewrite pr = a+'b+c<sn
lemma' : succ n +N (a +N result) <N succ n
lemma' = identityOfIndiscernablesLeft _ _ _ _<N_ lemma (transitivity (equalityCommutative (additionNIsAssociative a (succ n) result)) (transitivity (applyEquality (λ t t +N result) (additionNIsCommutative a (succ n))) (additionNIsAssociative (succ n) a result)))
lemma' = identityOfIndiscernablesLeft _ _ _ _<N_ lemma (transitivity (Semiring.+Associative Semiring a (succ n) result) (transitivity (applyEquality (λ t t +N result) (Semiring.commutative Semiring a (succ n))) (equalityCommutative (Semiring.+Associative Semiring (succ n) a result))))
false : False
false = cannotAddAndEnlarge' lemma'
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inr x) = equalityZn _ _ (exFalso (false {succ n} a+b+c<sn x))
where
false : {x : } (a +N b) +N c <N succ n succ n <N a +N (b +N c) False
false pr1 pr2 rewrite additionNIsAssociative a b c = lessIrreflexive (orderIsTransitive pr1 pr2)
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) = TotalOrder.irreflexive TotalOrder (orderIsTransitive pr1 pr2)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inr x = exFalso (false a+b+c<sn x)
where
false : (a +N b) +N c <N succ n a +N (b +N c) succ n False
false pr1 pr2 rewrite equalityCommutative pr2 | additionNIsAssociative a b c = lessIrreflexive pr1
false pr1 pr2 rewrite equalityCommutative pr2 | equalityCommutative (Semiring.+Associative Semiring a b c) = TotalOrder.irreflexive TotalOrder pr1
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inr sn=b+c = exFalso (false a+b+c<sn sn=b+c)
where
false : (a +N b) +N c <N succ n b +N c succ n False
false pr1 pr2 rewrite additionNIsAssociative a b c | pr2 | additionNIsCommutative a (succ n) = cannotAddAndEnlarge' a+b+c<sn
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | pr2 | Semiring.commutative Semiring a (succ n) = cannotAddAndEnlarge' a+b+c<sn
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) with orderIsTotal (b +N c) (succ n)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inl b+c<sn) with orderIsTotal (a +N (b +N c)) (succ n)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inl b+c<sn) | inl (inl a+'b+c<sn) = exFalso (false sn<a+b+c a+'b+c<sn)
where
false : (succ n <N (a +N b) +N c) a +N (b +N c) <N succ n False
false pr1 pr2 rewrite additionNIsAssociative a b c = lessIrreflexive (orderIsTransitive pr1 pr2)
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) = TotalOrder.irreflexive TotalOrder (orderIsTransitive pr1 pr2)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inl b+c<sn) | inl (inr sn<a+'b+c) = equalityZn _ _ ans
where
lemma : succ n +N ((a +N b) +N c) (a +N (b +N c)) +N succ n
lemma rewrite additionNIsAssociative a b c = additionNIsCommutative (succ n) _
lemma rewrite equalityCommutative (Semiring.+Associative Semiring a b c) = Semiring.commutative Semiring (succ n) _
ans : subtractionNResult.result (-N (inl sn<a+'b+c)) subtractionNResult.result (-N (inl sn<a+b+c))
ans = equivalentSubtraction _ _ _ _ sn<a+'b+c sn<a+b+c lemma
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inl b+c<sn) | inr x = exFalso (false sn<a+b+c x)
where
false : succ n <N (a +N b) +N c (a +N (b +N c)) succ n False
false pr1 pr2 rewrite additionNIsAssociative a b c | pr2 = lessIrreflexive sn<a+b+c
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | pr2 = TotalOrder.irreflexive TotalOrder sn<a+b+c
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inr sn<b+c) with orderIsTotal (a +N (b +N c)) (succ n)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inr sn<b+c) | inl (inl a+b+c<sn) = exFalso (false sn<a+b+c a+b+c<sn)
where
false : succ n <N (a +N b) +N c (a +N (b +N c)) <N succ n False
false pr1 pr2 rewrite additionNIsAssociative a b c = lessIrreflexive (orderIsTransitive pr1 pr2)
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) = TotalOrder.irreflexive TotalOrder (orderIsTransitive pr1 pr2)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inr sn<b+c) | inl (inr _) with orderIsTotal (a +N subtractionNResult.result (-N (inl sn<b+c))) (succ n)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inr sn<b+c) | inl (inr _) | inl (inl x) = equalityZn _ _ (help2 {a} {b} {c} {n} sn<b+c sn<a+b+c)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inr sn<b+c) | inl (inr _) | inl (inr x) = equalityZn _ _ (exFalso (help1 sn<b+c x a+b<sn aLess bLess cLess))
@@ -520,60 +523,60 @@ module IntegersModN where
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inr sn<b+c) | inr a+b+c=sn = exFalso (false sn<a+b+c a+b+c=sn)
where
false : (succ n <N (a +N b) +N c) (a +N (b +N c) succ n) False
false pr1 pr2 rewrite additionNIsAssociative a b c | pr2 = lessIrreflexive pr1
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | pr2 = TotalOrder.irreflexive TotalOrder pr1
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c with orderIsTotal (a +N (b +N c)) (succ n)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c | inl (inl x) = exFalso (false sn<a+b+c x)
where
false : succ n <N (a +N b) +N c a +N (b +N c) <N succ n False
false p1 p2 rewrite additionNIsAssociative a b c = lessIrreflexive (orderIsTransitive p1 p2)
false p1 p2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) = TotalOrder.irreflexive TotalOrder (orderIsTransitive p1 p2)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c | inl (inr _) with orderIsTotal (a +N 0) (succ n)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c | inl (inr _) | inl (inl x) = equalityZn _ _ (ans sn=b+c)
where
ans : b +N c succ n a +N 0 subtractionNResult.result (-N (inl sn<a+b+c))
ans pr with -N (inl sn<a+b+c)
ans b+c=sn | record { result = result ; pr = pr1 } rewrite additionNIsCommutative a 0 | additionNIsAssociative a b c | b+c=sn | additionNIsCommutative (succ n) result = equalityCommutative (canSubtractFromEqualityRight pr1)
ans b+c=sn | record { result = result ; pr = pr1 } rewrite Semiring.commutative Semiring a 0 | equalityCommutative (Semiring.+Associative Semiring a b c) | b+c=sn | Semiring.commutative Semiring (succ n) result = equalityCommutative (canSubtractFromEqualityRight pr1)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c | inl (inr _) | inl (inr x) = exFalso (false b a+b<sn x)
where
false : (b : ) a +N b <N succ n succ n <N a +N 0 False
false zero pr1 pr2 rewrite additionNIsCommutative a 0 = lessIrreflexive (orderIsTransitive pr1 pr2)
false (succ b) pr1 pr2 rewrite additionNIsCommutative a 0 = lessIrreflexive (orderIsTransitive pr2 (orderIsTransitive (le b (additionNIsCommutative (succ b) a)) pr1))
false zero pr1 pr2 rewrite Semiring.commutative Semiring a 0 = TotalOrder.irreflexive TotalOrder (orderIsTransitive pr1 pr2)
false (succ b) pr1 pr2 rewrite Semiring.commutative Semiring a 0 = TotalOrder.irreflexive TotalOrder (orderIsTransitive pr2 (orderIsTransitive (le b (Semiring.commutative Semiring (succ b) a)) pr1))
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c | inl (inr _) | inr x = exFalso (false b a+b<sn x)
where
false : (b : ) a +N b <N succ n a +N 0 succ n False
false zero pr1 pr2 rewrite pr2 = lessIrreflexive pr1
false (succ b) pr1 pr2 rewrite additionNIsCommutative a 0 | pr2 = cannotAddAndEnlarge' pr1
false zero pr1 pr2 rewrite pr2 = TotalOrder.irreflexive TotalOrder pr1
false (succ b) pr1 pr2 rewrite Semiring.commutative Semiring a 0 | pr2 = cannotAddAndEnlarge' pr1
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c | inr x = exFalso (false sn<a+b+c x)
where
false : succ n <N (a +N b) +N c a +N (b +N c) succ n False
false pr1 pr2 rewrite additionNIsAssociative a b c | pr2 = lessIrreflexive pr1
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | pr2 = TotalOrder.irreflexive TotalOrder pr1
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c with orderIsTotal (b +N c) (succ n)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inl (inl b+c<sn) with orderIsTotal (a +N (b +N c)) (succ n)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inl (inl b+c<sn) | inl (inl a+'b+c<sn) = exFalso (false sn=a+b+c a+'b+c<sn)
where
false : (a +N b) +N c succ n a +N (b +N c) <N succ n False
false pr1 pr2 rewrite additionNIsAssociative a b c | pr1 = lessIrreflexive pr2
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | pr1 = TotalOrder.irreflexive TotalOrder pr2
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inl (inl b+c<sn) | inl (inr sn<a+'b+c) = exFalso (false sn=a+b+c sn<a+'b+c)
where
false : (a +N b) +N c succ n succ n <N a +N (b +N c) False
false pr1 pr2 rewrite additionNIsAssociative a b c | pr1 = lessIrreflexive pr2
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | pr1 = TotalOrder.irreflexive TotalOrder pr2
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inl (inl b+c<sn) | inr _ = equalityZn _ _ refl
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inl (inr sn<b+c) = exFalso (false a sn=a+b+c sn<b+c)
where
false : (a : ) (a +N b) +N c succ n succ n <N b +N c False
false zero pr1 pr2 rewrite additionNIsAssociative a b c | equalityCommutative pr1 = lessIrreflexive pr2
false (succ a) pr1 pr2 rewrite additionNIsAssociative a b c | equalityCommutative pr1 = lessIrreflexive (orderIsTransitive pr2 (le a refl))
false zero pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | equalityCommutative pr1 = TotalOrder.irreflexive TotalOrder pr2
false (succ a) pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | equalityCommutative pr1 = TotalOrder.irreflexive TotalOrder (orderIsTransitive pr2 (le a refl))
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inr b+c=sn with orderIsTotal (a +N 0) (succ n)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inr b+c=sn | inl (inl a+0<sn) = equalityZn _ _ ans
where
a=0 : (a : ) (a +N b) +N c succ n b +N c succ n a 0
a=0 zero pr1 pr2 rewrite additionNIsAssociative a b c | pr2 = refl
a=0 (succ a) pr1 pr2 rewrite additionNIsAssociative a b c | pr2 = canSubtractFromEqualityRight pr1
a=0 zero pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | pr2 = refl
a=0 (succ a) pr1 pr2 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | pr2 = canSubtractFromEqualityRight pr1
ans : a +N 0 0
ans rewrite additionNIsCommutative a 0 = a=0 a sn=a+b+c b+c=sn
ans rewrite Semiring.commutative Semiring a 0 = a=0 a sn=a+b+c b+c=sn
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inr b+c=sn | inl (inr sn<a+0) = exFalso (false sn<a+0 sn=a+b+c b+c=sn)
where
false : succ n <N a +N 0 (a +N b) +N c succ n b +N c succ n False
false pr1 pr2 pr3 rewrite additionNIsAssociative a b c | pr3 | additionNIsCommutative a 0 = zeroNeverGreater {succ n} (identityOfIndiscernablesRight _ _ _ _<N_ pr1 (a=0 a pr2))
false pr1 pr2 pr3 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | pr3 | Semiring.commutative Semiring a 0 = zeroNeverGreater {succ n} (identityOfIndiscernablesRight _ _ _ _<N_ pr1 (a=0 a pr2))
where
a=0 : (a : ) (a +N succ n succ n) a 0
a=0 zero pr = refl
@@ -581,7 +584,7 @@ module IntegersModN where
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inr b+c=sn | inr sn=a+0 = exFalso (false sn=a+b+c b+c=sn sn=a+0)
where
false : (a +N b) +N c succ n b +N c succ n a +N 0 succ n False
false pr1 pr2 pr3 rewrite additionNIsAssociative a b c | pr2 | equalityCommutative pr3 | additionNIsCommutative a 0 = naughtE (identityOfIndiscernablesLeft _ _ _ _≡_ pr3 (a=0 a pr1))
false pr1 pr2 pr3 rewrite equalityCommutative (Semiring.+Associative Semiring a b c) | pr2 | equalityCommutative pr3 | Semiring.commutative Semiring a 0 = naughtE (identityOfIndiscernablesLeft _ _ _ _≡_ pr3 (a=0 a pr1))
where
a=0 : (a : ) (a +N a a) a 0
a=0 zero pr = refl
@@ -591,7 +594,7 @@ module IntegersModN where
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inl b+c<sn) | inl (inl a+'b+c<sn) = exFalso (false sn<a+b a+'b+c<sn)
where
false : succ n <N a +N b a +N (b +N c) <N succ n False
false pr1 pr2 rewrite equalityCommutative (additionNIsAssociative a b c) = cannotAddAndEnlarge' (orderIsTransitive pr2 pr1)
false pr1 pr2 rewrite Semiring.+Associative Semiring a b c = cannotAddAndEnlarge' (orderIsTransitive pr2 pr1)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inl b+c<sn) | inl (inr sn<a+'b+c) with orderIsTotal (subtractionNResult.result (-N (inl sn<a+b)) +N c) (succ n)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inl b+c<sn) | inl (inr sn<a+'b+c) | inl (inl x) = equalityZn _ _ (help4 {a} {b} {c} {succ n} sn<a+'b+c sn<a+b)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inl b+c<sn) | inl (inr sn<a+'b+c) | inl (inr x) = exFalso (help18 {a} {b} {c} {succ n} b+c<sn sn<a+b aLess x)
@@ -599,7 +602,7 @@ module IntegersModN where
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inl b+c<sn) | inr a+'b+c=sn = exFalso (false sn<a+b a+'b+c=sn)
where
false : (succ n <N a +N b) a +N (b +N c) succ n False
false pr1 pr2 rewrite equalityCommutative (additionNIsAssociative a b c) | equalityCommutative pr2 = cannotAddAndEnlarge' pr1
false pr1 pr2 rewrite Semiring.+Associative Semiring a b c | equalityCommutative pr2 = cannotAddAndEnlarge' pr1
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) with orderIsTotal (a +N subtractionNResult.result (-N (inl sn<b+c))) (succ n)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) | inl (inl x) with orderIsTotal (subtractionNResult.result (-N (inl sn<a+b)) +N c) (succ n)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) | inl (inl x) | inl (inl x₁) = equalityZn _ _ (help5 {a} {b} {c} {succ n} sn<b+c sn<a+b)
@@ -634,12 +637,12 @@ module IntegersModN where
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) | inr b+c=sn | inl (inl a+0<sn) = equalityZn _ _ (help25 {a} {b} {c} {succ n} sn=a+b b+c=sn)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) | inr b+c=sn | inl (inr sn<a+0) = exFalso (help24 {a} {succ n} aLess sn<a+0)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) | inr b+c=sn | inr a+0=sn = exFalso (help23 {a} {succ n} aLess a+0=sn)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inr sn<c) = exFalso (lessIrreflexive (orderIsTransitive sn<c cLess))
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inr sn<c) = exFalso (TotalOrder.irreflexive TotalOrder (orderIsTransitive sn<c cLess))
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inr c=sn with orderIsTotal (b +N c) (succ n)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inr c=sn | inl (inl b+c<sn) = exFalso (help b+c<sn)
where
help : (b +N c <N succ n) False
help b+c<sn rewrite equalityCommutative c=sn | additionNIsCommutative b c = cannotAddAndEnlarge' b+c<sn
help b+c<sn rewrite equalityCommutative c=sn | Semiring.commutative Semiring b c = cannotAddAndEnlarge' b+c<sn
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inr c=sn | inl (inr sn<b+c) with orderIsTotal (a +N subtractionNResult.result (-N (inl sn<b+c))) (succ n)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inr c=sn | inl (inr sn<b+c) | inl (inl x) = exFalso (help20 {a} {b} {c} {succ n} c=sn sn=a+b sn<b+c x)
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inr c=sn | inl (inr sn<b+c) | inl (inr x) = exFalso (help22 {a} {b} {c} {succ n} sn=a+b c=sn sn<b+c x)
@@ -647,7 +650,7 @@ module IntegersModN where
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inr c=sn | inr b+c=sn = exFalso (help21 {a} {b} {c} {succ n} sn=a+b b+c=sn c=sn aLess)
subLess : {a b : } (0<a : 0 <N a) (a<b : a <N b) subtractionNResult.result (-N (inl a<b)) <N b
subLess {zero} {b} 0<a a<b = exFalso (lessIrreflexive 0<a)
subLess {zero} {b} 0<a a<b = exFalso (TotalOrder.irreflexive TotalOrder 0<a)
subLess {succ a} {b} _ a<b with -N (inl a<b)
... | record { result = b-a ; pr = pr } = record { x = a ; proof = pr }
@@ -664,13 +667,13 @@ module IntegersModN where
h : subtractionNResult.result (-N (inl (le (_<N_.x aLess) (transitivity (equalityCommutative (succExtracts (_<N_.x aLess) a)) (succInjective (_<N_.proof aLess)))))) +N succ a succ n
h = addMinus {succ a} {succ n} (inl aLess)
f : False
f rewrite h = lessIrreflexive x
f rewrite h = TotalOrder.irreflexive TotalOrder x
... | inl (inr x) = exFalso f
where
h : subtractionNResult.result (-N (inl (subtractOneFromInequality aLess))) +N succ a succ n
h = addMinus {succ a} {succ n} (inl aLess)
f : False
f rewrite h = lessIrreflexive x
f rewrite h = TotalOrder.irreflexive TotalOrder x
... | (inr n-a+sa=sn) = equalityZn _ _ refl
nGroup : (n : ) (pr : 0 <N n) Group (reflSetoid (n n pr)) _+n_

View File

@@ -6,7 +6,8 @@ open import Maybe
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Vectors
open import Numbers.Naturals
open import Numbers.Naturals.Naturals
open import Numbers.Naturals.Order
module KeyValue where
record ReducedMap {a b c : _} (keys : Set a) (values : Set b) (keyOrder : TotalOrder {_} {c} keys) (min : keys) : Set (a b c)

View File

@@ -7,7 +7,7 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import KeyValue
open import Vectors
open import Numbers.Naturals
open import Numbers.Naturals.Naturals
module KeyValueWithDomain where

View File

@@ -2,7 +2,8 @@
open import LogicalFormulae
open import Functions
open import Numbers.Naturals -- for length
open import Numbers.Naturals.Naturals -- for length
open import Semirings.Definition
module Lists.Lists where
data List {a} (A : Set a) : Set a where
@@ -154,4 +155,4 @@ module Lists.Lists where
lengthRev : {a : _} {A : Set a} (l : List A) length (rev l) length l
lengthRev [] = refl
lengthRev (x :: l) rewrite lengthConcat (rev l) (x :: []) | lengthRev l | additionNIsCommutative (length l) 1 = refl
lengthRev (x :: l) rewrite lengthConcat (rev l) (x :: []) | lengthRev l | Semiring.commutative Semiring (length l) 1 = refl

View File

@@ -4,7 +4,7 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import LogicalFormulae
open import Logic.PropositionalLogic
open import Functions
open import Numbers.Naturals
open import Numbers.Naturals.Naturals
open import Vectors
module Logic.PropositionalAxiomsTautology where

View File

@@ -4,7 +4,7 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import LogicalFormulae
open import Functions
open import Numbers.Naturals
open import Numbers.Naturals.Naturals
open import Vectors
module Logic.PropositionalLogic where

View File

@@ -4,7 +4,7 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import LogicalFormulae
open import Logic.PropositionalLogic
open import Functions
open import Numbers.Naturals
open import Numbers.Naturals.Naturals
open import Vectors
module Logic.PropositionalLogicExamples where

13
Monoids/Definition.agda Normal file
View File

@@ -0,0 +1,13 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Monoids.Definition where
record Monoid {a : _} {A : Set a} (Zero : A) (_+_ : A A A) : Set a where
field
associative : (a b c : A) a + (b + c) (a + b) + c
idLeft : (a : A) Zero + a a
idRight : (a : A) a + Zero a

View File

@@ -3,9 +3,10 @@
open import LogicalFormulae
open import Functions
open import Lists.Lists
open import Numbers.Naturals
open import Groups.GroupDefinition
open import Numbers.Naturals.Naturals
open import Groups.Definition
open import Numbers.BinaryNaturals.Definition
open import Semirings.Definition
module Numbers.BinaryNaturals.Addition where
@@ -73,9 +74,9 @@ _+B_ : BinNat → BinNat → BinNat
ans2 rewrite bad | bad2 = refl
+BIsInherited [] b _ prB = +BIsInherited[] b prB
+BIsInherited (x :: a) [] prA _ = transitivity (applyEquality NToBinNat (additionNIsCommutative (binNatToN (x :: a)) 0)) (transitivity (binToBin (x :: a)) (equalityCommutative prA))
+BIsInherited (x :: a) [] prA _ = transitivity (applyEquality NToBinNat (Semiring.commutative Semiring (binNatToN (x :: a)) 0)) (transitivity (binToBin (x :: a)) (equalityCommutative prA))
+BIsInherited (zero :: as) (zero :: b) prA prB with orderIsTotal 0 (binNatToN as +N binNatToN b)
... | inl (inl 0<) rewrite additionNIsCommutative (binNatToN as) 0 | additionNIsCommutative (binNatToN b) 0 | equalityCommutative (additionNIsAssociative (binNatToN as +N binNatToN as) (binNatToN b) (binNatToN b)) | additionNIsAssociative (binNatToN as) (binNatToN as) (binNatToN b) | additionNIsCommutative (binNatToN as) (binNatToN b) | equalityCommutative (additionNIsAssociative (binNatToN as) (binNatToN b) (binNatToN as)) | additionNIsAssociative (binNatToN as +N binNatToN b) (binNatToN as) (binNatToN b) | additionNIsCommutative 0 ((binNatToN as +N binNatToN b) +N (binNatToN as +N binNatToN b)) | additionNIsAssociative (binNatToN as +N binNatToN b) (binNatToN as +N binNatToN b) 0 = transitivity (doubleIsBitShift (binNatToN as +N binNatToN b) (identityOfIndiscernablesRight _ _ _ _<N_ 0< (additionNIsCommutative (binNatToN b) _))) (applyEquality (zero ::_) (+BIsInherited as b (canonicalDescends as prA) (canonicalDescends b prB)))
... | inl (inl 0<) rewrite Semiring.commutative Semiring (binNatToN as) 0 | Semiring.commutative Semiring (binNatToN b) 0 | Semiring.+Associative Semiring (binNatToN as +N binNatToN as) (binNatToN b) (binNatToN b) | equalityCommutative (Semiring.+Associative Semiring (binNatToN as) (binNatToN as) (binNatToN b)) | Semiring.commutative Semiring (binNatToN as) (binNatToN b) | Semiring.+Associative Semiring (binNatToN as) (binNatToN b) (binNatToN as) | equalityCommutative (Semiring.+Associative Semiring (binNatToN as +N binNatToN b) (binNatToN as) (binNatToN b)) | Semiring.commutative Semiring 0 ((binNatToN as +N binNatToN b) +N (binNatToN as +N binNatToN b)) | equalityCommutative (Semiring.+Associative Semiring (binNatToN as +N binNatToN b) (binNatToN as +N binNatToN b) 0) = transitivity (doubleIsBitShift (binNatToN as +N binNatToN b) (identityOfIndiscernablesRight _ _ _ _<N_ 0< (Semiring.commutative Semiring (binNatToN b) _))) (applyEquality (zero ::_) (+BIsInherited as b (canonicalDescends as prA) (canonicalDescends b prB)))
+BIsInherited (zero :: as) (zero :: b) prA prB | inl (inr ())
... | inr p with sumZeroImpliesOperandsZero (binNatToN as) (equalityCommutative p)
+BIsInherited (zero :: as) (zero :: b) prA prB | inr p | as=0 ,, b=0 rewrite as=0 | b=0 = exFalso ans
@@ -93,9 +94,9 @@ _+B_ : BinNat → BinNat → BinNat
ans with inspect (canonical b)
ans | [] with x rewrite x = nono prB
ans | (x₁ :: y) with x = nono (transitivity (equalityCommutative x) (transitivity (equalityCommutative t) u))
+BIsInherited (zero :: as) (one :: b) prA prB rewrite additionNIsCommutative (binNatToN as +N (binNatToN as +N zero)) (succ (binNatToN b +N (binNatToN b +N zero))) | additionNIsCommutative (binNatToN b +N (binNatToN b +N zero)) (binNatToN as +N (binNatToN as +N zero)) | equalityCommutative (productDistributes 2 (binNatToN as) (binNatToN b)) = +BinheritLemma as b (canonicalDescends as prA) (canonicalDescends b prB)
+BIsInherited (one :: as) (zero :: bs) prA prB rewrite equalityCommutative (productDistributes 2 (binNatToN as) (binNatToN bs)) = +BinheritLemma as bs (canonicalDescends as prA) (canonicalDescends bs prB)
+BIsInherited (one :: as) (one :: bs) prA prB rewrite additionNIsCommutative (binNatToN as +N (binNatToN as +N zero)) (succ (binNatToN bs +N (binNatToN bs +N zero))) | additionNIsCommutative (binNatToN bs +N (binNatToN bs +N zero)) (2 *N binNatToN as) | equalityCommutative (productDistributes 2 (binNatToN as) (binNatToN bs)) | +BinheritLemma as bs (canonicalDescends as prA) (canonicalDescends bs prB) = refl
+BIsInherited (zero :: as) (one :: b) prA prB rewrite Semiring.commutative Semiring (binNatToN as +N (binNatToN as +N zero)) (succ (binNatToN b +N (binNatToN b +N zero))) | Semiring.commutative Semiring (binNatToN b +N (binNatToN b +N zero)) (binNatToN as +N (binNatToN as +N zero)) | equalityCommutative (Semiring.+DistributesOver* Semiring 2 (binNatToN as) (binNatToN b)) = +BinheritLemma as b (canonicalDescends as prA) (canonicalDescends b prB)
+BIsInherited (one :: as) (zero :: bs) prA prB rewrite equalityCommutative (Semiring.+DistributesOver* Semiring 2 (binNatToN as) (binNatToN bs)) = +BinheritLemma as bs (canonicalDescends as prA) (canonicalDescends bs prB)
+BIsInherited (one :: as) (one :: bs) prA prB rewrite Semiring.commutative Semiring (binNatToN as +N (binNatToN as +N zero)) (succ (binNatToN bs +N (binNatToN bs +N zero))) | Semiring.commutative Semiring (binNatToN bs +N (binNatToN bs +N zero)) (2 *N binNatToN as) | equalityCommutative (Semiring.+DistributesOver* Semiring 2 (binNatToN as) (binNatToN bs)) | +BinheritLemma as bs (canonicalDescends as prA) (canonicalDescends bs prB) = refl
+BIsInherited'[] : (b : BinNat) [] +Binherit b canonical ([] +B b)
+BIsInherited'[] [] = refl
@@ -126,7 +127,7 @@ _+B_ : BinNat → BinNat → BinNat
+BIsInherited' [] b = +BIsInherited'[] b
+BIsInherited' (zero :: a) [] with inspect (binNatToN a)
+BIsInherited' (zero :: a) [] | zero with x rewrite x | binNatToNZero a x = refl
+BIsInherited' (zero :: a) [] | succ y with x rewrite x | additionNIsCommutative (y +N succ (y +N 0)) 0 = transitivity (doubleIsBitShift' y) (transitivity (applyEquality (λ i (zero :: NToBinNat i)) (equalityCommutative x)) (transitivity (applyEquality (λ i zero :: i) (binToBin a)) (canonicalAscends' {zero} a bad)))
+BIsInherited' (zero :: a) [] | succ y with x rewrite x | Semiring.commutative Semiring (y +N succ (y +N 0)) 0 = transitivity (doubleIsBitShift' y) (transitivity (applyEquality (λ i (zero :: NToBinNat i)) (equalityCommutative x)) (transitivity (applyEquality (λ i zero :: i) (binToBin a)) (canonicalAscends' {zero} a bad)))
where
bad : canonical a [] False
bad pr with transitivity (equalityCommutative x) (transitivity (equalityCommutative (binNatToNIsCanonical a)) (applyEquality binNatToN pr))
@@ -136,8 +137,8 @@ _+B_ : BinNat → BinNat → BinNat
+BIsInherited' (one :: a) [] | succ n with x rewrite x | doubleIsBitShift' n = applyEquality incr {_} {zero :: canonical a} (transitivity {x = _} {NToBinNat (2 *N succ n)} bl (transitivity (doubleIsBitShift' n) (applyEquality (zero ::_) (transitivity (applyEquality NToBinNat (equalityCommutative x)) (binToBin a)))))
where
bl : incr (NToBinNat ((n +N succ (n +N 0)) +N 0)) NToBinNat (succ (n +N succ (n +N 0)))
bl rewrite equalityCommutative x | additionNIsCommutative (n +N succ (n +N 0)) 0 = refl
+BIsInherited' (zero :: as) (zero :: bs) rewrite equalityCommutative (productDistributes 2 (binNatToN as) (binNatToN bs)) = ans
bl rewrite equalityCommutative x | Semiring.commutative Semiring (n +N succ (n +N 0)) 0 = refl
+BIsInherited' (zero :: as) (zero :: bs) rewrite equalityCommutative (Semiring.+DistributesOver* Semiring 2 (binNatToN as) (binNatToN bs)) = ans
where
ans : NToBinNat (2 *N (binNatToN as +N binNatToN bs)) canonical (zero :: (as +B bs))
ans with inspect (binNatToN as +N binNatToN bs)
@@ -157,7 +158,7 @@ _+B_ : BinNat → BinNat → BinNat
u rewrite equalityCommutative (binNatToNIsCanonical (as +B bs)) | equalityCommutative (+BIsInherited' as bs) | x | nToN (succ y) = succIsPositive y
ans2 : zero :: NToBinNat (binNatToN as +N binNatToN bs) canonical (zero :: (as +B bs))
ans2 rewrite +BIsInherited' as bs = canonicalAscends (as +B bs) u
+BIsInherited' (zero :: as) (one :: bs) rewrite additionNIsCommutative (2 *N binNatToN as) (succ (2 *N binNatToN bs)) | additionNIsCommutative (2 *N binNatToN bs) (2 *N binNatToN as) | equalityCommutative (productDistributes 2 (binNatToN as) (binNatToN bs)) = ans2
+BIsInherited' (zero :: as) (one :: bs) rewrite Semiring.commutative Semiring (2 *N binNatToN as) (succ (2 *N binNatToN bs)) | Semiring.commutative Semiring (2 *N binNatToN bs) (2 *N binNatToN as) | equalityCommutative (Semiring.+DistributesOver* Semiring 2 (binNatToN as) (binNatToN bs)) = ans2
where
ans2 : incr (NToBinNat (2 *N (binNatToN as +N binNatToN bs))) one :: canonical (as +B bs)
ans2 with inspect (binNatToN as +N binNatToN bs)
@@ -167,7 +168,7 @@ _+B_ : BinNat → BinNat → BinNat
t : [] as +Binherit bs
t rewrite as=0 | bs=0 = refl
ans2 | succ y with x rewrite x | doubleIsBitShift' y = applyEquality (one ::_) (transitivity (applyEquality NToBinNat (equalityCommutative x)) (+BIsInherited' as bs))
+BIsInherited' (one :: as) (zero :: bs) rewrite equalityCommutative (productDistributes 2 (binNatToN as) (binNatToN bs)) = ans
+BIsInherited' (one :: as) (zero :: bs) rewrite equalityCommutative (Semiring.+DistributesOver* Semiring 2 (binNatToN as) (binNatToN bs)) = ans
where
ans : incr (NToBinNat (2 *N (binNatToN as +N binNatToN bs))) one :: canonical (as +B bs)
ans with inspect (binNatToN as +N binNatToN bs)
@@ -177,7 +178,7 @@ _+B_ : BinNat → BinNat → BinNat
t : [] NToBinNat (binNatToN as +N binNatToN bs)
t rewrite as=0 | bs=0 = refl
ans | succ y with x rewrite x | doubleIsBitShift' y = applyEquality (one ::_) (transitivity (applyEquality NToBinNat (equalityCommutative x)) (+BIsInherited' as bs))
+BIsInherited' (one :: as) (one :: bs) rewrite additionNIsCommutative (2 *N binNatToN as) (succ (2 *N binNatToN bs)) | additionNIsCommutative (2 *N binNatToN bs) (2 *N binNatToN as) | equalityCommutative (productDistributes 2 (binNatToN as) (binNatToN bs)) = ans
+BIsInherited' (one :: as) (one :: bs) rewrite Semiring.commutative Semiring (2 *N binNatToN as) (succ (2 *N binNatToN bs)) | Semiring.commutative Semiring (2 *N binNatToN bs) (2 *N binNatToN as) | equalityCommutative (Semiring.+DistributesOver* Semiring 2 (binNatToN as) (binNatToN bs)) = ans
where
ans : incr (incr (NToBinNat (2 *N (binNatToN as +N binNatToN bs)))) canonical (zero :: incr (as +B bs))
ans with inspect (binNatToN as +N binNatToN bs)

View File

@@ -3,8 +3,11 @@
open import LogicalFormulae
open import Functions
open import Lists.Lists
open import Numbers.Naturals
open import Groups.GroupDefinition
open import Numbers.Naturals.Naturals
open import Numbers.Naturals.Definition
open import Groups.Definition
open import Semirings.Definition
open import Orders
module Numbers.BinaryNaturals.Definition where
@@ -140,7 +143,7 @@ module Numbers.BinaryNaturals.Definition where
doubleIsBitShift' : (a : ) NToBinNat (2 *N succ a) zero :: NToBinNat (succ a)
doubleIsBitShift' zero = refl
doubleIsBitShift' (succ a) with doubleIsBitShift' a
... | bl rewrite additionNIsCommutative a (succ (succ (a +N 0))) | additionNIsCommutative (succ (a +N 0)) a | additionNIsCommutative a (succ (a +N 0)) | additionNIsCommutative (a +N 0) a | bl = refl
... | bl rewrite Semiring.commutative Semiring a (succ (succ (a +N 0))) | Semiring.commutative Semiring (succ (a +N 0)) a | Semiring.commutative Semiring a (succ (a +N 0)) | Semiring.commutative Semiring (a +N 0) a | bl = refl
doubleIsBitShift : (a : ) (0 <N a) NToBinNat (2 *N a) zero :: NToBinNat a
doubleIsBitShift zero ()
@@ -249,9 +252,9 @@ module Numbers.BinaryNaturals.Definition where
doubleInj (succ a) (succ b) pr = applyEquality succ (doubleInj a b u)
where
t : a +N a b +N b
t rewrite additionNIsCommutative (succ a) 0 | additionNIsCommutative (succ b) 0 | additionNIsCommutative a (succ a) | additionNIsCommutative b (succ b) = succInjective (succInjective pr)
t rewrite Semiring.commutative Semiring (succ a) 0 | Semiring.commutative Semiring (succ b) 0 | Semiring.commutative Semiring a (succ a) | Semiring.commutative Semiring b (succ b) = succInjective (succInjective pr)
u : a +N (a +N zero) b +N (b +N zero)
u rewrite additionNIsCommutative a 0 | additionNIsCommutative b 0 = t
u rewrite Semiring.commutative Semiring a 0 | Semiring.commutative Semiring b 0 = t
binNatToNZero [] pr = refl
binNatToNZero (zero :: xs) pr with inspect (canonical xs)
@@ -264,7 +267,7 @@ module Numbers.BinaryNaturals.Definition where
binNatToNSucc [] = refl
binNatToNSucc (zero :: n) = refl
binNatToNSucc (one :: n) rewrite additionNIsCommutative (binNatToN n) zero | additionNIsCommutative (binNatToN (incr n)) 0 | binNatToNSucc n = applyEquality succ (additionNIsCommutative (binNatToN n) (succ (binNatToN n)))
binNatToNSucc (one :: n) rewrite Semiring.commutative Semiring (binNatToN n) zero | Semiring.commutative Semiring (binNatToN (incr n)) 0 | binNatToNSucc n = applyEquality succ (Semiring.commutative Semiring (binNatToN n) (succ (binNatToN n)))
incrNonzero (one :: xs) pr with inspect (canonical (incr xs))
incrNonzero (one :: xs) pr | [] with x = incrNonzero xs x
@@ -282,11 +285,11 @@ module Numbers.BinaryNaturals.Definition where
ans rewrite t | pr = refl
nToN (succ x) | (bit :: xs) with pr = transitivity (binNatToNSucc (NToBinNat x)) (applyEquality succ (nToN x))
parity zero (succ b) pr rewrite additionNIsCommutative b (succ (b +N 0)) = bad pr
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 additionNIsCommutative b (succ (b +N 0)) | additionNIsCommutative a (succ (a +N 0)) | additionNIsCommutative (a +N 0) a | additionNIsCommutative (b +N 0) b = parity a b (succInjective (succInjective pr))
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))
binNatToNZero' [] pr = refl
binNatToNZero' (zero :: xs) pr with inspect (canonical xs)
@@ -311,13 +314,13 @@ module Numbers.BinaryNaturals.Definition where
v with inspect (binNatToN a)
v | zero with x = x
v | succ a' with x with inspect (binNatToN (zero :: a))
v | succ a' with x | zero with pr2 rewrite pr2 = exFalso (lessIrreflexive 0<a)
v | succ a' with x | succ y with pr2 rewrite u = exFalso (lessIrreflexive 0<a)
v | succ a' with x | zero with pr2 rewrite pr2 = exFalso (TotalOrder.irreflexive TotalOrder 0<a)
v | succ a' with x | succ y with pr2 rewrite u = exFalso (TotalOrder.irreflexive TotalOrder 0<a)
t : 0 <N binNatToN a
t with binNatToN a
t | succ bl rewrite additionNIsCommutative (succ bl) 0 = succIsPositive bl
t | succ bl rewrite Semiring.commutative Semiring (succ bl) 0 = succIsPositive bl
contr'' : (x : ) (0 <N x) (x 0) False
contr'' x 0<x x=0 rewrite x=0 = lessIrreflexive 0<x
contr'' x 0<x x=0 rewrite x=0 = TotalOrder.irreflexive TotalOrder 0<x
canonicalAscends {zero} a 0<a | (x₁ :: y) with x rewrite x = refl
canonicalAscends {one} a 0<a = refl

View File

@@ -3,10 +3,11 @@
open import LogicalFormulae
open import Functions
open import Lists.Lists
open import Numbers.Naturals
open import Groups.GroupDefinition
open import Numbers.Naturals.Naturals
open import Groups.Definition
open import Numbers.BinaryNaturals.Definition
open import Numbers.BinaryNaturals.Addition
open import Semirings.Definition
module Numbers.BinaryNaturals.Multiplication where
@@ -92,7 +93,7 @@ module Numbers.BinaryNaturals.Multiplication where
binNatToNDistributesPlus a b = transitivity (equalityCommutative (binNatToNIsCanonical (a +B b))) (transitivity (applyEquality binNatToN (equalityCommutative (+BIsInherited' a b))) (nToN (binNatToN a +N binNatToN b)))
+BAssociative : (a b c : BinNat) canonical (a +B (b +B c)) canonical ((a +B b) +B c)
+BAssociative a b c rewrite equalityCommutative (+BIsInherited' a (b +B c)) | equalityCommutative (+BIsInherited' (a +B b) c) | binNatToNDistributesPlus a b | binNatToNDistributesPlus b c | additionNIsAssociative (binNatToN a) (binNatToN b) (binNatToN c) = refl
+BAssociative a b c rewrite equalityCommutative (+BIsInherited' a (b +B c)) | equalityCommutative (+BIsInherited' (a +B b) c) | binNatToNDistributesPlus a b | binNatToNDistributesPlus b c | equalityCommutative (Semiring.+Associative Semiring (binNatToN a) (binNatToN b) (binNatToN c)) = refl
timesOne : (a b : BinNat) (canonical b) (one :: []) canonical (a *B b) canonical a
timesOne [] b pr = refl
@@ -173,7 +174,7 @@ module Numbers.BinaryNaturals.Multiplication where
where
ans : (a b : BinNat) binNatToN a *N binNatToN b binNatToN (a *B b)
ans [] b = refl
ans (zero :: a) b rewrite equalityCommutative (ans a b) = equalityCommutative (multiplicationNIsAssociative 2 (binNatToN a) (binNatToN b))
ans (one :: a) b rewrite binNatToNDistributesPlus (zero :: (a *B b)) b | additionNIsCommutative (binNatToN b) ((2 *N (binNatToN a)) *N (binNatToN b)) | equalityCommutative (ans a b) = applyEquality (_+N binNatToN b) (equalityCommutative (multiplicationNIsAssociative 2 (binNatToN a) (binNatToN b)))
ans (zero :: a) b rewrite equalityCommutative (ans a b) = equalityCommutative (Semiring.*Associative Semiring 2 (binNatToN a) (binNatToN b))
ans (one :: a) b rewrite binNatToNDistributesPlus (zero :: (a *B b)) b | Semiring.commutative Semiring (binNatToN b) ((2 *N (binNatToN a)) *N (binNatToN b)) | equalityCommutative (ans a b) = applyEquality (_+N binNatToN b) (equalityCommutative (Semiring.*Associative Semiring 2 (binNatToN a) (binNatToN b)))
t : (binNatToN a *N binNatToN b) binNatToN (canonical (a *B b))
t = transitivity (ans a b) (equalityCommutative (binNatToNIsCanonical (a *B b)))

View File

@@ -3,10 +3,11 @@
open import LogicalFormulae
open import Functions
open import Lists.Lists
open import Numbers.Naturals
open import Groups.GroupDefinition
open import Numbers.Naturals.Naturals
open import Groups.Definition
open import Numbers.BinaryNaturals.Definition
open import Orders
open import Semirings.Definition
module Numbers.BinaryNaturals.Order where
@@ -278,7 +279,7 @@ module Numbers.BinaryNaturals.Order where
chopDouble a b i | inr a=b | inr x = refl
succNotLess : {n : } succ n <N n False
succNotLess {succ n} (le x proof) = succNotLess {n} (le x (succInjective (transitivity (applyEquality succ (transitivity (additionNIsCommutative (succ x) (succ n)) (transitivity (applyEquality succ (transitivity (additionNIsCommutative n (succ x)) (applyEquality succ (additionNIsCommutative x n)))) (additionNIsCommutative (succ (succ n)) x)))) proof)))
succNotLess {succ n} (le x proof) = succNotLess {n} (le x (succInjective (transitivity (applyEquality succ (transitivity (Semiring.commutative Semiring (succ x) (succ n)) (transitivity (applyEquality succ (transitivity (Semiring.commutative Semiring n (succ x)) (applyEquality succ (Semiring.commutative Semiring x n)))) (Semiring.commutative Semiring (succ (succ n)) x)))) proof)))
<BIsInherited : (a b : BinNat) a <BInherited b a <B b
<BIsInherited [] b with orderIsTotal 0 (binNatToN b)

File diff suppressed because it is too large Load Diff

View File

@@ -0,0 +1,43 @@
{-# OPTIONS --warning=error --safe --without-K #-}
open import LogicalFormulae
open import Numbers.Naturals.Definition
module Numbers.Naturals.Addition where
infix 15 _+N_
_+N_ :
zero +N y = y
succ x +N y = succ (x +N y)
addzeroLeft : (x : ) (zero +N x) x
addzeroLeft x = refl
addZeroRight : (x : ) (x +N zero) x
addZeroRight zero = refl
addZeroRight (succ x) rewrite addZeroRight x = refl
succExtracts : (x y : ) (x +N succ y) (succ (x +N y))
succExtracts zero y = refl
succExtracts (succ x) y = applyEquality succ (succExtracts x y)
succCanMove : (x y : ) (x +N succ y) (succ x +N y)
succCanMove x y = transitivity (succExtracts x y) refl
additionNIsCommutative : (x y : ) (x +N y) (y +N x)
additionNIsCommutative zero y = transitivity (addzeroLeft y) (equalityCommutative (addZeroRight y))
additionNIsCommutative (succ x) zero = transitivity (addZeroRight (succ x)) refl
additionNIsCommutative (succ x) (succ y) = transitivity refl (applyEquality succ (transitivity (succCanMove x y) (additionNIsCommutative (succ x) y)))
addingPreservesEqualityRight : {a b : } (c : ) (a b) (a +N c b +N c)
addingPreservesEqualityRight {a} {b} c pr = applyEquality (λ n -> n +N c) pr
addingPreservesEqualityLeft : {a b : } (c : ) (a b) (c +N a c +N b)
addingPreservesEqualityLeft {a} {b} c pr = applyEquality (λ n -> c +N n) pr
additionNIsAssociative : (a b c : ) ((a +N b) +N c) (a +N (b +N c))
additionNIsAssociative zero b c = refl
additionNIsAssociative (succ a) zero c = transitivity (transitivity (applyEquality (λ n n +N c) (applyEquality succ (addZeroRight a))) refl) (transitivity refl refl)
additionNIsAssociative (succ a) (succ b) c = transitivity refl (transitivity refl (transitivity (applyEquality succ (additionNIsAssociative a (succ b) c)) refl))
succIsAddOne : (a : ) succ a a +N succ zero
succIsAddOne a = equalityCommutative (transitivity (additionNIsCommutative a (succ zero)) refl)

View File

@@ -0,0 +1,16 @@
{-# OPTIONS --warning=error --safe --without-K #-}
open import LogicalFormulae
module Numbers.Naturals.Definition where
data : Set where
zero :
succ :
infix 100 succ
{-# BUILTIN NATURAL #-}
succInjective : {a b : } (succ a succ b) a b
succInjective {a} {.a} refl = refl

View File

@@ -0,0 +1,82 @@
{-# OPTIONS --warning=error --safe --without-K #-}
open import LogicalFormulae
open import Numbers.Naturals.Definition
open import Numbers.Naturals.Addition
module Numbers.Naturals.Multiplication where
infix 25 _*N_
_*N_ :
zero *N y = zero
(succ x) *N y = y +N (x *N y)
productZeroIsZeroLeft : (a : ) (zero *N a zero)
productZeroIsZeroLeft a = refl
productZeroIsZeroRight : (a : ) (a *N zero zero)
productZeroIsZeroRight zero = refl
productZeroIsZeroRight (succ a) = productZeroIsZeroRight a
productWithOneLeft : (a : ) ((succ zero) *N a) a
productWithOneLeft a = transitivity refl (transitivity (applyEquality (λ { m -> a +N m }) refl) (additionNIsCommutative a zero))
productWithOneRight : (a : ) a *N succ zero a
productWithOneRight zero = refl
productWithOneRight (succ a) = transitivity refl (addingPreservesEqualityLeft (succ zero) (productWithOneRight a))
productDistributes : (a b c : ) (a *N (b +N c)) a *N b +N a *N c
productDistributes zero b c = refl
productDistributes (succ a) b c = transitivity refl
(transitivity (addingPreservesEqualityLeft (b +N c) (productDistributes a b c))
(transitivity (equalityCommutative (additionNIsAssociative (b +N c) (a *N b) (a *N c)))
(transitivity (addingPreservesEqualityRight (a *N c) (additionNIsCommutative (b +N c) (a *N b)))
(transitivity (addingPreservesEqualityRight (a *N c) (equalityCommutative (additionNIsAssociative (a *N b) b c)))
(transitivity (addingPreservesEqualityRight (a *N c) (addingPreservesEqualityRight c (additionNIsCommutative (a *N b) b)))
(transitivity (addingPreservesEqualityRight (a *N c) (addingPreservesEqualityRight {b +N a *N b} {(succ a) *N b} c (refl)))
(transitivity (additionNIsAssociative ((succ a) *N b) c (a *N c))
(transitivity (addingPreservesEqualityLeft (succ a *N b) refl)
refl)
)
))))))
productPreservesEqualityLeft : (a : ) {b c : } b c a *N b a *N c
productPreservesEqualityLeft a {b} {.b} refl = refl
aSucB : (a b : ) a *N succ b a *N b +N a
aSucB a b = transitivity {_} {} {a *N succ b} {a *N (b +N succ zero)} (productPreservesEqualityLeft a (succIsAddOne b)) (transitivity {_} {} {a *N (b +N succ zero)} {a *N b +N a *N succ zero} (productDistributes a b (succ zero)) (addingPreservesEqualityLeft (a *N b) (productWithOneRight a)))
aSucBRight : (a b : ) (succ a) *N b a *N b +N b
aSucBRight a b = additionNIsCommutative b (a *N b)
multiplicationNIsCommutative : (a b : ) (a *N b) (b *N a)
multiplicationNIsCommutative zero b = transitivity (productZeroIsZeroLeft b) (equalityCommutative (productZeroIsZeroRight b))
multiplicationNIsCommutative (succ a) zero = multiplicationNIsCommutative a zero
multiplicationNIsCommutative (succ a) (succ b) = transitivity refl
(transitivity (addingPreservesEqualityLeft (succ b) (aSucB a b))
(transitivity (additionNIsCommutative (succ b) (a *N b +N a))
(transitivity (additionNIsAssociative (a *N b) a (succ b))
(transitivity (addingPreservesEqualityLeft (a *N b) (succCanMove a b))
(transitivity (addingPreservesEqualityLeft (a *N b) (additionNIsCommutative (succ a) b))
(transitivity (equalityCommutative (additionNIsAssociative (a *N b) b (succ a)))
(transitivity (addingPreservesEqualityRight (succ a) (equalityCommutative (aSucBRight a b)))
(transitivity (addingPreservesEqualityRight (succ a) (multiplicationNIsCommutative (succ a) b))
(transitivity (additionNIsCommutative (b *N (succ a)) (succ a))
refl
)))))))))
productDistributes' : (a b c : ) (a +N b) *N c a *N c +N b *N c
productDistributes' a b c rewrite multiplicationNIsCommutative (a +N b) c | productDistributes c a b | multiplicationNIsCommutative c a | multiplicationNIsCommutative c b = refl
flipProductsWithinSum : (a b c : ) (c *N a +N c *N b a *N c +N b *N c)
flipProductsWithinSum a b c = transitivity (addingPreservesEqualityRight (c *N b) (multiplicationNIsCommutative c a)) ((addingPreservesEqualityLeft (a *N c) (multiplicationNIsCommutative c b)))
productDistributesRight : (a b c : ) (a +N b) *N c a *N c +N b *N c
productDistributesRight a b c = transitivity (multiplicationNIsCommutative (a +N b) c) (transitivity (productDistributes c a b) (flipProductsWithinSum a b c))
multiplicationNIsAssociative : (a b c : ) (a *N (b *N c)) ((a *N b) *N c)
multiplicationNIsAssociative zero b c = refl
multiplicationNIsAssociative (succ a) b c =
transitivity refl
(transitivity refl
(transitivity (applyEquality ((λ x b *N c +N x)) (multiplicationNIsAssociative a b c)) (transitivity (equalityCommutative (productDistributesRight b (a *N b) c)) refl)))

View File

@@ -5,196 +5,31 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import WellFoundedInduction
open import Functions
open import Orders
open import Numbers.Naturals.Definition
open import Numbers.Naturals.Addition
open import Numbers.Naturals.Order
open import Numbers.Naturals.Multiplication
open import Semirings.Definition
open import Monoids.Definition
module Numbers.Naturals where
data : Set where
zero :
succ :
module Numbers.Naturals.Naturals where
{-# BUILTIN NATURAL #-}
open Numbers.Naturals.Definition using ( ; zero ; succ; succInjective) public
open Numbers.Naturals.Addition using (_+N_) public
open Numbers.Naturals.Multiplication using (_*N_ ; multiplicationNIsCommutative) public
open Numbers.Naturals.Order using (_<N_ ; le; succPreservesInequality) public
infix 15 _+N_
infix 100 succ
_+N_ :
zero +N y = y
succ x +N y = succ (x +N y)
infix 25 _*N_
_*N_ :
zero *N y = zero
(succ x) *N y = y +N (x *N y)
infix 5 _<NLogical_
_<NLogical_ : Set
zero <NLogical zero = False
zero <NLogical (succ n) = True
(succ n) <NLogical zero = False
(succ n) <NLogical (succ m) = n <NLogical m
infix 5 _<N_
record _<N_ (a : ) (b : ) : Set where
constructor le
field
x :
proof : (succ x) +N a b
infix 5 _≤N_
_≤N_ : Set
a ≤N b = (a <N b) || (a b)
addzeroLeft : (x : ) (zero +N x) x
addzeroLeft x = refl
addZeroRight : (x : ) (x +N zero) x
addZeroRight zero = refl
addZeroRight (succ x) rewrite addZeroRight x = refl
succExtracts : (x y : ) (x +N succ y) (succ (x +N y))
succExtracts zero y = refl
succExtracts (succ x) y = applyEquality succ (succExtracts x y)
succCanMove : (x y : ) (x +N succ y) (succ x +N y)
succCanMove x y = transitivity (succExtracts x y) refl
aLessSucc : (a : ) (a <NLogical succ a)
aLessSucc zero = record {}
aLessSucc (succ a) = aLessSucc a
succPreservesInequalityLogical : {a b : } (a <NLogical b) (succ a <NLogical succ b)
succPreservesInequalityLogical {a} {b} prAB = prAB
lessTransitiveLogical : {a b c : } (a <NLogical b) (b <NLogical c) (a <NLogical c)
lessTransitiveLogical {a} {zero} {zero} prAB prBC = prAB
lessTransitiveLogical {a} {(succ b)} {zero} prAB ()
lessTransitiveLogical {zero} {zero} {(succ c)} prAB prBC = record {}
lessTransitiveLogical {(succ a)} {zero} {(succ c)} () prBC
lessTransitiveLogical {zero} {(succ b)} {(succ c)} prAB prBC = record {}
lessTransitiveLogical {(succ a)} {(succ b)} {(succ c)} prAB prBC = lessTransitiveLogical {a} {b} {c} prAB prBC
aLessXPlusSuccA : (a x : ) (a <NLogical x +N succ a)
aLessXPlusSuccA a zero = aLessSucc a
aLessXPlusSuccA zero (succ x) = record {}
aLessXPlusSuccA (succ a) (succ x) = lessTransitiveLogical {a} {succ a} {x +N succ (succ a)} (aLessXPlusSuccA a zero) (aLessXPlusSuccA (succ a) x)
leImpliesLogical<N : {a b : } (a <N b) (a <NLogical b)
leImpliesLogical<N {zero} {zero} (le x ())
leImpliesLogical<N {zero} {(succ b)} (le x proof) = record {}
leImpliesLogical<N {(succ a)} {zero} (le x ())
leImpliesLogical<N {(succ a)} {(succ .(succ a))} (le zero refl) = aLessSucc a
leImpliesLogical<N {(succ a)} {(succ .(succ (x +N succ a)))} (le (succ x) refl) = succPreservesInequalityLogical {a} {succ (x +N succ a)} (lessTransitiveLogical {a} {succ a} {succ (x +N succ a)} (aLessSucc a) (aLessXPlusSuccA a x))
logical<NImpliesLe : {a b : } (a <NLogical b) (a <N b)
logical<NImpliesLe {zero} {zero} ()
_<N_.x (logical<NImpliesLe {zero} {succ b} prAB) = b
_<N_.proof (logical<NImpliesLe {zero} {succ b} prAB) = applyEquality succ (addZeroRight b)
logical<NImpliesLe {(succ a)} {zero} ()
logical<NImpliesLe {(succ a)} {(succ b)} prAB with logical<NImpliesLe {a} prAB
logical<NImpliesLe {(succ a)} {(succ .(succ (x +N a)))} prAB | le x refl = le x (succCanMove (succ x) a)
lessTransitive : {a b c : } (a <N b) (b <N c) (a <N c)
lessTransitive {a} {b} {c} prab prbc = logical<NImpliesLe (lessTransitiveLogical {a} {b} {c} (leImpliesLogical<N prab) (leImpliesLogical<N prbc))
canRemoveSuccFrom<N : {a b : } (succ a <N succ b) (a <N b)
canRemoveSuccFrom<N {a} {b} prAB = logical<NImpliesLe (leImpliesLogical<N prAB)
succPreservesInequality : {a b : } (a <N b) (succ a <N succ b)
succPreservesInequality {a} {b} prAB = logical<NImpliesLe (succPreservesInequalityLogical {a} {b} (leImpliesLogical<N prAB))
lessIrreflexive : {a : } (a <N a) False
lessIrreflexive {zero} pr = leImpliesLogical<N pr
lessIrreflexive {succ a} pr = lessIrreflexive {a} (logical<NImpliesLe {a} {a} (leImpliesLogical<N {succ a} {succ a} pr))
additionNIsCommutative : (x y : ) (x +N y) (y +N x)
additionNIsCommutative zero y = transitivity (addzeroLeft y) (equalityCommutative (addZeroRight y))
additionNIsCommutative (succ x) zero = transitivity (addZeroRight (succ x)) refl
additionNIsCommutative (succ x) (succ y) =
transitivity refl (applyEquality succ (transitivity (succCanMove x y) (additionNIsCommutative (succ x) y)))
succInjective : {a b : } (succ a succ b) a b
succInjective {a} {.a} refl = refl
addingPreservesEqualityRight : {a b : } (c : ) (a b) (a +N c b +N c)
addingPreservesEqualityRight {a} {b} c pr = applyEquality (λ n -> n +N c) pr
addingPreservesEqualityLeft : {a b : } (c : ) (a b) (c +N a c +N b)
addingPreservesEqualityLeft {a} {b} c pr = applyEquality (λ n -> c +N n) pr
additionNIsAssociative : (a b c : ) ((a +N b) +N c) (a +N (b +N c))
additionNIsAssociative zero b c = refl
additionNIsAssociative (succ a) zero c = transitivity (transitivity (applyEquality (λ n n +N c) (applyEquality succ (addZeroRight a))) refl) (transitivity refl refl)
additionNIsAssociative (succ a) (succ b) c = transitivity refl (transitivity refl (transitivity (applyEquality succ (additionNIsAssociative a (succ b) c)) refl))
productZeroIsZeroLeft : (a : ) (zero *N a zero)
productZeroIsZeroLeft a = refl
productZeroIsZeroRight : (a : ) (a *N zero zero)
productZeroIsZeroRight zero = refl
productZeroIsZeroRight (succ a) = productZeroIsZeroRight a
productWithOneLeft : (a : ) ((succ zero) *N a) a
productWithOneLeft a = transitivity refl (transitivity (applyEquality (λ { m -> a +N m }) refl) (additionNIsCommutative a zero))
productWithOneRight : (a : ) a *N succ zero a
productWithOneRight zero = refl
productWithOneRight (succ a) = transitivity refl (addingPreservesEqualityLeft (succ zero) (productWithOneRight a))
productDistributes : (a b c : ) (a *N (b +N c)) a *N b +N a *N c
productDistributes zero b c = refl
productDistributes (succ a) b c = transitivity refl
(transitivity (addingPreservesEqualityLeft (b +N c) (productDistributes a b c))
(transitivity (equalityCommutative (additionNIsAssociative (b +N c) (a *N b) (a *N c)))
(transitivity (addingPreservesEqualityRight (a *N c) (additionNIsCommutative (b +N c) (a *N b)))
(transitivity (addingPreservesEqualityRight (a *N c) (equalityCommutative (additionNIsAssociative (a *N b) b c)))
(transitivity (addingPreservesEqualityRight (a *N c) (addingPreservesEqualityRight c (additionNIsCommutative (a *N b) b)))
(transitivity (addingPreservesEqualityRight (a *N c) (addingPreservesEqualityRight {b +N a *N b} {(succ a) *N b} c (refl)))
(transitivity (additionNIsAssociative ((succ a) *N b) c (a *N c))
(transitivity (addingPreservesEqualityLeft (succ a *N b) refl)
refl)
)
))))))
succIsAddOne : (a : ) succ a a +N succ zero
succIsAddOne a = equalityCommutative (transitivity (additionNIsCommutative a (succ zero)) refl)
productPreservesEqualityLeft : (a : ) {b c : } b c a *N b a *N c
productPreservesEqualityLeft a {b} {.b} refl = refl
aSucB : (a b : ) a *N succ b a *N b +N a
aSucB a b = transitivity {_} {} {a *N succ b} {a *N (b +N succ zero)} (productPreservesEqualityLeft a (succIsAddOne b)) (transitivity {_} {} {a *N (b +N succ zero)} {a *N b +N a *N succ zero} (productDistributes a b (succ zero)) (addingPreservesEqualityLeft (a *N b) (productWithOneRight a)))
aSucBRight : (a b : ) (succ a) *N b a *N b +N b
aSucBRight a b = additionNIsCommutative b (a *N b)
multiplicationNIsCommutative : (a b : ) (a *N b) (b *N a)
multiplicationNIsCommutative zero b = transitivity (productZeroIsZeroLeft b) (equalityCommutative (productZeroIsZeroRight b))
multiplicationNIsCommutative (succ a) zero = multiplicationNIsCommutative a zero
multiplicationNIsCommutative (succ a) (succ b) = transitivity refl
(transitivity (addingPreservesEqualityLeft (succ b) (aSucB a b))
(transitivity (additionNIsCommutative (succ b) (a *N b +N a))
(transitivity (additionNIsAssociative (a *N b) a (succ b))
(transitivity (addingPreservesEqualityLeft (a *N b) (succCanMove a b))
(transitivity (addingPreservesEqualityLeft (a *N b) (additionNIsCommutative (succ a) b))
(transitivity (equalityCommutative (additionNIsAssociative (a *N b) b (succ a)))
(transitivity (addingPreservesEqualityRight (succ a) (equalityCommutative (aSucBRight a b)))
(transitivity (addingPreservesEqualityRight (succ a) (multiplicationNIsCommutative (succ a) b))
(transitivity (additionNIsCommutative (b *N (succ a)) (succ a))
refl
)))))))))
productDistributes' : (a b c : ) (a +N b) *N c a *N c +N b *N c
productDistributes' a b c rewrite multiplicationNIsCommutative (a +N b) c | productDistributes c a b | multiplicationNIsCommutative c a | multiplicationNIsCommutative c b = refl
flipProductsWithinSum : (a b c : ) (c *N a +N c *N b a *N c +N b *N c)
flipProductsWithinSum a b c = transitivity (addingPreservesEqualityRight (c *N b) (multiplicationNIsCommutative c a)) ((addingPreservesEqualityLeft (a *N c) (multiplicationNIsCommutative c b)))
productDistributesRight : (a b c : ) (a +N b) *N c a *N c +N b *N c
productDistributesRight a b c = transitivity (multiplicationNIsCommutative (a +N b) c) (transitivity (productDistributes c a b) (flipProductsWithinSum a b c))
multiplicationNIsAssociative : (a b c : ) (a *N (b *N c)) ((a *N b) *N c)
multiplicationNIsAssociative zero b c = refl
multiplicationNIsAssociative (succ a) b c =
transitivity refl
(transitivity refl
(transitivity (applyEquality ((λ x b *N c +N x)) (multiplicationNIsAssociative a b c)) (transitivity (equalityCommutative (productDistributesRight b (a *N b) c)) refl)))
Semiring : Semiring 0 1 _+N_ _*N_
Monoid.associative (Semiring.monoid Semiring) a b c = equalityCommutative (additionNIsAssociative a b c)
Monoid.idLeft (Semiring.monoid Semiring) _ = refl
Monoid.idRight (Semiring.monoid Semiring) a = additionNIsCommutative a 0
Semiring.commutative Semiring = additionNIsCommutative
Monoid.associative (Semiring.multMonoid Semiring) = multiplicationNIsAssociative
Monoid.idLeft (Semiring.multMonoid Semiring) a = additionNIsCommutative a 0
Monoid.idRight (Semiring.multMonoid Semiring) a = transitivity (multiplicationNIsCommutative a 1) (additionNIsCommutative a 0)
Semiring.productZeroLeft Semiring _ = refl
Semiring.productZeroRight Semiring a = multiplicationNIsCommutative a 0
Semiring.+DistributesOver* Semiring = productDistributes
canAddZeroOnLeft : {a b : } (a <N b) (a +N zero) <N b
canAddZeroOnLeft {a} {b} prAB = identityOfIndiscernablesLeft a b (a +N zero) (λ x y -> x <N y) prAB (equalityCommutative (addZeroRight a))

View File

@@ -0,0 +1,73 @@
{-# OPTIONS --warning=error --safe --without-K #-}
open import LogicalFormulae
open import Numbers.Naturals.Definition
open import Numbers.Naturals.Addition
module Numbers.Naturals.Order where
infix 5 _<NLogical_
_<NLogical_ : Set
zero <NLogical zero = False
zero <NLogical (succ n) = True
(succ n) <NLogical zero = False
(succ n) <NLogical (succ m) = n <NLogical m
infix 5 _<N_
record _<N_ (a : ) (b : ) : Set where
constructor le
field
x :
proof : (succ x) +N a b
infix 5 _≤N_
_≤N_ : Set
a ≤N b = (a <N b) || (a b)
aLessSucc : (a : ) (a <NLogical succ a)
aLessSucc zero = record {}
aLessSucc (succ a) = aLessSucc a
succPreservesInequalityLogical : {a b : } (a <NLogical b) (succ a <NLogical succ b)
succPreservesInequalityLogical {a} {b} prAB = prAB
lessTransitiveLogical : {a b c : } (a <NLogical b) (b <NLogical c) (a <NLogical c)
lessTransitiveLogical {a} {zero} {zero} prAB prBC = prAB
lessTransitiveLogical {a} {(succ b)} {zero} prAB ()
lessTransitiveLogical {zero} {zero} {(succ c)} prAB prBC = record {}
lessTransitiveLogical {(succ a)} {zero} {(succ c)} () prBC
lessTransitiveLogical {zero} {(succ b)} {(succ c)} prAB prBC = record {}
lessTransitiveLogical {(succ a)} {(succ b)} {(succ c)} prAB prBC = lessTransitiveLogical {a} {b} {c} prAB prBC
aLessXPlusSuccA : (a x : ) (a <NLogical x +N succ a)
aLessXPlusSuccA a zero = aLessSucc a
aLessXPlusSuccA zero (succ x) = record {}
aLessXPlusSuccA (succ a) (succ x) = lessTransitiveLogical {a} {succ a} {x +N succ (succ a)} (aLessXPlusSuccA a zero) (aLessXPlusSuccA (succ a) x)
leImpliesLogical<N : {a b : } (a <N b) (a <NLogical b)
leImpliesLogical<N {zero} {zero} (le x ())
leImpliesLogical<N {zero} {(succ b)} (le x proof) = record {}
leImpliesLogical<N {(succ a)} {zero} (le x ())
leImpliesLogical<N {(succ a)} {(succ .(succ a))} (le zero refl) = aLessSucc a
leImpliesLogical<N {(succ a)} {(succ .(succ (x +N succ a)))} (le (succ x) refl) = succPreservesInequalityLogical {a} {succ (x +N succ a)} (lessTransitiveLogical {a} {succ a} {succ (x +N succ a)} (aLessSucc a) (aLessXPlusSuccA a x))
logical<NImpliesLe : {a b : } (a <NLogical b) (a <N b)
logical<NImpliesLe {zero} {zero} ()
_<N_.x (logical<NImpliesLe {zero} {succ b} prAB) = b
_<N_.proof (logical<NImpliesLe {zero} {succ b} prAB) = applyEquality succ (addZeroRight b)
logical<NImpliesLe {(succ a)} {zero} ()
logical<NImpliesLe {(succ a)} {(succ b)} prAB with logical<NImpliesLe {a} prAB
logical<NImpliesLe {(succ a)} {(succ .(succ (x +N a)))} prAB | le x refl = le x (succCanMove (succ x) a)
lessTransitive : {a b c : } (a <N b) (b <N c) (a <N c)
lessTransitive {a} {b} {c} prab prbc = logical<NImpliesLe (lessTransitiveLogical {a} {b} {c} (leImpliesLogical<N prab) (leImpliesLogical<N prbc))
lessIrreflexive : {a : } (a <N a) False
lessIrreflexive {zero} pr = leImpliesLogical<N pr
lessIrreflexive {succ a} pr = lessIrreflexive {a} (logical<NImpliesLe {a} {a} (leImpliesLogical<N {succ a} {succ a} pr))
succPreservesInequality : {a b : } (a <N b) (succ a <N succ b)
succPreservesInequality {a} {b} prAB = logical<NImpliesLe (succPreservesInequalityLogical {a} {b} (leImpliesLogical<N prAB))
canRemoveSuccFrom<N : {a b : } (succ a <N succ b) (a <N b)
canRemoveSuccFrom<N {a} {b} prAB = logical<NImpliesLe (leImpliesLogical<N prAB)

View File

@@ -5,9 +5,9 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import WellFoundedInduction
open import Functions
open import Orders
open import Numbers.Naturals
open import Numbers.Naturals.Naturals
module Numbers.NaturalsWithK where
module Numbers.Naturals.WithK where
<NRefl : {a b : } (p1 p2 : a <N b) (p1 p2)
<NRefl {a} {.(succ (p1 +N a))} (le p1 refl) (le p2 proof2) = help p1=p2 proof2

View File

@@ -1,10 +1,10 @@
{-# OPTIONS --safe --warning=error #-}
open import LogicalFormulae
open import Numbers.Naturals
open import Numbers.Naturals.Naturals
open import Numbers.Integers
open import Groups.Groups
open import Groups.GroupDefinition
open import Groups.Definition
open import Rings.Definition
open import Fields.Fields
open import PrimeNumbers

View File

@@ -1,12 +1,12 @@
{-# OPTIONS --safe --warning=error #-}
open import LogicalFormulae
open import Numbers.Naturals
open import Numbers.Naturals.Naturals
open import Numbers.Integers
open import Numbers.Rationals
open import Groups.Groups
open import Groups.GroupsLemmas
open import Groups.GroupDefinition
open import Groups.Lemmas
open import Groups.Definition
open import Rings.Definition
open import Fields.Fields
open import PrimeNumbers

View File

@@ -33,6 +33,9 @@ module Orders where
max a b | inl (inr b<a) = a
max a b | inr a=b = b
irreflexive = PartialOrder.irreflexive order
transitive = PartialOrder.transitive order
equivMin : {a b : _} {A : Set a} (order : TotalOrder {a} {b} A) {x y : A} (TotalOrder._<_ order x y) TotalOrder.min order x y x
equivMin order {x} {y} x<y with TotalOrder.totality order x y
equivMin order {x} {y} x<y | inl (inl x₁) = refl

View File

@@ -1,8 +1,11 @@
{-# OPTIONS --warning=error --safe #-}
open import LogicalFormulae
open import Numbers.Naturals
open import Numbers.NaturalsWithK
open import Numbers.Naturals.Naturals
open import Numbers.Naturals.Addition -- TODO - remove this dependency
open import Numbers.Naturals.Multiplication -- TODO - remove this dependency
open import Numbers.Naturals.Order -- TODO - remove this dependency
open import Numbers.Naturals.WithK
open import WellFoundedInduction
open import KeyValue
open import KeyValueWithDomain
@@ -10,6 +13,7 @@ open import Orders
open import Vectors
open import Maybe
open import WithK
open import Semirings.Definition
module PrimeNumbers where
@@ -22,15 +26,15 @@ module PrimeNumbers where
quotSmall : (0 <N a) || ((0 a) && (quot 0))
divAlgLessLemma : (a b : ) (0 <N a) (r : divisionAlgResult a b) (divisionAlgResult.quot r 0) || (divisionAlgResult.rem r <N b)
divAlgLessLemma zero b pr _ = exFalso (lessIrreflexive pr)
divAlgLessLemma zero b pr _ = exFalso (TotalOrder.irreflexive TotalOrder pr)
divAlgLessLemma (succ a) b _ record { quot = zero ; rem = a%b ; pr = pr ; remIsSmall = remIsSmall } = inl refl
divAlgLessLemma (succ a) b _ record { quot = (succ a/b) ; rem = a%b ; pr = pr ; remIsSmall = remIsSmall } = inr record { x = a/b +N a *N succ (a/b) ; proof = pr }
modUniqueLemma : {rem1 rem2 a : } (quot1 quot2 : ) rem1 <N a rem2 <N a a *N quot1 +N rem1 a *N quot2 +N rem2 rem1 rem2
modUniqueLemma {rem1} {rem2} {a} zero zero rem1<a rem2<a pr rewrite productZeroIsZeroRight a = pr
modUniqueLemma {rem1} {rem2} {a} zero (succ quot2) rem1<a rem2<a pr rewrite productZeroIsZeroRight a | pr | multiplicationNIsCommutative a (succ quot2) | additionNIsAssociative a (quot2 *N a) rem2 = exFalso (cannotAddAndEnlarge' {a} {quot2 *N a +N rem2} rem1<a)
modUniqueLemma {rem1} {rem2} {a} (succ quot1) zero rem1<a rem2<a pr rewrite productZeroIsZeroRight a | equalityCommutative pr | multiplicationNIsCommutative a (succ quot1) | additionNIsAssociative a (quot1 *N a) rem1 = exFalso (cannotAddAndEnlarge' {a} {quot1 *N a +N rem1} rem2<a)
modUniqueLemma {rem1} {rem2} {a} (succ quot1) (succ quot2) rem1<a rem2<a pr rewrite multiplicationNIsCommutative a (succ quot1) | multiplicationNIsCommutative a (succ quot2) | additionNIsAssociative a (quot1 *N a) rem1 | additionNIsAssociative a (quot2 *N a) rem2 = modUniqueLemma {rem1} {rem2} {a} quot1 quot2 rem1<a rem2<a (go {a}{quot1}{rem1}{quot2}{rem2} pr)
modUniqueLemma {rem1} {rem2} {a} zero zero rem1<a rem2<a pr rewrite Semiring.productZeroRight Semiring a = pr
modUniqueLemma {rem1} {rem2} {a} zero (succ quot2) rem1<a rem2<a pr rewrite Semiring.productZeroRight Semiring a | pr | multiplicationNIsCommutative a (succ quot2) | equalityCommutative (Semiring.+Associative Semiring a (quot2 *N a) rem2) = exFalso (cannotAddAndEnlarge' {a} {quot2 *N a +N rem2} rem1<a)
modUniqueLemma {rem1} {rem2} {a} (succ quot1) zero rem1<a rem2<a pr rewrite Semiring.productZeroRight Semiring a | equalityCommutative pr | multiplicationNIsCommutative a (succ quot1) | equalityCommutative (Semiring.+Associative Semiring a (quot1 *N a) rem1) = exFalso (cannotAddAndEnlarge' {a} {quot1 *N a +N rem1} rem2<a)
modUniqueLemma {rem1} {rem2} {a} (succ quot1) (succ quot2) rem1<a rem2<a pr rewrite multiplicationNIsCommutative a (succ quot1) | multiplicationNIsCommutative a (succ quot2) | equalityCommutative (Semiring.+Associative Semiring a (quot1 *N a) rem1) | equalityCommutative (Semiring.+Associative Semiring a (quot2 *N a) rem2) = modUniqueLemma {rem1} {rem2} {a} quot1 quot2 rem1<a rem2<a (go {a}{quot1}{rem1}{quot2}{rem2} pr)
where
go : {a quot1 rem1 quot2 rem2 : } (a +N (quot1 *N a +N rem1) a +N (quot2 *N a +N rem2)) a *N quot1 +N rem1 a *N quot2 +N rem2
go {a} {quot1} {rem1} {quot2} {rem2} pr rewrite multiplicationNIsCommutative quot1 a | multiplicationNIsCommutative quot2 a = canSubtractFromEqualityLeft {a} pr
@@ -42,16 +46,16 @@ module PrimeNumbers where
modIsUnique {succ a} {b} record { quot = quot1 ; rem = rem1 ; pr = pr1 ; remIsSmall = remIsSmall1 } record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = (inr ()) }
transferAddition : (a b c : ) (a +N b) +N c (a +N c) +N b
transferAddition a b c rewrite (additionNIsAssociative a b c) = p a b c
transferAddition a b c rewrite equalityCommutative (Semiring.+Associative Semiring a b c) = p a b c
where
p : (a b c : ) a +N (b +N c) (a +N c) +N b
p a b c rewrite (additionNIsCommutative b c) = equalityCommutative (additionNIsAssociative a c b)
p a b c = transitivity (applyEquality (a +N_) (Semiring.commutative Semiring b c)) (Semiring.+Associative Semiring a c b)
divisionAlgLemma : (x b : ) x *N zero +N b b
divisionAlgLemma x b rewrite (productZeroIsZeroRight x) = refl
divisionAlgLemma x b rewrite (Semiring.productZeroRight Semiring x) = refl
divisionAlgLemma2 : (x b : ) (x b) x *N succ zero +N zero b
divisionAlgLemma2 x b pr rewrite (productWithOneRight x) = equalityCommutative (transitivity (equalityCommutative pr) (equalityCommutative (addZeroRight x)))
divisionAlgLemma2 x b pr rewrite (Semiring.productOneRight Semiring x) = equalityCommutative (transitivity (equalityCommutative pr) (equalityCommutative (Semiring.sumZeroRight Semiring x)))
divisionAlgLemma3 : {a x : } (p : succ a <N succ x) (subtractionNResult.result (-N (inl p))) <N (succ x)
divisionAlgLemma3 {a} {x} p = -NIsDecreasing {a} {succ x} p
@@ -225,7 +229,7 @@ module PrimeNumbers where
numberLessThanOrder : (n : ) TotalOrder (numberLessThan n)
PartialOrder._<_ (TotalOrder.order (numberLessThanOrder n)) = λ a b (numberLessThan.a a) <N numberLessThan.a b
PartialOrder.irreflexive (TotalOrder.order (numberLessThanOrder n)) pr = lessIrreflexive pr
PartialOrder.irreflexive (TotalOrder.order (numberLessThanOrder n)) pr = TotalOrder.irreflexive TotalOrder pr
PartialOrder.transitive (TotalOrder.order (numberLessThanOrder n)) pr1 pr2 = orderIsTransitive pr1 pr2
TotalOrder.totality (numberLessThanOrder n) a b with TotalOrder.totality TotalOrder (numberLessThan.a a) (numberLessThan.a b)
TotalOrder.totality (numberLessThanOrder n) a b | inl (inl x) = inl (inl x)

View File

@@ -2,8 +2,8 @@
open import LogicalFormulae
open import Groups.Groups
open import Groups.GroupDefinition
open import Numbers.Naturals
open import Groups.Definition
open import Numbers.Naturals.Naturals
open import Setoids.Orders
open import Setoids.Setoids
open import Functions

View File

@@ -3,7 +3,7 @@
open import LogicalFormulae
open import Groups.Groups
open import Functions
open import Numbers.Naturals
open import Numbers.Naturals.Naturals
open import Numbers.Integers
open import IntegersModN
open import Rings.Examples.Proofs

View File

@@ -3,10 +3,10 @@
open import LogicalFormulae
open import Functions
open import Groups.Groups
open import Groups.GroupDefinition
open import Groups.Definition
open import Orders
open import Rings.Definition
open import Numbers.Naturals
open import Numbers.Naturals.Naturals
open import Numbers.Integers
open import PrimeNumbers
open import IntegersModN

View File

@@ -2,8 +2,8 @@
open import LogicalFormulae
open import Groups.Groups
open import Groups.GroupDefinition
open import Numbers.Naturals
open import Groups.Definition
open import Numbers.Naturals.Naturals
open import Orders
open import Setoids.Setoids
open import Functions

View File

@@ -3,8 +3,8 @@
open import LogicalFormulae
open import Functions
open import Groups.Groups
open import Groups.GroupDefinition
open import Groups.GroupsLemmas
open import Groups.Definition
open import Groups.Lemmas
open import Rings.Definition
open import Setoids.Setoids
open import Setoids.Orders

23
Semirings/Definition.agda Normal file
View File

@@ -0,0 +1,23 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Monoids.Definition
module Semirings.Definition where
record Semiring {a : _} {A : Set a} (Zero One : A) (_+_ : A A A) (_*_ : A A A) : Set a where
field
monoid : Monoid Zero _+_
commutative : (a b : A) a + b b + a
multMonoid : Monoid One _*_
productZeroLeft : (a : A) Zero * a Zero
productZeroRight : (a : A) a * Zero Zero
+DistributesOver* : (a b c : A) a * (b + c) (a * b) + (a * c)
+Associative = Monoid.associative monoid
*Associative = Monoid.associative multMonoid
productOneLeft = Monoid.idLeft multMonoid
productOneRight = Monoid.idRight multMonoid
sumZeroLeft = Monoid.idLeft monoid
sumZeroRight = Monoid.idRight monoid

View File

@@ -1,7 +1,7 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Numbers.Naturals
open import Numbers.Naturals.Naturals
open import Lists.Lists
open import Setoids.Setoids
open import Functions

View File

@@ -4,8 +4,9 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import LogicalFormulae
open import Functions
open import Numbers.Naturals
open import Numbers.Naturals.Naturals
open import Sets.FinSet
open import Semirings.Definition
module Sets.Cardinality where
record CountablyInfiniteSet {a : _} (A : Set a) : Set a where
@@ -25,13 +26,13 @@ module Sets.Cardinality where
evenCannotBeOdd : (a b : ) 2 *N a succ (2 *N b) False
evenCannotBeOdd zero b ()
evenCannotBeOdd (succ a) zero pr rewrite additionNIsCommutative a 0 | additionNIsCommutative a (succ a) = naughtE (equalityCommutative (succInjective pr))
evenCannotBeOdd (succ a) zero pr rewrite Semiring.commutative Semiring a 0 | Semiring.commutative Semiring a (succ a) = naughtE (equalityCommutative (succInjective pr))
evenCannotBeOdd (succ a) (succ b) pr = evenCannotBeOdd a b pr''
where
pr' : a +N a (b +N succ b)
pr' rewrite additionNIsCommutative a 0 | additionNIsCommutative b 0 | additionNIsCommutative a (succ a) = succInjective (succInjective pr)
pr' rewrite Semiring.commutative Semiring a 0 | Semiring.commutative Semiring b 0 | Semiring.commutative Semiring a (succ a) = succInjective (succInjective pr)
pr'' : 2 *N a succ (2 *N b)
pr'' rewrite additionNIsCommutative a 0 | additionNIsCommutative b 0 | additionNIsCommutative (succ b) b = pr'
pr'' rewrite Semiring.commutative Semiring a 0 | Semiring.commutative Semiring b 0 | Semiring.commutative Semiring (succ b) b = pr'
aMod2 : (a : ) Sg (λ i (2 *N i a) || (succ (2 *N i) a))
aMod2 zero = (0 , inl refl)
@@ -40,7 +41,7 @@ module Sets.Cardinality where
aMod2 (succ a) | b , inr x = (succ b) , inl pr
where
pr : succ (b +N succ (b +N 0)) succ a
pr rewrite additionNIsCommutative b (succ (b +N 0)) | additionNIsCommutative (b +N 0) b = applyEquality succ x
pr rewrite Semiring.commutative Semiring b (succ (b +N 0)) | Semiring.commutative Semiring (b +N 0) b = applyEquality succ x
sqrtFloor : (a : ) Sg ( && ) (λ pair ((_&&_.fst pair) *N (_&&_.fst pair) +N (_&&_.snd pair) a) && ((_&&_.snd pair) <N 2 *N (_&&_.fst pair) +N 1))
sqrtFloor zero = (0 ,, 0) , (refl ,, le zero refl)
@@ -49,16 +50,16 @@ module Sets.Cardinality where
sqrtFloor (succ n) | (a ,, b) , pr | inl (inl x) = (a ,, succ b) , (p ,, q)
where
p : a *N a +N succ b succ n
p rewrite additionNIsCommutative (a *N a) (succ b) | additionNIsCommutative b (a *N a) = applyEquality succ (_&&_.fst pr)
p rewrite Semiring.commutative Semiring (a *N a) (succ b) | Semiring.commutative Semiring b (a *N a) = applyEquality succ (_&&_.fst pr)
q : succ b <N (a +N (a +N 0)) +N 1
q rewrite additionNIsCommutative (a +N (a +N 0)) (succ 0) | additionNIsCommutative a 0 = succPreservesInequality x
sqrtFloor (succ n) | (a ,, b) , (_ ,, pr) | inl (inr x) rewrite additionNIsCommutative (a +N (a +N 0)) (succ 0) = exFalso (noIntegersBetweenXAndSuccX (a +N (a +N 0)) x pr)
q rewrite Semiring.commutative Semiring (a +N (a +N 0)) (succ 0) | Semiring.commutative Semiring a 0 = succPreservesInequality x
sqrtFloor (succ n) | (a ,, b) , (_ ,, pr) | inl (inr x) rewrite Semiring.commutative Semiring (a +N (a +N 0)) (succ 0) = exFalso (noIntegersBetweenXAndSuccX (a +N (a +N 0)) x pr)
sqrtFloor (succ n) | (a ,, b) , pr | inr x = (succ a ,, 0) , (q ,, succIsPositive _)
where
p : a +N a *N succ a n
p rewrite x | additionNIsCommutative a 0 | additionNIsCommutative (a +N a) (succ 0) | additionNIsCommutative (a *N a) (succ a +N a) | multiplicationNIsCommutative a (succ a) | additionNIsCommutative (a *N a) (a +N a) | additionNIsAssociative a a (a *N a) = _&&_.fst pr
p rewrite x | Semiring.commutative Semiring a 0 | Semiring.commutative Semiring (a +N a) (succ 0) | Semiring.commutative Semiring (a *N a) (succ a +N a) | multiplicationNIsCommutative a (succ a) | Semiring.commutative Semiring (a *N a) (a +N a) | Semiring.+Associative Semiring a a (a *N a) = _&&_.fst pr
q : succ ((a +N a *N succ a) +N 0) succ n
q rewrite additionNIsCommutative (a +N a *N succ a) 0 = applyEquality succ p
q rewrite Semiring.commutative Semiring (a +N a *N succ a) 0 = applyEquality succ p
{-
triangle : (a : ) →
@@ -69,11 +70,11 @@ module Sets.Cardinality where
cantorPairing' a b zero pr = 0
cantorPairing' zero zero (succ s) ()
cantorPairing' zero (succ b) (succ s) pr = succ (cantorPairing' 1 b (succ s) pr)
cantorPairing' (succ a) zero (succ s) pr = succ (cantorPairing' 0 a s (succInjective (identityOfIndiscernablesRight _ _ _ _≡_ pr (applyEquality (λ i → succ i) (identityOfIndiscernablesLeft _ _ _ _≡_ (additionNIsCommutative a 0) refl)))))
cantorPairing' (succ a) zero (succ s) pr = succ (cantorPairing' 0 a s (succInjective (identityOfIndiscernablesRight _ _ _ _≡_ pr (applyEquality (λ i → succ i) (identityOfIndiscernablesLeft _ _ _ _≡_ (Semiring.commutative Semiring a 0) refl)))))
cantorPairing' (succ a) (succ b) (succ zero) pr = naughtE f
where
f : 0 ≡ succ b +N a
f rewrite additionNIsCommutative (succ b) a = succInjective pr
f rewrite Semiring.commutative Semiring (succ b) a = succInjective pr
cantorPairing' (succ a) (succ b) (succ (succ s)) pr = succ (cantorPairing' (succ (succ a)) b (succ (succ s)) (identityOfIndiscernablesRight _ _ _ _≡_ pr (applyEquality succ (succExtracts a b))))
cantorPairing'Zero : (x y s : ) → (pr : s ≡ x +N y) → cantorPairing' x y s pr ≡ 0 → (x ≡ 0) && (y ≡ 0)
@@ -81,7 +82,7 @@ module Sets.Cardinality where
cantorPairing'Zero zero zero (succ s) () pr'
cantorPairing'Zero zero (succ y) (succ s) pr ()
cantorPairing'Zero (succ x) zero (succ s) pr ()
cantorPairing'Zero (succ x) (succ y) (succ zero) pr pr' = naughtE (identityOfIndiscernablesRight _ _ _ _≡_ (succInjective pr) (additionNIsCommutative x (succ y)))
cantorPairing'Zero (succ x) (succ y) (succ zero) pr pr' = naughtE (identityOfIndiscernablesRight _ _ _ _≡_ (succInjective pr) (Semiring.commutative Semiring x (succ y)))
cantorPairing'Zero (succ x) (succ y) (succ (succ s)) pr ()
cantorPairingInj' : (s1 s2 : ) → (x1 x2 y1 y2 : ) → (pr1 : s1 ≡ x1 +N y1) → (pr2 : s2 ≡ x2 +N y2) → (cantorPairing' x1 y1 s1 pr1) ≡ (cantorPairing' x2 y2 s2 pr2) → (x1 ≡ x2) && (y1 ≡ y2)
@@ -130,7 +131,7 @@ module Sets.Cardinality where
rec = succInjective injF
rec' : (1 ≡ zero)
rec' = _&&_.fst (cantorPairingInj' (succ s1) s2 1 zero s1 x2 refl _ rec)
cantorPairingInj' (succ s1) (succ zero) zero (succ x2) (succ .s1) (succ y2) refl pr2 injF = naughtE (identityOfIndiscernablesRight _ _ _ _≡_ (succInjective pr2) (additionNIsCommutative x2 (succ y2)))
cantorPairingInj' (succ s1) (succ zero) zero (succ x2) (succ .s1) (succ y2) refl pr2 injF = naughtE (identityOfIndiscernablesRight _ _ _ _≡_ (succInjective pr2) (Semiring.commutative Semiring x2 (succ y2)))
cantorPairingInj' (succ s1) (succ (succ s2)) zero (succ x2) (succ .s1) (succ y2) refl pr2 injF = naughtE (succInjective rec')
where
rec : cantorPairing' 1 s1 (succ s1) refl ≡ cantorPairing' (succ (succ x2)) y2 (succ (succ s2)) _
@@ -178,7 +179,7 @@ module Sets.Cardinality where
pairUnionIsCountable : {a b : _} {A : Set a} {B : Set b} (X : CountablyInfiniteSet A) (Y : CountablyInfiniteSet B) CountablyInfiniteSet (A || B)
CountablyInfiniteSet.counting (pairUnionIsCountable X Y) (inl r) = 2 *N (CountablyInfiniteSet.counting X r)
CountablyInfiniteSet.counting (pairUnionIsCountable X Y) (inr s) = succ (2 *N (CountablyInfiniteSet.counting Y s))
Injection.property (Bijection.inj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y))) {inl x} {inl y} pr rewrite additionNIsCommutative (CountablyInfiniteSet.counting X x) 0 | additionNIsCommutative (CountablyInfiniteSet.counting X y) 0 | doubleIsAddTwo (CountablyInfiniteSet.counting X x) | doubleIsAddTwo (CountablyInfiniteSet.counting X y) = applyEquality inl (Injection.property (Bijection.inj (CountablyInfiniteSet.countingIsBij X)) inter)
Injection.property (Bijection.inj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y))) {inl x} {inl y} pr rewrite Semiring.commutative Semiring (CountablyInfiniteSet.counting X x) 0 | Semiring.commutative Semiring (CountablyInfiniteSet.counting X y) 0 | doubleIsAddTwo (CountablyInfiniteSet.counting X x) | doubleIsAddTwo (CountablyInfiniteSet.counting X y) = applyEquality inl (Injection.property (Bijection.inj (CountablyInfiniteSet.countingIsBij X)) inter)
where
inter : CountablyInfiniteSet.counting X x CountablyInfiniteSet.counting X y
inter = doubleLemma (CountablyInfiniteSet.counting X x) (CountablyInfiniteSet.counting X y) pr

View File

@@ -3,7 +3,9 @@
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import LogicalFormulae
open import Numbers.Naturals
open import Numbers.Naturals.Naturals
open import Numbers.Naturals.Definition
open import Numbers.Naturals.Order
open import Orders
module Sets.FinSet where

View File

@@ -1,9 +1,10 @@
{-# OPTIONS --safe --warning=error #-}
open import Numbers.Naturals
open import Numbers.Naturals.Naturals
open import Numbers.Naturals.Order
open import Sets.FinSet
open import LogicalFormulae
open import Numbers.NaturalsWithK
open import Numbers.Naturals.WithK
module Sets.FinSetWithK where

View File

@@ -1,9 +1,11 @@
{-# OPTIONS --warning=error --safe #-}
open import LogicalFormulae
open import Numbers.Naturals
open import Numbers.NaturalsWithK
open import Numbers.Naturals.Naturals
open import Numbers.Naturals.Order
open import Numbers.Naturals.WithK
open import Functions
open import Semirings.Definition
data Vec {a : _} (X : Set a) : -> Set a where
[] : Vec X zero
@@ -67,7 +69,7 @@ vecSolelyContains {m} n record { index = zero ; index<m = _ ; isHere = isHere }
vecSolelyContains {m} n record { index = (succ index) ; index<m = (le x proof) ; isHere = isHere } = exFalso (f {x} {index} proof)
where
f : {x index : } succ (x +N succ index) 1 False
f {x} {index} pr rewrite additionNIsCommutative x (succ index) = naughtE (equalityCommutative (succInjective pr))
f {x} {index} pr rewrite Semiring.commutative Semiring x (succ index) = naughtE (equalityCommutative (succInjective pr))
vecChop : {a : _} {X : Set a} (m : ) {n : } Vec X (m +N n) Vec X m && Vec X n
_&&_.fst (vecChop zero xs) = []