Add Everything (#34)

This commit is contained in:
Patrick Stevens
2019-08-18 10:35:15 +01:00
committed by GitHub
parent a31ae0d1ea
commit e7c54fa48a
45 changed files with 415 additions and 378 deletions

View File

@@ -2,9 +2,11 @@
open import LogicalFormulae
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Naturals
open import Numbers.Naturals
open import Vectors
module Categories.Category where
postulate
extensionality : {a b : _} {S : Set a}{T : S Set b} {f g : (x : S) T x} ((x : S) f x g x) f g

View File

@@ -1,4 +1,4 @@
{-# OPTIONS --safe --warning=error #-}
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)

View File

@@ -1,11 +1,11 @@
{-# OPTIONS --safe --warning=error #-}
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Groups.Groups
open import Groups.GroupDefinition
open import Groups.GroupsLemmas
open import Rings.RingDefinition
open import Rings.RingLemmas
open import Rings.Definition
open import Rings.Lemmas
open import Rings.IntegralDomains
open import Fields.Fields
open import Functions

View File

@@ -1,11 +1,11 @@
{-# OPTIONS --safe --warning=error #-}
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Groups.Groups
open import Groups.GroupDefinition
open import Groups.GroupsLemmas
open import Rings.RingDefinition
open import Rings.RingLemmas
open import Rings.Definition
open import Rings.Lemmas
open import Rings.IntegralDomains
open import Fields.Fields
open import Functions

View File

@@ -1,10 +1,10 @@
{-# OPTIONS --safe --warning=error #-}
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Groups.Groups
open import Groups.GroupDefinition
open import Rings.RingDefinition
open import Rings.RingLemmas
open import Rings.Definition
open import Rings.Lemmas
open import Setoids.Setoids
open import Setoids.Orders
open import Orders

View File

@@ -1,4 +1,4 @@
{-# OPTIONS --safe --warning=error #-}
{-# OPTIONS --safe --warning=error --without-K #-}
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)

View File

@@ -1,15 +1,16 @@
{-# 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 Numbers.Naturals
open import Numbers.Integers
open import Sets.FinSet
open import Groups.Groups
open import Groups.GroupDefinition
module CyclicGroups where
module Groups.CyclicGroups where
positiveEltPower : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A A A} (G : Group S _·_) (x : A) (i : ) A
positiveEltPower G x 0 = Group.identity G

View File

@@ -1,4 +1,4 @@
{-# OPTIONS --safe --warning=error #-}
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Numbers.Naturals -- for length

View File

@@ -1,10 +1,12 @@
{-# OPTIONS --safe --warning=error #-}
open import LogicalFormulae
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 Sets.FinSet
open import Groups.GroupDefinition
open import Groups.SymmetryGroups
@@ -26,15 +28,15 @@ module Groups.FreeGroups where
ofInvInjective refl = refl
decidableFreeCompletion : {a : _} {A : Set a} DecidableSet A DecidableSet (FreeCompletion A)
decidableFreeCompletion {A = A} record { eq = eq } = record { eq = pr }
decidableFreeCompletion {A = A} record { eq = dec } = record { eq = pr }
where
pr : (a b : FreeCompletion A) (a b) || (a b False)
pr (ofLetter x) (ofLetter y) with eq x y
pr (ofLetter x) (ofLetter y) with dec x y
... | inl refl = inl refl
... | inr x!=y = inr λ p x!=y (ofLetterInjective p)
pr (ofLetter x) (ofInv y) = inr λ ()
pr (ofInv x) (ofLetter y) = inr λ ()
pr (ofInv x) (ofInv y) with eq x y
pr (ofInv x) (ofInv y) with dec x y
... | inl refl = inl refl
... | inr x!=y = inr λ p x!=y (ofInvInjective p)
@@ -117,7 +119,7 @@ module Groups.FreeGroups where
where
ans : freeCompletionEqual decA (ofLetter x) (ofLetter startLetter) BoolFalse
ans with DecidableSet.eq (decidableFreeCompletion decA) (ofLetter x) (ofInv startLetter)
ans | bl with eq x startLetter
ans | bl with DecidableSet.eq decA x startLetter
ans | bl | inl x=sl = exFalso (startLetter!=x (equalityCommutative x=sl))
ans | bl | inr x!=sl = refl
bij : SetoidBijection (reflSetoid (ReducedWord decA)) (reflSetoid (ReducedWord decA)) f
@@ -136,7 +138,7 @@ module Groups.FreeGroups where
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofLetter x) w1 prX} {empty} fx=fy with prependLetterInjective fx=fy
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofLetter x) w1 prX} {empty} fx=fy | ()
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofLetter x) w1 prX} {prependLetter (ofLetter y) w2 prY} fx=fy = prependLetterInjective fx=fy
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofLetter a) w1 prX} {prependLetter (ofInv y) w2 prY} fx=fy with eq y x
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofLetter a) w1 prX} {prependLetter (ofInv y) w2 prY} fx=fy with DecidableSet.eq decA y x
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofLetter a) w1 prX} {prependLetter (ofInv y) empty prY} () | inl y=x
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofLetter a) w1 prX} {prependLetter (ofInv y) (prependLetter (ofLetter b) w2 x₁) prY} fx=fy | inl y=x with ofLetterInjective (prependLetterInjective' fx=fy)
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofLetter a) w1 prX} {prependLetter (ofInv y) (prependLetter (ofLetter b) w2 x₁) (wordEmpty ())} fx=fy | inl y=x | x=b
@@ -145,7 +147,7 @@ module Groups.FreeGroups where
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofLetter a) w1 prX} {prependLetter (ofInv y) (prependLetter (ofInv b) w2 x₁) prY} fx=fy | inl y=x | ()
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofLetter a) w1 prX} {prependLetter (ofInv y) w2 prY} fx=fy | inr y!=x with prependLetterInjective fx=fy
... | bl = bl
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofInv a) w1 prX} {empty} fx=fy with eq a x
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofInv a) w1 prX} {empty} fx=fy with DecidableSet.eq decA a x
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofInv a) empty prX} {empty} () | inl a=x
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofInv a) (prependLetter (ofLetter x₂) w1 x₁) prX} {empty} fx=fy | inl a=x with ofLetterInjective (prependLetterInjective' fx=fy)
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofInv a) (prependLetter (ofLetter x₂) w1 x₁) (wordEmpty ())} {empty} fx=fy | inl a=x | x2=x
@@ -153,21 +155,21 @@ module Groups.FreeGroups where
... | bl = exFalso (bl (applyEquality ofInv (transitivity a=x (equalityCommutative x2=x))))
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofInv a) (prependLetter (ofInv x₂) w1 x₁) prX} {empty} () | inl a=x
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofInv a) w1 prX} {empty} () | inr a!=x
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofInv a) w1 prX} {prependLetter (ofLetter x₁) y x₂} fx=fy with eq a x
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofInv a) w1 prX} {prependLetter (ofLetter x₁) y x₂} fx=fy with DecidableSet.eq decA a x
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofInv a) empty prX} {prependLetter (ofLetter b) y x₂} () | inl a=x
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofInv a) (prependLetter (ofLetter x₃) w1 x₁) prX} {prependLetter (ofLetter b) y x₂} fx=fy | inl a=x with ofLetterInjective (prependLetterInjective' fx=fy)
... | x3=x rewrite a=x | x3=x = exFalso (badPrepend' prX)
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofInv a) (prependLetter (ofInv x₃) w1 x₁) prX} {prependLetter (ofLetter b) y x₂} fx=fy | inl a=x with prependLetterInjective' fx=fy
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofInv a) (prependLetter (ofInv x₃) w1 x₁) prX} {prependLetter (ofLetter b) y x₂} fx=fy | inl a=x | ()
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofInv a) w1 prX} {prependLetter (ofLetter x₁) y x₂} () | inr a!=x
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofInv a) w1 prX} {prependLetter (ofInv x₁) y x₂} fx=fy with eq a x
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofInv a) w1 prX} {prependLetter (ofInv b) y x₂} fx=fy | inl a=x with eq b x
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofInv a) w1 prX} {prependLetter (ofInv x₁) y x₂} fx=fy with DecidableSet.eq decA a x
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofInv a) w1 prX} {prependLetter (ofInv b) y x₂} fx=fy | inl a=x with DecidableSet.eq decA b x
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofInv a) w1 prX} {prependLetter (ofInv b) y x₂} fx=fy | inl a=x | inl b=x rewrite fx=fy | a=x | b=x = prependLetterRefl
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofInv a) empty prX} {prependLetter (ofInv b) y x₂} () | inl a=x | inr b!=x
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofInv a) (prependLetter (ofLetter x₃) w1 x₁) prX} {prependLetter (ofInv b) y x₂} fx=fy | inl a=x | inr b!=x with ofLetterInjective (prependLetterInjective' fx=fy)
... | x3=x rewrite a=x | x3=x = exFalso (badPrepend' prX)
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofInv a) (prependLetter (ofInv x₃) w1 x₁) prX} {prependLetter (ofInv b) y x₂} () | inl a=x | inr b!=x
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofInv a) w1 prX} {prependLetter (ofInv b) y x₂} fx=fy | inr a!=x with eq b x
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofInv a) w1 prX} {prependLetter (ofInv b) y x₂} fx=fy | inr a!=x with DecidableSet.eq decA b x
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofInv a) w1 prX} {prependLetter (ofInv b) empty x₂} () | inr a!=x | inl b=x
SetoidInjection.injective (SetoidBijection.inj bij) {prependLetter (ofInv a) w1 prX} {prependLetter (ofInv b) (prependLetter (ofLetter x₃) y x₁) x2} fx=fy | inr a!=x | inl b=x with ofLetterInjective (prependLetterInjective' fx=fy)
... | x3=x rewrite (equalityCommutative x3=x) | b=x = exFalso (badPrepend' x2)
@@ -200,12 +202,12 @@ module Groups.FreeGroups where
SetoidSurjection.surjective (SetoidBijection.surj bij) {prependLetter (ofLetter l) (prependLetter letter w pr2) pr} | inr l!=x = prependLetter (ofInv x) (prependLetter (ofLetter l) (prependLetter letter w pr2) pr) (wordEnding (succIsPositive _) (freeCompletionEqualFalse decA λ p l!=x (ofInvInjective (equalityCommutative p)))) , needed
where
needed : f (prependLetter (ofInv x) (prependLetter (ofLetter l) (prependLetter letter w pr2) pr) (wordEnding (succIsPositive (succ (wordLength w))) (freeCompletionEqualFalse decA (λ p l!=x (ofInvInjective (equalityCommutative p)))))) prependLetter (ofLetter l) (prependLetter letter w pr2) pr
needed with eq x x
needed with DecidableSet.eq decA x x
needed | inl x₁ = refl
needed | inr x!=x = exFalso (x!=x refl)
SetoidSurjection.surjective (SetoidBijection.surj bij) {prependLetter (ofInv l) w pr} = prependLetter (ofInv x) (prependLetter (ofInv l) w pr) (wordEnding (succIsPositive _) (freeCompletionEqualFalse decA {ofInv x} {ofLetter l} λ ())) , needed
where
needed : f (prependLetter (ofInv x) (prependLetter (ofInv l) w pr) (wordEnding (succIsPositive (wordLength w)) (freeCompletionEqualFalse decA {ofInv x} {ofLetter l} λ ()))) (prependLetter (ofInv l) w pr)
needed with eq x x
needed with DecidableSet.eq decA x x
needed | inl x₁ = refl
needed | inr x!=x = exFalso (x!=x refl)

View File

@@ -8,7 +8,7 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals
open import Sets.FinSet
open import Groups.GroupDefinition
open import Lists
open import Lists.Lists
open import Orders
open import Groups.Groups
@@ -65,7 +65,7 @@ module Groups.GeneratedGroup where
evalWord {_+_ = _+_} G record { word = (ofInv x :: w) } = (Group.inverse G x) + evalWord G record { word = w }
mapWord : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} (f : A B) (w : Word S) Word T
mapWord f record { word = word } = record { word = listMap (freeCompletionMap f) word }
mapWord f record { word = word } = record { word = map (freeCompletionMap f) word }
subgroupSetoid : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} (T : Setoid {b} {d} B) {_+_ : A A A} (G : Group S _+_) {f : B A} (fInj : SetoidInjection T S f) Setoid (Word T)
Setoid.__ (subgroupSetoid {S = S} T G {f} fInj) w1 w2 = Setoid.__ S (evalWord G (mapWord f w1)) (evalWord G (mapWord f w2))
@@ -85,7 +85,7 @@ module Groups.GeneratedGroup where
mapWordWellDefined : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} (f : A B) ({x y : A} (Setoid.__ S x y) (Setoid.__ T (f x) (f y))) {w1 : Word S} {w2 : Word S} Setoid.__ (wordSetoid S) w1 w2 Setoid.__ (wordSetoid T) (mapWord f w1) (mapWord f w2)
mapWordWellDefined {S = S} {T = T} f fWD {w1} {w2} w1=w2 = p
where
p : listEquality (freeCompletionSetoid T) (listMap (freeCompletionMap {S = S} {T = T} f) (Word.word w1)) (listMap (freeCompletionMap f) (Word.word w2))
p : listEquality (freeCompletionSetoid T) (map (freeCompletionMap {S = S} {T = T} f) (Word.word w1)) (map (freeCompletionMap f) (Word.word w2))
p = listEqualityRespectsMap (freeCompletionSetoid S) (freeCompletionSetoid T) (freeCompletionMap {S = S} {T} f) (λ {x} {y} freeCompletionMapWellDefined {S = S} {T = T} f fWD {x} {y}) {w1 = Word.word w1} {w2 = Word.word w2} w1=w2
subgroupOp : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} (T : Setoid {b} {d} B) {_+_ : A A A} (G : Group S _+_) {f : B A} (fInj : SetoidInjection T S f) Word T Word T Word T
@@ -97,7 +97,7 @@ module Groups.GeneratedGroup where
need : Setoid.__ (subgroupSetoid T G fInj) (subgroupOp T G fInj (record { word = w }) (record { word = x })) (subgroupOp T G fInj (record { word = y }) (record { word = z }))
need = {!!}
Group.identity (generatedSubgroup T G fInj) = record { word = [] }
Group.inverse (generatedSubgroup T G fInj) record { word = word } = record { word = listMap freeInverse (listRev word) }
Group.inverse (generatedSubgroup T G fInj) record { word = word } = record { word = map freeInverse (rev word) }
Group.multAssoc (generatedSubgroup T G fInj) = {!!}
Group.multIdentRight (generatedSubgroup {S = S} T G {f = f} fInj) {a = record { word = word }} = need
where
@@ -111,6 +111,6 @@ module Groups.GeneratedGroup where
where
open Setoid S
open Group G
need : Setoid.__ (subgroupSetoid T G fInj) (record { word = (listMap freeInverse (listRev (x :: word))) ++ (x :: word) }) (Group.identity (generatedSubgroup T G fInj))
need : Setoid.__ (subgroupSetoid T G fInj) (record { word = (map freeInverse (rev (x :: word))) ++ (x :: word) }) (Group.identity (generatedSubgroup T G fInj))
need = {!!}
Group.invRight (generatedSubgroup T G fInj) = {!!}

