Problem 1 of Project Euler (#85)

This commit is contained in:
Patrick Stevens
2019-12-04 19:53:34 +00:00
committed by GitHub
parent 2ed7bd8044
commit cfd9787bb8
10 changed files with 526 additions and 447 deletions

37
Decidable/Relations.agda Normal file
View File

@@ -0,0 +1,37 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Decidable.Relations where
DecidableRelation : {a b : _} {A : Set a} (f : A Set b) Set (a b)
DecidableRelation {A = A} f = (x : A) (f x) || (f x False)
orDecidable : {a b c : _} {A : Set a} {f : A Set b} {g : A Set c} DecidableRelation f DecidableRelation g DecidableRelation (λ x (f x) || (g x))
orDecidable fDec gDec x with fDec x
orDecidable fDec gDec x | inl fx = inl (inl fx)
orDecidable fDec gDec x | inr !fx with gDec x
orDecidable fDec gDec x | inr !fx | inl gx = inl (inr gx)
orDecidable {f = f} {g} fDec gDec x | inr !fx | inr !gx = inr t
where
t : (f x || g x) False
t (inl x) = !fx x
t (inr x) = !gx x
andDecidable : {a b c : _} {A : Set a} {f : A Set b} {g : A Set c} DecidableRelation f DecidableRelation g DecidableRelation (λ x (f x) && (g x))
andDecidable fDec gDec x with fDec x
andDecidable fDec gDec x | inl fx with gDec x
andDecidable fDec gDec x | inl fx | inl gx = inl (fx ,, gx)
andDecidable fDec gDec x | inl fx | inr !gx = inr (λ x !gx (_&&_.snd x))
andDecidable fDec gDec x | inr !fx = inr (λ x !fx (_&&_.fst x))
notDecidable : {a b : _} {A : Set a} {f : A Set b} DecidableRelation f DecidableRelation (λ x (f x) False)
notDecidable fDec x with fDec x
notDecidable fDec x | inl fx = inr (λ x x fx)
notDecidable fDec x | inr !fx = inl !fx
doubleNegation : {a b : _} {A : Set a} {f : A Set b} DecidableRelation f (x : A) (((f x) False) False) f x
doubleNegation fDec x with fDec x
doubleNegation fDec x | inl fx = λ _ fx
doubleNegation fDec x | inr !fx = λ p exFalso (p !fx)

10
Decidable/Sets.agda Normal file
View File

@@ -0,0 +1,10 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Decidable.Sets where
record DecidableSet {a : _} (A : Set a) : Set a where
field
eq : (a b : A) ((a b) || ((a b) False))

View File

@@ -1,10 +0,0 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module DecidableSet where
record DecidableSet {a : _} (A : Set a) : Set a where
field
eq : (a b : A) ((a b) || ((a b) False))

View File

@@ -79,7 +79,8 @@ open import Setoids.Functions.Extension
open import Sets.Cardinality open import Sets.Cardinality
open import Sets.FinSet open import Sets.FinSet
open import DecidableSet open import Decidable.Sets
open import Decidable.Relations
open import Vectors open import Vectors
open import Vectors.VectorSpace open import Vectors.VectorSpace

View File

@@ -27,4 +27,6 @@ open import Groups.LectureNotes.Lecture1
open import LectureNotes.NumbersAndSets.Lecture1 open import LectureNotes.NumbersAndSets.Lecture1
open import LectureNotes.Groups.Lecture1 open import LectureNotes.Groups.Lecture1
open import ProjectEuler.Problem1
module Everything.WithK where module Everything.WithK where

View File

@@ -11,7 +11,7 @@ open import Numbers.Naturals.WithK
open import Sets.FinSet open import Sets.FinSet
open import Groups.Definition open import Groups.Definition
open import Groups.SymmetricGroups.Definition open import Groups.SymmetricGroups.Definition
open import DecidableSet open import Decidable.Sets
module Groups.FreeGroups where module Groups.FreeGroups where
@@ -107,7 +107,7 @@ badPrepend' {decA = decA} {x} (wordEnding pr x₁) | inr pr2 = pr2 refl
freeGroupGenerators : {a : _} (A : Set a) (decA : DecidableSet A) (w : FreeGroupGenerators A) SymmetryGroupElements (reflSetoid (ReducedWord decA)) freeGroupGenerators : {a : _} (A : Set a) (decA : DecidableSet A) (w : FreeGroupGenerators A) SymmetryGroupElements (reflSetoid (ReducedWord decA))
freeGroupGenerators A decA (χ x) = sym {f = f} bij freeGroupGenerators A decA (χ x) = sym {f = f} bij
where where
open DecidableSet.DecidableSet decA open DecidableSet decA
f : ReducedWord decA ReducedWord decA f : ReducedWord decA ReducedWord decA
f empty = prependLetter (ofLetter x) empty (wordEmpty refl) f empty = prependLetter (ofLetter x) empty (wordEmpty refl)
f (prependLetter (ofLetter startLetter) w pr) = prependLetter (ofLetter x) (prependLetter (ofLetter startLetter) w pr) (wordEnding (succIsPositive _) ans) f (prependLetter (ofLetter startLetter) w pr) = prependLetter (ofLetter x) (prependLetter (ofLetter startLetter) w pr) (wordEnding (succIsPositive _) ans)

View File

@@ -94,6 +94,10 @@ boolAnd : Bool → Bool → Bool
boolAnd BoolTrue y = y boolAnd BoolTrue y = y
boolAnd BoolFalse y = BoolFalse boolAnd BoolFalse y = BoolFalse
boolOr : Bool Bool Bool
boolOr BoolTrue y = BoolTrue
boolOr BoolFalse y = y
typeCast : {a : _} {A : Set a} {B : Set a} (el : A) (pr : A B) B typeCast : {a : _} {A : Set a} {B : Set a} (el : A) (pr : A B) B
typeCast {a} {A} {.A} elt refl = elt typeCast {a} {A} {.A} elt refl = elt

View File

@@ -12,102 +12,102 @@ open import Orders
module Numbers.BinaryNaturals.Definition where module Numbers.BinaryNaturals.Definition where
data Bit : Set where data Bit : Set where
zero : Bit zero : Bit
one : Bit one : Bit
BinNat : Set BinNat : Set
BinNat = List Bit BinNat = List Bit
::Inj : {xs ys : BinNat} {i : Bit} i :: xs i :: ys xs ys ::Inj : {xs ys : BinNat} {i : Bit} i :: xs i :: ys xs ys
::Inj {i = zero} refl = refl ::Inj {i = zero} refl = refl
::Inj {i = one} refl = refl ::Inj {i = one} refl = refl
nonEmptyNotEmpty : {a : _} {A : Set a} {l1 : List A} {i : A} i :: l1 [] False nonEmptyNotEmpty : {a : _} {A : Set a} {l1 : List A} {i : A} i :: l1 [] False
nonEmptyNotEmpty {l1 = l1} {i} () nonEmptyNotEmpty {l1 = l1} {i} ()
-- TODO - maybe we should do the floating-point style of assuming there's a leading bit and not storing it. -- TODO - maybe we should do the floating-point style of assuming there's a leading bit and not storing it.
-- That way, everything is already canonical. -- That way, everything is already canonical.
canonical : BinNat BinNat canonical : BinNat BinNat
canonical [] = [] canonical [] = []
canonical (zero :: n) with canonical n canonical (zero :: n) with canonical n
canonical (zero :: n) | [] = [] canonical (zero :: n) | [] = []
canonical (zero :: n) | x :: bl = zero :: x :: bl canonical (zero :: n) | x :: bl = zero :: x :: bl
canonical (one :: n) = one :: canonical n canonical (one :: n) = one :: canonical n
Canonicalised : Set Canonicalised : Set
Canonicalised = Sg BinNat (λ i canonical i i) Canonicalised = Sg BinNat (λ i canonical i i)
binNatToN : BinNat binNatToN : BinNat
binNatToN [] = 0 binNatToN [] = 0
binNatToN (zero :: b) = 2 *N binNatToN b binNatToN (zero :: b) = 2 *N binNatToN b
binNatToN (one :: b) = 1 +N (2 *N binNatToN b) binNatToN (one :: b) = 1 +N (2 *N binNatToN b)
incr : BinNat BinNat incr : BinNat BinNat
incr [] = one :: [] incr [] = one :: []
incr (zero :: n) = one :: n incr (zero :: n) = one :: n
incr (one :: n) = zero :: (incr n) incr (one :: n) = zero :: (incr n)
incrNonzero : (x : BinNat) canonical (incr x) [] False incrNonzero : (x : BinNat) canonical (incr x) [] False
incrPreservesCanonical : (x : BinNat) (canonical x x) canonical (incr x) incr x incrPreservesCanonical : (x : BinNat) (canonical x x) canonical (incr x) incr x
incrPreservesCanonical [] pr = refl incrPreservesCanonical [] pr = refl
incrPreservesCanonical (zero :: xs) pr with canonical xs incrPreservesCanonical (zero :: xs) pr with canonical xs
incrPreservesCanonical (zero :: xs) pr | x :: t = applyEquality (one ::_) (::Inj pr) incrPreservesCanonical (zero :: xs) pr | x :: t = applyEquality (one ::_) (::Inj pr)
incrPreservesCanonical (one :: xs) pr with inspect (canonical (incr xs)) incrPreservesCanonical (one :: xs) pr with inspect (canonical (incr xs))
incrPreservesCanonical (one :: xs) pr | [] with x = exFalso (incrNonzero xs x) incrPreservesCanonical (one :: xs) pr | [] with x = exFalso (incrNonzero xs x)
incrPreservesCanonical (one :: xs) pr | (x₁ :: y) with x rewrite x = applyEquality (zero ::_) (transitivity (equalityCommutative x) (incrPreservesCanonical xs (::Inj pr))) incrPreservesCanonical (one :: xs) pr | (x₁ :: y) with x rewrite x = applyEquality (zero ::_) (transitivity (equalityCommutative x) (incrPreservesCanonical xs (::Inj pr)))
incrPreservesCanonical' : (x : BinNat) canonical (incr x) incr (canonical x) incrPreservesCanonical' : (x : BinNat) canonical (incr x) incr (canonical x)
incrC : Canonicalised Canonicalised incrC : Canonicalised Canonicalised
incrC (a , b) = incr a , incrPreservesCanonical a b incrC (a , b) = incr a , incrPreservesCanonical a b
NToBinNat : BinNat NToBinNat : BinNat
NToBinNat zero = [] NToBinNat zero = []
NToBinNat (succ n) with NToBinNat n NToBinNat (succ n) with NToBinNat n
NToBinNat (succ n) | t = incr t NToBinNat (succ n) | t = incr t
NToBinNatC : Canonicalised NToBinNatC : Canonicalised
NToBinNatC zero = [] , refl NToBinNatC zero = [] , refl
NToBinNatC (succ n) = incrC (NToBinNatC n) NToBinNatC (succ n) = incrC (NToBinNatC n)
incrInj : {x y : BinNat} incr x incr y canonical x canonical y incrInj : {x y : BinNat} incr x incr y canonical x canonical y
incrNonzero' : (x : BinNat) (incr x) [] False incrNonzero' : (x : BinNat) (incr x) [] False
incrNonzero' (zero :: xs) () incrNonzero' (zero :: xs) ()
incrNonzero' (one :: xs) () incrNonzero' (one :: xs) ()
canonicalRespectsIncr' : {x y : BinNat} canonical (incr x) canonical (incr y) canonical x canonical y canonicalRespectsIncr' : {x y : BinNat} canonical (incr x) canonical (incr y) canonical x canonical y
binNatToNSucc : (n : BinNat) binNatToN (incr n) succ (binNatToN n) binNatToNSucc : (n : BinNat) binNatToN (incr n) succ (binNatToN n)
NToBinNatSucc : (n : ) incr (NToBinNat n) NToBinNat (succ n) NToBinNatSucc : (n : ) incr (NToBinNat n) NToBinNat (succ n)
binNatToNZero : (x : BinNat) binNatToN x 0 canonical x [] binNatToNZero : (x : BinNat) binNatToN x 0 canonical x []
binNatToNZero' : (x : BinNat) canonical x [] binNatToN x 0 binNatToNZero' : (x : BinNat) canonical x [] binNatToN x 0
NToBinNatZero : (n : ) NToBinNat n [] n 0 NToBinNatZero : (n : ) NToBinNat n [] n 0
NToBinNatZero zero pr = refl NToBinNatZero zero pr = refl
NToBinNatZero (succ n) pr with NToBinNat n NToBinNatZero (succ n) pr with NToBinNat n
NToBinNatZero (succ n) pr | zero :: bl = exFalso (nonEmptyNotEmpty pr) NToBinNatZero (succ n) pr | zero :: bl = exFalso (nonEmptyNotEmpty pr)
NToBinNatZero (succ n) pr | one :: bl = exFalso (nonEmptyNotEmpty pr) NToBinNatZero (succ n) pr | one :: bl = exFalso (nonEmptyNotEmpty pr)
canonicalAscends : {i : Bit} (a : BinNat) 0 <N binNatToN a i :: canonical a canonical (i :: a) canonicalAscends : {i : Bit} (a : BinNat) 0 <N binNatToN a i :: canonical a canonical (i :: a)
canonicalAscends' : {i : Bit} (a : BinNat) (canonical a [] False) i :: canonical a canonical (i :: a) canonicalAscends' : {i : Bit} (a : BinNat) (canonical a [] False) i :: canonical a canonical (i :: a)
canonicalAscends' {i} a pr = canonicalAscends {i} a (t a pr) canonicalAscends' {i} a pr = canonicalAscends {i} a (t a pr)
where where
t : (a : BinNat) (canonical a [] False) 0 <N binNatToN a t : (a : BinNat) (canonical a [] False) 0 <N binNatToN a
t a pr with TotalOrder.totality TotalOrder 0 (binNatToN a) t a pr with TotalOrder.totality TotalOrder 0 (binNatToN a)
t a pr | inl (inl x) = x t a pr | inl (inl x) = x
t a pr | inr x = exFalso (pr (binNatToNZero a (equalityCommutative x))) t a pr | inr x = exFalso (pr (binNatToNZero a (equalityCommutative x)))
canonicalIdempotent : (a : BinNat) canonical a canonical (canonical a) canonicalIdempotent : (a : BinNat) canonical a canonical (canonical a)
canonicalIdempotent [] = refl canonicalIdempotent [] = refl
canonicalIdempotent (zero :: a) with inspect (canonical a) canonicalIdempotent (zero :: a) with inspect (canonical a)
canonicalIdempotent (zero :: a) | [] with y rewrite y = refl canonicalIdempotent (zero :: a) | [] with y rewrite y = refl
canonicalIdempotent (zero :: a) | (x :: bl) with y = transitivity (equalityCommutative (canonicalAscends' {zero} a λ p contr p y)) (transitivity (applyEquality (zero ::_) (canonicalIdempotent a)) (equalityCommutative v)) canonicalIdempotent (zero :: a) | (x :: bl) with y = transitivity (equalityCommutative (canonicalAscends' {zero} a λ p contr p y)) (transitivity (applyEquality (zero ::_) (canonicalIdempotent a)) (equalityCommutative v))
where where
contr : {a : _} {A : Set a} {l1 l2 : List A} {x : A} l1 [] l1 x :: l2 False contr : {a : _} {A : Set a} {l1 l2 : List A} {x : A} l1 [] l1 x :: l2 False
contr {l1 = []} p1 () contr {l1 = []} p1 ()
@@ -116,204 +116,204 @@ module Numbers.BinaryNaturals.Definition where
u = applyEquality canonical (equalityCommutative (canonicalAscends' {zero} a λ p contr p y)) u = applyEquality canonical (equalityCommutative (canonicalAscends' {zero} a λ p contr p y))
v : canonical (canonical (zero :: a)) zero :: canonical (canonical a) v : canonical (canonical (zero :: a)) zero :: canonical (canonical a)
v = transitivity u (equalityCommutative (canonicalAscends' {zero} (canonical a) λ p contr (transitivity (canonicalIdempotent a) p) y)) v = transitivity u (equalityCommutative (canonicalAscends' {zero} (canonical a) λ p contr (transitivity (canonicalIdempotent a) p) y))
canonicalIdempotent (one :: a) rewrite equalityCommutative (canonicalIdempotent a) = refl canonicalIdempotent (one :: a) rewrite equalityCommutative (canonicalIdempotent a) = refl
canonicalAscends'' : {i : Bit} (a : BinNat) canonical (i :: canonical a) canonical (i :: a) canonicalAscends'' : {i : Bit} (a : BinNat) canonical (i :: canonical a) canonical (i :: a)
binNatToNInj : (x y : BinNat) binNatToN x binNatToN y canonical x canonical y binNatToNInj : (x y : BinNat) binNatToN x binNatToN y canonical x canonical y
NToBinNatInj : (x y : ) canonical (NToBinNat x) canonical (NToBinNat y) x y NToBinNatInj : (x y : ) canonical (NToBinNat x) canonical (NToBinNat y) x y
NToBinNatIsCanonical : (x : ) NToBinNat x canonical (NToBinNat x) NToBinNatIsCanonical : (x : ) NToBinNat x canonical (NToBinNat x)
NToBinNatIsCanonical zero = refl NToBinNatIsCanonical zero = refl
NToBinNatIsCanonical (succ x) with NToBinNatC x NToBinNatIsCanonical (succ x) with NToBinNatC x
NToBinNatIsCanonical (succ x) | a , b = equalityCommutative (incrPreservesCanonical (NToBinNat x) (equalityCommutative (NToBinNatIsCanonical x))) NToBinNatIsCanonical (succ x) | a , b = equalityCommutative (incrPreservesCanonical (NToBinNat x) (equalityCommutative (NToBinNatIsCanonical x)))
contr' : {a : _} {A : Set a} {l1 l2 : List A} {x : A} l1 [] l1 x :: l2 False contr' : {a : _} {A : Set a} {l1 l2 : List A} {x : A} l1 [] l1 x :: l2 False
contr' {l1 = []} p1 () contr' {l1 = []} p1 ()
contr' {l1 = x :: l1} () p2 contr' {l1 = x :: l1} () p2
binNatToNIsCanonical : (x : BinNat) binNatToN (canonical x) binNatToN x binNatToNIsCanonical : (x : BinNat) binNatToN (canonical x) binNatToN x
binNatToNIsCanonical [] = refl binNatToNIsCanonical [] = refl
binNatToNIsCanonical (zero :: x) with inspect (canonical x) binNatToNIsCanonical (zero :: x) with inspect (canonical x)
binNatToNIsCanonical (zero :: x) | [] with t rewrite t | binNatToNZero' x t = refl binNatToNIsCanonical (zero :: x) | [] with t rewrite t | binNatToNZero' x t = refl
binNatToNIsCanonical (zero :: x) | (x₁ :: bl) with t rewrite (equalityCommutative (canonicalAscends' {zero} x λ p contr' p t)) | binNatToNIsCanonical x = refl binNatToNIsCanonical (zero :: x) | (x₁ :: bl) with t rewrite (equalityCommutative (canonicalAscends' {zero} x λ p contr' p t)) | binNatToNIsCanonical x = refl
binNatToNIsCanonical (one :: x) rewrite binNatToNIsCanonical x = refl binNatToNIsCanonical (one :: x) rewrite binNatToNIsCanonical x = refl
-- The following two theorems demonstrate that Canonicalised is isomorphic to -- The following two theorems demonstrate that Canonicalised is isomorphic to
nToN : (x : ) binNatToN (NToBinNat x) x nToN : (x : ) binNatToN (NToBinNat x) x
binToBin : (x : BinNat) NToBinNat (binNatToN x) canonical x binToBin : (x : BinNat) NToBinNat (binNatToN x) canonical x
binToBin x = transitivity (NToBinNatIsCanonical (binNatToN x)) (binNatToNInj (NToBinNat (binNatToN x)) x (nToN (binNatToN x))) binToBin x = transitivity (NToBinNatIsCanonical (binNatToN x)) (binNatToNInj (NToBinNat (binNatToN x)) x (nToN (binNatToN x)))
doubleIsBitShift' : (a : ) NToBinNat (2 *N succ a) zero :: NToBinNat (succ a) doubleIsBitShift' : (a : ) NToBinNat (2 *N succ a) zero :: NToBinNat (succ a)
doubleIsBitShift' zero = refl doubleIsBitShift' zero = refl
doubleIsBitShift' (succ a) with doubleIsBitShift' a doubleIsBitShift' (succ a) with doubleIsBitShift' a
... | 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 ... | 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 : (a : ) (0 <N a) NToBinNat (2 *N a) zero :: NToBinNat a
doubleIsBitShift zero () doubleIsBitShift zero ()
doubleIsBitShift (succ a) _ = doubleIsBitShift' a doubleIsBitShift (succ a) _ = doubleIsBitShift' a
canonicalDescends : {a : Bit} (as : BinNat) (prA : a :: as canonical (a :: as)) as canonical as canonicalDescends : {a : Bit} (as : BinNat) (prA : a :: as canonical (a :: as)) as canonical as
canonicalDescends {zero} as pr with canonical as canonicalDescends {zero} as pr with canonical as
canonicalDescends {zero} as pr | x :: bl = ::Inj pr canonicalDescends {zero} as pr | x :: bl = ::Inj pr
canonicalDescends {one} as pr = ::Inj pr canonicalDescends {one} as pr = ::Inj pr
--- Proofs --- Proofs
parity : (a b : ) succ (2 *N a) 2 *N b False parity : (a b : ) succ (2 *N a) 2 *N b False
doubleInj : (a b : ) (2 *N a) (2 *N b) a b doubleInj : (a b : ) (2 *N a) (2 *N b) a b
incrNonzeroTwice : (x : BinNat) canonical (incr (incr x)) one :: [] False incrNonzeroTwice : (x : BinNat) canonical (incr (incr x)) one :: [] False
incrNonzeroTwice (zero :: xs) pr with canonical (incr xs) incrNonzeroTwice (zero :: xs) pr with canonical (incr xs)
incrNonzeroTwice (zero :: xs) () | [] incrNonzeroTwice (zero :: xs) () | []
incrNonzeroTwice (zero :: xs) () | x :: bl incrNonzeroTwice (zero :: xs) () | x :: bl
incrNonzeroTwice (one :: xs) pr = exFalso (incrNonzero xs (::Inj pr)) incrNonzeroTwice (one :: xs) pr = exFalso (incrNonzero xs (::Inj pr))
canonicalRespectsIncr' {[]} {[]} pr = refl canonicalRespectsIncr' {[]} {[]} pr = refl
canonicalRespectsIncr' {[]} {zero :: y} pr with canonical y canonicalRespectsIncr' {[]} {zero :: y} pr with canonical y
canonicalRespectsIncr' {[]} {zero :: y} pr | [] = refl canonicalRespectsIncr' {[]} {zero :: y} pr | [] = refl
canonicalRespectsIncr' {[]} {one :: y} pr with canonical (incr y) canonicalRespectsIncr' {[]} {one :: y} pr with canonical (incr y)
canonicalRespectsIncr' {[]} {one :: y} () | [] canonicalRespectsIncr' {[]} {one :: y} () | []
canonicalRespectsIncr' {[]} {one :: y} () | x :: bl canonicalRespectsIncr' {[]} {one :: y} () | x :: bl
canonicalRespectsIncr' {zero :: xs} {y} pr with canonical xs canonicalRespectsIncr' {zero :: xs} {y} pr with canonical xs
canonicalRespectsIncr' {zero :: xs} {[]} pr | [] = refl canonicalRespectsIncr' {zero :: xs} {[]} pr | [] = refl
canonicalRespectsIncr' {zero :: xs} {zero :: y} pr | [] with canonical y canonicalRespectsIncr' {zero :: xs} {zero :: y} pr | [] with canonical y
canonicalRespectsIncr' {zero :: xs} {zero :: y} pr | [] | [] = refl canonicalRespectsIncr' {zero :: xs} {zero :: y} pr | [] | [] = refl
canonicalRespectsIncr' {zero :: xs} {one :: y} pr | [] with canonical (incr y) canonicalRespectsIncr' {zero :: xs} {one :: y} pr | [] with canonical (incr y)
canonicalRespectsIncr' {zero :: xs} {one :: y} () | [] | [] canonicalRespectsIncr' {zero :: xs} {one :: y} () | [] | []
canonicalRespectsIncr' {zero :: xs} {one :: y} () | [] | x :: bl canonicalRespectsIncr' {zero :: xs} {one :: y} () | [] | x :: bl
canonicalRespectsIncr' {zero :: xs} {zero :: ys} pr | x :: bl with canonical ys canonicalRespectsIncr' {zero :: xs} {zero :: ys} pr | x :: bl with canonical ys
canonicalRespectsIncr' {zero :: xs} {zero :: ys} pr | x :: bl | x₁ :: th = applyEquality (zero ::_) (::Inj pr) canonicalRespectsIncr' {zero :: xs} {zero :: ys} pr | x :: bl | x₁ :: th = applyEquality (zero ::_) (::Inj pr)
canonicalRespectsIncr' {zero :: xs} {one :: ys} pr | x :: bl with canonical (incr ys) canonicalRespectsIncr' {zero :: xs} {one :: ys} pr | x :: bl with canonical (incr ys)
canonicalRespectsIncr' {zero :: xs} {one :: ys} () | x :: bl | [] canonicalRespectsIncr' {zero :: xs} {one :: ys} () | x :: bl | []
canonicalRespectsIncr' {zero :: xs} {one :: ys} () | x :: bl | x₁ :: th canonicalRespectsIncr' {zero :: xs} {one :: ys} () | x :: bl | x₁ :: th
canonicalRespectsIncr' {one :: xs} {[]} pr with canonical (incr xs) canonicalRespectsIncr' {one :: xs} {[]} pr with canonical (incr xs)
canonicalRespectsIncr' {one :: xs} {[]} () | [] canonicalRespectsIncr' {one :: xs} {[]} () | []
canonicalRespectsIncr' {one :: xs} {[]} () | x :: bl canonicalRespectsIncr' {one :: xs} {[]} () | x :: bl
canonicalRespectsIncr' {one :: xs} {zero :: ys} pr with canonical ys canonicalRespectsIncr' {one :: xs} {zero :: ys} pr with canonical ys
canonicalRespectsIncr' {one :: xs} {zero :: ys} pr | [] with canonical (incr xs) canonicalRespectsIncr' {one :: xs} {zero :: ys} pr | [] with canonical (incr xs)
canonicalRespectsIncr' {one :: xs} {zero :: ys} () | [] | [] canonicalRespectsIncr' {one :: xs} {zero :: ys} () | [] | []
canonicalRespectsIncr' {one :: xs} {zero :: ys} () | [] | x :: t canonicalRespectsIncr' {one :: xs} {zero :: ys} () | [] | x :: t
canonicalRespectsIncr' {one :: xs} {zero :: ys} pr | x :: bl with canonical (incr xs) canonicalRespectsIncr' {one :: xs} {zero :: ys} pr | x :: bl with canonical (incr xs)
canonicalRespectsIncr' {one :: xs} {zero :: ys} () | x :: bl | [] canonicalRespectsIncr' {one :: xs} {zero :: ys} () | x :: bl | []
canonicalRespectsIncr' {one :: xs} {zero :: ys} () | x :: bl | x₁ :: t canonicalRespectsIncr' {one :: xs} {zero :: ys} () | x :: bl | x₁ :: t
canonicalRespectsIncr' {one :: xs} {one :: ys} pr with inspect (canonical (incr xs)) canonicalRespectsIncr' {one :: xs} {one :: ys} pr with inspect (canonical (incr xs))
canonicalRespectsIncr' {one :: xs} {one :: ys} pr | [] with x = exFalso (incrNonzero xs x) canonicalRespectsIncr' {one :: xs} {one :: ys} pr | [] with x = exFalso (incrNonzero xs x)
canonicalRespectsIncr' {one :: xs} {one :: ys} pr | (x+1 :: x+1s) with bad rewrite bad = applyEquality (one ::_) ans canonicalRespectsIncr' {one :: xs} {one :: ys} pr | (x+1 :: x+1s) with bad rewrite bad = applyEquality (one ::_) ans
where where
ans : canonical xs canonical ys ans : canonical xs canonical ys
ans with inspect (canonical (incr ys)) ans with inspect (canonical (incr ys))
ans | [] with x = exFalso (incrNonzero ys x) ans | [] with x = exFalso (incrNonzero ys x)
ans | (x₁ :: y) with x rewrite x = canonicalRespectsIncr' {xs} {ys} (transitivity bad (transitivity (::Inj pr) (equalityCommutative x))) ans | (x₁ :: y) with x rewrite x = canonicalRespectsIncr' {xs} {ys} (transitivity bad (transitivity (::Inj pr) (equalityCommutative x)))
binNatToNInj[] : (y : BinNat) 0 binNatToN y [] canonical y binNatToNInj[] : (y : BinNat) 0 binNatToN y [] canonical y
binNatToNInj[] [] pr = refl binNatToNInj[] [] pr = refl
binNatToNInj[] (zero :: y) pr with productZeroImpliesOperandZero {2} {binNatToN y} (equalityCommutative pr) binNatToNInj[] (zero :: y) pr with productZeroImpliesOperandZero {2} {binNatToN y} (equalityCommutative pr)
binNatToNInj[] (zero :: y) pr | inr x with inspect (canonical y) binNatToNInj[] (zero :: y) pr | inr x with inspect (canonical y)
binNatToNInj[] (zero :: y) pr | inr x | [] with canYPr rewrite canYPr = refl binNatToNInj[] (zero :: y) pr | inr x | [] with canYPr rewrite canYPr = refl
binNatToNInj[] (zero :: ys) pr | inr x | (y :: canY) with canYPr with binNatToNZero ys x binNatToNInj[] (zero :: ys) pr | inr x | (y :: canY) with canYPr with binNatToNZero ys x
... | r with canonical ys ... | r with canonical ys
binNatToNInj[] (zero :: ys) pr | inr x | (y :: canY) with canYPr | r | [] = refl binNatToNInj[] (zero :: ys) pr | inr x | (y :: canY) with canYPr | r | [] = refl
binNatToNInj[] (zero :: ys) pr | inr x | (y :: canY) with canYPr | () | x₁ :: t binNatToNInj[] (zero :: ys) pr | inr x | (y :: canY) with canYPr | () | x₁ :: t
binNatToNInj [] y pr = binNatToNInj[] y pr binNatToNInj [] y pr = binNatToNInj[] y pr
binNatToNInj (zero :: xs) [] pr = equalityCommutative (binNatToNInj[] (zero :: xs) (equalityCommutative pr)) binNatToNInj (zero :: xs) [] pr = equalityCommutative (binNatToNInj[] (zero :: xs) (equalityCommutative pr))
binNatToNInj (zero :: xs) (zero :: ys) pr with doubleInj (binNatToN xs) (binNatToN ys) pr binNatToNInj (zero :: xs) (zero :: ys) pr with doubleInj (binNatToN xs) (binNatToN ys) pr
... | x=y with binNatToNInj xs ys x=y ... | x=y with binNatToNInj xs ys x=y
... | t with canonical xs ... | t with canonical xs
binNatToNInj (zero :: xs) (zero :: ys) pr | x=y | t | [] with canonical ys binNatToNInj (zero :: xs) (zero :: ys) pr | x=y | t | [] with canonical ys
binNatToNInj (zero :: xs) (zero :: ys) pr | x=y | t | [] | [] = refl binNatToNInj (zero :: xs) (zero :: ys) pr | x=y | t | [] | [] = refl
binNatToNInj (zero :: xs) (zero :: ys) pr | x=y | t | x :: cxs with canonical ys binNatToNInj (zero :: xs) (zero :: ys) pr | x=y | t | x :: cxs with canonical ys
binNatToNInj (zero :: xs) (zero :: ys) pr | x=y | t | x :: cxs | x₁ :: cys = applyEquality (zero ::_) t binNatToNInj (zero :: xs) (zero :: ys) pr | x=y | t | x :: cxs | x₁ :: cys = applyEquality (zero ::_) t
binNatToNInj (zero :: xs) (one :: ys) pr = exFalso (parity (binNatToN ys) (binNatToN xs) (equalityCommutative pr)) binNatToNInj (zero :: xs) (one :: ys) pr = exFalso (parity (binNatToN ys) (binNatToN xs) (equalityCommutative pr))
binNatToNInj (one :: xs) (zero :: ys) pr = exFalso (parity (binNatToN xs) (binNatToN ys) pr) binNatToNInj (one :: xs) (zero :: ys) pr = exFalso (parity (binNatToN xs) (binNatToN ys) pr)
binNatToNInj (one :: xs) (one :: ys) pr = applyEquality (one ::_) (binNatToNInj xs ys (doubleInj (binNatToN xs) (binNatToN ys) (succInjective pr))) binNatToNInj (one :: xs) (one :: ys) pr = applyEquality (one ::_) (binNatToNInj xs ys (doubleInj (binNatToN xs) (binNatToN ys) (succInjective pr)))
NToBinNatInj zero zero pr = refl NToBinNatInj zero zero pr = refl
NToBinNatInj zero (succ y) pr with NToBinNat y NToBinNatInj zero (succ y) pr with NToBinNat y
NToBinNatInj zero (succ y) pr | bl = exFalso (incrNonzero bl (equalityCommutative pr)) NToBinNatInj zero (succ y) pr | bl = exFalso (incrNonzero bl (equalityCommutative pr))
NToBinNatInj (succ x) zero pr with NToBinNat x NToBinNatInj (succ x) zero pr with NToBinNat x
... | bl = exFalso (incrNonzero bl pr) ... | bl = exFalso (incrNonzero bl pr)
NToBinNatInj (succ x) (succ y) pr with inspect (NToBinNat x) NToBinNatInj (succ x) (succ y) pr with inspect (NToBinNat x)
NToBinNatInj (succ zero) (succ zero) pr | [] with nToBinXPr = refl NToBinNatInj (succ zero) (succ zero) pr | [] with nToBinXPr = refl
NToBinNatInj (succ zero) (succ (succ y)) pr | [] with nToBinXPr with NToBinNat y NToBinNatInj (succ zero) (succ (succ y)) pr | [] with nToBinXPr with NToBinNat y
NToBinNatInj (succ zero) (succ (succ y)) pr | [] with nToBinXPr | y' = exFalso (incrNonzeroTwice y' (equalityCommutative pr)) NToBinNatInj (succ zero) (succ (succ y)) pr | [] with nToBinXPr | y' = exFalso (incrNonzeroTwice y' (equalityCommutative pr))
NToBinNatInj (succ (succ x)) (succ y) pr | [] with nToBinXPr with NToBinNat x NToBinNatInj (succ (succ x)) (succ y) pr | [] with nToBinXPr with NToBinNat x
NToBinNatInj (succ (succ x)) (succ y) pr | [] with nToBinXPr | t = exFalso (incrNonzero' t nToBinXPr) NToBinNatInj (succ (succ x)) (succ y) pr | [] with nToBinXPr | t = exFalso (incrNonzero' t nToBinXPr)
NToBinNatInj (succ x) (succ y) pr | (x₁ :: nToBinX) with nToBinXPr with NToBinNatInj x y (canonicalRespectsIncr' {NToBinNat x} {NToBinNat y} pr) NToBinNatInj (succ x) (succ y) pr | (x₁ :: nToBinX) with nToBinXPr with NToBinNatInj x y (canonicalRespectsIncr' {NToBinNat x} {NToBinNat y} pr)
NToBinNatInj (succ x) (succ .x) pr | (x₁ :: nToBinX) with nToBinXPr | refl = refl NToBinNatInj (succ x) (succ .x) pr | (x₁ :: nToBinX) with nToBinXPr | refl = refl
incrInj {[]} {[]} pr = refl incrInj {[]} {[]} pr = refl
incrInj {[]} {zero :: ys} pr rewrite equalityCommutative (::Inj pr) = refl incrInj {[]} {zero :: ys} pr rewrite equalityCommutative (::Inj pr) = refl
incrInj {zero :: xs} {[]} pr rewrite ::Inj pr = refl incrInj {zero :: xs} {[]} pr rewrite ::Inj pr = refl
incrInj {zero :: xs} {zero :: .xs} refl = refl incrInj {zero :: xs} {zero :: .xs} refl = refl
incrInj {one :: xs} {one :: ys} pr = applyEquality (one ::_) (incrInj {xs} {ys} (l (incr xs) (incr ys) pr)) incrInj {one :: xs} {one :: ys} pr = applyEquality (one ::_) (incrInj {xs} {ys} (l (incr xs) (incr ys) pr))
where where
l : (a : BinNat) (b : BinNat) zero :: a zero :: b a b l : (a : BinNat) (b : BinNat) zero :: a zero :: b a b
l a .a refl = refl l a .a refl = refl
doubleInj zero zero pr = refl doubleInj zero zero pr = refl
doubleInj (succ a) (succ b) pr = applyEquality succ (doubleInj a b u) doubleInj (succ a) (succ b) pr = applyEquality succ (doubleInj a b u)
where where
t : a +N a b +N b t : a +N a b +N b
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) 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 : a +N (a +N zero) b +N (b +N zero)
u rewrite Semiring.commutative Semiring a 0 | Semiring.commutative Semiring b 0 = t u rewrite Semiring.commutative Semiring a 0 | Semiring.commutative Semiring b 0 = t
binNatToNZero [] pr = refl binNatToNZero [] pr = refl
binNatToNZero (zero :: xs) pr with inspect (canonical xs) binNatToNZero (zero :: xs) pr with inspect (canonical xs)
binNatToNZero (zero :: xs) pr | [] with x rewrite x = refl binNatToNZero (zero :: xs) pr | [] with x rewrite x = refl
binNatToNZero (zero :: xs) pr | (x₁ :: y) with x with productZeroImpliesOperandZero {2} {binNatToN xs} pr binNatToNZero (zero :: xs) pr | (x₁ :: y) with x with productZeroImpliesOperandZero {2} {binNatToN xs} pr
binNatToNZero (zero :: xs) pr | (x₁ :: y) with x | inr pr' with binNatToNZero xs pr' binNatToNZero (zero :: xs) pr | (x₁ :: y) with x | inr pr' with binNatToNZero xs pr'
... | bl with canonical xs ... | bl with canonical xs
binNatToNZero (zero :: xs) pr | (x₁ :: y) with x | inr pr' | bl | [] = refl binNatToNZero (zero :: xs) pr | (x₁ :: y) with x | inr pr' | bl | [] = refl
binNatToNZero (zero :: xs) pr | (x₁ :: y) with x | inr pr' | () | x₂ :: t binNatToNZero (zero :: xs) pr | (x₁ :: y) with x | inr pr' | () | x₂ :: t
binNatToNSucc [] = refl binNatToNSucc [] = refl
binNatToNSucc (zero :: n) = refl binNatToNSucc (zero :: n) = refl
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))) 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 inspect (canonical (incr xs))
incrNonzero (one :: xs) pr | [] with x = incrNonzero xs x incrNonzero (one :: xs) pr | [] with x = incrNonzero xs x
incrNonzero (one :: xs) pr | (y :: ys) with x with canonical (incr xs) incrNonzero (one :: xs) pr | (y :: ys) with x with canonical (incr xs)
incrNonzero (one :: xs) pr | (y :: ys) with () | [] incrNonzero (one :: xs) pr | (y :: ys) with () | []
incrNonzero (one :: xs) () | (.x₁ :: .th) with refl | x₁ :: th incrNonzero (one :: xs) () | (.x₁ :: .th) with refl | x₁ :: th
nToN zero = refl nToN zero = refl
nToN (succ x) with inspect (NToBinNat x) nToN (succ x) with inspect (NToBinNat x)
nToN (succ x) | [] with pr = ans nToN (succ x) | [] with pr = ans
where where
t : x zero t : x zero
t = NToBinNatInj x 0 (applyEquality canonical pr) t = NToBinNatInj x 0 (applyEquality canonical pr)
ans : binNatToN (incr (NToBinNat x)) succ x ans : binNatToN (incr (NToBinNat x)) succ x
ans rewrite t | pr = refl ans rewrite t | pr = refl
nToN (succ x) | (bit :: xs) with pr = transitivity (binNatToNSucc (NToBinNat x)) (applyEquality succ (nToN x)) nToN (succ x) | (bit :: xs) with pr = transitivity (binNatToNSucc (NToBinNat x)) (applyEquality succ (nToN x))
parity zero (succ b) pr rewrite Semiring.commutative Semiring b (succ (b +N 0)) = bad pr parity zero (succ b) pr rewrite Semiring.commutative Semiring b (succ (b +N 0)) = bad pr
where where
bad : (1 succ (succ ((b +N 0) +N b))) False bad : (1 succ (succ ((b +N 0) +N b))) False
bad () bad ()
parity (succ a) (succ b) pr rewrite Semiring.commutative Semiring b (succ (b +N 0)) | Semiring.commutative Semiring a (succ (a +N 0)) | Semiring.commutative Semiring (a +N 0) a | Semiring.commutative Semiring (b +N 0) b = parity a b (succInjective (succInjective pr)) 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' [] pr = refl
binNatToNZero' (zero :: xs) pr with inspect (canonical xs) binNatToNZero' (zero :: xs) pr with inspect (canonical xs)
binNatToNZero' (zero :: xs) pr | [] with p2 = ans binNatToNZero' (zero :: xs) pr | [] with p2 = ans
where where
t : binNatToN xs 0 t : binNatToN xs 0
t = binNatToNZero' xs p2 t = binNatToNZero' xs p2
ans : 2 *N binNatToN xs 0 ans : 2 *N binNatToN xs 0
ans rewrite t = refl ans rewrite t = refl
binNatToNZero' (zero :: xs) pr | (y :: ys) with p rewrite p = exFalso (bad pr) binNatToNZero' (zero :: xs) pr | (y :: ys) with p rewrite p = exFalso (bad pr)
where where
bad : zero :: y :: ys [] False bad : zero :: y :: ys [] False
bad () bad ()
binNatToNZero' (one :: xs) () binNatToNZero' (one :: xs) ()
canonicalAscends {zero} a 0<a with inspect (canonical a) canonicalAscends {zero} a 0<a with inspect (canonical a)
canonicalAscends {zero} (zero :: a) 0<a | [] with x = exFalso (contr'' (binNatToN a) t v) canonicalAscends {zero} (zero :: a) 0<a | [] with x = exFalso (contr'' (binNatToN a) t v)
where where
u : binNatToN (zero :: a) 0 u : binNatToN (zero :: a) 0
u = binNatToNZero' (zero :: a) x u = binNatToNZero' (zero :: a) x
@@ -328,27 +328,27 @@ module Numbers.BinaryNaturals.Definition where
t | succ bl rewrite Semiring.commutative Semiring (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 <N x) (x 0) False
contr'' x 0<x x=0 rewrite x=0 = TotalOrder.irreflexive TotalOrder 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 {zero} a 0<a | (x₁ :: y) with x rewrite x = refl
canonicalAscends {one} a 0<a = refl canonicalAscends {one} a 0<a = refl
canonicalAscends'' {i} a with inspect (canonical a) canonicalAscends'' {i} a with inspect (canonical a)
canonicalAscends'' {zero} a | [] with x rewrite x = refl canonicalAscends'' {zero} a | [] with x rewrite x = refl
canonicalAscends'' {one} a | [] with x rewrite x = refl canonicalAscends'' {one} a | [] with x rewrite x = refl
canonicalAscends'' {i} a | (x₁ :: y) with x = transitivity (applyEquality canonical (canonicalAscends' {i} a λ p contr p x)) (equalityCommutative (canonicalIdempotent (i :: a))) canonicalAscends'' {i} a | (x₁ :: y) with x = transitivity (applyEquality canonical (canonicalAscends' {i} a λ p contr p x)) (equalityCommutative (canonicalIdempotent (i :: a)))
where where
contr : {a : _} {A : Set a} {l1 l2 : List A} {x : A} l1 [] l1 x :: l2 False contr : {a : _} {A : Set a} {l1 l2 : List A} {x : A} l1 [] l1 x :: l2 False
contr {l1 = []} p1 () contr {l1 = []} p1 ()
contr {l1 = x :: l1} () p2 contr {l1 = x :: l1} () p2
incrPreservesCanonical' [] = refl incrPreservesCanonical' [] = refl
incrPreservesCanonical' (zero :: xs) with inspect (canonical xs) incrPreservesCanonical' (zero :: xs) with inspect (canonical xs)
incrPreservesCanonical' (zero :: xs) | [] with x rewrite x = refl incrPreservesCanonical' (zero :: xs) | [] with x rewrite x = refl
incrPreservesCanonical' (zero :: xs) | (x₁ :: y) with x rewrite x = refl incrPreservesCanonical' (zero :: xs) | (x₁ :: y) with x rewrite x = refl
incrPreservesCanonical' (one :: xs) with inspect (canonical (incr xs)) incrPreservesCanonical' (one :: xs) with inspect (canonical (incr xs))
... | [] with pr = exFalso (incrNonzero xs pr) ... | [] with pr = exFalso (incrNonzero xs pr)
... | (_ :: _) with pr rewrite pr = applyEquality (zero ::_) (transitivity (equalityCommutative pr) (incrPreservesCanonical' xs)) ... | (_ :: _) with pr rewrite pr = applyEquality (zero ::_) (transitivity (equalityCommutative pr) (incrPreservesCanonical' xs))
NToBinNatSucc zero = refl NToBinNatSucc zero = refl
NToBinNatSucc (succ n) with NToBinNat n NToBinNatSucc (succ n) with NToBinNat n
... | [] = refl ... | [] = refl
... | a :: as = refl ... | a :: as = refl