View File

@@ -1,15 +1,16 @@
{-# 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 Numbers.Naturals
open import Numbers.Integers
open import Sets.FinSet
open import Groups.GroupDefinition
open import Groups.Groups
module GroupCosets where
module Groups.GroupCosets where
data GroupCoset {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+A_ : A A A} {_+B_ : B B B} {G : Group S _+A_} {H : Group T _+B_} {f : B A} {fHom : GroupHom H G f} (subgrp : Subgroup G H fHom) : Set (a b c d) where
cosetOfElt : (x : A) GroupCoset subgrp

View File

@@ -1,4 +1,4 @@
{-# OPTIONS --safe --warning=error #-}
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Setoids.Setoids

View File

@@ -1,4 +1,4 @@
{-# OPTIONS --safe --warning=error #-}
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Setoids.Setoids

View File

@@ -10,7 +10,7 @@ open import Sets.FinSet
open import Functions
open import Numbers.Naturals
open import IntegersModN
open import Rings.RingExamples
open import Rings.Examples.Examples
open import PrimeNumbers
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)

View File

@@ -1,4 +1,4 @@
{-# OPTIONS --safe --warning=error #-}
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Setoids.Setoids

View File

@@ -1,4 +1,4 @@
{-# OPTIONS --safe --warning=error #-}
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Setoids.Setoids

View File

@@ -7,7 +7,9 @@ open import Numbers.Naturals
open import PrimeNumbers
open import Setoids.Setoids
open import Sets.FinSet
open import Sets.FinSetWithK
open import Functions
open import Numbers.NaturalsWithK
module IntegersModN where
record n (n : ) (pr : 0 <N n) : Set where

View File

@@ -2,11 +2,11 @@
open import LogicalFormulae
open import Orders
open import Groups
open import Naturals
open import Groups.Groups
open import Numbers.Naturals
open import PrimeNumbers
open import Rings
open import Setoids
open import Rings.Definition
open import Setoids.Setoids
open import IntegersModN
module IntegersModNRing where
@@ -54,12 +54,11 @@ module IntegersModNRing where
nMultIdent {succ (succ n)} {pr} record { x = a ; xLess = aLess } | inl (inr ssn<a) = exFalso (PartialOrder.irreflexive (TotalOrder.order TotalOrder) (PartialOrder.transitive (TotalOrder.order TotalOrder) aLess ssn<a))
nMultIdent {succ (succ n)} {pr} record { x = .(succ (succ n)) ; xLess = aLess } | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order TotalOrder) aLess)
nRing : (n : ) (pr : 0 <N n) Ring {_} (n n pr)
Ring.additiveGroup (nRing n 0<n) = AbelianGroup.grp (nGroup n 0<n)
Ring._*_ (nRing n 0<n) a b = a *n b
nRing : (n : ) (pr : 0 <N n) Ring (reflSetoid (n n pr)) (_+n_) (_*n_)
Ring.additiveGroup (nRing n 0<n) = (nGroup n 0<n)
Ring.multWellDefined (nRing n 0<n) = reflGroupWellDefined
Ring.1R (nRing n pr) = nIdent n pr
Ring.groupIsAbelian (nRing n 0<n) = AbelianGroup.commutative (nGroup n 0<n)
Ring.groupIsAbelian (nRing n 0<n) = AbelianGroup.commutative (nAbGroup n 0<n)
Ring.multAssoc (nRing n 0<n) = {!!}
Ring.multCommutative (nRing n 0<n) {a} {b} = nMultCommutative a b
Ring.multDistributes (nRing n 0<n) = {!!}

View File

@@ -1,4 +1,4 @@
{-# OPTIONS --safe --warning=error #-}
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Functions

View File

@@ -1,4 +1,4 @@
{-# OPTIONS --safe --warning=error #-}
{-# OPTIONS --safe --warning=error --without-K #-}
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
@@ -94,9 +94,6 @@ boolAnd BoolFalse y = BoolFalse
typeCast : {a : _} {A : Set a} {B : Set a} (el : A) (pr : A B) B
typeCast {a} {A} {.A} elt refl = elt
typeCastEqual : {a : _} {A : Set a} {B : Set a} {el : A} (pr1 pr2 : A B) (typeCast el pr1) (typeCast el pr2)
typeCastEqual {a} {A} {.A} {el} refl refl = refl
data Singleton {a} {A : Set a} (x : A) : Set a where
_with≡_ : (y : A) x y Singleton x
@@ -113,6 +110,3 @@ decidableOr : {a b : _} → (A : Set a) → (B : Set b) → (A || (A → False))
decidableOr {a} {b} A B decidable (inl x) = inl x
decidableOr {a} {b} A B (inl y) (inr x) = inl y
decidableOr {a} {b} A B (inr y) (inr x) = inr (record { fst = y ; snd = x})
reflRefl : {a : _} {A : Set a} {r s : A} (pr1 pr2 : r s) pr1 pr2
reflRefl refl refl = refl

View File

@@ -1,4 +1,4 @@
{-# OPTIONS --warning=error --safe #-}
{-# OPTIONS --warning=error --safe --without-K #-}
open import LogicalFormulae
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)

View File

@@ -0,0 +1,198 @@
{-# OPTIONS --warning=error --safe --without-K #-}
open import LogicalFormulae
open import Functions
open import Lists.Lists
open import Numbers.Naturals
open import Groups.GroupDefinition
open import Numbers.BinaryNaturals.Definition
module Numbers.BinaryNaturals.Addition where
-- Define the monoid structure, and show that it's the same as 's
_+Binherit_ : BinNat BinNat BinNat
a +Binherit b = NToBinNat (binNatToN a +N binNatToN b)
_+B_ : BinNat BinNat BinNat
[] +B b = b
(x :: a) +B [] = x :: a
(zero :: xs) +B (y :: ys) = y :: (xs +B ys)
(one :: xs) +B (zero :: ys) = one :: (xs +B ys)
(one :: xs) +B (one :: ys) = zero :: incr (xs +B ys)
+BCommutative : (a b : BinNat) a +B b b +B a
+BCommutative [] [] = refl
+BCommutative [] (x :: b) = refl
+BCommutative (x :: a) [] = refl
+BCommutative (zero :: as) (zero :: bs) rewrite +BCommutative as bs = refl
+BCommutative (zero :: as) (one :: bs) rewrite +BCommutative as bs = refl
+BCommutative (one :: as) (zero :: bs) rewrite +BCommutative as bs = refl
+BCommutative (one :: as) (one :: bs) rewrite +BCommutative as bs = refl
+BIsInherited[] : (b : BinNat) (prB : b canonical b) [] +Binherit b [] +B b
+BIsInherited[] [] prB = refl
+BIsInherited[] (zero :: b) prB = t
where
refine : (b : BinNat) zero :: b canonical (zero :: b) b canonical b
refine b pr with canonical b
refine b pr | x :: bl = ::Inj pr
t : NToBinNat (0 +N binNatToN (zero :: b)) zero :: b
t with orderIsTotal 0 (binNatToN b)
t | inl (inl pos) = transitivity (doubleIsBitShift (binNatToN b) pos) (applyEquality (zero ::_) (transitivity (binToBin b) (equalityCommutative (refine b prB))))
t | inl (inr ())
... | inr eq with binNatToNZero b (equalityCommutative eq)
... | u with canonical b
t | inr eq | u | [] = exFalso (bad b prB)
where
bad : (c : BinNat) zero :: c [] False
bad c ()
t | inr eq | () | x :: bl
+BIsInherited[] (one :: b) prB = ans
where
ans : NToBinNat (binNatToN (one :: b)) one :: b
ans = transitivity (binToBin (one :: b)) (equalityCommutative prB)
-- Show that the monoid structure of is the same as that of BinNat
+BIsInherited : (a b : BinNat) (prA : a canonical a) (prB : b canonical b) a +Binherit b a +B b
+BinheritLemma : (a : BinNat) (b : BinNat) (prA : a canonical a) (prB : b canonical b) incr (NToBinNat ((binNatToN a +N binNatToN b) +N ((binNatToN a +N binNatToN b) +N zero))) one :: (a +B b)
+BIsInherited' : (a b : BinNat) a +Binherit b canonical (a +B b)
+BinheritLemma a b prA prB with orderIsTotal 0 (binNatToN a +N binNatToN b)
+BinheritLemma a b prA prB | inl (inl x) rewrite doubleIsBitShift (binNatToN a +N binNatToN b) x = applyEquality (one ::_) (+BIsInherited a b prA prB)
+BinheritLemma a b prA prB | inr x with sumZeroImpliesOperandsZero (binNatToN a) (equalityCommutative x)
+BinheritLemma a b prA prB | inr x | fst ,, snd = ans2
where
bad : b []
bad = transitivity prB (binNatToNZero b snd)
bad2 : a []
bad2 = transitivity prA (binNatToNZero a fst)
ans2 : incr (NToBinNat ((binNatToN a +N binNatToN b) +N ((binNatToN a +N binNatToN b) +N zero))) one :: (a +B b)
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 (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)))
+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
where
bad : (b : BinNat) (pr : b canonical b) (pr2 : binNatToN b 0) b []
bad b pr pr2 = transitivity pr (binNatToNZero b pr2)
t : b canonical b
t with canonical b
t | x :: bl = ::Inj prB
u : b []
u = bad b t b=0
nono : {A : Set} {a : A} {as : List A} a :: as [] False
nono ()
ans : False
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'[] : (b : BinNat) [] +Binherit b canonical ([] +B b)
+BIsInherited'[] [] = refl
+BIsInherited'[] (zero :: b) with inspect (canonical b)
+BIsInherited'[] (zero :: b) | [] with pr rewrite binNatToNZero' b pr | pr = refl
+BIsInherited'[] (zero :: b) | (x :: bl) with pr rewrite pr = ans
where
contr : {a : _} {A : Set a} {l1 l2 : List A} {x : A} l1 [] l1 x :: l2 False
contr {l1 = []} p1 ()
contr {l1 = x :: l1} () p2
ans : NToBinNat (binNatToN b +N (binNatToN b +N zero)) zero :: x :: bl
ans with inspect (binNatToN b)
ans | zero with th rewrite th = exFalso (contr (binNatToNZero b th) pr)
ans | succ th with blah rewrite blah | doubleIsBitShift' th = applyEquality (zero ::_) (transitivity (equalityCommutative u) pr)
where
u : canonical b incr (NToBinNat th)
u = transitivity (equalityCommutative (binToBin b)) (applyEquality NToBinNat blah)
+BIsInherited'[] (one :: b) with inspect (binNatToN b)
... | zero with pr rewrite pr = applyEquality (one ::_) (equalityCommutative (binNatToNZero b pr))
... | (succ bl) with pr = ans
where
u : NToBinNat (2 *N binNatToN b) zero :: canonical b
u with doubleIsBitShift' bl
... | t = transitivity (identityOfIndiscernablesLeft _ _ _ _≡_ t (applyEquality (λ i NToBinNat (2 *N i)) (equalityCommutative pr))) (applyEquality (zero ::_) (transitivity (applyEquality NToBinNat (equalityCommutative pr)) (binToBin b)))
ans : incr (NToBinNat (binNatToN b +N (binNatToN b +N zero))) one :: canonical b
ans = applyEquality incr u
+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)))
where
bad : canonical a [] False
bad pr with transitivity (equalityCommutative x) (transitivity (equalityCommutative (binNatToNIsCanonical a)) (applyEquality binNatToN pr))
bad pr | ()
+BIsInherited' (one :: a) [] with inspect (binNatToN a)
+BIsInherited' (one :: a) [] | 0 with x rewrite x | binNatToNZero a x = refl
+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
where
ans : NToBinNat (2 *N (binNatToN as +N binNatToN bs)) canonical (zero :: (as +B bs))
ans with inspect (binNatToN as +N binNatToN bs)
ans | zero with x with sumZeroImpliesOperandsZero (binNatToN as) x
... | as=0 ,, bs=0 rewrite as=0 | bs=0 = foo
where
u : canonical (as +Binherit bs) []
u rewrite as=0 | bs=0 = refl
foo : [] canonical (zero :: (as +B bs))
foo = transitivity (transitivity b (applyEquality (λ i canonical (zero :: i)) (+BIsInherited' as bs))) (canonicalAscends'' {zero} (as +B bs))
where
b : [] canonical (zero :: (as +Binherit bs))
b rewrite u = refl
ans | succ y with x rewrite x | doubleIsBitShift' y = transitivity (applyEquality (λ i zero :: NToBinNat i) (equalityCommutative x)) ans2
where
u : 0 <N binNatToN (as +B bs)
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
where
ans2 : incr (NToBinNat (2 *N (binNatToN as +N binNatToN bs))) one :: canonical (as +B bs)
ans2 with inspect (binNatToN as +N binNatToN bs)
ans2 | zero with x with sumZeroImpliesOperandsZero (binNatToN as) x
ans2 | zero with x | as=0 ,, bs=0 rewrite as=0 | bs=0 = applyEquality (one ::_) (transitivity t (+BIsInherited' as bs))
where
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
where
ans : incr (NToBinNat (2 *N (binNatToN as +N binNatToN bs))) one :: canonical (as +B bs)
ans with inspect (binNatToN as +N binNatToN bs)
ans | zero with x with sumZeroImpliesOperandsZero (binNatToN as) x
... | as=0 ,, bs=0 rewrite as=0 | bs=0 = applyEquality (one ::_) (transitivity t (+BIsInherited' as bs))
where
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
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)
... | zero with x with sumZeroImpliesOperandsZero (binNatToN as) x
ans | zero with x | as=0 ,, bs=0 rewrite as=0 | bs=0 = bar
where
u' : canonical (as +Binherit bs) []
u' rewrite as=0 | bs=0 = refl
u : canonical (as +B bs) []
u rewrite equalityCommutative (+BIsInherited' as bs) = transitivity (NToBinNatIsCanonical (binNatToN as +N binNatToN bs)) u'
t : canonical (incr (as +B bs)) one :: []
t rewrite incrPreservesCanonical' (as +B bs) | u = refl
bar : zero :: one :: [] canonical (zero :: incr (as +B bs))
bar rewrite t = refl
ans | succ y with x rewrite x | doubleIsBitShift' y = transitivity (applyEquality (λ i zero :: incr (NToBinNat i)) (equalityCommutative x)) ans2
where
ans2 : zero :: incr (as +Binherit bs) canonical (zero :: incr (as +B bs))
ans2 rewrite +BIsInherited' as bs | equalityCommutative (incrPreservesCanonical' (as +B bs)) | canonicalAscends' {zero} (incr (as +B bs)) (incrNonzero (as +B bs)) = refl

View File

@@ -1,4 +1,4 @@
{-# OPTIONS --warning=error --safe #-}
{-# OPTIONS --warning=error --safe --without-K #-}
open import LogicalFormulae
open import Functions
@@ -16,7 +16,8 @@ module Numbers.BinaryNaturals.Definition where
BinNat = List Bit
::Inj : {xs ys : BinNat} {i : Bit} i :: xs i :: ys xs ys
::Inj refl = refl
::Inj {i = zero} refl = refl
::Inj {i = one} refl = refl
nonEmptyNotEmpty : {a : _} {A : Set a} {l1 : List A} {i : A} i :: l1 [] False
nonEmptyNotEmpty {l1 = l1} {i} ()
@@ -136,18 +137,6 @@ module Numbers.BinaryNaturals.Definition where
binToBin : (x : BinNat) NToBinNat (binNatToN x) canonical x
binToBin x = transitivity (NToBinNatIsCanonical (binNatToN x)) (binNatToNInj (NToBinNat (binNatToN x)) x (nToN (binNatToN x)))
-- Define the monoid structure, and show that it's the same as 's
_+Binherit_ : BinNat BinNat BinNat
a +Binherit b = NToBinNat (binNatToN a +N binNatToN b)
_+B_ : BinNat BinNat BinNat
[] +B b = b
(x :: a) +B [] = x :: a
(zero :: xs) +B (y :: ys) = y :: (xs +B ys)
(one :: xs) +B (zero :: ys) = one :: (xs +B ys)
(one :: xs) +B (one :: ys) = zero :: incr (xs +B ys)
doubleIsBitShift' : (a : ) NToBinNat (2 *N succ a) zero :: NToBinNat (succ a)
doubleIsBitShift' zero = refl
doubleIsBitShift' (succ a) with doubleIsBitShift' a
@@ -157,87 +146,11 @@ module Numbers.BinaryNaturals.Definition where
doubleIsBitShift zero ()
doubleIsBitShift (succ a) _ = doubleIsBitShift' a
+BCommutative : (a b : BinNat) a +B b b +B a
+BCommutative [] [] = refl
+BCommutative [] (x :: b) = refl
+BCommutative (x :: a) [] = refl
+BCommutative (zero :: as) (zero :: bs) rewrite +BCommutative as bs = refl
+BCommutative (zero :: as) (one :: bs) rewrite +BCommutative as bs = refl
+BCommutative (one :: as) (zero :: bs) rewrite +BCommutative as bs = refl
+BCommutative (one :: as) (one :: bs) rewrite +BCommutative as bs = refl
+BIsInherited[] : (b : BinNat) (prB : b canonical b) [] +Binherit b [] +B b
+BIsInherited[] [] prB = refl
+BIsInherited[] (zero :: b) prB = t
where
refine : (b : BinNat) zero :: b canonical (zero :: b) b canonical b
refine b pr with canonical b
refine b pr | x :: bl = ::Inj pr
t : NToBinNat (0 +N binNatToN (zero :: b)) zero :: b
t with orderIsTotal 0 (binNatToN b)
t | inl (inl pos) = transitivity (doubleIsBitShift (binNatToN b) pos) (applyEquality (zero ::_) (transitivity (binToBin b) (equalityCommutative (refine b prB))))
t | inl (inr ())
... | inr eq with binNatToNZero b (equalityCommutative eq)
... | u with canonical b
t | inr eq | u | [] = exFalso (bad b prB)
where
bad : (c : BinNat) zero :: c [] False
bad c ()
t | inr eq | () | x :: bl
+BIsInherited[] (one :: b) prB = ans
where
ans : NToBinNat (binNatToN (one :: b)) one :: b
ans = transitivity (binToBin (one :: b)) (equalityCommutative prB)
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 | x :: bl = ::Inj pr
canonicalDescends {one} as pr = ::Inj pr
-- Show that the monoid structure of is the same as that of BinNat
+BIsInherited : (a b : BinNat) (prA : a canonical a) (prB : b canonical b) a +Binherit b a +B b
+BinheritLemma : (a : BinNat) (b : BinNat) (prA : a canonical a) (prB : b canonical b) incr (NToBinNat ((binNatToN a +N binNatToN b) +N ((binNatToN a +N binNatToN b) +N zero))) one :: (a +B b)
+BIsInherited' : (a b : BinNat) a +Binherit b canonical (a +B b)
+BinheritLemma a b prA prB with orderIsTotal 0 (binNatToN a +N binNatToN b)
+BinheritLemma a b prA prB | inl (inl x) rewrite doubleIsBitShift (binNatToN a +N binNatToN b) x = applyEquality (one ::_) (+BIsInherited a b prA prB)
+BinheritLemma a b prA prB | inr x with sumZeroImpliesOperandsZero (binNatToN a) (equalityCommutative x)
+BinheritLemma a b prA prB | inr x | fst ,, snd = ans2
where
bad : b []
bad = transitivity prB (binNatToNZero b snd)
bad2 : a []
bad2 = transitivity prA (binNatToNZero a fst)
ans2 : incr (NToBinNat ((binNatToN a +N binNatToN b) +N ((binNatToN a +N binNatToN b) +N zero))) one :: (a +B b)
ans2 rewrite bad | bad2 = refl
+BIsInherited [] b _ prB = +BIsInherited[] b prB
+BIsInherited (x :: a) [] prA _ rewrite +BCommutative (x :: a) [] | additionNIsCommutative (binNatToN (x :: a)) (binNatToN []) = +BIsInherited[] (x :: a) 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)))
+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
where
bad : (b : BinNat) (pr : b canonical b) (pr2 : binNatToN b 0) b []
bad b pr pr2 = transitivity pr (binNatToNZero b pr2)
t : b canonical b
t with canonical b
t | x :: bl = ::Inj prB
u : b []
u = bad b t b=0
nono : {A : Set} {a : A} {as : List A} a :: as [] False
nono ()
ans : False
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
--- Proofs
parity : (a b : ) succ (2 *N a) 2 *N b False
@@ -425,94 +338,6 @@ module Numbers.BinaryNaturals.Definition where
... | [] with pr = exFalso (incrNonzero xs pr)
... | (_ :: _) with pr rewrite pr = applyEquality (zero ::_) (transitivity (equalityCommutative pr) (incrPreservesCanonical' xs))
+BIsInherited'[] : (b : BinNat) [] +Binherit b canonical ([] +B b)
+BIsInherited'[] [] = refl
+BIsInherited'[] (zero :: b) with inspect (canonical b)
+BIsInherited'[] (zero :: b) | [] with pr rewrite binNatToNZero' b pr | pr = refl
+BIsInherited'[] (zero :: b) | (x :: bl) with pr rewrite pr = ans
where
contr : {a : _} {A : Set a} {l1 l2 : List A} {x : A} l1 [] l1 x :: l2 False
contr {l1 = []} p1 ()
contr {l1 = x :: l1} () p2
ans : NToBinNat (binNatToN b +N (binNatToN b +N zero)) zero :: x :: bl
ans with inspect (binNatToN b)
ans | zero with th rewrite th = exFalso (contr (binNatToNZero b th) pr)
ans | succ th with blah rewrite blah | doubleIsBitShift' th = applyEquality (zero ::_) (transitivity (equalityCommutative u) pr)
where
u : canonical b incr (NToBinNat th)
u = transitivity (equalityCommutative (binToBin b)) (applyEquality NToBinNat blah)
+BIsInherited'[] (one :: b) with inspect (binNatToN b)
... | zero with pr rewrite pr = applyEquality (one ::_) (equalityCommutative (binNatToNZero b pr))
... | (succ bl) with pr = ans
where
u : NToBinNat (2 *N binNatToN b) zero :: canonical b
u with doubleIsBitShift' bl
... | t = transitivity (identityOfIndiscernablesLeft _ _ _ _≡_ t (applyEquality (λ i NToBinNat (2 *N i)) (equalityCommutative pr))) (applyEquality (zero ::_) (transitivity (applyEquality NToBinNat (equalityCommutative pr)) (binToBin b)))
ans : incr (NToBinNat (binNatToN b +N (binNatToN b +N zero))) one :: canonical b
ans = applyEquality incr u
+BIsInherited' [] b = +BIsInherited'[] b
+BIsInherited' (x :: a) [] rewrite +BCommutative (x :: a) [] | additionNIsCommutative (binNatToN (x :: a)) 0 = binToBin (x :: a)
+BIsInherited' (zero :: as) (zero :: bs) rewrite equalityCommutative (productDistributes 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)
ans | zero with x with sumZeroImpliesOperandsZero (binNatToN as) x
... | as=0 ,, bs=0 rewrite as=0 | bs=0 = foo
where
u : canonical (as +Binherit bs) []
u rewrite as=0 | bs=0 = refl
foo : [] canonical (zero :: (as +B bs))
foo = transitivity (transitivity b (applyEquality (λ i canonical (zero :: i)) (+BIsInherited' as bs))) (canonicalAscends'' {zero} (as +B bs))
where
b : [] canonical (zero :: (as +Binherit bs))
b rewrite u = refl
ans | succ y with x rewrite x | doubleIsBitShift' y = transitivity (applyEquality (λ i zero :: NToBinNat i) (equalityCommutative x)) ans2
where
u : 0 <N binNatToN (as +B bs)
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
where
ans2 : incr (NToBinNat (2 *N (binNatToN as +N binNatToN bs))) one :: canonical (as +B bs)
ans2 with inspect (binNatToN as +N binNatToN bs)
ans2 | zero with x with sumZeroImpliesOperandsZero (binNatToN as) x
ans2 | zero with x | as=0 ,, bs=0 rewrite as=0 | bs=0 = applyEquality (one ::_) (transitivity t (+BIsInherited' as bs))
where
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
where
ans : incr (NToBinNat (2 *N (binNatToN as +N binNatToN bs))) one :: canonical (as +B bs)
ans with inspect (binNatToN as +N binNatToN bs)
ans | zero with x with sumZeroImpliesOperandsZero (binNatToN as) x
... | as=0 ,, bs=0 rewrite as=0 | bs=0 = applyEquality (one ::_) (transitivity t (+BIsInherited' as bs))
where
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
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)
... | zero with x with sumZeroImpliesOperandsZero (binNatToN as) x
ans | zero with x | as=0 ,, bs=0 rewrite as=0 | bs=0 = bar
where
u' : canonical (as +Binherit bs) []
u' rewrite as=0 | bs=0 = refl
u : canonical (as +B bs) []
u rewrite equalityCommutative (+BIsInherited' as bs) = transitivity (NToBinNatIsCanonical (binNatToN as +N binNatToN bs)) u'
t : canonical (incr (as +B bs)) one :: []
t rewrite incrPreservesCanonical' (as +B bs) | u = refl
bar : zero :: one :: [] canonical (zero :: incr (as +B bs))
bar rewrite t = refl
ans | succ y with x rewrite x | doubleIsBitShift' y = transitivity (applyEquality (λ i zero :: incr (NToBinNat i)) (equalityCommutative x)) ans2
where
ans2 : zero :: incr (as +Binherit bs) canonical (zero :: incr (as +B bs))
ans2 rewrite +BIsInherited' as bs | equalityCommutative (incrPreservesCanonical' (as +B bs)) | canonicalAscends' {zero} (incr (as +B bs)) (incrNonzero (as +B bs)) = refl
NToBinNatSucc zero = refl
NToBinNatSucc (succ n) with NToBinNat n
... | [] = refl

View File

@@ -6,6 +6,7 @@ open import Lists.Lists
open import Numbers.Naturals
open import Groups.GroupDefinition
open import Numbers.BinaryNaturals.Definition
open import Numbers.BinaryNaturals.Addition
module Numbers.BinaryNaturals.Multiplication where

View File

@@ -2,9 +2,10 @@
open import LogicalFormulae
open import Numbers.Naturals
open import Numbers.NaturalsWithK -- TODO: remove this dependency, it's baked into ZSimple
open import Groups.Groups
open import Groups.GroupDefinition
open import Rings.RingDefinition
open import Rings.Definition
open import Functions
open import Orders
open import Setoids.Setoids

View File

@@ -1,4 +1,4 @@
{-# OPTIONS --warning=error --safe #-}
{-# OPTIONS --warning=error --safe --without-K #-}
open import LogicalFormulae
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
@@ -810,14 +810,6 @@ module Numbers.Naturals where
aIsNotSuccA zero pr = naughtE pr
aIsNotSuccA (succ a) pr = aIsNotSuccA a (succInjective pr)
<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
where
p1=p2 : p1 p2
p1=p2 = equalityCommutative (canSubtractFromEqualityRight {p2} {a} {p1} (succInjective proof2))
help : (p1 p2) (pr2 : succ (p2 +N a) succ (p1 +N a)) le p1 refl le p2 pr2
help refl refl = refl
TotalOrder : TotalOrder
PartialOrder._<_ (TotalOrder.order TotalOrder) = _<N_
PartialOrder.irreflexive (TotalOrder.order TotalOrder) = lessIrreflexive

View File

@@ -0,0 +1,18 @@
{-# OPTIONS --warning=error --safe #-}
open import LogicalFormulae
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import WellFoundedInduction
open import Functions
open import Orders
open import Numbers.Naturals
module Numbers.NaturalsWithK 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
where
p1=p2 : p1 p2
p1=p2 = equalityCommutative (canSubtractFromEqualityRight {p2} {a} {p1} (succInjective proof2))
help : (p1 p2) (pr2 : succ (p2 +N a) succ (p1 +N a)) le p1 refl le p2 pr2
help refl refl = refl

View File

@@ -5,7 +5,7 @@ open import Numbers.Naturals
open import Numbers.Integers
open import Groups.Groups
open import Groups.GroupDefinition
open import Rings.RingDefinition
open import Rings.Definition
open import Fields.Fields
open import PrimeNumbers
open import Setoids.Setoids

View File

@@ -7,7 +7,7 @@ open import Numbers.Rationals
open import Groups.Groups
open import Groups.GroupsLemmas
open import Groups.GroupDefinition
open import Rings.RingDefinition
open import Rings.Definition
open import Fields.Fields
open import PrimeNumbers
open import Setoids.Setoids
@@ -16,7 +16,7 @@ open import Functions
open import Fields.FieldOfFractions
open import Fields.FieldOfFractionsOrder
open import Rings.IntegralDomains
open import Rings.RingLemmas
open import Rings.Lemmas
module Numbers.RationalsLemmas where

View File

@@ -1,4 +1,4 @@
{-# OPTIONS --safe --warning=error #-}
{-# OPTIONS --safe --warning=error --without-K #-}
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)

View File

@@ -2,12 +2,14 @@
open import LogicalFormulae
open import Numbers.Naturals
open import Numbers.NaturalsWithK
open import WellFoundedInduction
open import KeyValue
open import KeyValueWithDomain
open import Orders
open import Vectors
open import Maybe
open import WithK
module PrimeNumbers where

View File

@@ -1,4 +1,4 @@
{-# OPTIONS --safe --warning=error #-}
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Groups.Groups
@@ -11,7 +11,7 @@ open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
-- Following Part IB's course Groups, Rings, and Modules, we take rings to be commutative with one.
module Rings.RingDefinition where
module Rings.Definition where
record Ring {n m} {A : Set n} (S : Setoid {n} {m} A) (_+_ : A A A) (_*_ : A A A) : Set (lsuc n m) where
field
additiveGroup : Group S _+_

View File

@@ -6,10 +6,10 @@ open import Functions
open import Numbers.Naturals
open import Numbers.Integers
open import IntegersModN
open import Rings.RingExamplesProofs
open import Rings.Examples.Proofs
open import PrimeNumbers
module Rings.RingExamples where
module Rings.Examples.Examples where
nToZn : (n : ) (pr : 0 <N n) (x : ) n n pr
nToZn n pr x = nToZn' n pr x

View File

@@ -5,13 +5,13 @@ open import Functions
open import Groups.Groups
open import Groups.GroupDefinition
open import Orders
open import Rings.RingDefinition
open import Rings.Definition
open import Numbers.Naturals
open import Numbers.Integers
open import PrimeNumbers
open import IntegersModN
module Rings.RingExamplesProofs where
module Rings.Examples.Proofs where
nToZn' : (n : ) (pr : 0 <N n) (x : ) n n pr
nToZn' 0 ()
nToZn' (succ n) pr x with divisionAlg (succ n) x

View File

@@ -1,4 +1,4 @@
{-# OPTIONS --safe --warning=error #-}
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Groups.Groups
@@ -7,8 +7,8 @@ open import Numbers.Naturals
open import Orders
open import Setoids.Setoids
open import Functions
open import Rings.RingDefinition
open import Rings.RingLemmas
open import Rings.Definition
open import Rings.Lemmas
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)

View File

@@ -1,15 +1,15 @@
{-# OPTIONS --safe --warning=error #-}
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Functions
open import Groups.Groups
open import Groups.GroupDefinition
open import Groups.GroupsLemmas
open import Rings.RingDefinition
open import Rings.Definition
open import Setoids.Setoids
open import Setoids.Orders
module Rings.RingLemmas where
module Rings.Lemmas where
ringTimesZero : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} {_*_ : A A A} (R : Ring S _+_ _*_) {x : A} Setoid.__ S (x * (Ring.0R R)) (Ring.0R R)
ringTimesZero {S = S} {_+_ = _+_} {_*_ = _*_} R {x} = symmetric (transitive blah'' (transitive (Group.multAssoc additiveGroup) (transitive (wellDefined invLeft reflexive) multIdentLeft)))
where

View File

@@ -1,8 +1,8 @@
{-# OPTIONS --safe --warning=error #-}
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Numbers.Naturals
open import Lists
open import Lists.Lists
open import Setoids.Setoids
open import Functions
@@ -31,7 +31,7 @@ module Setoids.Lists where
listEqualityTransitive S {w1 = x :: w1} {y :: ys} {[]} w1=w2 ()
listEqualityTransitive S {w1 = x :: w1} {y :: w2} {z :: w3} (pr1 ,, pr2) (pr3 ,, pr4) = Transitive.transitive (Equivalence.transitiveEq (Setoid.eq S)) pr1 pr3 ,, listEqualityTransitive S pr2 pr4
listEqualityRespectsMap : {a b c d : _} {A : Set a} {B : Set b} (S : Setoid {a} {c} A) (T : Setoid {b} {d} B) (f : A B) (fWD : {x y : A} Setoid.__ S x y Setoid.__ T (f x) (f y)) {w1 w2 : List A} (w1=w2 : listEquality S w1 w2) listEquality T (listMap f w1) (listMap f w2)
listEqualityRespectsMap : {a b c d : _} {A : Set a} {B : Set b} (S : Setoid {a} {c} A) (T : Setoid {b} {d} B) (f : A B) (fWD : {x y : A} Setoid.__ S x y Setoid.__ T (f x) (f y)) {w1 w2 : List A} (w1=w2 : listEquality S w1 w2) listEquality T (map f w1) (map f w2)
listEqualityRespectsMap S T f fWD {[]} {[]} w1=w2 = record {}
listEqualityRespectsMap S T f fWD {[]} {x :: w2} ()
listEqualityRespectsMap S T f fWD {x :: w1} {[]} ()

View File

@@ -1,4 +1,4 @@
{-# OPTIONS --safe --warning=error #-}
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Orders

View File

@@ -1,4 +1,4 @@
{-# OPTIONS --safe --warning=error #-}
{-# OPTIONS --safe --warning=error --without-K #-}
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)

View File

@@ -1,13 +1,13 @@
{-# OPTIONS --safe --warning=error #-}
{-# OPTIONS --safe --warning=error --without-K #-}
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import LogicalFormulae
open import Functions
open import Naturals
open import FinSet
open import Numbers.Naturals
open import Sets.FinSet
module Cardinality where
module Sets.Cardinality where
record CountablyInfiniteSet {a : _} (A : Set a) : Set a where
field
counting : A
@@ -227,5 +227,5 @@ module Cardinality where
-}
injectionToNMeansCountable : {a : _} {A : Set a} {f : A } Injection f InfiniteSet A Countable A
injectionToNMeansCountable {f = f} inj record { isInfinite = isInfinite } = {!!}
-- injectionToNMeansCountable : {a : _} {A : Set a} {f : A → } → Injection f → InfiniteSet A → Countable A
-- injectionToNMeansCountable {f = f} inj record { isInfinite = isInfinite } = {!!}

View File

@@ -1,4 +1,4 @@
{-# OPTIONS --safe --warning=error #-}
{-# OPTIONS --safe --warning=error --without-K #-}
open import Functions
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
@@ -32,73 +32,6 @@ module Sets.FinSet where
sgEq {l} {m} {L} {prop} {(a , b1)} {(.a , b)} refl pr2 with pr2 {a} b b1
sgEq {l} {m} {L} {prop} {(a , b1)} {(.a , .b1)} refl pr2 | refl = refl
finNotEqualsRefl : {n : } {a b : FinSet (succ n)} (p1 p2 : FinNotEquals a b) p1 p2
finNotEqualsRefl {.1} {.fzero} {.(fsucc fzero)} (fne2 .fzero .(fsucc fzero) (inl (refl ,, refl))) (fne2 .fzero .(fsucc fzero) (inl (refl ,, refl))) = refl
finNotEqualsRefl {.1} {.fzero} {.(fsucc fzero)} (fne2 .fzero .(fsucc fzero) (inl (refl ,, refl))) (fne2 .fzero .(fsucc fzero) (inr (() ,, snd)))
finNotEqualsRefl {.1} {.(fsucc fzero)} {.fzero} (fne2 .(fsucc fzero) .fzero (inr (refl ,, refl))) (fne2 .(fsucc fzero) .fzero (inl (() ,, snd)))
finNotEqualsRefl {.1} {.(fsucc fzero)} {.fzero} (fne2 .(fsucc fzero) .fzero (inr (refl ,, refl))) (fne2 .(fsucc fzero) .fzero (inr (refl ,, refl))) = refl
finNotEqualsRefl {.(succ (succ _))} {.fzero} {.(fsucc a)} (fneN .fzero .(fsucc a) (inl (inl (refl ,, (a , refl))))) (fneN .fzero .(fsucc a) (inl (inl (refl ,, (.a , refl))))) = refl
finNotEqualsRefl {.(succ (succ _))} {.fzero} {.(fsucc a)} (fneN .fzero .(fsucc a) (inl (inl (refl ,, (a , refl))))) (fneN .fzero .(fsucc a) (inl (inr ((a₁ , ()) ,, snd))))
finNotEqualsRefl {.(succ (succ _))} {.fzero} {.(fsucc a)} (fneN .fzero .(fsucc a) (inl (inl (refl ,, (a , refl))))) (fneN .fzero .(fsucc a) (inr ((fst ,, snd) , b))) = exFalso q
where
p : fzero fsucc fst
p = _&_&_.one b
q : False
q with p
... | ()
finNotEqualsRefl {.(succ (succ _))} {.(fsucc a)} {.fzero} (fneN .(fsucc a) .fzero (inl (inr ((a , refl) ,, refl)))) (fneN .(fsucc a) .fzero (inl (inl (() ,, snd))))
finNotEqualsRefl {.(succ (succ _))} {.(fsucc a)} {.fzero} (fneN .(fsucc a) .fzero (inl (inr ((a , refl) ,, refl)))) (fneN .(fsucc a) .fzero (inl (inr ((.a , refl) ,, refl)))) = refl
finNotEqualsRefl {.(succ (succ _))} {.(fsucc a)} {.fzero} (fneN .(fsucc a) .fzero (inl (inr ((a , refl) ,, refl)))) (fneN .(fsucc a) .fzero (inr ((fst ,, snd) , b))) = exFalso q
where
p : fzero fsucc snd
p = _&_&_.two b
q : False
q with p
... | ()
finNotEqualsRefl {.(succ (succ _))} {a} {b} (fneN a b (inr (record { fst = fst ; snd = snd₁ } , snd))) (fneN .a .b (inl (inl (fst1 ,, snd₂)))) = exFalso q
where
p : fzero fsucc fst
p = transitivity (equalityCommutative fst1) (_&_&_.one snd)
q : False
q with p
... | ()
finNotEqualsRefl {.(succ (succ _))} {a} {b} (fneN a b (inr (record { fst = fst ; snd = snd1 } , snd))) (fneN .a .b (inl (inr (fst₁ ,, snd2)))) = exFalso q
where
p : fzero fsucc snd1
p = transitivity (equalityCommutative snd2) (_&_&_.two snd)
q : False
q with p
... | ()
finNotEqualsRefl {.(succ (succ _))} {.fzero} {b} (fneN fzero b (inr (record { fst = fst ; snd = snd₁ } , snd))) (fneN .fzero .b (inr ((fst1 ,, snd₂) , b1))) = exFalso q
where
p : fzero fsucc fst1
p = _&_&_.one b1
q : False
q with p
... | ()
finNotEqualsRefl {.(succ (succ _))} {.(fsucc a)} {.fzero} (fneN (fsucc a) fzero (inr (record { fst = fst ; snd = snd₁ } , snd))) (fneN .(fsucc a) .fzero (inr ((fst₁ ,, snd2) , b1))) = exFalso q
where
p : fzero fsucc snd2
p = _&_&_.two b1
q : False
q with p
... | ()
finNotEqualsRefl {.(succ (succ _))} {.(fsucc a)} {.(fsucc b)} (fneN (fsucc a) (fsucc b) (inr (record { fst = fst ; snd = snd1 } , snd))) (fneN .(fsucc a) .(fsucc b) (inr ((fst1 ,, snd2) , b1))) = ans
where
t : a fst1
t = fsuccInjective (_&_&_.one b1)
t' : a fst
t' = fsuccInjective (_&_&_.one snd)
u : b snd1
u = fsuccInjective (_&_&_.two snd)
u' : b snd2
u' = fsuccInjective (_&_&_.two b1)
equality : {c : FinSet (succ (succ _)) && FinSet (succ (succ _))} (r1 s : (fsucc a fsucc (_&&_.fst c)) & (fsucc b fsucc (_&&_.snd c)) & FinNotEquals (_&&_.fst c) (_&&_.snd c)) r1 s
equality record { one = refl ; two = refl ; three = q } record { one = refl ; two = refl ; three = q' } = applyEquality (λ t record { one = refl ; two = refl ; three = t }) (finNotEqualsRefl q q')
r : (fst ,, snd1) (fst1 ,, snd2)
r rewrite equalityCommutative t | equalityCommutative t' | equalityCommutative u | equalityCommutative u' = refl
ans : fneN (fsucc a) (fsucc b) (inr ((fst ,, snd1) , snd)) fneN (fsucc a) (fsucc b) (inr ((fst1 ,, snd2) , b1))
ans = applyEquality (λ t fneN (fsucc a) (fsucc b) (inr t)) (sgEq r equality)
finSetNotEquals : {n : } {a b : FinSet (succ n)} (a b False) FinNotEquals {n} a b
finSetNotEquals {zero} {fzero} {fzero} pr = exFalso (pr refl)
@@ -228,14 +161,6 @@ module Sets.FinSet where
toNatOfNat {succ n} zero a<n = refl
toNatOfNat {succ n} (succ a) a<n = applyEquality succ (toNatOfNat a (canRemoveSuccFrom<N a<n))
ofNatToNat : {n : } (a : FinSet n) ofNat (toNat a) (toNatSmaller a) a
ofNatToNat {zero} ()
ofNatToNat {succ n} fzero = refl
ofNatToNat {succ n} (fsucc a) = applyEquality fsucc indHyp
where
indHyp : ofNat (toNat a) (canRemoveSuccFrom<N (succPreservesInequality (toNatSmaller a))) a
indHyp rewrite <NRefl (canRemoveSuccFrom<N (succPreservesInequality (toNatSmaller a))) (toNatSmaller a) = ofNatToNat a
intoSmallerLemm : {n m : } {n<m : n <N (succ m)} {m<sm : m <N succ m} (b : FinSet n) intoSmaller n<m b ofNat m m<sm False
intoSmallerLemm {.(succ _)} {m} {n<m} fzero pr with intoSmaller (canRemoveSuccFrom<N n<m)
intoSmallerLemm {.(succ _)} {zero} {n<m} fzero refl | bl = zeroNeverGreater (canRemoveSuccFrom<N n<m)

84
Sets/FinSetWithK.agda Normal file
View File

@@ -0,0 +1,84 @@
{-# OPTIONS --safe --warning=error #-}
open import Numbers.Naturals
open import Sets.FinSet
open import LogicalFormulae
open import Numbers.NaturalsWithK
module Sets.FinSetWithK where
finNotEqualsRefl : {n : } {a b : FinSet (succ n)} (p1 p2 : FinNotEquals a b) p1 p2
finNotEqualsRefl {.1} {.fzero} {.(fsucc fzero)} (fne2 .fzero .(fsucc fzero) (inl (refl ,, refl))) (fne2 .fzero .(fsucc fzero) (inl (refl ,, refl))) = refl
finNotEqualsRefl {.1} {.fzero} {.(fsucc fzero)} (fne2 .fzero .(fsucc fzero) (inl (refl ,, refl))) (fne2 .fzero .(fsucc fzero) (inr (() ,, snd)))
finNotEqualsRefl {.1} {.(fsucc fzero)} {.fzero} (fne2 .(fsucc fzero) .fzero (inr (refl ,, refl))) (fne2 .(fsucc fzero) .fzero (inl (() ,, snd)))
finNotEqualsRefl {.1} {.(fsucc fzero)} {.fzero} (fne2 .(fsucc fzero) .fzero (inr (refl ,, refl))) (fne2 .(fsucc fzero) .fzero (inr (refl ,, refl))) = refl
finNotEqualsRefl {.(succ (succ _))} {.fzero} {.(fsucc a)} (fneN .fzero .(fsucc a) (inl (inl (refl ,, (a , refl))))) (fneN .fzero .(fsucc a) (inl (inl (refl ,, (.a , refl))))) = refl
finNotEqualsRefl {.(succ (succ _))} {.fzero} {.(fsucc a)} (fneN .fzero .(fsucc a) (inl (inl (refl ,, (a , refl))))) (fneN .fzero .(fsucc a) (inl (inr ((a₁ , ()) ,, snd))))
finNotEqualsRefl {.(succ (succ _))} {.fzero} {.(fsucc a)} (fneN .fzero .(fsucc a) (inl (inl (refl ,, (a , refl))))) (fneN .fzero .(fsucc a) (inr ((fst ,, snd) , b))) = exFalso q
where
p : fzero fsucc fst
p = _&_&_.one b
q : False
q with p
... | ()
finNotEqualsRefl {.(succ (succ _))} {.(fsucc a)} {.fzero} (fneN .(fsucc a) .fzero (inl (inr ((a , refl) ,, refl)))) (fneN .(fsucc a) .fzero (inl (inl (() ,, snd))))
finNotEqualsRefl {.(succ (succ _))} {.(fsucc a)} {.fzero} (fneN .(fsucc a) .fzero (inl (inr ((a , refl) ,, refl)))) (fneN .(fsucc a) .fzero (inl (inr ((.a , refl) ,, refl)))) = refl
finNotEqualsRefl {.(succ (succ _))} {.(fsucc a)} {.fzero} (fneN .(fsucc a) .fzero (inl (inr ((a , refl) ,, refl)))) (fneN .(fsucc a) .fzero (inr ((fst ,, snd) , b))) = exFalso q
where
p : fzero fsucc snd
p = _&_&_.two b
q : False
q with p
... | ()
finNotEqualsRefl {.(succ (succ _))} {a} {b} (fneN a b (inr (record { fst = fst ; snd = snd₁ } , snd))) (fneN .a .b (inl (inl (fst1 ,, snd₂)))) = exFalso q
where
p : fzero fsucc fst
p = transitivity (equalityCommutative fst1) (_&_&_.one snd)
q : False
q with p
... | ()
finNotEqualsRefl {.(succ (succ _))} {a} {b} (fneN a b (inr (record { fst = fst ; snd = snd1 } , snd))) (fneN .a .b (inl (inr (fst₁ ,, snd2)))) = exFalso q
where
p : fzero fsucc snd1
p = transitivity (equalityCommutative snd2) (_&_&_.two snd)
q : False
q with p
... | ()
finNotEqualsRefl {.(succ (succ _))} {.fzero} {b} (fneN fzero b (inr (record { fst = fst ; snd = snd₁ } , snd))) (fneN .fzero .b (inr ((fst1 ,, snd₂) , b1))) = exFalso q
where
p : fzero fsucc fst1
p = _&_&_.one b1
q : False
q with p
... | ()
finNotEqualsRefl {.(succ (succ _))} {.(fsucc a)} {.fzero} (fneN (fsucc a) fzero (inr (record { fst = fst ; snd = snd₁ } , snd))) (fneN .(fsucc a) .fzero (inr ((fst₁ ,, snd2) , b1))) = exFalso q
where
p : fzero fsucc snd2
p = _&_&_.two b1
q : False
q with p
... | ()
finNotEqualsRefl {.(succ (succ _))} {.(fsucc a)} {.(fsucc b)} (fneN (fsucc a) (fsucc b) (inr (record { fst = fst ; snd = snd1 } , snd))) (fneN .(fsucc a) .(fsucc b) (inr ((fst1 ,, snd2) , b1))) = ans
where
t : a fst1
t = fsuccInjective (_&_&_.one b1)
t' : a fst
t' = fsuccInjective (_&_&_.one snd)
u : b snd1
u = fsuccInjective (_&_&_.two snd)
u' : b snd2
u' = fsuccInjective (_&_&_.two b1)
equality : {c : FinSet (succ (succ _)) && FinSet (succ (succ _))} (r1 s : (fsucc a fsucc (_&&_.fst c)) & (fsucc b fsucc (_&&_.snd c)) & FinNotEquals (_&&_.fst c) (_&&_.snd c)) r1 s
equality record { one = refl ; two = refl ; three = q } record { one = refl ; two = refl ; three = q' } = applyEquality (λ t record { one = refl ; two = refl ; three = t }) (finNotEqualsRefl q q')
r : (fst ,, snd1) (fst1 ,, snd2)
r rewrite equalityCommutative t | equalityCommutative t' | equalityCommutative u | equalityCommutative u' = refl
ans : fneN (fsucc a) (fsucc b) (inr ((fst ,, snd1) , snd)) fneN (fsucc a) (fsucc b) (inr ((fst1 ,, snd2) , b1))
ans = applyEquality (λ t fneN (fsucc a) (fsucc b) (inr t)) (sgEq r equality)
ofNatToNat : {n : } (a : FinSet n) ofNat (toNat a) (toNatSmaller a) a
ofNatToNat {zero} ()
ofNatToNat {succ n} fzero = refl
ofNatToNat {succ n} (fsucc a) = applyEquality fsucc indHyp
where
indHyp : ofNat (toNat a) (canRemoveSuccFrom<N (succPreservesInequality (toNatSmaller a))) a
indHyp rewrite <NRefl (canRemoveSuccFrom<N (succPreservesInequality (toNatSmaller a))) (toNatSmaller a) = ofNatToNat a

View File

@@ -2,6 +2,7 @@
open import LogicalFormulae
open import Numbers.Naturals
open import Numbers.NaturalsWithK
open import Functions
data Vec {a : _} (X : Set a) : -> Set a where
@@ -251,25 +252,3 @@ o' th o>> o' th' = o' (o' th o>> th')
eqHelper : {a : _} {X : Set a} (x : X) {n : } {r s : Vec X n} (pr : r s) (x ,- r) (x ,- s)
eqHelper x pr rewrite pr = refl
{-
cp-<?= : {p n m : } (th : p <= n) (th' : n <= m) → {a : _} {X : Set a} (xs : Vec X m) → ((th o>> th') <?= xs) ≡ (th <?= (th' <?= xs))
cp-<?= oz oz [] = refl
cp-<?= oz (o' th') xs = refl
cp-<?= (os th) (os th') (x ,- xs) = eqHelper x (cp-<?= th th' xs)
cp-<?= (o' th) (os th') (x ,- xs) = cp-<?= th th' xs
cp-<?= (o' th) (o' th') (x ,- xs) = cp-<?= (o' th) th' xs
cp-<?= (os th) (o' th') (x ,- xs) = {!!}
idThen-o>> : {n m : } (th : n <= m) → (oi o>> th) ≡ th
idThen-o>> oz = refl
idThen-o>> (os th) = applyEquality os (idThen-o>> th)
idThen-o>> {zero} (o' th) = refl
idThen-o>> {succ n} (o' th) with idThen-o>> th
... | bl = {!!}
idAfter-o>> : {n m : } (th : n <= m) → (th o>> oi) ≡ th
idAfter-o>> th = {!!}
-}

View File

@@ -1,4 +1,4 @@
{-# OPTIONS --safe --warning=error #-}
{-# OPTIONS --safe --warning=error --without-K #-}
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Functions

11
WithK.agda Normal file
View File

@@ -0,0 +1,11 @@
{-# OPTIONS --safe --warning=error #-}
open import LogicalFormulae
module WithK where
typeCastEqual : {a : _} {A : Set a} {B : Set a} {el : A} (pr1 pr2 : A B) (typeCast el pr1) (typeCast el pr2)
typeCastEqual {a} {A} {.A} {el} refl refl = refl
reflRefl : {a : _} {A : Set a} {r s : A} (pr1 pr2 : r s) pr1 pr2
reflRefl refl refl = refl