View File

@@ -12,68 +12,68 @@ open import Semirings.Definition
module Numbers.BinaryNaturals.Multiplication where module Numbers.BinaryNaturals.Multiplication where
_*Binherit_ : BinNat BinNat BinNat _*Binherit_ : BinNat BinNat BinNat
a *Binherit b = NToBinNat (binNatToN a *N binNatToN b) a *Binherit b = NToBinNat (binNatToN a *N binNatToN b)
_*B_ : BinNat BinNat BinNat _*B_ : BinNat BinNat BinNat
[] *B b = [] [] *B b = []
(zero :: a) *B b = zero :: (a *B b) (zero :: a) *B b = zero :: (a *B b)
(one :: a) *B b = (zero :: (a *B b)) +B b (one :: a) *B b = (zero :: (a *B b)) +B b
contr : {a : _} {A : Set a} {l1 l2 : List A} {x : A} l1 [] l1 x :: l2 False contr : {a : _} {A : Set a} {l1 l2 : List A} {x : A} l1 [] l1 x :: l2 False
contr {l1 = []} p1 () contr {l1 = []} p1 ()
contr {l1 = x :: l1} () p2 contr {l1 = x :: l1} () p2
*BEmpty : (a : BinNat) canonical (a *B []) [] *BEmpty : (a : BinNat) canonical (a *B []) []
*BEmpty [] = refl *BEmpty [] = refl
*BEmpty (zero :: a) rewrite *BEmpty a = refl *BEmpty (zero :: a) rewrite *BEmpty a = refl
*BEmpty (one :: a) rewrite *BEmpty a = refl *BEmpty (one :: a) rewrite *BEmpty a = refl
canonicalDistributesPlus : (a b : BinNat) canonical (a +B b) canonical a +B canonical b canonicalDistributesPlus : (a b : BinNat) canonical (a +B b) canonical a +B canonical b
canonicalDistributesPlus a b = transitivity ans (+BIsInherited (canonical a) (canonical b) (canonicalIdempotent a) (canonicalIdempotent b)) canonicalDistributesPlus a b = transitivity ans (+BIsInherited (canonical a) (canonical b) (canonicalIdempotent a) (canonicalIdempotent b))
where where
ans : canonical (a +B b) NToBinNat (binNatToN (canonical a) +N binNatToN (canonical b)) ans : canonical (a +B b) NToBinNat (binNatToN (canonical a) +N binNatToN (canonical b))
ans rewrite binNatToNIsCanonical a | binNatToNIsCanonical b = equalityCommutative (+BIsInherited' a b) ans rewrite binNatToNIsCanonical a | binNatToNIsCanonical b = equalityCommutative (+BIsInherited' a b)
incrPullsOut : (a b : BinNat) incr (a +B b) (incr a) +B b incrPullsOut : (a b : BinNat) incr (a +B b) (incr a) +B b
incrPullsOut [] [] = refl incrPullsOut [] [] = refl
incrPullsOut [] (zero :: b) = refl incrPullsOut [] (zero :: b) = refl
incrPullsOut [] (one :: b) = refl incrPullsOut [] (one :: b) = refl
incrPullsOut (zero :: a) [] = refl incrPullsOut (zero :: a) [] = refl
incrPullsOut (zero :: a) (zero :: b) = refl incrPullsOut (zero :: a) (zero :: b) = refl
incrPullsOut (zero :: a) (one :: b) = refl incrPullsOut (zero :: a) (one :: b) = refl
incrPullsOut (one :: a) [] = refl incrPullsOut (one :: a) [] = refl
incrPullsOut (one :: a) (zero :: b) = applyEquality (zero ::_) (incrPullsOut a b) incrPullsOut (one :: a) (zero :: b) = applyEquality (zero ::_) (incrPullsOut a b)
incrPullsOut (one :: a) (one :: b) = applyEquality (one ::_) (incrPullsOut a b) incrPullsOut (one :: a) (one :: b) = applyEquality (one ::_) (incrPullsOut a b)
timesZero : (a b : BinNat) canonical a [] canonical (a *B b) [] timesZero : (a b : BinNat) canonical a [] canonical (a *B b) []
timesZero [] b pr = refl timesZero [] b pr = refl
timesZero (zero :: a) b pr with inspect (canonical a) timesZero (zero :: a) b pr with inspect (canonical a)
timesZero (zero :: a) b pr | [] with prA rewrite prA | timesZero a b prA = refl timesZero (zero :: a) b pr | [] with prA rewrite prA | timesZero a b prA = refl
timesZero (zero :: a) b pr | (a1 :: as) with prA rewrite prA = exFalso (nonEmptyNotEmpty pr) timesZero (zero :: a) b pr | (a1 :: as) with prA rewrite prA = exFalso (nonEmptyNotEmpty pr)
timesZero'' : (a b : BinNat) canonical (a *B b) [] (canonical a []) || (canonical b []) timesZero'' : (a b : BinNat) canonical (a *B b) [] (canonical a []) || (canonical b [])
timesZero'' [] b pr = inl refl timesZero'' [] b pr = inl refl
timesZero'' (x :: a) [] pr = inr refl timesZero'' (x :: a) [] pr = inr refl
timesZero'' (zero :: as) (zero :: bs) pr with inspect (canonical as) timesZero'' (zero :: as) (zero :: bs) pr with inspect (canonical as)
timesZero'' (zero :: as) (zero :: bs) pr | y with x with inspect (canonical bs) timesZero'' (zero :: as) (zero :: bs) pr | y with x with inspect (canonical bs)
timesZero'' (zero :: as) (zero :: bs) pr | [] with prAs | y₁ with prBs rewrite prAs = inl refl timesZero'' (zero :: as) (zero :: bs) pr | [] with prAs | y₁ with prBs rewrite prAs = inl refl
timesZero'' (zero :: as) (zero :: bs) pr | (a1 :: a's) with prAs | [] with prBs rewrite prBs = inr refl timesZero'' (zero :: as) (zero :: bs) pr | (a1 :: a's) with prAs | [] with prBs rewrite prBs = inr refl
timesZero'' (zero :: as) (zero :: bs) pr | (a1 :: a's) with prAs | (x :: y) with prBs with inspect (canonical (as *B (zero :: bs))) timesZero'' (zero :: as) (zero :: bs) pr | (a1 :: a's) with prAs | (x :: y) with prBs with inspect (canonical (as *B (zero :: bs)))
timesZero'' (zero :: as) (zero :: bs) pr | (a1 :: a's) with prAs | (b1 :: b's) with prBs | [] with pr' with timesZero'' as (zero :: bs) pr' timesZero'' (zero :: as) (zero :: bs) pr | (a1 :: a's) with prAs | (b1 :: b's) with prBs | [] with pr' with timesZero'' as (zero :: bs) pr'
timesZero'' (zero :: as) (zero :: bs) pr | (a1 :: a's) with prAs | (b1 :: b's) with prBs | [] with pr' | inl x rewrite prAs | prBs | pr' | x = exFalso (nonEmptyNotEmpty x) timesZero'' (zero :: as) (zero :: bs) pr | (a1 :: a's) with prAs | (b1 :: b's) with prBs | [] with pr' | inl x rewrite prAs | prBs | pr' | x = exFalso (nonEmptyNotEmpty x)
timesZero'' (zero :: as) (zero :: bs) pr | (a1 :: a's) with prAs | (b1 :: b's) with prBs | [] with pr' | inr x rewrite prBs | pr' | x = exFalso (nonEmptyNotEmpty x) timesZero'' (zero :: as) (zero :: bs) pr | (a1 :: a's) with prAs | (b1 :: b's) with prBs | [] with pr' | inr x rewrite prBs | pr' | x = exFalso (nonEmptyNotEmpty x)
timesZero'' (zero :: as) (zero :: bs) pr | (a1 :: a's) with prAs | (b1 :: b's) with prBs | (x :: y) with pr' rewrite prAs | prBs | pr' = exFalso (nonEmptyNotEmpty pr) timesZero'' (zero :: as) (zero :: bs) pr | (a1 :: a's) with prAs | (b1 :: b's) with prBs | (x :: y) with pr' rewrite prAs | prBs | pr' = exFalso (nonEmptyNotEmpty pr)
timesZero'' (zero :: as) (one :: bs) pr with inspect (canonical as) timesZero'' (zero :: as) (one :: bs) pr with inspect (canonical as)
timesZero'' (zero :: as) (one :: bs) pr | [] with x rewrite x = inl refl timesZero'' (zero :: as) (one :: bs) pr | [] with x rewrite x = inl refl
timesZero'' (zero :: as) (one :: bs) pr | (a1 :: a's) with x with inspect (canonical (as *B (one :: bs))) timesZero'' (zero :: as) (one :: bs) pr | (a1 :: a's) with x with inspect (canonical (as *B (one :: bs)))
timesZero'' (zero :: as) (one :: bs) pr | (a1 :: a's) with x | [] with pr' with timesZero'' as (one :: bs) pr' timesZero'' (zero :: as) (one :: bs) pr | (a1 :: a's) with x | [] with pr' with timesZero'' as (one :: bs) pr'
timesZero'' (zero :: as) (one :: bs) pr | (a1 :: a's) with x | [] with pr' | inl pr'' rewrite x | pr' | pr'' = exFalso (nonEmptyNotEmpty pr'') timesZero'' (zero :: as) (one :: bs) pr | (a1 :: a's) with x | [] with pr' | inl pr'' rewrite x | pr' | pr'' = exFalso (nonEmptyNotEmpty pr'')
timesZero'' (zero :: as) (one :: bs) pr | (a1 :: a's) with x | (x₁ :: y) with pr' rewrite x | pr' = exFalso (nonEmptyNotEmpty pr) timesZero'' (zero :: as) (one :: bs) pr | (a1 :: a's) with x | (x₁ :: y) with pr' rewrite x | pr' = exFalso (nonEmptyNotEmpty pr)
timesZero'' (one :: as) (zero :: bs) pr with inspect (canonical bs) timesZero'' (one :: as) (zero :: bs) pr with inspect (canonical bs)
timesZero'' (one :: as) (zero :: bs) pr | [] with x rewrite x = inr refl timesZero'' (one :: as) (zero :: bs) pr | [] with x rewrite x = inr refl
timesZero'' (one :: as) (zero :: bs) pr | (b1 :: b's) with prB with inspect (canonical ((as *B (zero :: bs)) +B bs)) timesZero'' (one :: as) (zero :: bs) pr | (b1 :: b's) with prB with inspect (canonical ((as *B (zero :: bs)) +B bs))
timesZero'' (one :: as) (zero :: bs) pr | (b1 :: b's) with prB | [] with x rewrite equalityCommutative (+BIsInherited' (as *B (zero :: bs)) bs) = exFalso (nonEmptyNotEmpty (transitivity (equalityCommutative prB) v)) timesZero'' (one :: as) (zero :: bs) pr | (b1 :: b's) with prB | [] with x rewrite equalityCommutative (+BIsInherited' (as *B (zero :: bs)) bs) = exFalso (nonEmptyNotEmpty (transitivity (equalityCommutative prB) v))
where where
t : binNatToN (as *B (zero :: bs)) +N binNatToN bs 0 t : binNatToN (as *B (zero :: bs)) +N binNatToN bs 0
t = transitivity (equalityCommutative (nToN _)) (applyEquality binNatToN x) t = transitivity (equalityCommutative (nToN _)) (applyEquality binNatToN x)
@@ -82,98 +82,98 @@ module Numbers.BinaryNaturals.Multiplication where
v : canonical bs [] v : canonical bs []
v with u v with u
... | fst ,, snd = binNatToNZero bs snd ... | fst ,, snd = binNatToNZero bs snd
timesZero'' (one :: as) (zero :: bs) pr | (b1 :: b's) with prB | (x₁ :: y) with x rewrite prB | x = exFalso (nonEmptyNotEmpty pr) timesZero'' (one :: as) (zero :: bs) pr | (b1 :: b's) with prB | (x₁ :: y) with x rewrite prB | x = exFalso (nonEmptyNotEmpty pr)
lemma : {i : Bit} (a b : BinNat) canonical a canonical b canonical (i :: a) canonical (i :: b) lemma : {i : Bit} (a b : BinNat) canonical a canonical b canonical (i :: a) canonical (i :: b)
lemma {zero} a b pr with inspect (canonical a) lemma {zero} a b pr with inspect (canonical a)
lemma {zero} a b pr | [] with x rewrite x | equalityCommutative pr = refl lemma {zero} a b pr | [] with x rewrite x | equalityCommutative pr = refl
lemma {zero} a b pr | (a1 :: as) with x rewrite x | equalityCommutative pr = refl lemma {zero} a b pr | (a1 :: as) with x rewrite x | equalityCommutative pr = refl
lemma {one} a b pr = applyEquality (one ::_) pr lemma {one} a b pr = applyEquality (one ::_) pr
binNatToNDistributesPlus : (a b : BinNat) binNatToN (a +B b) binNatToN a +N binNatToN b binNatToNDistributesPlus : (a b : BinNat) binNatToN (a +B b) binNatToN a +N binNatToN b
binNatToNDistributesPlus a b = transitivity (equalityCommutative (binNatToNIsCanonical (a +B b))) (transitivity (applyEquality binNatToN (equalityCommutative (+BIsInherited' a b))) (nToN (binNatToN a +N binNatToN b))) 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 : 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 | equalityCommutative (Semiring.+Associative Semiring (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 : (a b : BinNat) (canonical b) (one :: []) canonical (a *B b) canonical a
timesOne [] b pr = refl timesOne [] b pr = refl
timesOne (zero :: a) b pr with inspect (canonical (a *B b)) timesOne (zero :: a) b pr with inspect (canonical (a *B b))
timesOne (zero :: a) b pr | [] with prAB with timesZero'' a b prAB timesOne (zero :: a) b pr | [] with prAB with timesZero'' a b prAB
timesOne (zero :: a) b pr | [] with prAB | inl a=0 rewrite a=0 | prAB = refl timesOne (zero :: a) b pr | [] with prAB | inl a=0 rewrite a=0 | prAB = refl
timesOne (zero :: a) b pr | [] with prAB | inr b=0 = exFalso (nonEmptyNotEmpty (transitivity (equalityCommutative pr) b=0)) timesOne (zero :: a) b pr | [] with prAB | inr b=0 = exFalso (nonEmptyNotEmpty (transitivity (equalityCommutative pr) b=0))
timesOne (zero :: a) b pr | (ab1 :: abs) with prAB with inspect (canonical a) timesOne (zero :: a) b pr | (ab1 :: abs) with prAB with inspect (canonical a)
timesOne (zero :: a) b pr | (ab1 :: abs) with prAB | [] with x = exFalso (nonEmptyNotEmpty (transitivity (equalityCommutative prAB) (timesZero a b x))) timesOne (zero :: a) b pr | (ab1 :: abs) with prAB | [] with x = exFalso (nonEmptyNotEmpty (transitivity (equalityCommutative prAB) (timesZero a b x)))
timesOne (zero :: a) b pr | (ab1 :: abs) with prAB | (x₁ :: y) with x rewrite prAB | x | timesOne a b pr = applyEquality (zero ::_) (transitivity (equalityCommutative prAB) x) timesOne (zero :: a) b pr | (ab1 :: abs) with prAB | (x₁ :: y) with x rewrite prAB | x | timesOne a b pr = applyEquality (zero ::_) (transitivity (equalityCommutative prAB) x)
timesOne (one :: a) (zero :: b) pr with canonical b timesOne (one :: a) (zero :: b) pr with canonical b
timesOne (one :: a) (zero :: b) () | [] timesOne (one :: a) (zero :: b) () | []
timesOne (one :: a) (zero :: b) () | x :: bl timesOne (one :: a) (zero :: b) () | x :: bl
timesOne (one :: a) (one :: b) pr rewrite canonicalDistributesPlus (a *B (one :: b)) b | ::Inj pr | +BCommutative (canonical (a *B (one :: b))) [] | timesOne a (one :: b) pr = refl timesOne (one :: a) (one :: b) pr rewrite canonicalDistributesPlus (a *B (one :: b)) b | ::Inj pr | +BCommutative (canonical (a *B (one :: b))) [] | timesOne a (one :: b) pr = refl
timesZero' : (a b : BinNat) canonical b [] canonical (a *B b) [] timesZero' : (a b : BinNat) canonical b [] canonical (a *B b) []
timesZero' [] b pr = refl timesZero' [] b pr = refl
timesZero' (zero :: a) b pr with inspect (canonical (a *B b)) timesZero' (zero :: a) b pr with inspect (canonical (a *B b))
timesZero' (zero :: a) b pr | [] with x rewrite x = refl timesZero' (zero :: a) b pr | [] with x rewrite x = refl
timesZero' (zero :: a) b pr | (ab1 :: abs) with prAB rewrite prAB = exFalso (nonEmptyNotEmpty (transitivity (equalityCommutative prAB) (timesZero' a b pr))) timesZero' (zero :: a) b pr | (ab1 :: abs) with prAB rewrite prAB = exFalso (nonEmptyNotEmpty (transitivity (equalityCommutative prAB) (timesZero' a b pr)))
timesZero' (one :: a) b pr rewrite canonicalDistributesPlus (zero :: (a *B b)) b | pr | +BCommutative (canonical (zero :: (a *B b))) [] = ans timesZero' (one :: a) b pr rewrite canonicalDistributesPlus (zero :: (a *B b)) b | pr | +BCommutative (canonical (zero :: (a *B b))) [] = ans
where where
ans : canonical (zero :: (a *B b)) [] ans : canonical (zero :: (a *B b)) []
ans with inspect (canonical (a *B b)) ans with inspect (canonical (a *B b))
ans | [] with x rewrite x = refl ans | [] with x rewrite x = refl
ans | (x₁ :: y) with x rewrite x = exFalso (nonEmptyNotEmpty (transitivity (equalityCommutative x) (timesZero' a b pr))) ans | (x₁ :: y) with x rewrite x = exFalso (nonEmptyNotEmpty (transitivity (equalityCommutative x) (timesZero' a b pr)))
canonicalDistributesTimes : (a b : BinNat) canonical (a *B b) canonical ((canonical a) *B (canonical b)) canonicalDistributesTimes : (a b : BinNat) canonical (a *B b) canonical ((canonical a) *B (canonical b))
canonicalDistributesTimes [] b = refl canonicalDistributesTimes [] b = refl
canonicalDistributesTimes (zero :: a) b with inspect (canonical a) canonicalDistributesTimes (zero :: a) b with inspect (canonical a)
canonicalDistributesTimes (zero :: a) b | [] with x rewrite timesZero a b x | x = refl canonicalDistributesTimes (zero :: a) b | [] with x rewrite timesZero a b x | x = refl
canonicalDistributesTimes (zero :: a) b | (a1 :: as) with prA with inspect (canonical b) canonicalDistributesTimes (zero :: a) b | (a1 :: as) with prA with inspect (canonical b)
canonicalDistributesTimes (zero :: a) b | (a1 :: as) with prA | [] with prB rewrite prA | prB = ans canonicalDistributesTimes (zero :: a) b | (a1 :: as) with prA | [] with prB rewrite prA | prB = ans
where where
ans : canonical (zero :: (a *B b)) canonical (zero :: ((a1 :: as) *B [])) ans : canonical (zero :: (a *B b)) canonical (zero :: ((a1 :: as) *B []))
ans with inspect (canonical ((a1 :: as) *B [])) ans with inspect (canonical ((a1 :: as) *B []))
ans | [] with x rewrite x | timesZero' a b prB = refl ans | [] with x rewrite x | timesZero' a b prB = refl
ans | (x₁ :: y) with x = exFalso (nonEmptyNotEmpty (equalityCommutative (transitivity (equalityCommutative (timesZero' (a1 :: as) [] refl)) x))) ans | (x₁ :: y) with x = exFalso (nonEmptyNotEmpty (equalityCommutative (transitivity (equalityCommutative (timesZero' (a1 :: as) [] refl)) x)))
canonicalDistributesTimes (zero :: a) b | (a1 :: as) with prA | (b1 :: bs) with prB with inspect (canonical (a *B b)) canonicalDistributesTimes (zero :: a) b | (a1 :: as) with prA | (b1 :: bs) with prB with inspect (canonical (a *B b))
canonicalDistributesTimes (zero :: a) b | (a1 :: as) with prA | (b1 :: bs) with prB | [] with x with timesZero'' a b x canonicalDistributesTimes (zero :: a) b | (a1 :: as) with prA | (b1 :: bs) with prB | [] with x with timesZero'' a b x
canonicalDistributesTimes (zero :: a) b | (a1 :: as) with prA | (b1 :: bs) with prB | [] with x | inl a=0 = exFalso (nonEmptyNotEmpty (transitivity (equalityCommutative prA) a=0)) canonicalDistributesTimes (zero :: a) b | (a1 :: as) with prA | (b1 :: bs) with prB | [] with x | inl a=0 = exFalso (nonEmptyNotEmpty (transitivity (equalityCommutative prA) a=0))
canonicalDistributesTimes (zero :: a) b | (a1 :: as) with prA | (b1 :: bs) with prB | [] with x | inr b=0 = exFalso (nonEmptyNotEmpty (transitivity (equalityCommutative prB) b=0)) canonicalDistributesTimes (zero :: a) b | (a1 :: as) with prA | (b1 :: bs) with prB | [] with x | inr b=0 = exFalso (nonEmptyNotEmpty (transitivity (equalityCommutative prB) b=0))
canonicalDistributesTimes (zero :: a) b | (a1 :: as) with prA | (b1 :: bs) with prB | (ab1 :: abs) with x rewrite prA | prB | x = ans canonicalDistributesTimes (zero :: a) b | (a1 :: as) with prA | (b1 :: bs) with prB | (ab1 :: abs) with x rewrite prA | prB | x = ans
where where
ans : zero :: ab1 :: abs canonical (zero :: ((a1 :: as) *B (b1 :: bs))) ans : zero :: ab1 :: abs canonical (zero :: ((a1 :: as) *B (b1 :: bs)))
ans rewrite equalityCommutative prA | equalityCommutative prB | equalityCommutative (canonicalDistributesTimes a b) | x = refl ans rewrite equalityCommutative prA | equalityCommutative prB | equalityCommutative (canonicalDistributesTimes a b) | x = refl
canonicalDistributesTimes (one :: a) b = transitivity (canonicalDistributesPlus (zero :: (a *B b)) b) (transitivity (transitivity (applyEquality (_+B canonical b) (transitivity (equalityCommutative (canonicalAscends'' (a *B b))) (transitivity (applyEquality (λ i canonical (zero :: i)) (canonicalDistributesTimes a b)) (canonicalAscends'' (canonical a *B canonical b))))) (applyEquality (λ i canonical (zero :: (canonical a *B canonical b)) +B i) (canonicalIdempotent b))) (equalityCommutative (canonicalDistributesPlus (zero :: (canonical a *B canonical b)) (canonical b)))) canonicalDistributesTimes (one :: a) b = transitivity (canonicalDistributesPlus (zero :: (a *B b)) b) (transitivity (transitivity (applyEquality (_+B canonical b) (transitivity (equalityCommutative (canonicalAscends'' (a *B b))) (transitivity (applyEquality (λ i canonical (zero :: i)) (canonicalDistributesTimes a b)) (canonicalAscends'' (canonical a *B canonical b))))) (applyEquality (λ i canonical (zero :: (canonical a *B canonical b)) +B i) (canonicalIdempotent b))) (equalityCommutative (canonicalDistributesPlus (zero :: (canonical a *B canonical b)) (canonical b))))
NToBinNatDistributesPlus : (a b : ) NToBinNat (a +N b) NToBinNat a +B NToBinNat b NToBinNatDistributesPlus : (a b : ) NToBinNat (a +N b) NToBinNat a +B NToBinNat b
NToBinNatDistributesPlus zero b = refl NToBinNatDistributesPlus zero b = refl
NToBinNatDistributesPlus (succ a) b with inspect (NToBinNat a) NToBinNatDistributesPlus (succ a) b with inspect (NToBinNat a)
... | bl with prA with inspect (NToBinNat (a +N b)) ... | bl with prA with inspect (NToBinNat (a +N b))
... | q with prAB = transitivity (applyEquality incr (NToBinNatDistributesPlus a b)) (incrPullsOut (NToBinNat a) (NToBinNat b)) ... | q with prAB = transitivity (applyEquality incr (NToBinNatDistributesPlus a b)) (incrPullsOut (NToBinNat a) (NToBinNat b))
timesCommutative : (a b : BinNat) canonical (a *B b) canonical (b *B a) timesCommutative : (a b : BinNat) canonical (a *B b) canonical (b *B a)
timesCommLemma : (a b : BinNat) canonical (zero :: (b *B a)) canonical (b *B (zero :: a)) timesCommLemma : (a b : BinNat) canonical (zero :: (b *B a)) canonical (b *B (zero :: a))
timesCommLemma a b rewrite timesCommutative b (zero :: a) | equalityCommutative (canonicalAscends'' {zero} (b *B a)) | equalityCommutative (canonicalAscends'' {zero} (a *B b)) | timesCommutative b a = refl timesCommLemma a b rewrite timesCommutative b (zero :: a) | equalityCommutative (canonicalAscends'' {zero} (b *B a)) | equalityCommutative (canonicalAscends'' {zero} (a *B b)) | timesCommutative b a = refl
timesCommutative [] b rewrite timesZero' b [] refl = refl timesCommutative [] b rewrite timesZero' b [] refl = refl
timesCommutative (x :: a) [] rewrite timesZero' (x :: a) [] refl = refl timesCommutative (x :: a) [] rewrite timesZero' (x :: a) [] refl = refl
timesCommutative (zero :: as) (zero :: b) rewrite equalityCommutative (canonicalAscends'' {zero} (as *B (zero :: b))) | timesCommutative as (zero :: b) | canonicalAscends'' {zero} (zero :: b *B as) | equalityCommutative (canonicalAscends'' {zero} (b *B (zero :: as))) | timesCommutative b (zero :: as) | canonicalAscends'' {zero} (zero :: (as *B b)) = ans timesCommutative (zero :: as) (zero :: b) rewrite equalityCommutative (canonicalAscends'' {zero} (as *B (zero :: b))) | timesCommutative as (zero :: b) | canonicalAscends'' {zero} (zero :: b *B as) | equalityCommutative (canonicalAscends'' {zero} (b *B (zero :: as))) | timesCommutative b (zero :: as) | canonicalAscends'' {zero} (zero :: (as *B b)) = ans
where where
ans : canonical (zero :: zero :: (b *B as)) canonical (zero :: zero :: (as *B b)) ans : canonical (zero :: zero :: (b *B as)) canonical (zero :: zero :: (as *B b))
ans rewrite equalityCommutative (canonicalAscends'' {zero} (zero :: (b *B as))) | equalityCommutative (canonicalAscends'' {zero} (b *B as)) | timesCommutative b as | canonicalAscends'' {zero} (as *B b) | canonicalAscends'' {zero} (zero :: (as *B b)) = refl ans rewrite equalityCommutative (canonicalAscends'' {zero} (zero :: (b *B as))) | equalityCommutative (canonicalAscends'' {zero} (b *B as)) | timesCommutative b as | canonicalAscends'' {zero} (as *B b) | canonicalAscends'' {zero} (zero :: (as *B b)) = refl
timesCommutative (zero :: as) (one :: b) = transitivity (equalityCommutative (canonicalAscends'' (as *B (one :: b)))) (transitivity (applyEquality (λ i canonical (zero :: i)) (timesCommutative as (one :: b))) (transitivity (applyEquality (λ i canonical (zero :: i)) ans) (canonicalAscends'' ((b *B (zero :: as)) +B as)))) timesCommutative (zero :: as) (one :: b) = transitivity (equalityCommutative (canonicalAscends'' (as *B (one :: b)))) (transitivity (applyEquality (λ i canonical (zero :: i)) (timesCommutative as (one :: b))) (transitivity (applyEquality (λ i canonical (zero :: i)) ans) (canonicalAscends'' ((b *B (zero :: as)) +B as))))
where where
ans : canonical ((zero :: (b *B as)) +B as) canonical ((b *B (zero :: as)) +B as) ans : canonical ((zero :: (b *B as)) +B as) canonical ((b *B (zero :: as)) +B as)
ans rewrite canonicalDistributesPlus (zero :: (b *B as)) as | canonicalDistributesPlus (b *B (zero :: as)) as = applyEquality (_+B canonical as) (timesCommLemma as b) ans rewrite canonicalDistributesPlus (zero :: (b *B as)) as | canonicalDistributesPlus (b *B (zero :: as)) as = applyEquality (_+B canonical as) (timesCommLemma as b)
timesCommutative (one :: as) (zero :: bs) = transitivity (equalityCommutative (canonicalAscends'' ((as *B (zero :: bs)) +B bs))) (transitivity (applyEquality (λ i canonical (zero :: i)) ans) (canonicalAscends'' (bs *B (one :: as)))) timesCommutative (one :: as) (zero :: bs) = transitivity (equalityCommutative (canonicalAscends'' ((as *B (zero :: bs)) +B bs))) (transitivity (applyEquality (λ i canonical (zero :: i)) ans) (canonicalAscends'' (bs *B (one :: as))))
where where
ans : canonical ((as *B (zero :: bs)) +B bs) canonical (bs *B (one :: as)) ans : canonical ((as *B (zero :: bs)) +B bs) canonical (bs *B (one :: as))
ans rewrite timesCommutative bs (one :: as) | canonicalDistributesPlus (as *B (zero :: bs)) bs | canonicalDistributesPlus (zero :: (as *B bs)) bs = applyEquality (_+B canonical bs) (equalityCommutative (timesCommLemma bs as)) ans rewrite timesCommutative bs (one :: as) | canonicalDistributesPlus (as *B (zero :: bs)) bs | canonicalDistributesPlus (zero :: (as *B bs)) bs = applyEquality (_+B canonical bs) (equalityCommutative (timesCommLemma bs as))
timesCommutative (one :: as) (one :: bs) = applyEquality (one ::_) (transitivity (canonicalDistributesPlus (as *B (one :: bs)) bs) (transitivity (transitivity (applyEquality (_+B canonical bs) (timesCommutative as (one :: bs))) (transitivity ans (equalityCommutative (applyEquality (_+B canonical as) (timesCommutative bs (one :: as)))))) (equalityCommutative (canonicalDistributesPlus (bs *B (one :: as)) as)))) timesCommutative (one :: as) (one :: bs) = applyEquality (one ::_) (transitivity (canonicalDistributesPlus (as *B (one :: bs)) bs) (transitivity (transitivity (applyEquality (_+B canonical bs) (timesCommutative as (one :: bs))) (transitivity ans (equalityCommutative (applyEquality (_+B canonical as) (timesCommutative bs (one :: as)))))) (equalityCommutative (canonicalDistributesPlus (bs *B (one :: as)) as))))
where where
ans : canonical ((zero :: (bs *B as)) +B as) +B canonical bs canonical ((zero :: (as *B bs)) +B bs) +B canonical as ans : canonical ((zero :: (bs *B as)) +B as) +B canonical bs canonical ((zero :: (as *B bs)) +B bs) +B canonical as
ans rewrite equalityCommutative (canonicalDistributesPlus ((zero :: (bs *B as)) +B as) bs) | equalityCommutative (canonicalDistributesPlus ((zero :: (as *B bs)) +B bs) as) | equalityCommutative (+BAssociative (zero :: (bs *B as)) as bs) | equalityCommutative (+BAssociative (zero :: (as *B bs)) bs as) | canonicalDistributesPlus (zero :: (bs *B as)) (as +B bs) | canonicalDistributesPlus (zero :: (as *B bs)) (bs +B as) | equalityCommutative (canonicalAscends'' {zero} (bs *B as)) | timesCommutative bs as | canonicalAscends'' {zero} (as *B bs) | +BCommutative as bs = refl ans rewrite equalityCommutative (canonicalDistributesPlus ((zero :: (bs *B as)) +B as) bs) | equalityCommutative (canonicalDistributesPlus ((zero :: (as *B bs)) +B bs) as) | equalityCommutative (+BAssociative (zero :: (bs *B as)) as bs) | equalityCommutative (+BAssociative (zero :: (as *B bs)) bs as) | canonicalDistributesPlus (zero :: (bs *B as)) (as +B bs) | canonicalDistributesPlus (zero :: (as *B bs)) (bs +B as) | equalityCommutative (canonicalAscends'' {zero} (bs *B as)) | timesCommutative bs as | canonicalAscends'' {zero} (as *B bs) | +BCommutative as bs = refl
*BIsInherited : (a b : BinNat) a *Binherit b canonical (a *B b) *BIsInherited : (a b : BinNat) a *Binherit b canonical (a *B b)
*BIsInherited a b = transitivity (applyEquality NToBinNat t) (transitivity (binToBin (canonical (a *B b))) (equalityCommutative (canonicalIdempotent (a *B b)))) *BIsInherited a b = transitivity (applyEquality NToBinNat t) (transitivity (binToBin (canonical (a *B b))) (equalityCommutative (canonicalIdempotent (a *B b))))
where where
ans : (a b : BinNat) binNatToN a *N binNatToN b binNatToN (a *B b) ans : (a b : BinNat) binNatToN a *N binNatToN b binNatToN (a *B b)
ans [] b = refl ans [] b = refl

View File

@@ -0,0 +1,35 @@
{-# OPTIONS --warning=error --safe #-}
open import LogicalFormulae
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals.Semiring
open import Numbers.Naturals.Order
open import Lists.Lists
open import Numbers.Primes.PrimeNumbers
open import Decidable.Relations
open import Numbers.BinaryNaturals.Definition
open import Numbers.BinaryNaturals.Addition
module ProjectEuler.Problem1 where
numbers<N : (N : ) List
numbers<N zero = []
numbers<N (succ N) = N :: numbers<N N
filter' : {a b : _} {A : Set a} {f : A Set b} (decidable : (x : A) (f x) || (f x False)) List A List A
filter' {f} decid [] = []
filter' {f} decid (x :: xs) with decid x
filter' {f} decid (x :: xs) | inl fx = x :: filter' decid xs
filter' {f} decid (x :: xs) | inr Notfx = filter' decid xs
filtered : (N : ) List
filtered N = filter' (orDecidable (divisionDecidable 3) (divisionDecidable 5)) (numbers<N N)
ans :
ans n = binNatToN (fold _+B_ (NToBinNat 0) (map NToBinNat (filtered n)))
t : ans 10 23
t = refl
--q : ans 1000 0 -- takes about 15secs for me to reduce the term that fills this hole
--q = refl