├── .gitignore ├── Algebra.agda ├── BankersQueue.agda ├── Bool.agda ├── Equality.agda ├── Fin.agda ├── FingerTree.agda ├── Function.agda ├── FunctorApplicativeMonad.agda ├── Homotopy.agda ├── Int.agda ├── List.agda ├── Logic.agda ├── Maybe.agda ├── NQueens.agda ├── Nat.agda ├── Sign.agda ├── Stream.agda ├── TalkDemo.agda ├── Termination.agda ├── Test.agda ├── TypeOf.agda └── Vector.agda /.gitignore: -------------------------------------------------------------------------------- 1 | *.agdai 2 | 3 | -------------------------------------------------------------------------------- /Algebra.agda: -------------------------------------------------------------------------------- 1 | module Algebra where 2 | 3 | open import Equality 4 | open Equality.≡-Reasoning 5 | open import Agda.Primitive 6 | 7 | Associative : ∀ {α} {A : Set α} (_·_ : A → A → A) -> Set α 8 | Associative _·_ = ∀ x y z → ((x · y) · z) ≡ (x · (y · z)) 9 | 10 | record Semigroup {α : _} {A : Set α} (_·_ : A → A → A) : Set α where 11 | field 12 | assoc : Associative _·_ 13 | 14 | RightIdentity : ∀ {α} {A : Set α} (_·_ : A → A → A) (ε : A) -> Set α 15 | RightIdentity _·_ ε = ∀ x → (x · ε) ≡ x 16 | 17 | LeftIdentity : ∀ {α} {A : Set α} (_·_ : A → A → A) (ε : A) -> Set α 18 | LeftIdentity _·_ ε = ∀ x → (ε · x) ≡ x 19 | 20 | record Identity {α} {A : Set α} (_·_ : A → A → A) (ε : A) : Set α where 21 | field 22 | left-id : LeftIdentity _·_ ε 23 | right-id : RightIdentity _·_ ε 24 | 25 | record Monoid {α} {A : Set α} (_·_ : A → A → A) (ε : A) : Set α where 26 | field 27 | isSemigroup : Semigroup _·_ 28 | identity : Identity _·_ ε 29 | open Semigroup isSemigroup public 30 | open Identity identity public 31 | 32 | Commutative : ∀ {α} {A : Set α} (_·_ : A → A → A) -> Set α 33 | Commutative _·_ = ∀ x y → (x · y) ≡ (y · x) 34 | 35 | record CommutativeMonoid {α} {A : Set α} (_·_ : A → A → A) (ε : A) : Set α where 36 | field 37 | isMonoid : Monoid _·_ ε 38 | comm : Commutative _·_ 39 | open Monoid isMonoid public 40 | 41 | LeftInverse : ∀ {α} {A : Set α} (_·_ : A → A → A) (ε : A) (_⁻¹ : A → A) -> Set α 42 | LeftInverse _·_ ε _⁻¹ = ∀ a → ((a ⁻¹) · a) ≡ ε 43 | 44 | RightInverse : ∀ {α} {A : Set α} (_·_ : A → A → A) (ε : A) (_⁻¹ : A → A) -> Set α 45 | RightInverse _·_ ε _⁻¹ = ∀ a → (a · (a ⁻¹)) ≡ ε 46 | 47 | record Group {α} {A : Set α} (_·_ : A → A → A) (ε : A) (_⁻¹ : A → A) : Set α where 48 | field 49 | isMonoid : Monoid _·_ ε 50 | inverse-l : LeftInverse _·_ ε _⁻¹ 51 | inverse-r : RightInverse _·_ ε _⁻¹ 52 | open Monoid isMonoid public 53 | 54 | Involutive : ∀ {α} {A : Set α} (f : A → A) -> Set α 55 | Involutive f = ∀ x → f (f x) ≡ x 56 | 57 | [x⁻¹]⁻¹≡x 58 | : ∀ {α} {A : Set α} 59 | → (_·_ : A → A → A) (ε : A) (_⁻¹ : A → A) (g : Group _·_ ε _⁻¹) 60 | → Involutive _⁻¹ 61 | [x⁻¹]⁻¹≡x _·_ ε _⁻¹ g x = 62 | begin 63 | (x ⁻¹) ⁻¹ 64 | ≡⟨ symm (Group.right-id g _) ⟩ 65 | ((x ⁻¹) ⁻¹) · ε 66 | ≡⟨ cong (λ e → _ · e) (symm (Group.inverse-l g x)) ⟩ 67 | ((x ⁻¹) ⁻¹) · ((x ⁻¹) · x) 68 | ≡⟨ symm (Group.assoc g _ _ _) ⟩ 69 | (((x ⁻¹) ⁻¹) · (x ⁻¹)) · x 70 | ≡⟨ cong (λ e → e · _) (Group.inverse-l g _) ⟩ 71 | ε · x 72 | ≡⟨ Group.left-id g x ⟩ 73 | x 74 | qed 75 | 76 | record AbelianGroup {α} {A : Set α} (_·_ : A → A → A) (ε : A) (_⁻¹ : A → A) : Set α where 77 | field 78 | isGroup : Group _·_ ε _⁻¹ 79 | comm : Commutative _·_ 80 | open Group isGroup public 81 | 82 | abelianToCMonoid 83 | : ∀ {α} {A : Set α} {_·_} {ε : A} {_⁻¹} 84 | → AbelianGroup _·_ ε _⁻¹ 85 | → CommutativeMonoid _·_ ε 86 | abelianToCMonoid g = record 87 | { isMonoid = AbelianGroup.isMonoid g 88 | ; comm = AbelianGroup.comm g } 89 | 90 | Rel : ∀ {α} (A : Set α) → (β : Level) → Set (α ⊔ lsuc β) 91 | Rel A β = A → A → Set β 92 | 93 | Transitive : ∀ {α} {A : Set α} → Rel A α → Set α 94 | Transitive _~_ = ∀ {a b c} → a ~ b → b ~ c → a ~ c 95 | 96 | Symmetric : ∀ {α} {A : Set α} → Rel A α → Set α 97 | Symmetric _~_ = ∀ {x y} → x ~ y → y ~ x 98 | 99 | Reflexive : ∀ {α} {A : Set α} → Rel A α → Set α 100 | Reflexive _~_ = ∀ {a} → a ~ a 101 | 102 | record Equivalence {α} {A : Set α} (_~_ : Rel A α) : Set α where 103 | field 104 | isTransitive : Transitive _~_ 105 | isSymmetric : Symmetric _~_ 106 | isReflexive : Reflexive _~_ 107 | -------------------------------------------------------------------------------- /BankersQueue.agda: -------------------------------------------------------------------------------- 1 | module BankersQueue where 2 | 3 | open import List 4 | open import Logic 5 | 6 | data BankersQueue {α} (A : Set α) : Set α where 7 | BQ : (front back : List A) → BankersQueue A 8 | 9 | pushFront : ∀ {α} {A : Set α} → A → BankersQueue A → BankersQueue A 10 | pushFront x (BQ front back) = BQ (x ∷ front) back 11 | 12 | popBack : ∀ {α} {A : Set α} → BankersQueue A → A ∧ BankersQueue A ∨ ⊤ 13 | popBack (BQ [] []) = inr tt 14 | popBack (BQ (f ∷ ront) []) with reverse ront 15 | popBack (BQ (f ∷ _) []) | [] = inl (f , BQ [] []) 16 | popBack (BQ (f ∷ _) []) | t ∷ nor = inl (t , BQ [] nor) 17 | popBack (BQ front (x ∷ back)) = inl (x , BQ back back) 18 | -------------------------------------------------------------------------------- /Bool.agda: -------------------------------------------------------------------------------- 1 | module Bool where 2 | 3 | 4 | 5 | open import Function 6 | open import Logic 7 | open import Algebra 8 | open import Equality 9 | 10 | open Logic.Decidable.Binary 11 | 12 | 13 | 14 | data Bool : Set where 15 | true : Bool 16 | false : Bool 17 | {-# BUILTIN BOOL Bool #-} 18 | {-# BUILTIN TRUE true #-} 19 | {-# BUILTIN FALSE false #-} 20 | 21 | private 22 | uniqueness-bool : ∀ x → (x ≡ true) ∨ (x ≡ false) 23 | uniqueness-bool true = inl refl 24 | uniqueness-bool false = inr refl 25 | 26 | not : Bool → Bool 27 | not true = false 28 | not false = true 29 | 30 | not²≡id : ∀ {P} → not (not P) ≡ P 31 | not²≡id {true} = refl 32 | not²≡id {false} = refl 33 | 34 | T : Bool → Set 35 | T true = ⊤ 36 | T false = ⊥ 37 | 38 | ⌊_⌋ : ∀ {α} {P : Set α} → Dec P → Bool 39 | ⌊ yes _ ⌋ = true 40 | ⌊ no _ ⌋ = false 41 | 42 | ¬Dec : ∀ {α} {P : Set α} → Dec P → Dec (¬ P) 43 | ¬Dec (yes p) = no (λ ¬p → ¬p p) 44 | ¬Dec (no ¬p) = yes ¬p 45 | 46 | -- Inhabited iff the proposition is true 47 | True : ∀ {α} {P : Set α} → Dec P → Set 48 | True p = T ⌊ p ⌋ 49 | 50 | -- Inhabited iff the proposition is false 51 | False : ∀ {α} {P : Set α} → Dec P → Set 52 | False ¬p = T (not ⌊ ¬p ⌋) 53 | 54 | data So : Bool → Set where 55 | Oh : So true 56 | 57 | _≟_ : Decidable {A = Bool} _≡_ 58 | true ≟ true = yes refl 59 | true ≟ false = no (λ ()) 60 | false ≟ true = no (λ ()) 61 | false ≟ false = yes refl 62 | 63 | ind-Bool : ∀ {α} {C : Bool → Set α} → C true → C false → (x : Bool) → C x 64 | ind-Bool t _ true = t 65 | ind-Bool _ f false = f 66 | 67 | rec-Bool : ∀ {α} {C : Set α} → C → C → Bool → C 68 | rec-Bool {C = C} = ind-Bool {C = const C} 69 | 70 | private 71 | -- This is how I thought this had to be written 72 | rec-via-ind : ∀ {α} {C : Set α} → C → C → Bool → C 73 | rec-via-ind {C = C} = ind-Bool {C = λ _ → C} 74 | 75 | -- This works as well, but I don’t understand how the inference works. C 76 | -- isn’t equivalent to λ _ → C, after all. 77 | rec-via-ind₂ : ∀ {α} {C : Set α} → C → C → Bool → C 78 | rec-via-ind₂ = ind-Bool 79 | 80 | _||_ : Bool → Bool → Bool 81 | true || _ = true 82 | false || x = x 83 | 84 | commutative-monoid-|| : CommutativeMonoid _||_ false 85 | commutative-monoid-|| = record 86 | { isMonoid = record 87 | { isSemigroup = record { assoc = assoc } 88 | ; identity = record { left-id = left-id ; right-id = right-id } } 89 | ; comm = comm } 90 | where 91 | comm : Commutative _||_ 92 | comm true true = refl 93 | comm true false = refl 94 | comm false true = refl 95 | comm false false = refl 96 | 97 | left-id : LeftIdentity _||_ false 98 | left-id _ = refl 99 | 100 | right-id : RightIdentity _||_ false 101 | right-id x rewrite comm x false = refl 102 | 103 | assoc : Associative _||_ 104 | assoc true y z = refl 105 | assoc false y z = refl 106 | 107 | _&&_ : Bool → Bool → Bool 108 | true && x = x 109 | false && _ = false 110 | 111 | commutative-monoid-&& : CommutativeMonoid _&&_ true 112 | commutative-monoid-&& = record 113 | { isMonoid = record 114 | { isSemigroup = record { assoc = assoc } 115 | ; identity = record { left-id = left-id ; right-id = right-id } } 116 | ; comm = comm } 117 | where 118 | comm : Commutative _&&_ 119 | comm true true = refl 120 | comm true false = refl 121 | comm false true = refl 122 | comm false false = refl 123 | 124 | left-id : LeftIdentity _&&_ true 125 | left-id _ = refl 126 | 127 | right-id : RightIdentity _&&_ true 128 | right-id x rewrite comm x true = refl 129 | 130 | assoc : Associative _&&_ 131 | assoc true y z = refl 132 | assoc false y z = refl 133 | 134 | _⊕_ : Bool → Bool → Bool 135 | true ⊕ true = false 136 | true ⊕ false = true 137 | false ⊕ true = true 138 | false ⊕ false = false 139 | 140 | group-⊕ : AbelianGroup _⊕_ false (λ x → false ⊕ x) 141 | group-⊕ = record 142 | { isGroup = record 143 | { isMonoid = record 144 | { isSemigroup = record 145 | { assoc = assoc } 146 | ; identity = record 147 | { left-id = left-id 148 | ; right-id = right-id } } 149 | ; inverse-l = inverse-l ; inverse-r = inverse-r } 150 | ; comm = comm } 151 | where 152 | right-id : RightIdentity _⊕_ false 153 | right-id true = refl 154 | right-id false = refl 155 | 156 | left-id : LeftIdentity _⊕_ false 157 | left-id true = refl 158 | left-id false = refl 159 | 160 | comm : Commutative _⊕_ 161 | comm true true = refl 162 | comm true false = refl 163 | comm false true = refl 164 | comm false false = refl 165 | 166 | -- Urgh. 167 | assoc : Associative _⊕_ 168 | assoc true true true = refl 169 | assoc true true false = refl 170 | assoc true false true = refl 171 | assoc true false false = refl 172 | assoc false true true = refl 173 | assoc false true false = refl 174 | assoc false false true = refl 175 | assoc false false false = refl 176 | 177 | inverse-l : LeftInverse _⊕_ false (λ x → false ⊕ x) 178 | inverse-l true = refl 179 | inverse-l false = refl 180 | 181 | inverse-r : RightInverse _⊕_ false (λ x → false ⊕ x) 182 | inverse-r true = refl 183 | inverse-r false = refl 184 | -------------------------------------------------------------------------------- /Equality.agda: -------------------------------------------------------------------------------- 1 | module Equality where 2 | 3 | open import Agda.Primitive 4 | 5 | infix 1 _≡_ 6 | data _≡_ {α} {A : Set α} (x : A) : A → Set α where 7 | refl : x ≡ x 8 | {-# BUILTIN EQUALITY _≡_ #-} 9 | 10 | -- a ≡ b, but with explicit type argument. 11 | Id : ∀ {α} (A : Set α) (x y : A) → Set α 12 | Id _ x y = x ≡ y 13 | 14 | -- refl, but with explicit arguments. 15 | refl' : ∀ {α} (A : Set α) (x : A) → x ≡ x 16 | refl' _ _ = refl 17 | 18 | -- Neatly auto-derivable: case, match on equalty proof, auto. 19 | ind-≡ 20 | : ∀ {α β} {A : Set α} 21 | → (C : (a₁ a₂ : A) → a₁ ≡ a₂ → Set β) 22 | → (c : (a : A) → C a a refl) 23 | → (x y : A) 24 | → (p : x ≡ y) 25 | → C x y p 26 | ind-≡ C c x .x refl = c x 27 | 28 | symm 29 | : ∀ {α} {A : Set α} {x y : A} 30 | → x ≡ y 31 | → y ≡ x 32 | symm refl = refl 33 | 34 | private 35 | 36 | symm-via-ind : ∀ {α} {A : Set α} {x y : A} → x ≡ y → y ≡ x 37 | symm-via-ind {x = x} {y = y} x≡y 38 | = ind-≡ (λ a₁ a₂ _ → a₂ ≡ a₁) 39 | (λ _ → refl) 40 | x 41 | y 42 | x≡y 43 | 44 | trans-via-ind : ∀ {α} {A : Set α} {x y z : A} → x ≡ y → y ≡ z → x ≡ z 45 | trans-via-ind {x = x} {y = y} {z = z} x≡y 46 | = ind-≡ (λ a₁ a₂ _ → a₂ ≡ z → a₁ ≡ z) 47 | (λ _ refl → refl) 48 | x 49 | y 50 | x≡y 51 | 52 | subst-via-ind : ∀ {α β} {A : Set α} → {x y : A} → (P : A → Set β) → x ≡ y → P x → P y 53 | subst-via-ind {x = x} {y = y} p x≡y 54 | = ind-≡ (λ a₁ a₂ _ → p a₁ → p a₂) 55 | (λ _ refl → refl) 56 | x 57 | y 58 | x≡y 59 | 60 | 61 | trans 62 | : ∀ {α} {A : Set α} {x y z : A} 63 | → x ≡ y 64 | → y ≡ z 65 | → x ≡ z 66 | trans refl refl = refl 67 | 68 | -- Congruency allows us to »scope into« definitions. In order to prove 69 | -- 70 | -- > a + (b + c) ≡ a + (c + b) 71 | -- 72 | -- for example, we can use »λ e → a + e« to scope into the »b+c« part, 73 | -- 74 | -- > proof = cong (λ e → a + e) {! !} 75 | -- 76 | -- The hole now has type »b+c ≡ c+b«. We thus scoped into the sum and isolated 77 | -- the location we’d like to apply commutativity to. 78 | cong 79 | : ∀ {α β} {A : Set α} {B : Set β} {x y : A} 80 | → (f : A → B) 81 | → x ≡ y 82 | → f x ≡ f y 83 | cong _ refl = refl 84 | 85 | subst 86 | : ∀ {α β} {A : Set α} 87 | → {x y : A} 88 | → (P : A → Set β) 89 | → x ≡ y 90 | → P x 91 | → P y 92 | subst _ refl x = x 93 | 94 | private 95 | trans-via-subst : ∀ {α} {A : Set α} {x y z : A} → x ≡ y → y ≡ z → x ≡ z 96 | trans-via-subst {z = z} eq₁ eq₂ = subst (λ e → e ≡ z) (symm eq₁) eq₂ 97 | 98 | subst₂ 99 | : ∀ {α β} {A : Set α} {B : Set β} 100 | → {a₁ a₂ : A} {b₁ b₂ : B} 101 | → (P : A → B → Set β) 102 | → a₁ ≡ a₂ 103 | → b₁ ≡ b₂ 104 | → P a₁ b₁ 105 | → P a₂ b₂ 106 | subst₂ _ refl refl Pa₁a₂ = Pa₁a₂ 107 | 108 | subst₃ 109 | : ∀ {α β γ} {A : Set α} {B : Set β} {C : Set γ} 110 | → {a₁ a₂ : A} {b₁ b₂ : B} {c₁ c₂ : C} 111 | → (P : A → B → C → Set β) 112 | → a₁ ≡ a₂ 113 | → b₁ ≡ b₂ 114 | → c₁ ≡ c₂ 115 | → P a₁ b₁ c₁ 116 | → P a₂ b₂ c₂ 117 | subst₃ _ refl refl refl Pa₁a₂a₃ = Pa₁a₂a₃ 118 | 119 | -- Nice way to write chains of equality proofs, courtesy of the Agda standard lib 120 | module ≡-Reasoning {α} {A : Set α} where 121 | infix 3 _qed 122 | infixr 2 _≡⟨⟩_ _≡⟨_⟩_ 123 | infix 1 begin_ 124 | 125 | begin_ : {x y : A} → x ≡ y → x ≡ y 126 | begin_ x≡y = x≡y 127 | 128 | _≡⟨⟩_ : (x {y} : A) → x ≡ y → x ≡ y 129 | x ≡⟨⟩ x≡y = x≡y 130 | 131 | _≡⟨_⟩_ : (x {y z} : A) → x ≡ y → y ≡ z → x ≡ z 132 | _ ≡⟨ x≡y ⟩ y≡z = trans x≡y y≡z 133 | 134 | _qed : (x : A) → x ≡ x 135 | _ qed = refl 136 | 137 | unify : ∀ {α} {a : Set α} → a → a → a 138 | unify x _ = x 139 | 140 | -- Simple inspection type, taken from 141 | -- http://agda.readthedocs.io/en/latest/language/with-abstraction.html#id16 142 | data Inspect {α} {A : Set α} (x : A) : Set α where 143 | _with≡_ : (y : A) → x ≡ y → Inspect x 144 | 145 | inspect : ∀ {α} {A : Set α} (x : A) → Inspect x 146 | inspect x = x with≡ refl 147 | -------------------------------------------------------------------------------- /Fin.agda: -------------------------------------------------------------------------------- 1 | module Fin where 2 | 3 | 4 | 5 | open import Logic 6 | open import Equality 7 | 8 | open Logic.Decidable.Binary 9 | 10 | open import Nat as Nat 11 | using (ℕ; zero; succ; z≤n; s≤s; _+_) 12 | 13 | data Fin : ℕ → Set where 14 | zero : {n : ℕ} → Fin (succ n) 15 | succ : {n : ℕ} → (i : Fin n) → Fin (succ n) 16 | 17 | toℕ : ∀ {n} → Fin n → ℕ 18 | toℕ zero = 0 19 | toℕ (succ x) = succ (toℕ x) 20 | 21 | fromℕ : (n : ℕ) → Fin (succ n) 22 | fromℕ zero = zero 23 | fromℕ (succ n) = succ (fromℕ n) 24 | 25 | -- »The finite x ∈ 0..n is also a finite x ∈ 0..(n+Δn).« 26 | extendBy : ∀ {n} Δn → Fin n → Fin (n + Δn) 27 | extendBy {zero} _ () 28 | extendBy {succ _} _ zero = zero 29 | extendBy {succ _} Δn (succ fin) = succ (extendBy Δn fin) 30 | 31 | -- Flipped sum in the result type; if Δn is known, then this reduces more easily 32 | extendBy' : ∀ {n} Δn → Fin n → Fin (Δn + n) 33 | extendBy' {n} Δn fin = subst Fin (Nat.comm-+ n Δn) (extendBy Δn fin) 34 | 35 | fromℕ' : (n : ℕ) (Δn : ℕ) → Fin (succ n + Δn) 36 | fromℕ' n Δn = extendBy Δn (fromℕ n) 37 | 38 | module extendBy-proofs where 39 | extendBy-toℕ 40 | : ∀ {n : ℕ} Δn (fin : Fin n) 41 | → toℕ fin ≡ toℕ (extendBy Δn fin) 42 | extendBy-toℕ _ zero = refl 43 | extendBy-toℕ Δn (succ fin) rewrite extendBy-toℕ Δn fin = refl 44 | 45 | test₁ : fromℕ' 4 0 ≡ fromℕ 4 46 | test₁ = refl 47 | 48 | -- The finite 4 ∈ 0..(4+1) 49 | test₂ : fromℕ' 4 1 ≡ succ (succ (succ (succ zero))) 50 | test₂ = refl 51 | 52 | -- Largest representable number 53 | finMax : ∀ {n} → Fin (succ n) → ℕ 54 | finMax {n} _ = n 55 | 56 | -- Number of representable numbers. 57 | finSize : ∀ {n} → Fin n → ℕ 58 | finSize {n} _ = n 59 | 60 | module fin-max-size-tests where 61 | 62 | test₁ : ∀ {n} {fin : Fin n} → finMax (extendBy' 1 fin) ≡ finSize fin 63 | test₁ = refl 64 | -------------------------------------------------------------------------------- /FingerTree.agda: -------------------------------------------------------------------------------- 1 | module FingerTree where 2 | 3 | open import Algebra 4 | open import List hiding (foldr) 5 | open import Equality 6 | open import Fin 7 | open import Function 8 | open import Logic 9 | open import Nat 10 | open import Vector hiding (_++_; foldr; toList) 11 | 12 | -- Vectors of length 0..n 13 | Vec≤ : ∀ {α} → Set α → ℕ → Set α 14 | Vec≤ A l = Σ (Fin (succ l)) (Vec A ∘ toℕ) 15 | 16 | infixr 6 _◁_ 17 | infixr 6 _◁◁_ 18 | infixl 6 _▷_ 19 | infixl 6 _▷▷_ 20 | infixr 5 _⋈_ 21 | 22 | module Examples where 23 | examples : Vec (Vec≤ ℕ 3) _ 24 | examples = (fromℕ' 0 3 , []) 25 | ∷ (fromℕ' 1 2 , 1 ∷ []) 26 | ∷ (fromℕ' 2 1 , 1 ∷ 2 ∷ []) 27 | ∷ (fromℕ' 3 0 , 1 ∷ 2 ∷ 3 ∷ []) 28 | ∷ [] 29 | 30 | data Digit {α} (A : Set α) : Set α where 31 | Digit1 : A → Digit A 32 | Digit2 : A → A → Digit A 33 | Digit3 : A → A → A → Digit A 34 | Digit4 : A → A → A → A → Digit A 35 | 36 | data Node {α} (A : Set α) : Set α where 37 | Node2 : A → A → Node A 38 | Node3 : A → A → A → Node A 39 | 40 | data FingerTree23 {α} (A : Set α) : Set α where 41 | Empty : FingerTree23 A 42 | Single : (x : A) → FingerTree23 A 43 | Deep : (ld : Digit A) → (sub : FingerTree23 (Node A)) → (rd : Digit A) → FingerTree23 A 44 | 45 | _◁_ : ∀ {α} {A : Set α} → (x : A) → (xs : FingerTree23 A) → FingerTree23 A 46 | x ◁ Empty = Single x 47 | x ◁ Single y = Deep (Digit1 x) Empty (Digit1 y) 48 | x ◁ Deep (Digit1 d₁) xs rd = Deep (Digit2 x d₁) xs rd 49 | x ◁ Deep (Digit2 d₁ d₂) xs rd = Deep (Digit3 x d₁ d₂) xs rd 50 | x ◁ Deep (Digit3 d₁ d₂ d₃) xs rd = Deep (Digit4 x d₁ d₂ d₃) xs rd 51 | x ◁ Deep (Digit4 d₁ d₂ d₃ d₄) xs rd = Deep (Digit2 x d₁) (Node3 d₂ d₃ d₄ ◁ xs) rd 52 | 53 | data ViewL {α} (A : Set α) (B : Set α → Set α) : Set α where 54 | EmptyL : ViewL A B 55 | _◁∷_ : (x : A) → (xs : B A) → ViewL A B 56 | 57 | viewL : ∀ {α} {A : Set α} → FingerTree23 A → ViewL A FingerTree23 58 | viewL Empty = EmptyL 59 | viewL (Single x) = x ◁∷ Empty 60 | viewL (Deep (Digit1 _) deep _) with viewL deep 61 | viewL (Deep (Digit1 x) _ (Digit1 d₁)) | EmptyL = x ◁∷ Single d₁ 62 | viewL (Deep (Digit1 x) _ (Digit2 d₁ d₂)) | EmptyL = x ◁∷ Deep (Digit1 d₁) Empty (Digit1 d₂) 63 | viewL (Deep (Digit1 x) _ (Digit3 d₁ d₂ d₃)) | EmptyL = x ◁∷ Deep (Digit1 d₁) Empty (Digit2 d₂ d₃) 64 | viewL (Deep (Digit1 x) _ (Digit4 d₁ d₂ d₃ d₄)) | EmptyL = x ◁∷ Deep (Digit1 d₁) Empty (Digit3 d₂ d₃ d₄) 65 | viewL (Deep (Digit1 x) _ dr) | Node2 n₁ n₂ ◁∷ xs = x ◁∷ Deep (Digit2 n₁ n₂) xs dr 66 | viewL (Deep (Digit1 x) _ dr) | Node3 n₁ n₂ n₃ ◁∷ xs = x ◁∷ Deep (Digit3 n₁ n₂ n₃) xs dr 67 | viewL (Deep (Digit2 x x₁) x₂ rd) = x ◁∷ Deep (Digit1 x₁) x₂ rd 68 | viewL (Deep (Digit3 x x₁ x₂) x₃ rd) = x ◁∷ Deep (Digit2 x₁ x₂) x₃ rd 69 | viewL (Deep (Digit4 x x₁ x₂ x₃) x₄ rd) = x ◁∷ Deep (Digit3 x₁ x₂ x₃) x₄ rd 70 | 71 | _◁◁_ : ∀ {α} {A : Set α} → (ds : Digit A) → (xs : FingerTree23 A) → FingerTree23 A 72 | Digit1 x₁ ◁◁ xs = x₁ ◁ xs 73 | Digit2 x₁ x₂ ◁◁ xs = x₁ ◁ x₂ ◁ xs 74 | Digit3 x₁ x₂ x₃ ◁◁ xs = x₁ ◁ x₂ ◁ x₃ ◁ xs 75 | Digit4 x₁ x₂ x₃ x₄ ◁◁ xs = x₁ ◁ x₂ ◁ x₃ ◁ x₄ ◁ xs 76 | 77 | _▷_ : ∀ {α} {A : Set α} → (xs : FingerTree23 A) → (x : A) → FingerTree23 A 78 | Empty ▷ x = Single x 79 | Single x ▷ y = Deep (Digit1 x) Empty (Digit1 y) 80 | Deep ld xs (Digit1 d₁) ▷ x = Deep ld xs (Digit2 d₁ x) 81 | Deep ld xs (Digit2 d₁ d₂) ▷ x = Deep ld xs (Digit3 d₁ d₂ x) 82 | Deep ld xs (Digit3 d₁ d₂ d₃) ▷ x = Deep ld xs (Digit4 d₁ d₂ d₃ x) 83 | Deep ld xs (Digit4 d₁ d₂ d₃ d₄) ▷ x = Deep ld (xs ▷ Node3 d₁ d₂ d₃) (Digit2 d₄ x) 84 | 85 | data ViewR {α} (A : Set α) (B : Set α → Set α) : Set α where 86 | EmptyR : ViewR A B 87 | _∷▷_ : (xs : B A) → (x : A) → ViewR A B 88 | 89 | viewR : ∀ {α} {A : Set α} → FingerTree23 A → ViewR A FingerTree23 90 | viewR Empty = EmptyR 91 | viewR (Single x) = Empty ∷▷ x 92 | viewR (Deep _ deep (Digit1 _)) with viewR deep 93 | viewR (Deep (Digit1 d₁) _ (Digit1 x)) | EmptyR = Single d₁ ∷▷ x 94 | viewR (Deep (Digit2 d₁ d₂) _ (Digit1 x)) | EmptyR = Deep (Digit1 d₁) Empty (Digit1 d₂) ∷▷ x 95 | viewR (Deep (Digit3 d₁ d₂ d₃) _ (Digit1 x)) | EmptyR = Deep (Digit2 d₁ d₂) Empty (Digit1 d₃) ∷▷ x 96 | viewR (Deep (Digit4 d₁ d₂ d₃ d₄) _ (Digit1 x)) | EmptyR = Deep (Digit3 d₁ d₂ d₃) Empty (Digit1 d₄) ∷▷ x 97 | viewR (Deep ld _ (Digit1 x)) | xs ∷▷ Node2 n₁ n₂ = Deep ld xs (Digit2 n₁ n₂) ∷▷ x 98 | viewR (Deep ld _ (Digit1 x)) | xs ∷▷ Node3 n₁ n₂ n₃ = Deep ld xs (Digit3 n₁ n₂ n₃) ∷▷ x 99 | viewR (Deep ld deep (Digit2 d₁ x)) = Deep ld deep (Digit1 d₁) ∷▷ x 100 | viewR (Deep ld deep (Digit3 d₂ d₁ x)) = Deep ld deep (Digit2 d₂ d₁) ∷▷ x 101 | viewR (Deep ld deep (Digit4 d₃ d₂ d₁ x)) = Deep ld deep (Digit3 d₃ d₂ d₁) ∷▷ x 102 | 103 | _▷▷_ : ∀ {α} {A : Set α} → (xs : FingerTree23 A) → (ds : Digit A) → FingerTree23 A 104 | xs ▷▷ Digit1 x₁ = xs ▷ x₁ 105 | xs ▷▷ Digit2 x₁ x₂ = xs ▷ x₁ ▷ x₂ 106 | xs ▷▷ Digit3 x₁ x₂ x₃ = xs ▷ x₁ ▷ x₂ ▷ x₃ 107 | xs ▷▷ Digit4 x₁ x₂ x₃ x₄ = xs ▷ x₁ ▷ x₂ ▷ x₃ ▷ x₄ 108 | 109 | -- ⋈ = \bowtie 110 | _⋈_ : ∀ {α} {A : Set α} → (xs : FingerTree23 A) → (ys : FingerTree23 A) → FingerTree23 A 111 | Empty ⋈ ys = ys 112 | Single x ⋈ ys = x ◁ ys 113 | xs ⋈ Empty = xs 114 | xs ⋈ Single x = xs ▷ x 115 | Deep xld xs (Digit1 lx₁) ⋈ Deep (Digit1 rx₁) ys yrd = Deep xld {! !} yrd 116 | Deep xld xs (Digit1 lx₁) ⋈ Deep (Digit2 rx₁ rx₂) ys yrd = Deep xld {! !} yrd 117 | Deep xld xs (Digit1 lx₁) ⋈ Deep (Digit3 rx₁ rx₂ rx₃) ys yrd = Deep xld {! !} yrd 118 | Deep xld xs (Digit1 lx₁) ⋈ Deep (Digit4 rx₁ rx₂ rx₃ rx₄) ys yrd = Deep xld {! !} yrd 119 | Deep xld xs (Digit2 lx₁ lx₂) ⋈ Deep (Digit1 rx₁) ys yrd = Deep xld {! !} yrd 120 | Deep xld xs (Digit2 lx₁ lx₂) ⋈ Deep (Digit2 rx₁ rx₂) ys yrd = Deep xld {! !} yrd 121 | Deep xld xs (Digit2 lx₁ lx₂) ⋈ Deep (Digit3 rx₁ rx₂ rx₃) ys yrd = Deep xld {! !} yrd 122 | Deep xld xs (Digit2 lx₁ lx₂) ⋈ Deep (Digit4 rx₁ rx₂ rx₃ rx₄) ys yrd = Deep xld {! !} yrd 123 | Deep xld xs (Digit3 lx₁ lx₂ x₃) ⋈ Deep (Digit1 rx₁) ys yrd = Deep xld {! !} yrd 124 | Deep xld xs (Digit3 lx₁ lx₂ x₃) ⋈ Deep (Digit2 rx₁ rx₂) ys yrd = Deep xld {! !} yrd 125 | Deep xld xs (Digit3 lx₁ lx₂ x₃) ⋈ Deep (Digit3 rx₁ rx₂ rx₃) ys yrd = Deep xld {! !} yrd 126 | Deep xld xs (Digit3 lx₁ lx₂ x₃) ⋈ Deep (Digit4 rx₁ rx₂ rx₃ rx₄) ys yrd = Deep xld {! !} yrd 127 | Deep xld xs (Digit4 lx₁ lx₂ x₃ x₄) ⋈ Deep (Digit1 rx₁) ys yrd = Deep xld {! !} yrd 128 | Deep xld xs (Digit4 lx₁ lx₂ x₃ x₄) ⋈ Deep (Digit2 rx₁ rx₂) ys yrd = Deep xld {! !} yrd 129 | Deep xld xs (Digit4 lx₁ lx₂ x₃ x₄) ⋈ Deep (Digit3 rx₁ rx₂ rx₃) ys yrd = Deep xld {! !} yrd 130 | Deep xld xs (Digit4 lx₁ lx₂ x₃ x₄) ⋈ Deep (Digit4 rx₁ rx₂ rx₃ rx₄) ys yrd = Deep xld {! !} yrd 131 | 132 | module Induction where 133 | ind-FingerTree23 134 | : ∀ {α γ} {A : Set α} {C : FingerTree23 A → Set γ} 135 | → (empty : C Empty) 136 | → (single : (x : A) → C (Single x)) 137 | → (deep : (d← : Digit A) → (xs : FingerTree23 (Node A)) → (d→ : Digit A) → C (Deep d← xs d→)) 138 | → (xs : FingerTree23 A) 139 | → C xs 140 | ind-FingerTree23 empty _ _ Empty = empty 141 | ind-FingerTree23 _ single _ (Single x) = single x 142 | ind-FingerTree23 _ _ deep (Deep ld xs rd) = deep ld xs rd 143 | 144 | ind-Digit 145 | : ∀ {α γ} {A : Set α} {C : Digit A → Set γ} 146 | → (digit1 : (x₁ : A) → C (Digit1 x₁)) 147 | → (digit2 : (x₁ x₂ : A) → C (Digit2 x₁ x₂)) 148 | → (digit3 : (x₁ x₂ x₃ : A) → C (Digit3 x₁ x₂ x₃)) 149 | → (digit4 : (x₁ x₂ x₃ x₄ : A) → C (Digit4 x₁ x₂ x₃ x₄)) 150 | → (d : Digit A) 151 | → C d 152 | ind-Digit digit1 _ _ _ (Digit1 x₁) = digit1 x₁ 153 | ind-Digit _ digit2 _ _ (Digit2 x₁ x₂) = digit2 x₁ x₂ 154 | ind-Digit _ _ digit3 _ (Digit3 x₁ x₂ x₃) = digit3 x₁ x₂ x₃ 155 | ind-Digit _ _ _ digit4 (Digit4 x₁ x₂ x₃ x₄) = digit4 x₁ x₂ x₃ x₄ 156 | 157 | ind-Node 158 | : ∀ {α γ} {A : Set α} {C : Node A → Set γ} 159 | → (node2 : (x₁ x₂ : A) → C (Node2 x₁ x₂)) 160 | → (node3 : (x₁ x₂ x₃ : A) → C (Node3 x₁ x₂ x₃)) 161 | → (d : Node A) 162 | → C d 163 | ind-Node node2 _ (Node2 x₁ x₂) = node2 x₁ x₂ 164 | ind-Node _ node3 (Node3 x₁ x₂ x₃) = node3 x₁ x₂ x₃ 165 | 166 | rec-FingerTree23 167 | : ∀ {α γ} {A : Set α} {C : Set γ} 168 | → (empty : C) 169 | → (single : A → C) 170 | → (deep : Digit A → FingerTree23 (Node A) → Digit A → C) 171 | → FingerTree23 A 172 | → C 173 | rec-FingerTree23 = ind-FingerTree23 174 | 175 | toList : ∀ {α} {A : Set α} → FingerTree23 A → List A 176 | toList Empty = [] 177 | toList (Single x) = x ∷ [] 178 | toList (Deep ld xs rd) = {! !} 179 | -------------------------------------------------------------------------------- /Function.agda: -------------------------------------------------------------------------------- 1 | module Function where 2 | 3 | 4 | 5 | open import Equality 6 | 7 | 8 | 9 | id : ∀ {α} {a : Set α} → a → a 10 | id x = x 11 | 12 | -- Dependent function composition 13 | _∘_ 14 | : ∀ {α β γ} {A : Set α} {B : A → Set β} {C : {x : A} → B x → Set γ} 15 | → (f : {x : A} (y : B x) → C y) 16 | → (g : (x : A) → B x) 17 | → (x : A) 18 | → C (g x) 19 | f ∘ g = λ x → f (g x) 20 | 21 | -- Ordinary function composition 22 | _∘'_ 23 | : ∀ {α β γ} {a : Set α} {b : Set β} {c : Set γ} 24 | → (f : b → c) 25 | → (g : a → b) 26 | → (a → c) 27 | f ∘' g = f ∘ g 28 | 29 | assoc-∘ 30 | : ∀ {α β γ δ} {a : Set α} {b : Set β} {c : Set γ} {d : Set δ} 31 | → (f : c → d) (g : b → c) (h : a → b) (x : a) 32 | → ((f ∘ g) ∘ h) x ≡ (f ∘ (g ∘ h)) x 33 | assoc-∘ _ _ _ _ = refl 34 | 35 | f∘id≡f : ∀ {α β} {a : Set α} {b : Set β} {x : a} (f : a → b) → (f ∘ id) x ≡ f x 36 | f∘id≡f _ = refl 37 | 38 | id∘f≡f : ∀ {α β} {a : Set α} {b : Set β} {x : a} (f : a → b) → (id ∘ f) x ≡ f x 39 | id∘f≡f _ = refl 40 | 41 | const : ∀ {α β} {A : Set α} {B : Set β} → A → B → A 42 | const a _ = a 43 | 44 | the : ∀ {α} (A : Set α) → A → A 45 | the _ x = x 46 | -------------------------------------------------------------------------------- /FunctorApplicativeMonad.agda: -------------------------------------------------------------------------------- 1 | module FunctorApplicativeMonad where 2 | 3 | open import Function 4 | open import Agda.Primitive 5 | open import Equality 6 | 7 | record Functor {α} (F : Set α → Set α) : Set (lsuc α) where 8 | field 9 | fmap : ∀ {A B : Set α} → (A → B) → F A → F B 10 | fmap[id]≡id : ∀ {A : Set α} {x : F A} → fmap id x ≡ x 11 | 12 | record Applicative {α} (F : Set α → Set α) : Set (lsuc α) where 13 | infixl 4 _⟨*⟩_ 14 | field 15 | pure : ∀ {A : Set α} → A → F A 16 | _⟨*⟩_ : ∀ {A B : Set α} → F (A → B) → F A → F B 17 | 18 | -- Laws 19 | isFunctor : Functor F 20 | fmap-from-Applicative 21 | : {A B : Set α} {f : A → B} {x : F A} 22 | → pure f ⟨*⟩ x ≡ Functor.fmap isFunctor f x 23 | homomorphism 24 | : {A B : Set α} {f : A → B} {x : A} 25 | → pure f ⟨*⟩ pure x ≡ pure (f x) 26 | interchange 27 | : {A B : Set α} {u : F (A → B)} {y : A} 28 | → u ⟨*⟩ pure y ≡ pure (λ f → f y) ⟨*⟩ u 29 | composition 30 | : {A B C : Set α} {u : F (B → C)} {v : F (A → B)} {w : F A} 31 | → u ⟨*⟩ (v ⟨*⟩ w) ≡ pure _∘_ ⟨*⟩ u ⟨*⟩ v ⟨*⟩ w 32 | 33 | open Functor isFunctor public 34 | 35 | -- LAWS! 36 | 37 | record Monad {α} (F : Set α → Set α) : Set (lsuc α) where 38 | field 39 | isApplicative : Applicative F 40 | join : ∀ {A : Set α} → F (F A) → F A 41 | -- fmap[join]∘join≡join∘join : {A : Set α} {mmx : F (F A)} → Functor.fmap isApplicative join (join mmx) ≡ join (join mmx) 42 | 43 | open Applicative isApplicative public 44 | 45 | -- fmap[id]≡id : ∀ {A : Set α} {x : F A} → fmap id x ≡ x 46 | -------------------------------------------------------------------------------- /Homotopy.agda: -------------------------------------------------------------------------------- 1 | module Homotopy where 2 | 3 | open import Agda.Primitive 4 | open import Equality 5 | open import Logic 6 | open import Function 7 | 8 | -- Homotopy 9 | _~_ : ∀ {α} {A : Set α} {P : A → Set α} → (f : (x : A) → P x) → (g : (y : A) → P y) → Set α 10 | f ~ g = ∀ x → f x ≡ g x 11 | 12 | -- Left-invertible functions 13 | linv : ∀ {α β} {A : Set α} {B : Set β} → (A → B) → Set (α ⊔ β) 14 | linv f = ∃ (λ g → (g ∘ f) ~ id) 15 | 16 | -- Right-invertible functions 17 | rinv : ∀ {α β} {A : Set α} {B : Set β} → (A → B) → Set (α ⊔ β) 18 | rinv f = ∃ (λ g → (f ∘ g) ~ id) 19 | 20 | -- Type equivalence 21 | _≃_ : ∀ {α β} → (A : Set α) → (B : Set β) → Set (α ⊔ β) 22 | A ≃ B = Σ (A → B) (λ f → rinv f × linv f) 23 | 24 | private 25 | 26 | open import Bool 27 | 28 | _∨'_ : ∀ {α} → (A : Set α) → (B : Set α) → Set α 29 | A ∨' B = ∃ (rec-Bool A B) 30 | 31 | ⟶ : ∀ {α} {A : Set α} {B : Set α} → A ∨ B → A ∨' B 32 | ⟶ (inl l) = (true , l) 33 | ⟶ (inr r) = (false , r) 34 | 35 | ⟵ : ∀ {α} {A : Set α} {B : Set α} → A ∨' B → A ∨ B 36 | ⟵ (true , l) = inl l 37 | ⟵ (false , r) = inr r 38 | 39 | ⟵∘⟶=id : ∀ {α} {A : Set α} {B : Set α} (x : A ∨ B) → (⟵ ∘ ⟶) x ≡ id x 40 | ⟵∘⟶=id (inl l) = refl 41 | ⟵∘⟶=id (inr r) = refl 42 | 43 | ⟶∘⟵=id : ∀ {α} {A : Set α} {B : Set α} (x : A ∨' B) → (⟶ ∘ ⟵) x ≡ id x 44 | ⟶∘⟵=id (true , x) = refl 45 | ⟶∘⟵=id (false , x) = refl 46 | 47 | theorem 48 | : ∀ {α} {A : Set α} {B : Set α} 49 | → (A ∨ B) ≃ (A ∨' B) 50 | theorem = (⟶ , ((⟵ , ⟶∘⟵=id) , (⟵ , ⟵∘⟶=id))) 51 | -------------------------------------------------------------------------------- /Int.agda: -------------------------------------------------------------------------------- 1 | module Int where 2 | 3 | open import Nat 4 | as ℕ 5 | using (ℕ) 6 | renaming (_+_ to _ℕ+_; _*_ to _ℕ*_) 7 | open import Equality 8 | open import Sign 9 | as Sign 10 | using (Sign) 11 | renaming (_*_ to _S*_) 12 | open import Algebra 13 | open import Logic 14 | 15 | open Equality.≡-Reasoning 16 | 17 | 18 | 19 | module first-attempt where 20 | 21 | infix 7 +_ 22 | data ℤ : Set where 23 | +_ : (n : ℕ) → ℤ 24 | -[1+_] : (n : ℕ) → ℤ 25 | {-# BUILTIN INTEGER ℤ #-} 26 | {-# BUILTIN INTEGERPOS +_ #-} 27 | {-# BUILTIN INTEGERNEGSUC -[1+_] #-} 28 | 29 | ∣_∣ : ℤ → ℕ 30 | ∣ -[1+ x ] ∣ = ℕ.succ x 31 | ∣ + x ∣ = x 32 | 33 | sign : ℤ → Sign 34 | sign (+ _) = Sign.+ 35 | sign -[1+ _ ] = Sign.- 36 | 37 | _◁_ : Sign → ℕ → ℤ 38 | Sign.+ ◁ n = + n 39 | Sign.- ◁ ℕ.zero = + 0 40 | Sign.- ◁ ℕ.succ n = -[1+ n ] 41 | 42 | ◁-inverse : ∀ n → sign n ◁ ∣ n ∣ ≡ n 43 | ◁-inverse (+ x) = refl 44 | ◁-inverse -[1+ x ] = refl 45 | 46 | sign²[≠0] : ∀ ±1 x → sign (±1 ◁ ℕ.succ x) ≡ ±1 47 | sign²[≠0] Sign.+ x = refl 48 | sign²[≠0] Sign.- x = refl 49 | 50 | sign²[=0] : ∀ ±1 → sign (±1 ◁ 0) ≡ Sign.+ 51 | sign²[=0] Sign.+ = refl 52 | sign²[=0] Sign.- = refl 53 | 54 | infix 6 _ℕ-_ 55 | _ℕ-_ : ℕ → ℕ → ℤ 56 | x ℕ- ℕ.zero = + x 57 | ℕ.zero ℕ- ℕ.succ y = -[1+ y ] 58 | ℕ.succ x ℕ- ℕ.succ y = x ℕ- y 59 | 60 | infix 8 -_ 61 | -_ : ℤ → ℤ 62 | - (+ ℕ.zero) = + ℕ.zero 63 | - (+ ℕ.succ n) = -[1+ n ] 64 | - -[1+ n ] = + (ℕ.succ n) 65 | 66 | infix 6 _+_ 67 | _+_ : ℤ → ℤ → ℤ 68 | (+ x) + (+ y) = + (x ℕ+ y) 69 | (+ x) + -[1+ y ] = x ℕ- ℕ.succ y 70 | -[1+ x ] + (+ y) = y ℕ- ℕ.succ x 71 | -[1+ x ] + -[1+ y ] = -[1+ ℕ.succ (x ℕ+ y) ] 72 | 73 | infix 6 _-_ 74 | _-_ : ℤ → ℤ → ℤ 75 | x - y = x + - y 76 | 77 | -[-x]=x : ∀ x → - (- x) ≡ x 78 | -[-x]=x (+ ℕ.zero) = refl 79 | -[-x]=x (+ ℕ.succ x) = refl 80 | -[-x]=x -[1+ x ] = refl 81 | 82 | -x≡[0-x] : ∀ x → - x ≡ (+ 0) - x 83 | -x≡[0-x] (+ ℕ.zero) = refl 84 | -x≡[0-x] (+ ℕ.succ x) = refl 85 | -x≡[0-x] -[1+ x ] = refl 86 | 87 | module test+- where 88 | test₁ : (+ 1) + (+ 2) ≡ (+ 3) 89 | test₁ = refl 90 | 91 | test₂ : (+ 1) + (- (+ 2)) ≡ (- (+ 1)) 92 | test₂ = refl 93 | 94 | x-x≡0 : ∀ x → x - x ≡ + 0 95 | x-x≡0 (+ ℕ.zero) = refl 96 | x-x≡0 (+ ℕ.succ n) = x-x≡0 -[1+ n ] 97 | x-x≡0 -[1+ ℕ.zero ] = refl 98 | x-x≡0 -[1+ ℕ.succ n ] = x-x≡0 -[1+ n ] 99 | 100 | x-y=-[y-x] : ∀ x y → x - y ≡ - (y - x) 101 | x-y=-[y-x] (+ ℕ.zero) (+ ℕ.zero) = refl 102 | x-y=-[y-x] (+ ℕ.zero) (+ ℕ.succ y) rewrite ℕ.x+0≡x y = refl 103 | x-y=-[y-x] (+ ℕ.succ x) (+ ℕ.zero) rewrite ℕ.x+0≡x x = refl 104 | x-y=-[y-x] (+ ℕ.succ x) (+ ℕ.succ y) = x-y=-[y-x] -[1+ y ] -[1+ x ] 105 | x-y=-[y-x] (+ ℕ.zero) -[1+ ℕ.zero ] = refl 106 | x-y=-[y-x] (+ ℕ.zero) -[1+ ℕ.succ y ] = refl 107 | x-y=-[y-x] (+ ℕ.succ x) -[1+ ℕ.zero ] rewrite ℕ.comm-+ x 1 = refl 108 | x-y=-[y-x] (+ ℕ.succ x) -[1+ ℕ.succ y ] 109 | rewrite ℕ.x+[1+y]≡[1+x]+y x (ℕ.succ y) 110 | | ℕ.x+[1+y]≡[1+x]+y x y 111 | | ℕ.comm-+ x y 112 | = refl 113 | x-y=-[y-x] -[1+ ℕ.zero ] (+ ℕ.zero) = refl 114 | x-y=-[y-x] -[1+ ℕ.zero ] (+ ℕ.succ y) rewrite ℕ.comm-+ y 1 = refl 115 | x-y=-[y-x] -[1+ ℕ.succ x ] (+ ℕ.zero) = refl 116 | x-y=-[y-x] -[1+ ℕ.succ x ] (+ ℕ.succ y) 117 | rewrite ℕ.x+[1+y]≡[1+x]+y y (ℕ.succ x) 118 | | ℕ.x+[1+y]≡[1+x]+y y x 119 | | ℕ.comm-+ x y 120 | = refl 121 | x-y=-[y-x] -[1+ ℕ.zero ] -[1+ ℕ.zero ] = refl 122 | x-y=-[y-x] -[1+ ℕ.zero ] -[1+ ℕ.succ n ] = refl 123 | x-y=-[y-x] -[1+ ℕ.succ x ] -[1+ ℕ.zero ] = refl 124 | x-y=-[y-x] -[1+ ℕ.succ x ] -[1+ ℕ.succ y ] = x-y=-[y-x] -[1+ x ] -[1+ y ] 125 | 126 | module test- where 127 | 5-3=2 : (+ 5) - (+ 3) ≡ (+ 2) 128 | 5-3=2 = refl 129 | 130 | 5-7=-2 : (+ 5) - (+ 7) ≡ ((+ 0) - (+ 2)) 131 | 5-7=-2 = refl 132 | 133 | comm-+ : Commutative _+_ 134 | comm-+ (+ x) (+ y) rewrite ℕ.comm-+ x y = refl 135 | comm-+ (+ x) -[1+ y ] = refl 136 | comm-+ -[1+ x ] (+ y) = refl 137 | comm-+ -[1+ x ] -[1+ y ] rewrite ℕ.comm-+ x y = refl 138 | 139 | 0+x≡x : LeftIdentity _+_ (+ 0) 140 | 0+x≡x (+ _) = refl 141 | 0+x≡x -[1+ _ ] = refl 142 | 143 | x+0≡x : RightIdentity _+_ (+ 0) 144 | x+0≡x x rewrite comm-+ x (+ 0) = 0+x≡x x 145 | 146 | assoc₁ : ∀ x y z → (x ℕ+ y) ℕ- z ≡ + x + (y ℕ- z) 147 | assoc₁ _ _ ℕ.zero = refl 148 | assoc₁ ℕ.zero ℕ.zero (ℕ.succ _) = refl 149 | assoc₁ ℕ.zero (ℕ.succ y) (ℕ.succ z) = assoc₁ ℕ.zero y z 150 | assoc₁ (ℕ.succ x) ℕ.zero (ℕ.succ z) rewrite ℕ.x+0≡x x = refl 151 | assoc₁ (ℕ.succ x) (ℕ.succ y) (ℕ.succ z) = {! !} 152 | 153 | assoc₂ : ∀ x y z → (y ℕ- x) + + z ≡ (y ℕ+ z) ℕ- x 154 | assoc₂ ℕ.zero _ _ = refl 155 | assoc₂ (ℕ.succ _) ℕ.zero _ = refl 156 | assoc₂ (ℕ.succ x) (ℕ.succ y) z = assoc₂ x y z 157 | 158 | assoc₃ : ∀ x y z → (x ℕ- y) + + z ≡ + x + (z ℕ- y) 159 | assoc₃ ℕ.zero ℕ.zero _ = refl 160 | assoc₃ ℕ.zero (ℕ.succ _) ℕ.zero = refl 161 | assoc₃ ℕ.zero (ℕ.succ y) (ℕ.succ z) = symm (0+x≡x (z ℕ- y)) 162 | assoc₃ (ℕ.succ _) ℕ.zero _ = refl 163 | assoc₃ (ℕ.succ x) (ℕ.succ y) ℕ.zero = x+0≡x (x ℕ- y) 164 | assoc₃ (ℕ.succ x) (ℕ.succ y) (ℕ.succ z) = {! !} 165 | 166 | assoc₄ : ∀ x y z → (x ℕ- y) + -[1+ z ] ≡ x ℕ- (ℕ.succ y ℕ+ z) 167 | assoc₄ ℕ.zero ℕ.zero z = refl 168 | assoc₄ ℕ.zero (ℕ.succ _) _ = refl 169 | assoc₄ (ℕ.succ _) ℕ.zero _ = refl 170 | assoc₄ (ℕ.succ x) (ℕ.succ y) z = assoc₄ x y z 171 | 172 | assoc₅ : ∀ x y z → (y ℕ- ℕ.succ x) + -[1+ z ] ≡ -[1+ x ] + (y ℕ- ℕ.succ z) 173 | assoc₅ _ ℕ.zero _ = refl 174 | assoc₅ ℕ.zero (ℕ.succ _) ℕ.zero = refl 175 | assoc₅ ℕ.zero (ℕ.succ y) (ℕ.succ z) = {! !} 176 | assoc₅ (ℕ.succ x) (ℕ.succ y) ℕ.zero = {! !} 177 | assoc₅ (ℕ.succ x) (ℕ.succ y) (ℕ.succ z) = {! !} 178 | 179 | assoc₆ : ∀ x y z → z ℕ- ℕ.succ (ℕ.succ x ℕ+ y) ≡ -[1+ x ] + (z ℕ- ℕ.succ y) 180 | assoc₆ _ _ ℕ.zero = refl 181 | assoc₆ ℕ.zero ℕ.zero (ℕ.succ z) = refl 182 | assoc₆ ℕ.zero (ℕ.succ y) (ℕ.succ z) = assoc₆ ℕ.zero y z 183 | assoc₆ (ℕ.succ x) ℕ.zero (ℕ.succ z) = {! !} 184 | assoc₆ (ℕ.succ x) (ℕ.succ y) (ℕ.succ z) = {! !} 185 | 186 | assoc-+ : Associative _+_ 187 | assoc-+ (+ x) (+ y) (+ z) rewrite ℕ.assoc-+ x y z = refl 188 | assoc-+ (+ x) (+ y) -[1+ z ] = assoc₁ x y (ℕ.succ z) 189 | assoc-+ (+ x) -[1+ y ] (+ z) = assoc₃ x (ℕ.succ y) z 190 | assoc-+ (+ x) -[1+ y ] -[1+ z ] = assoc₄ x (ℕ.succ y) z 191 | assoc-+ -[1+ x ] (+ y) (+ z) = assoc₂ (ℕ.succ x) y z 192 | assoc-+ -[1+ x ] (+ y) -[1+ z ] = assoc₅ x y z 193 | assoc-+ -[1+ x ] -[1+ y ] (+ z) = assoc₆ x y z 194 | assoc-+ -[1+ x ] -[1+ y ] -[1+ z ] rewrite ℕ.x+[1+y]≡[1+x]+y x (y ℕ+ z) | ℕ.assoc-+ x y z = refl 195 | 196 | -- Auto-derive 197 | left-inverse-+0 : LeftInverse _+_ (+ 0) (-_) 198 | left-inverse-+0 (+ ℕ.zero) = refl 199 | left-inverse-+0 (+ ℕ.succ n) = left-inverse-+0 -[1+ n ] 200 | left-inverse-+0 -[1+ ℕ.zero ] = refl 201 | left-inverse-+0 -[1+ ℕ.succ n ] = left-inverse-+0 -[1+ n ] 202 | 203 | -- Auto-derive 204 | right-inverse-+0 : RightInverse _+_ (+ 0) (-_) 205 | right-inverse-+0 (+ ℕ.zero) = refl 206 | right-inverse-+0 (+ ℕ.succ n) = right-inverse-+0 -[1+ n ] 207 | right-inverse-+0 -[1+ ℕ.zero ] = refl 208 | right-inverse-+0 -[1+ ℕ.succ n ] = right-inverse-+0 -[1+ n ] 209 | 210 | ℤ-+0-abelianGroup : AbelianGroup _+_ (+ 0) (-_) 211 | ℤ-+0-abelianGroup = record 212 | { isGroup = record { 213 | isMonoid = record 214 | { isSemigroup = record { assoc = assoc-+ } 215 | ; identity = record 216 | { left-id = 0+x≡x 217 | ; right-id = x+0≡x } } 218 | ; inverse-l = left-inverse-+0 219 | ; inverse-r = right-inverse-+0 } 220 | ; comm = comm-+ } 221 | 222 | 223 | _*_ : ℤ → ℤ → ℤ 224 | x * y = (sign x S* sign y) ◁ (∣ x ∣ ℕ* ∣ y ∣) 225 | 226 | comm-* : Commutative _*_ 227 | comm-* x y = begin 228 | (sign x S* sign y) ◁ (∣ x ∣ ℕ* ∣ y ∣) ≡⟨ cong (λ e → (sign x S* sign y) ◁ e) (ℕ.comm-* ∣ x ∣ ∣ y ∣) ⟩ 229 | (sign x S* sign y) ◁ (∣ y ∣ ℕ* ∣ x ∣) ≡⟨ cong (λ e → e ◁ (∣ y ∣ ℕ* ∣ x ∣)) (Sign.comm-* (sign x) (sign y)) ⟩ 230 | (sign y S* sign x) ◁ (∣ y ∣ ℕ* ∣ x ∣) qed 231 | 232 | assoc-* : Associative _*_ 233 | assoc-* = {! !} 234 | 235 | 1*x≡x : LeftIdentity _*_ (+ 1) 236 | 1*x≡x (+ ℕ.zero) = refl 237 | 1*x≡x (+ ℕ.succ x) = 238 | begin 239 | (+ ℕ.succ (x ℕ+ ℕ.zero)) 240 | ≡⟨ cong (λ e → + ℕ.succ e) (ℕ.x+0≡x x) ⟩ 241 | (+ ℕ.succ x) 242 | qed 243 | 1*x≡x -[1+ ℕ.zero ] = refl 244 | 1*x≡x -[1+ ℕ.succ x ] = 245 | begin 246 | -[1+ ℕ.succ (x ℕ+ ℕ.zero) ] 247 | ≡⟨ cong (λ e → -[1+ ℕ.succ e ]) (ℕ.x+0≡x x) ⟩ 248 | -[1+ ℕ.succ x ] 249 | qed 250 | 251 | x*1≡x : RightIdentity _*_ (+ 1) 252 | x*1≡x x = 253 | begin 254 | (x * (+ 1)) 255 | ≡⟨ symm (comm-* (+ ℕ.succ ℕ.zero) x) ⟩ 256 | ((+ 1) * x) 257 | ≡⟨ 1*x≡x x ⟩ 258 | x 259 | qed 260 | 261 | module second-attempt where 262 | 263 | signℤ : ℕ → Set 264 | signℤ ℕ.zero = ⊤ 265 | signℤ (ℕ.succ x) = Sign 266 | 267 | data ℤ : Set where 268 | mkℤ : (n : ℕ) → (s : signℤ n) → ℤ 269 | 270 | ∣_∣ : ℤ → ℕ 271 | ∣ mkℤ n _ ∣ = n 272 | 273 | sign : ℤ → Sign 274 | sign (mkℤ ℕ.zero tt) = Sign.+ 275 | sign (mkℤ (ℕ.succ _) Sign.+) = Sign.+ 276 | sign (mkℤ (ℕ.succ _) Sign.-) = Sign.- 277 | 278 | _◁_ : Sign → ℕ → ℤ 279 | _ ◁ ℕ.zero = mkℤ ℕ.zero tt 280 | s ◁ ℕ.succ n = mkℤ (ℕ.succ n) s 281 | 282 | fromℕ : ℕ → ℤ 283 | fromℕ x = Sign.+ ◁ x 284 | 285 | ◁-left-inverse : ∀ n → sign n ◁ ∣ n ∣ ≡ n 286 | ◁-left-inverse (mkℤ ℕ.zero tt) = refl 287 | ◁-left-inverse (mkℤ (ℕ.succ n) Sign.+) = refl 288 | ◁-left-inverse (mkℤ (ℕ.succ n) Sign.-) = refl 289 | 290 | _ℕ-_ : ℕ → ℕ → ℤ 291 | x ℕ- ℕ.zero = Sign.+ ◁ x 292 | ℕ.zero ℕ- ℕ.succ y = Sign.- ◁ y 293 | ℕ.succ x ℕ- ℕ.succ y = x ℕ- y 294 | 295 | xℕ-x≡0 : ∀ {n} → n ℕ- n ≡ fromℕ 0 296 | xℕ-x≡0 {ℕ.zero} = refl 297 | xℕ-x≡0 {ℕ.succ n} = xℕ-x≡0 {n} 298 | 299 | -_ : ℤ → ℤ 300 | - mkℤ ℕ.zero tt = mkℤ ℕ.zero tt 301 | - mkℤ (ℕ.succ n) s = mkℤ (ℕ.succ n) (s Sign.* Sign.-) 302 | 303 | _+_ : ℤ → ℤ → ℤ 304 | mkℤ ℕ.zero _ + y = y 305 | x + mkℤ ℕ.zero _ = x 306 | mkℤ (ℕ.succ x) Sign.+ + mkℤ (ℕ.succ y) Sign.+ = Sign.+ ◁ (ℕ.succ x ℕ+ ℕ.succ y) 307 | mkℤ (ℕ.succ x) Sign.+ + mkℤ (ℕ.succ y) Sign.- = ℕ.succ x ℕ- ℕ.succ y 308 | mkℤ (ℕ.succ x) Sign.- + mkℤ (ℕ.succ y) Sign.+ = - (ℕ.succ y ℕ- ℕ.succ x) 309 | mkℤ (ℕ.succ x) Sign.- + mkℤ (ℕ.succ y) Sign.- = Sign.- ◁ (ℕ.succ x ℕ+ ℕ.succ y) 310 | 311 | _-_ : ℤ → ℤ → ℤ 312 | x - mkℤ ℕ.zero tt = x 313 | x - mkℤ (ℕ.succ n) Sign.+ = x + mkℤ (ℕ.succ n) Sign.- 314 | x - mkℤ (ℕ.succ n) Sign.- = x + mkℤ (ℕ.succ n) Sign.+ 315 | 316 | -[-x]=x : ∀ {x} → - (- x) ≡ x 317 | -[-x]=x {mkℤ ℕ.zero tt} = refl 318 | -[-x]=x {mkℤ (ℕ.succ n) Sign.+} = refl 319 | -[-x]=x {mkℤ (ℕ.succ n) Sign.-} = refl 320 | 321 | -x≡[0-x] : ∀ {x} → - x ≡ fromℕ 0 - x 322 | -x≡[0-x] {mkℤ ℕ.zero tt} = refl 323 | -x≡[0-x] {mkℤ (ℕ.succ n) Sign.+} = refl 324 | -x≡[0-x] {mkℤ (ℕ.succ n) Sign.-} = refl 325 | 326 | -- Ha! My first use of »rewrite«. Pretty handy, although I don’t fully 327 | -- understand its inner workings yet (despite having seen the with-clause 328 | -- it desugars to). 329 | x-x≡0 : ∀ {x} → x - x ≡ fromℕ 0 330 | x-x≡0 {mkℤ ℕ.zero tt} = refl 331 | x-x≡0 {mkℤ (ℕ.succ n) Sign.+} rewrite xℕ-x≡0 {n} = refl 332 | x-x≡0 {mkℤ (ℕ.succ n) Sign.-} rewrite xℕ-x≡0 {n} = refl 333 | 334 | x+0≡x : RightIdentity _+_ (fromℕ 0) 335 | x+0≡x (mkℤ ℕ.zero tt) = refl 336 | x+0≡x (mkℤ (ℕ.succ n) Sign.+) = refl 337 | x+0≡x (mkℤ (ℕ.succ n) Sign.-) = refl 338 | 339 | 0+x≡x : LeftIdentity _+_ (fromℕ 0) 340 | 0+x≡x (mkℤ ℕ.zero tt) = refl 341 | 0+x≡x (mkℤ (ℕ.succ n) Sign.+) = refl 342 | 0+x≡x (mkℤ (ℕ.succ n) Sign.-) = refl 343 | 344 | comm-+ : Commutative _+_ 345 | comm-+ (mkℤ ℕ.zero tt) (mkℤ ℕ.zero tt) = refl 346 | comm-+ (mkℤ ℕ.zero tt) (mkℤ (ℕ.succ _) _) = refl 347 | comm-+ (mkℤ (ℕ.succ _) _) (mkℤ ℕ.zero tt) = refl 348 | comm-+ (mkℤ (ℕ.succ n) Sign.+) (mkℤ (ℕ.succ n₁) Sign.+) = {! !} 349 | comm-+ (mkℤ (ℕ.succ n) Sign.+) (mkℤ (ℕ.succ n₁) Sign.-) = {! !} 350 | comm-+ (mkℤ (ℕ.succ n) Sign.-) (mkℤ (ℕ.succ n₁) Sign.+) = {! !} 351 | comm-+ (mkℤ (ℕ.succ n) Sign.-) (mkℤ (ℕ.succ n₁) Sign.-) = {! !} 352 | 353 | _*_ : ℤ → ℤ → ℤ 354 | mkℤ ℕ.zero _ * _ = mkℤ ℕ.zero tt 355 | _ * mkℤ ℕ.zero _ = mkℤ ℕ.zero tt 356 | mkℤ (ℕ.succ n₁) s₁ * mkℤ (ℕ.succ n₂) s₂ = (s₁ Sign.* s₂) ◁ (ℕ.succ n₁ ℕ.* ℕ.succ n₂) 357 | 358 | x*1≡x : RightIdentity _*_ (fromℕ 1) 359 | x*1≡x (mkℤ ℕ.zero tt) = refl 360 | x*1≡x (mkℤ (ℕ.succ n) s) rewrite ℕ.x*1≡x n | Sign.x*+≡x s = refl 361 | 362 | 1*x≡x : LeftIdentity _*_ (fromℕ 1) 363 | 1*x≡x (mkℤ ℕ.zero tt) = refl 364 | 1*x≡x (mkℤ (ℕ.succ n) s) rewrite ℕ.1*x≡x n | Sign.+*x≡x s = refl 365 | 366 | comm-* : Commutative _*_ 367 | comm-* (mkℤ ℕ.zero tt) (mkℤ ℕ.zero tt) = refl 368 | comm-* (mkℤ ℕ.zero _) (mkℤ (ℕ.succ _) _) = refl 369 | comm-* (mkℤ (ℕ.succ _) _) (mkℤ ℕ.zero _) = refl 370 | comm-* (mkℤ (ℕ.succ a) Sign.+) (mkℤ (ℕ.succ b) Sign.+) = {! !} 371 | comm-* (mkℤ (ℕ.succ a) Sign.+) (mkℤ (ℕ.succ b) Sign.-) = {! !} 372 | comm-* (mkℤ (ℕ.succ a) Sign.-) (mkℤ (ℕ.succ b) Sign.+) = {! !} 373 | comm-* (mkℤ (ℕ.succ a) Sign.-) (mkℤ (ℕ.succ b) Sign.-) = {! !} 374 | -------------------------------------------------------------------------------- /List.agda: -------------------------------------------------------------------------------- 1 | module List where 2 | 3 | 4 | 5 | open import Nat 6 | open import Function 7 | open import Equality 8 | open import Bool 9 | open import Logic 10 | open Equality.≡-Reasoning 11 | 12 | 13 | 14 | infixr 6 _∷_ 15 | data List {α} (a : Set α) : Set α where 16 | [] : List a 17 | _∷_ : (x : a) → (xs : List a) → List a 18 | {-# BUILTIN LIST List #-} 19 | 20 | foldr : ∀ {α β} {a : Set α} {b : Set β} → b → (a → b → b) → List a → b 21 | foldr z f [] = z 22 | foldr z f (x ∷ xs) = f x (foldr z f xs) 23 | 24 | [_] : ∀ {α} {a : Set α} → a → List a 25 | [ x ] = x ∷ [] 26 | 27 | infixr 5 _++_ 28 | _++_ : ∀ {α} {a : Set α} → List a → List a → List a 29 | [] ++ ys = ys 30 | (x ∷ xs) ++ ys = x ∷ (xs ++ ys) 31 | 32 | NonEmpty : ∀ {α} {A : Set α} (xs : List A) → Set 33 | NonEmpty [] = ⊥ 34 | NonEmpty (_ ∷ _) = ⊤ 35 | 36 | head : ∀ {α} {A : Set α} (xs : List A) → NonEmpty xs → A 37 | head [] () 38 | head (x ∷ _) _ = x 39 | 40 | module head-example where 41 | myList : List ℕ 42 | myList = 0  ∷ 1 ∷ 2 ∷ [] 43 | 44 | myList-head : ℕ 45 | myList-head = head myList (record {}) 46 | 47 | tail : ∀ {α} {A : Set α} (xs : List A) → NonEmpty xs → List A 48 | tail [] () 49 | tail (_ ∷ xs) _ = xs 50 | 51 | private 52 | -- We cannot write the `head` function, proof: 53 | notHead : ¬ (∀ {A} → List A → A) 54 | notHead head = head [] 55 | -- The [] is of type [⊥], and since we claim we can get the first element 56 | -- out of *any* list, we just take the first ⊥ out of that list. 57 | 58 | length : ∀ {α} {a : Set α} → List a → ℕ 59 | length [] = 0 60 | length (_ ∷ xs) = 1 + length xs 61 | 62 | length[xs++ys]≡length[xs]+length[ys] 63 | : ∀ {α} {a : Set α} (xs ys : List a) 64 | → length (xs ++ ys) ≡ length xs + length ys 65 | length[xs++ys]≡length[xs]+length[ys] [] ys = refl 66 | length[xs++ys]≡length[xs]+length[ys] (x ∷ xs) ys = cong succ (length[xs++ys]≡length[xs]+length[ys] xs ys) 67 | 68 | length[xs++ys]≡length[ys++xs] 69 | : ∀ {α} {a : Set α} (xs ys : List a) 70 | → length (xs ++ ys) ≡ length (ys ++ xs) 71 | length[xs++ys]≡length[ys++xs] xs ys = 72 | begin 73 | length (xs ++ ys) 74 | ≡⟨ length[xs++ys]≡length[xs]+length[ys] xs ys ⟩ 75 | length xs + length ys 76 | ≡⟨ comm-+ (length xs) (length ys) ⟩ 77 | length ys + length xs 78 | ≡⟨ symm (length[xs++ys]≡length[xs]+length[ys] ys xs) ⟩ 79 | length (ys ++ xs) 80 | qed 81 | 82 | -- Computationally inefficient, but easier for theorem proving :-) 83 | reverse : ∀ {α} {a : Set α} → List a → List a 84 | reverse [] = [] 85 | reverse (x ∷ xs) = reverse xs ++ [ x ] 86 | 87 | module reverse-equivalence where 88 | 89 | reverse-helper : ∀ {α} {a : Set α} → List a → List a → List a 90 | reverse-helper xs [] = xs 91 | reverse-helper xs (y ∷ ys) = reverse-helper (y ∷ xs) ys 92 | 93 | -- The usual linear reverse implementation 94 | reverse' : ∀ {α} {a : Set α} → List a → List a 95 | reverse' = reverse-helper [] 96 | 97 | length[reverse[xs]]≡length[xs] 98 | : ∀ {α} {a : Set α} (xs : List a) 99 | → length (reverse xs) ≡ length xs 100 | length[reverse[xs]]≡length[xs] [] = refl 101 | length[reverse[xs]]≡length[xs] (x ∷ xs) = 102 | begin 103 | length (reverse xs ++ (x ∷ [])) 104 | ≡⟨ length[xs++ys]≡length[xs]+length[ys] (reverse xs) [ x ] ⟩ 105 | length (reverse xs) + length [ x ] 106 | ≡⟨ refl ⟩ 107 | length (reverse xs) + 1 108 | ≡⟨ comm-+ (length (reverse xs)) 1 ⟩ 109 | 1 + length (reverse xs) 110 | ≡⟨ cong succ (length[reverse[xs]]≡length[xs] xs) ⟩ 111 | succ (length xs) 112 | qed 113 | 114 | map 115 | : ∀ {α β} {a : Set α} {b : Set β} 116 | → (f : a → b) → (xs : List a) → List b 117 | map _ [] = [] 118 | map f (x ∷ xs) = f x ∷ map f xs 119 | 120 | length[map[f][xs]]≡length[xs] 121 | : ∀ {α β} {a : Set α} {b : Set β} (f : a → b) (xs : List a) 122 | → length (map f xs) ≡ length xs 123 | length[map[f][xs]]≡length[xs] _ [] = refl 124 | length[map[f][xs]]≡length[xs] f (x ∷ xs) = cong succ (length[map[f][xs]]≡length[xs] f xs) 125 | 126 | take 127 | : ∀ {α} {a : Set α} 128 | → ℕ → List a → List a 129 | take zero _ = [] 130 | take _ [] = [] 131 | take (succ n) (x ∷ xs) = x ∷ take n xs 132 | 133 | drop' 134 | : ∀ {α} {a : Set α} 135 | → ℕ → List a → List a 136 | drop' zero xs = xs 137 | drop' _ [] = [] 138 | drop' (succ n) (x ∷ xs) = drop' n xs 139 | 140 | splitAt 141 | : ∀ {a : Set} 142 | → ℕ → List a → List a ∧ List a 143 | splitAt zero ys = ([] , ys) 144 | splitAt n [] = ([] , []) 145 | splitAt (succ n) (x ∷ xs) with splitAt n xs 146 | splitAt (succ n) (x ∷ _) | (ys , zs) = (x ∷ ys , zs) 147 | 148 | infixr 3 _⊆_ 149 | infixr 3 _⊈_ 150 | data _⊆_ {a : Set} : List a → List a → Set where 151 | stop : [] ⊆ [] 152 | drop : ∀ {x xs ys} → xs ⊆ ys → xs ⊆ x ∷ ys 153 | keep : ∀ {x xs ys} → xs ⊆ ys → x ∷ xs ⊆ x ∷ ys 154 | 155 | _⊈_ : ∀ {a} → List a → List a → Set 156 | xs ⊈ ys = ¬ (xs ⊆ ys) 157 | 158 | []⊆xs : ∀ {a} {xs : List a} → [] ⊆ xs 159 | []⊆xs {xs = []} = stop 160 | []⊆xs {xs = x ∷ xs} = drop []⊆xs 161 | 162 | drop-head 163 | : ∀ {a} {x : a} {xs ys : List a} 164 | → (_∷_ x xs) ⊆ ys → xs ⊆ ys 165 | drop-head {xs = []} _ = []⊆xs 166 | drop-head {ys = _ ∷ _} (drop ξ∷x∷xs⊆ys) = drop (drop-head ξ∷x∷xs⊆ys) 167 | drop-head {ys = _ ∷ _} (keep x∷xs⊆ys) = drop x∷xs⊆ys 168 | 169 | drop-tail 170 | : ∀ {a} {x : a} {xs ys : List a} 171 | → (_∷_ x xs) ⊆ ys → [ x ] ⊆ ys 172 | drop-tail (drop xs) = drop (drop-tail xs) 173 | drop-tail (keep _) = keep []⊆xs 174 | 175 | -- Case -> autoderive 176 | refl-⊆ : ∀ {a} {xs : List a} → xs ⊆ xs 177 | refl-⊆ {xs = []} = stop 178 | refl-⊆ {xs = _ ∷ _} = keep refl-⊆ 179 | 180 | trans-⊆ 181 | : ∀ {a} {xs ys zs : List a} 182 | → xs ⊆ ys 183 | → ys ⊆ zs 184 | → xs ⊆ zs 185 | trans-⊆ stop stop = stop 186 | trans-⊆ stop (drop q) = drop q 187 | trans-⊆ (drop p) (drop q) = drop (trans-⊆ p (drop-head q)) 188 | trans-⊆ (drop p) (keep q) = drop (trans-⊆ p q) 189 | trans-⊆ (keep p) (drop q) = drop (trans-⊆ (keep p) q) 190 | trans-⊆ (keep p) (keep q) = keep (trans-⊆ p q) 191 | 192 | module Sublist where 193 | infixr 6 _∷_ 194 | data Sublist {A : Set} : List A → Set where 195 | [] : Sublist [] 196 | _∷_ : ∀ x {xs} → Sublist xs → Sublist (x ∷ xs) 197 | skip : ∀ {x xs} → Sublist xs → Sublist (x ∷ xs) 198 | 199 | forget : ∀ {A : Set} {xs : List A} → Sublist xs → List A 200 | forget [] = [] 201 | forget (x ∷ xs) = x ∷ forget xs 202 | forget (skip xs) = forget xs 203 | -- Interesting (caught) bug here if we add the x to the result in error. 204 | -- Type system at work! 205 | 206 | inject : {A : Set} → (xs : List A) → Sublist xs 207 | inject [] = [] 208 | inject (x ∷ xs) = x ∷ inject xs 209 | 210 | forget∘inject≡id : {A : Set} {xs : List A} → forget (inject xs) ≡ xs 211 | forget∘inject≡id {xs = []} = refl 212 | forget∘inject≡id {xs = x ∷ xs} = cong (λ e → x ∷ e) forget∘inject≡id 213 | 214 | head-subset : ∀ {A} {x : A} {xs ys} → x ∷ xs ⊆ ys → [ x ] ⊆ ys 215 | head-subset (drop xs) = drop (head-subset xs) 216 | head-subset (keep _) = keep []⊆xs 217 | 218 | sublist-implies-⊆ 219 | : ∀ {A : Set} {xs : List A} 220 | → (ys : Sublist xs) 221 | → forget ys ⊆ xs 222 | sublist-implies-⊆ [] = stop 223 | sublist-implies-⊆ (_ ∷ ys) = keep (sublist-implies-⊆ ys) 224 | sublist-implies-⊆ (skip ys) = drop (sublist-implies-⊆ ys) 225 | 226 | filter : ∀ {A : Set} → (p : A → Bool) → (xs : List A) → Sublist xs 227 | filter _ [] = [] 228 | filter p (x ∷ xs) with p x 229 | … | true = x ∷ filter p xs 230 | … | false = skip (filter p xs) 231 | 232 | complement 233 | : {A : Set} {xs : List A} 234 | → Sublist xs 235 | → Sublist xs 236 | complement [] = [] 237 | complement (x ∷ ys) = skip (complement ys) 238 | complement (skip {x = x} ys) = x ∷ complement ys 239 | 240 | complement²≡id 241 | : {A : Set} {xs : List A} 242 | → (ys : Sublist xs) 243 | → complement (complement ys) ≡ ys 244 | complement²≡id [] = refl 245 | complement²≡id (_ ∷ ys) rewrite complement²≡id ys = refl 246 | complement²≡id (skip ys) rewrite complement²≡id ys = refl 247 | 248 | sublists : {A : Set} → (xs : List A) → List (Sublist xs) 249 | sublists [] = [] 250 | sublists (x ∷ xs) = map skip (sublists xs) ++ map (_∷_ x) (sublists xs) 251 | 252 | module Element where 253 | 254 | infixr 5 _∈_ 255 | data _∈_ {A : Set} : A → List A → Set where 256 | here : ∀ {e xs} → e ∈ e ∷ xs 257 | there : ∀ {e x xs} → e ∈ xs → e ∈ x ∷ xs 258 | 259 | open Sublist 260 | 261 | inject-∈→⊆ : ∀ {A} {x : A} {xs} → x ∈ xs → [ x ] ⊆ xs 262 | inject-∈→⊆ here = keep []⊆xs 263 | inject-∈→⊆ (there x∈xs) = drop (inject-∈→⊆ x∈xs) 264 | 265 | extract-⊆→∈ : ∀ {A} {x : A} {xs} → [ x ] ⊆ xs → x ∈ xs 266 | extract-⊆→∈ (drop [x]⊆xs) = there (extract-⊆→∈ [x]⊆xs) 267 | extract-⊆→∈ (keep _) = here 268 | 269 | trans-∈ : ∀ {A} {x : A} {xs} {ys} → x ∈ xs → xs ⊆ ys → x ∈ ys 270 | trans-∈ x∈xs xs⊆ys = extract-⊆→∈ (trans-⊆ (inject-∈→⊆ x∈xs) xs⊆ys) 271 | 272 | -- I can’t get this to typecheck without the »as« parameter in the »f«. But 273 | -- doesn’t this make the induction on lists more powerful than the usual 274 | -- »foldr«, since it has the whole rest of the list available on recursion? 275 | -- What’s this called again, a histomorphism or something? 276 | ind-List 277 | : ∀ {α} {A : Set α} 278 | → (C : List A → Set α) 279 | → (z : C []) 280 | → (f : (a : A) → (as : List A) → C as → C (a ∷ as)) 281 | → (xs : List A) 282 | → C xs 283 | ind-List C z f [] = z 284 | ind-List C z f (x ∷ xs) = f x xs (ind-List C z f xs) 285 | 286 | rec-List : ∀ {α} {A : Set α} {C : Set α} → C → (A → C → C) → List A → C 287 | rec-List z _ [] = z 288 | rec-List z f (x ∷ xs) = f x (rec-List z f xs) 289 | 290 | rec-List-ind : ∀ {α} {A : Set α} {C : Set α} → C → (A → C → C) → List A → C 291 | rec-List-ind {C = C} z f xs = ind-List (const C) z (λ a as x → f a x) xs 292 | -------------------------------------------------------------------------------- /Logic.agda: -------------------------------------------------------------------------------- 1 | module Logic where 2 | 3 | open import Algebra 4 | open import Function 5 | open import Agda.Primitive 6 | open import Equality 7 | 8 | module Top where 9 | 10 | record ⊤ : Set where 11 | {-# BUILTIN UNIT ⊤ #-} 12 | 13 | tt : ⊤ 14 | tt = record {} 15 | 16 | ind-⊤ : ∀ {α} {C : (x : ⊤) → Set α} → C tt → (x : ⊤) → C x 17 | ind-⊤ x _ = x 18 | 19 | private 20 | uniqueness-⊤ : (x : ⊤) → tt ≡ x 21 | uniqueness-⊤ _ = refl 22 | 23 | uniqueness-⊤-via-ind : ∀ x → tt ≡ x 24 | uniqueness-⊤-via-ind = ind-⊤ refl 25 | 26 | open Top public 27 | 28 | module Bottom where 29 | 30 | data ⊥ : Set where 31 | 32 | ind-⊥ : ∀ {α} {C : (x : ⊥) → Set α} → (x : ⊥) → C x 33 | ind-⊥ () 34 | 35 | open Bottom public 36 | 37 | exFalso : ∀ {α} {A : Set α} → ⊥ → A 38 | exFalso () 39 | 40 | infix 5 ¬_ 41 | ¬_ : ∀ {α} → Set α → Set α 42 | ¬ a = a → ⊥ 43 | 44 | cancel-¬ : ∀ {α} {P : Set α} → ¬ ¬ ¬ P → ¬ P 45 | cancel-¬ ¬¬¬p = λ p → ¬¬¬p (λ ¬p → ¬p p) 46 | 47 | add-¬ : ∀ {α} {P : Set α} → P → ¬ ¬ P 48 | add-¬ p ¬p = ¬p p 49 | 50 | infixr 2 _,_ 51 | record Σ {α β} (A : Set α) (B : A → Set β) : Set (α ⊔ β) where 52 | constructor _,_ 53 | field 54 | π₁ : A 55 | π₂ : B π₁ 56 | 57 | rec-Σ 58 | : ∀ {α β γ} {A : Set α} {B : A → Set β} 59 | → (C : Set γ) 60 | → (f : (a : A) → (b : B a) → C) 61 | → (x : Σ A B) 62 | → C 63 | rec-Σ C f (x , y) = f x y 64 | 65 | ind-Σ 66 | : ∀ {α β γ} {A : Set α} {B : A → Set β} 67 | → (C : Σ A B → Set γ) 68 | → ((a : A) → (b : B a) → C (a , b)) 69 | → (x : Σ A B) 70 | → C x 71 | ind-Σ C f (x , y) = f x y 72 | 73 | -- Σ constructor, but infer the witness type. 74 | ∃ : ∀ {α β} {A : Set α} → (A → Set β) → Set (α ⊔ β) 75 | ∃ = Σ _ 76 | 77 | module existential-via-universal where 78 | -- »An existential quantification can be expressed as two universal 79 | -- quantifications« 80 | ∃-to-∀ 81 | : ∀ {α β γ} {A : Set α} {P : A → Set β} 82 | → ∃ P 83 | → ∀ (y : Set γ) → (f : ∀ x → P x → y) → y 84 | ∃-to-∀ (a , Pa) _ f = f a Pa 85 | 86 | -- …and vice versa 87 | ∀-to-∃ 88 | : ∀ {α β} {A : Set α} {P : A → Set β} 89 | → (f : ∀ (y : Set (α ⊔ β)) → (∀ (x : A) → P x → y) → y) 90 | → ∃ P 91 | ∀-to-∃ {A = A} {P = P} f = f (Σ A P) (λ x Px → x , Px) 92 | -- Shorter with type inference: ∀-to-∃ f = f _ (_,_) 93 | 94 | -- ==> : ∀-to-∃ (∃-to-∀ x) ≡ x 95 | -- ==> = ? 96 | 97 | -- <== : ∃-to-∀ (∀-to-∃ x) ≡ x 98 | -- <== = ? 99 | 100 | 101 | infix 3 _∧_ 102 | _∧_ : ∀ {α β} → Set α → Set β → Set (α ⊔ β) 103 | A ∧ B = Σ A (const B) 104 | 105 | infix 2 _×_ 106 | _×_ : ∀ {α β} → Set α → Set β → Set (α ⊔ β) 107 | A × B = Σ A (const B) 108 | 109 | infix 2 _∨_ 110 | data _∨_ {α β} (A : Set α) (B : Set β) : Set (α ⊔ β) where 111 | inl : (l : A) → A ∨ B 112 | inr : (r : B) → A ∨ B 113 | 114 | -- »A Σ type has exactly two fields as values« 115 | uniqueness-Σ 116 | : ∀ {α β} {A : Set α} {B : A → Set β} {⟨a,b⟩ : Σ A B} 117 | → (Σ.π₁ ⟨a,b⟩ , Σ.π₂ ⟨a,b⟩) ≡ ⟨a,b⟩ 118 | uniqueness-Σ = refl 119 | 120 | private 121 | uniqueness-via-ind-Σ 122 | : ∀ {α β} {A : Set α} {B : A → Set β} {⟨a,b⟩ : Σ A B} 123 | → (Σ.π₁ ⟨a,b⟩ , Σ.π₂ ⟨a,b⟩) ≡ ⟨a,b⟩ 124 | uniqueness-via-ind-Σ {⟨a,b⟩ = ⟨a,b⟩} 125 | = ind-Σ 126 | (λ ⟨a,b⟩ → (Σ.π₁ ⟨a,b⟩ , Σ.π₂ ⟨a,b⟩ ≡ ⟨a,b⟩)) 127 | (λ _ _ → refl) 128 | ⟨a,b⟩ 129 | 130 | data Dec {α} (P : Set α) : Set α where 131 | yes : ( p : P) → Dec P 132 | no : (¬p : ¬ P) → Dec P 133 | 134 | DNE : ∀ {α} {P : Set α} → Dec P → ¬ ¬ P → P 135 | DNE (yes p) _ = p 136 | DNE (no ¬p) ¬¬p = exFalso (¬¬p ¬p) 137 | 138 | LEM : ∀ {α} {P : Set α} → Dec P → P ∨ ¬ P 139 | LEM (yes p) = inl p 140 | LEM (no ¬p) = inr ¬p 141 | 142 | module Decidable where 143 | 144 | module Unary where 145 | -- »LEM holds for this predicate« 146 | Decidable 147 | : ∀ {α β} {A : Set α} 148 | → (A → Set β) 149 | → Set (α ⊔ β) 150 | Decidable P = ∀ x → Dec (P x) 151 | 152 | module Binary where 153 | -- »LEM holds for this binary relation« 154 | Decidable 155 | : ∀ {α β γ} {A : Set α} {B : Set β} 156 | → (A → B → Set γ) 157 | → Set (α ⊔ β ⊔ γ) 158 | Decidable P = ∀ x y → Dec (P x y) 159 | 160 | module Ternary where 161 | -- »LEM holds for this ternary relation« 162 | Decidable 163 | : ∀ {α β γ δ} {A : Set α} {B : Set β} {C : Set γ} 164 | → (A → B → C → Set δ) 165 | → Set (α ⊔ β ⊔ γ ⊔ δ) 166 | Decidable P = ∀ x y z → Dec (P x y z) 167 | 168 | ∧-assoc-l : ∀ {α β γ} {P : Set α} {Q : Set β} {R : Set γ} → P ∧ (Q ∧ R) → (P ∧ Q) ∧ R 169 | ∧-assoc-l (p , (q , r)) = ((p , q) , r) 170 | 171 | ∧-assoc-r : ∀ {α β γ} {P : Set α} {Q : Set β} {R : Set γ} → (P ∧ Q) ∧ R → P ∧ (Q ∧ R) 172 | ∧-assoc-r ((p , q) , r) = (p , (q , r)) 173 | 174 | ∧-commute : ∀ {α β} {P : Set α} {Q : Set β} → P ∧ Q → Q ∧ P 175 | ∧-commute (p , q) = (q , p) 176 | 177 | ∨-assoc-l : ∀ {α} {P Q R : Set α} → P ∨ (Q ∨ R) → (P ∨ Q) ∨ R 178 | ∨-assoc-l (inl p) = inl (inl p) 179 | ∨-assoc-l (inr (inl q)) = inl (inr q) 180 | ∨-assoc-l (inr (inr r)) = inr r 181 | 182 | ∨-assoc-r : ∀ {α} {P Q R : Set α} → (P ∨ Q) ∨ R → P ∨ (Q ∨ R) 183 | ∨-assoc-r (inl (inl p)) = inl p 184 | ∨-assoc-r (inl (inr q)) = inr (inl q) 185 | ∨-assoc-r (inr r) = inr (inr r) 186 | 187 | ∨-commute : ∀ {α} {P Q : Set α} → P ∨ Q → Q ∨ P 188 | ∨-commute (inl p) = inr p 189 | ∨-commute (inr q) = inl q 190 | 191 | deMorgan₁ : ∀ {α β} {P : Set α} {Q : Set β} → ¬ (P ∨ Q) → (¬ P ∧ ¬ Q) 192 | deMorgan₁ ¬⟨p∨q⟩ = ((λ p → ¬⟨p∨q⟩ (inl p)) , (λ q → ¬⟨p∨q⟩ (inr q))) 193 | 194 | deMorgan₂ : ∀ {α β} {P : Set α} {Q : Set β} → (¬ P ∧ ¬ Q) → ¬ (P ∨ Q) 195 | deMorgan₂ (¬p , _) (inl p) = ¬p p 196 | deMorgan₂ (_ , ¬q) (inr q) = ¬q q 197 | 198 | deMorgan₃ : ∀ {α β} {P : Set α} {Q : Set β} → (¬ P ∨ ¬ Q) → ¬ (P ∧ Q) 199 | deMorgan₃ (inl ¬p) (p , _) = ¬p p 200 | deMorgan₃ (inr ¬q) (_ , q) = ¬q q 201 | 202 | distr-∨∧ : ∀ {α β γ} {P : Set α} {Q : Set β} {R : Set γ} → (P ∨ Q) ∧ R → (P ∧ R) ∨ (Q ∧ R) 203 | distr-∨∧ (inl p , r) = inl (p , r) 204 | distr-∨∧ (inr q , r) = inr (q , r) 205 | 206 | distr-∧∨ : ∀ {α β γ} {P : Set α} {Q : Set β} {R : Set γ} → (P ∧ Q) ∨ R → (P ∨ R) ∧ (Q ∨ R) 207 | distr-∧∨ (inl (p , q)) = (inl p , inl q) 208 | distr-∧∨ (inr r) = (inr r , inr r) 209 | 210 | rdistr-∧∨∧ : ∀ {α β γ} {P : Set α} {Q : Set β} {R : Set γ} → (P ∧ R) ∨ (Q ∧ R) → (P ∨ Q) ∧ R 211 | rdistr-∧∨∧ (inl (p , r)) = (inl p , r) 212 | rdistr-∧∨∧ (inr (q , r)) = (inr q , r) 213 | 214 | rdistr-∨∧∨ : ∀ {α β γ} {P : Set α} {Q : Set β} {R : Set γ} → (P ∨ R) ∧ (Q ∨ R) → (P ∧ Q) ∨ R 215 | rdistr-∨∧∨ (inl p , inl q) = inl (p , q) 216 | rdistr-∨∧∨ (inl p , inr r) = inr r 217 | rdistr-∨∧∨ (inr r , _) = inr r 218 | 219 | module LEM_and_DNE where 220 | -- Some logical exercises from 221 | -- https://www.cs.uoregon.edu/research/summerschool/summer15/notes/AgdaExercises.pdf 222 | 223 | ∀⟨LEM⇒DNE⟩ : ∀ {α} {A : Set α} → (A ∨ ¬ A) → (¬ ¬ A → A) 224 | ∀⟨LEM⇒DNE⟩ (inl a) _ = a 225 | ∀⟨LEM⇒DNE⟩ (inr ¬a) ¬¬a = exFalso (¬¬a ¬a) 226 | 227 | -- Should work: (∀ a. DNE a) → (∀ a. LEM a) 228 | -- Holy shit, autoderive completely solves this 229 | ∀DNE⇒∀LEM : (∀ {α} {A : Set α} → ¬ ¬ A → A) → (∀ {β} {B : Set β} → (B ∨ ¬ B)) 230 | ∀DNE⇒∀LEM = λ ¬¬a⇒a → ¬¬a⇒a (λ ¬⟨b∨¬b⟩ → ¬⟨b∨¬b⟩ (inr (λ b → ¬⟨b∨¬b⟩ (inl b)))) 231 | 232 | -- Should work: (∀ a. LEM a) → (∀ a. DNE a) 233 | -- Clever exercise! We can just commute all the ∀ to the very beginning, and 234 | -- then this becomes a special case of LEM⇒DNE. I don’t fully understand how 235 | -- this works, I think there’s something left to be learned about ∀ here. 236 | ∀LEM⇒∀DNE : (∀ {α} {A : Set α} → (A ∨ ¬ A)) → (∀ {β} {B : Set β} → ¬ ¬ B → B) 237 | ∀LEM⇒∀DNE x = ∀⟨LEM⇒DNE⟩ x 238 | 239 | private 240 | ∀LEM⇒∀DNE-with-with : (∀ (A : Set) → (A ∨ ¬ A)) → (∀ {B : Set} → ¬ ¬ B → B) 241 | ∀LEM⇒∀DNE-with-with x {b} _ with x b 242 | ∀LEM⇒∀DNE-with-with _ _ | inl b = b 243 | ∀LEM⇒∀DNE-with-with _ ¬¬b | inr ¬b = exFalso (¬¬b ¬b) 244 | 245 | -- Woo I’m doing modules! 246 | open LEM_and_DNE 247 | 248 | -- Exercise given as an aside in »how many is two«, a nice article about sets of 249 | -- truth values. 250 | -- http://math.andrej.com/2005/05/16/how-many-is-two/ 251 | andrejsTheorem : ∀ {α} {P : Set α} → ¬ (¬ P ∧ ¬ ¬ P) 252 | andrejsTheorem (¬p , ¬¬p) = ¬¬p ¬p 253 | 254 | -- Π type (dependent function), for when we want to make Agda as awkward to read 255 | -- as HoTT ;-) 256 | Π : ∀ {α β} (A : Set α) → (B : A → Set β) → Set (α ⊔ β) 257 | Π A B = (x : A) → B x 258 | 259 | private 260 | -- Recover the function arrow from the dependent function type Π 261 | _⟶_ : ∀ {α β} → Set α → Set β → Set (α ⊔ β) 262 | A ⟶ B = Π A (const B) 263 | 264 | -- Currying in HoTT notation 265 | :curry: : {A B C : Set} → Π (A × B → C) (λ f → Σ (A → B → C) (λ g → Π A (λ x → Π B (λ y → f (x , y) ≡ g x y)))) 266 | :curry: = λ x → (λ x₁ x₂ → x (x₁ , x₂)) , (λ x₁ x₂ → refl) 267 | -------------------------------------------------------------------------------- /Maybe.agda: -------------------------------------------------------------------------------- 1 | module Maybe where 2 | 3 | open import Agda.Primitive 4 | open import Equality 5 | open import Function 6 | open import FunctorApplicativeMonad 7 | 8 | data Maybe {α} (A : Set α) : Set α where 9 | nothing : Maybe A 10 | just : A → Maybe A 11 | 12 | fmap : ∀ {α} {A B : Set α} → (A → B) → Maybe A → Maybe B 13 | fmap f nothing = nothing 14 | fmap f (just x) = just (f x) 15 | 16 | fmap[id]≡id : ∀ {α} {A : Set α} {x : Maybe A} → fmap id x ≡ x 17 | fmap[id]≡id {x = nothing} = refl 18 | fmap[id]≡id {x = just _} = refl 19 | 20 | Maybe-Functor : ∀ {α} → Functor {α} Maybe 21 | Maybe-Functor = record 22 | { fmap = fmap 23 | ; fmap[id]≡id = fmap[id]≡id } 24 | 25 | pure : ∀ {α} {A : Set α} → A → Maybe A 26 | pure = just 27 | 28 | _⟨*⟩_ : ∀ {α} {A B : Set α} → Maybe (A → B) → Maybe A → Maybe B 29 | just f ⟨*⟩ just x = just (f x) 30 | _ ⟨*⟩ _ = nothing 31 | 32 | fmap-from-Applicative 33 | : ∀ {α} {A B : Set α} {f : A → B} {x : Maybe A} 34 | → pure f ⟨*⟩ x ≡ fmap f x 35 | fmap-from-Applicative {x = nothing} = refl 36 | fmap-from-Applicative {x = just _} = refl 37 | 38 | homomorphism 39 | : ∀ {α} {A B : Set α} {f : A → B} {x : A} 40 | → pure f ⟨*⟩ pure x ≡ pure (f x) 41 | homomorphism = refl 42 | 43 | Maybe-Applicative : ∀ {α} → Applicative {α} Maybe 44 | Maybe-Applicative = record 45 | { isFunctor = Maybe-Functor 46 | ; pure = pure 47 | ; _⟨*⟩_ = _⟨*⟩_ 48 | ; fmap-from-Applicative = fmap-from-Applicative 49 | ; homomorphism = homomorphism 50 | } 51 | -------------------------------------------------------------------------------- /NQueens.agda: -------------------------------------------------------------------------------- 1 | -- A solution of the N-Queens-Problem in Agda. 2 | -- Completely standalone. 3 | 4 | module NQueens where 5 | 6 | data List A : Set where 7 | [] : List A 8 | _∷_ : (x : A) → (xs : List A) → List A 9 | 10 | -- »listConcat« in the blogpost 11 | _++_ : ∀ {A} → List A → List A → List A 12 | [] ++ ys = ys 13 | (x ∷ xs) ++ ys = x ∷ (xs ++ ys) 14 | 15 | -- »listConcatAll« in the post 16 | concat : ∀ {A} → List (List A) → List A 17 | concat [] = [] 18 | concat (xs ∷ xss) = xs ++ concat xss 19 | 20 | module Bool where 21 | data Bool : Set where 22 | true : Bool 23 | false : Bool 24 | 25 | infixr 20 _or_ 26 | _or_ : Bool → Bool → Bool 27 | true or _ = true 28 | false or y = y 29 | 30 | open Bool 31 | 32 | module Nat where 33 | data ℕ : Set where 34 | zero : ℕ 35 | succ : (n : ℕ) → ℕ 36 | -- To recover nice notation like 3 instead of succ (succ (succ zero)) 37 | {-# BUILTIN NATURAL ℕ #-} 38 | 39 | _≟_ : ℕ → ℕ → Bool 40 | zero ≟ zero = true 41 | succ x ≟ succ y = x ≟ y 42 | _ ≟ _ = false 43 | 44 | ∣_-_∣ : ℕ → ℕ → ℕ 45 | ∣ zero - n₂ ∣ = n₂ 46 | ∣ n₁ - zero ∣ = n₁ 47 | ∣ succ n₁ - succ n₂ ∣ = ∣ n₁ - n₂ ∣ 48 | 49 | open Nat 50 | 51 | module Fin where 52 | data Fin : ℕ → Set where 53 | fzero : {n : ℕ} → Fin (succ n) 54 | fsucc : {n : ℕ} → (i : Fin n) → Fin (succ n) 55 | 56 | toℕ : ∀ {n} → Fin n → ℕ 57 | toℕ fzero = 0 58 | toℕ (fsucc n) = succ (toℕ n) 59 | 60 | fromℕ : ∀ n → Fin (succ n) 61 | fromℕ zero = fzero 62 | fromℕ (succ n) = fsucc (fromℕ n) 63 | 64 | raise : ∀ {n} → Fin n → Fin (succ n) 65 | raise fzero = fzero 66 | raise (fsucc x) = fsucc (raise x) 67 | 68 | open Fin 69 | 70 | map : ∀ {A B} → (A → B) → List A → List B 71 | map _ [] = [] 72 | map f (x ∷ xs) = f x ∷ map f xs 73 | 74 | range2 : (n : ℕ) → List (Fin (succ n)) 75 | range2 zero = fzero ∷ [] 76 | range2 (succ n) = fromℕ (succ n) ∷ map raise (range2 n) 77 | 78 | range : (n : ℕ) → List (Fin (succ n)) 79 | range zero = [] 80 | range (succ n) = fzero ∷ map fsucc (range n) 81 | 82 | -- »catMap« in the post 83 | concatMap : ∀ {A B} → (A → List B) → List A → List B 84 | concatMap f xs = concat (map f xs) 85 | 86 | filter : ∀ {A} → (A → Bool) → List A → List A 87 | filter _ [] = [] 88 | filter p (x ∷ xs) with p x 89 | … | true = x ∷ filter p xs 90 | … | false = filter p xs 91 | 92 | none : ∀ {A} → (A → Bool) → List A → Bool 93 | none p xs with filter p xs 94 | … | [] = true 95 | … | _ = false 96 | 97 | -- »Queen n« is a queen on an n×n chess board. 98 | data Queen : ℕ → Set where 99 | queen : ∀ {n} (x y : Fin (succ n)) → Queen (succ n) 100 | 101 | Configuration : ℕ → Set 102 | Configuration n = List (Queen n) 103 | 104 | queensInRow : {n : ℕ} → Fin (succ n) → List (Queen (succ n)) 105 | queensInRow {n} x = map (queen x) (range n) 106 | 107 | threatens : ∀ {n} → Queen (succ n) → Queen (succ n) → Bool 108 | threatens (queen x₁ y₁) (queen x₂ y₂) = (Δx ≟ 0) or (Δy ≟ 0) or (Δx ≟ Δy) 109 | where 110 | Δx = ∣ toℕ x₁ - toℕ x₂ ∣ 111 | Δy = ∣ toℕ y₁ - toℕ y₂ ∣ 112 | 113 | safeToAdd : ∀ {n} → Configuration (succ n) → Queen (succ n) → Bool 114 | safeToAdd c q = none (threatens q) c 115 | 116 | -- Add a queen to an n×n chess board in the x-th row, for all n column positions 117 | -- in that row. Return all working configurations. 118 | addQueenToRow : ∀ {n} → (row : Fin (succ n)) → Configuration (succ n) → List (Configuration (succ n)) 119 | addQueenToRow x c = map (λ q → q ∷ c) (filter (safeToAdd c) (queensInRow x)) 120 | 121 | addQueensToRow : ∀ {n} → (row : Fin (succ n)) → List (Configuration (succ n)) → List (Configuration (succ n)) 122 | addQueensToRow x = concatMap (addQueenToRow x) 123 | 124 | addAllQueensToRows : ∀ {n} → (rows : List (Fin (succ n))) → List (Configuration (succ n)) → List (Configuration (succ n)) 125 | addAllQueensToRows [] c = c 126 | addAllQueensToRows (row ∷ rs) c = addAllQueensToRows rs (addQueensToRow row c) 127 | 128 | data Solution : ∀ {n} → Configuration (succ n) → Set where 129 | validConfiguration : ∀ {n} (c : Configuration (succ n)) → Solution c 130 | 131 | data Unsolvable : Set where 132 | 133 | NQueens : ℕ → Set 134 | NQueens n with addAllQueensToRows (range n) ([] ∷ []) 135 | … | [] = Unsolvable 136 | … | c ∷ _ = Solution c 137 | 138 | solution : NQueens 4 139 | solution = {! !} 140 | -------------------------------------------------------------------------------- /Nat.agda: -------------------------------------------------------------------------------- 1 | module Nat where 2 | 3 | 4 | 5 | open import Algebra 6 | open import Bool using (Bool; ⌊_⌋) 7 | open import Equality 8 | open import Function 9 | open import Logic 10 | open import Termination 11 | 12 | open Equality.≡-Reasoning 13 | open Logic.Decidable.Binary 14 | 15 | 16 | 17 | data ℕ : Set where 18 | zero : ℕ 19 | succ : (n : ℕ) → ℕ 20 | {-# BUILTIN NATURAL ℕ #-} 21 | 22 | equality-of-successors : ∀ {m n} → succ m ≡ succ n → m ≡ n 23 | equality-of-successors refl = refl 24 | 25 | _≟_ : Decidable {A = ℕ} _≡_ 26 | zero ≟ zero = yes refl 27 | zero ≟ succ y = no (λ ()) 28 | succ x ≟ zero = no (λ ()) 29 | succ x ≟ succ y with x ≟ y 30 | succ x ≟ succ .x | yes refl = yes refl 31 | succ x ≟ succ y | no x≢y = no (λ pSucc → x≢y (equality-of-successors pSucc)) 32 | 33 | _==_ : ℕ → ℕ → Bool 34 | zero == zero = Bool.true 35 | succ x == succ y = x == y 36 | _ == _ = Bool.false 37 | -- x == y = ⌊ x ≟ y ⌋ -- This ought to work, but Agda rejects the NATEQUALS definition below. Huh? 38 | {-# BUILTIN NATEQUALS _==_ #-} 39 | 40 | ind-ℕ : ∀ {α} → (C : ℕ → Set α) → C 0 → ((n : ℕ) → C n → C (succ n)) → (x : ℕ) → C x 41 | ind-ℕ _ z _ zero = z 42 | ind-ℕ C z s (succ n) = s n (ind-ℕ C z s n) 43 | 44 | rec-ℕ : {C : Set} → C → (ℕ → C → C) → ℕ → C 45 | rec-ℕ {C} = ind-ℕ (const C) 46 | 47 | module test-fromNat where 48 | test₁ : 1 ≡ succ zero 49 | test₁ = refl 50 | 51 | test₂ : 2 ≡ succ (succ zero) 52 | test₂ = refl 53 | 54 | test₃ : 5 ≡ succ (succ (succ (succ (succ zero)))) 55 | test₃ = refl 56 | 57 | module addition where 58 | infix 6 _+_ 59 | _+_ : ℕ → ℕ → ℕ 60 | zero + n = n 61 | succ m + n = succ (m + n) 62 | {-# BUILTIN NATPLUS _+_ #-} 63 | 64 | module test-addition where 65 | test₁ : 1 + 2 ≡ 3 66 | test₁ = refl 67 | 68 | module rec-addition where 69 | _+rec_ : ℕ → ℕ → ℕ 70 | x +rec y = rec-ℕ y (λ _n rest → succ rest) x 71 | 72 | test₁ : 1 +rec 2 ≡ 3 73 | test₁ = refl 74 | 75 | test₂ : 12 +rec 23 ≡ 35 76 | test₂ = refl 77 | 78 | +rec-adds : ∀ x y → x + y ≡ x +rec y 79 | +rec-adds zero y = refl 80 | +rec-adds (succ x) y rewrite +rec-adds x y = refl 81 | 82 | _+ind_ : ℕ → ℕ → ℕ 83 | _+ind_ = ind-ℕ (const (ℕ → ℕ)) id (λ _n rest → succ ∘ rest) 84 | 85 | +ind-adds : ∀ x y → x + y ≡ x +ind y 86 | +ind-adds zero y = refl 87 | +ind-adds (succ x) y rewrite +ind-adds x y = refl 88 | 89 | module alternative-+ where 90 | 91 | _+'_ : ℕ → ℕ → ℕ 92 | zero +' n = n 93 | succ n +' m = n +' succ m 94 | 95 | +'-succ-right : ∀ m n → succ (m +' n) ≡ m +' succ n 96 | +'-succ-right zero n = refl 97 | +'-succ-right (succ m) n = +'-succ-right m (succ n) 98 | 99 | equivalent-to-+ : ∀ m n → m + n ≡ m +' n 100 | equivalent-to-+ zero _ = refl 101 | equivalent-to-+ (succ m) n = begin 102 | succ (m + n) 103 | ≡⟨ cong succ (equivalent-to-+ m n) ⟩ 104 | succ (m +' n) 105 | ≡⟨ +'-succ-right m n ⟩ 106 | (m +' succ n) 107 | qed 108 | 109 | assoc-+ : Associative _+_ 110 | assoc-+ zero _ _ = refl 111 | assoc-+ (succ x) y z = cong succ (assoc-+ x y z) 112 | 113 | private 114 | assoc-+-rewrite : Associative _+_ 115 | assoc-+-rewrite zero _ _ = refl 116 | assoc-+-rewrite (succ x) y z rewrite assoc-+-rewrite x y z = refl 117 | 118 | ℕ-+0-semigroup : Semigroup _+_ 119 | ℕ-+0-semigroup = record { assoc = assoc-+ } 120 | 121 | x+0≡x : RightIdentity _+_ 0 122 | x+0≡x zero = refl 123 | x+0≡x (succ x) = cong succ (x+0≡x x) 124 | 125 | private 126 | x+0≡x-rewrite : RightIdentity _+_ 0 127 | x+0≡x-rewrite zero = refl 128 | x+0≡x-rewrite (succ x) rewrite x+0≡x-rewrite x = refl 129 | 130 | 131 | 0+x≡x : LeftIdentity _+_ 0 132 | 0+x≡x _ = refl 133 | 134 | ℕ-+0-monoid : Monoid _+_ 0 135 | ℕ-+0-monoid = record 136 | { isSemigroup = ℕ-+0-semigroup 137 | ; identity = record { left-id = 0+x≡x ; right-id = x+0≡x } } 138 | 139 | x+[1+y]≡[1+x]+y : ∀ x y → (x + succ y) ≡ (succ x + y) 140 | x+[1+y]≡[1+x]+y zero _ = refl 141 | x+[1+y]≡[1+x]+y (succ x) y rewrite x+[1+y]≡[1+x]+y x y = refl 142 | 143 | comm-+ : Commutative _+_ 144 | comm-+ zero y = symm (x+0≡x y) 145 | comm-+ (succ x) y = begin 146 | succ x + y ≡⟨ refl ⟩ 147 | (1 + x) + y ≡⟨ refl ⟩ 148 | 1 + (x + y) ≡⟨ cong succ (comm-+ x y) ⟩ 149 | 1 + (y + x) ≡⟨ refl ⟩ 150 | (1 + y) + x ≡⟨ refl ⟩ 151 | succ y + x ≡⟨ symm (x+[1+y]≡[1+x]+y y x) ⟩ 152 | y + succ x 153 | qed 154 | 155 | private 156 | comm-+-oneline : Commutative _+_ 157 | comm-+-oneline zero y = symm (x+0≡x y) 158 | comm-+-oneline (succ x) y = trans (cong succ (comm-+-oneline x y)) (symm (x+[1+y]≡[1+x]+y y x)) 159 | 160 | comm-+-rewrite : Commutative _+_ 161 | comm-+-rewrite zero y rewrite comm-+ y 0 = refl 162 | comm-+-rewrite (succ x) y rewrite comm-+ x y | x+[1+y]≡[1+x]+y y x = refl 163 | 164 | ℕ-+0-commutative-monoid : CommutativeMonoid _+_ 0 165 | ℕ-+0-commutative-monoid = record 166 | { isMonoid = ℕ-+0-monoid 167 | ; comm = comm-+ } 168 | 169 | open addition public 170 | 171 | module subtraction where 172 | infix 6 _∸_ 173 | _∸_ : ℕ → ℕ → ℕ 174 | zero ∸ x = 0 175 | x ∸ zero = x 176 | succ a ∸ succ b = a ∸ b 177 | {-# BUILTIN NATMINUS _∸_ #-} 178 | 179 | x∸x≡0 : ∀ x → x ∸ x ≡ 0 180 | x∸x≡0 zero = refl 181 | x∸x≡0 (succ x) = x∸x≡0 x 182 | 183 | x∸0≡x : RightIdentity _∸_ 0 184 | x∸0≡x zero = refl 185 | x∸0≡x (succ x) = refl 186 | 187 | open subtraction public 188 | 189 | [x+y]∸y≡x : ∀ x y → (x + y) ∸ y ≡ x 190 | [x+y]∸y≡x x zero rewrite x+0≡x x | x∸0≡x x = refl 191 | [x+y]∸y≡x x (succ y) rewrite x+[1+y]≡[1+x]+y x y | [x+y]∸y≡x x y = refl 192 | 193 | module inequality where 194 | infix 1 _≤_ 195 | data _≤_ : ℕ → ℕ → Set where 196 | z≤n : ∀ {n} → zero ≤ n 197 | s≤s : ∀ {m n} → m ≤ n → succ m ≤ succ n 198 | 199 | -- Apparently better for well-founded recursion, says the actual stdlib. 200 | -- module ≤' where 201 | -- 202 | -- data _≤'_ (m : ℕ) : ℕ → Set where 203 | -- ≤'-refl : m ≤' m 204 | -- ≤'-step : ∀ {n} (m≤'n : m ≤' n) → m ≤' succ n 205 | -- 206 | -- ≤⇒≤' : ∀ {a b} → a ≤ b → a ≤' b 207 | -- ≤⇒≤' z≤n = {! !} 208 | -- ≤⇒≤' (s≤s x) = {! !} 209 | -- 210 | -- ≤'⇒≤ : ∀ {a b} → a ≤' b → a ≤ b 211 | -- ≤'⇒≤ = {! !} 212 | -- 213 | -- open ≤' 214 | 215 | module comparizon-zoo where 216 | 217 | infix 1 _≰_ 218 | 219 | _≰_ : ℕ → ℕ → Set 220 | x ≰ y = ¬ (x ≤ y) 221 | 222 | _<_ : ℕ → ℕ → Set 223 | a < b = succ a ≤ b 224 | 225 | module untested where 226 | 227 | _≮_ : ℕ → ℕ → Set 228 | a ≮ b = succ b ≤ a 229 | 230 | _>_ : ℕ → ℕ → Set 231 | a > b = succ b ≤ a 232 | 233 | _≯_ : ℕ → ℕ → Set 234 | a ≯ b = a ≤ b 235 | 236 | _≥_ : ℕ → ℕ → Set 237 | a ≥ b = b < a 238 | 239 | _≱_ : ℕ → ℕ → Set 240 | a ≱ b = succ b ≤ a 241 | 242 | open untested public 243 | 244 | open comparizon-zoo public 245 | 246 | pred-≤ : ∀ {m n} → succ m ≤ succ n → m ≤ n 247 | pred-≤ (s≤s x) = x 248 | 249 | n+1>n : ∀ n → ¬ (succ n ≤ n) 250 | n+1>n zero () 251 | n+1>n (succ n) (s≤s x) = n+1>n n x 252 | 253 | refl-≤ : Reflexive _≤_ 254 | refl-≤ {zero} = z≤n 255 | refl-≤ {succ n} = s≤s refl-≤ 256 | 257 | trans-≤ : Transitive _≤_ 258 | trans-≤ z≤n _ = z≤n 259 | trans-≤ (s≤s a≤b) (s≤s b≤c) = s≤s (trans-≤ a≤b b≤c) 260 | 261 | asymm-≤ : ∀ {a b} → a ≤ b → b ≤ a → a ≡ b 262 | asymm-≤ z≤n z≤n = refl 263 | asymm-≤ (s≤s a) (s≤s b) rewrite asymm-≤ a b = refl 264 | 265 | m≡n⇒m≤n : ∀ m n → m ≡ n → m ≤ n 266 | m≡n⇒m≤n m n x = subst (λ e → e ≤ n) (symm x) refl-≤ 267 | 268 | -- Supremely well auto-derivable :-) 269 | m≤m+n : ∀ m n → m ≤ m + n 270 | m≤m+n zero n = z≤n 271 | m≤m+n (succ m) n = s≤s (m≤m+n m n) 272 | 273 | ¬⟨1+m+n≤m⟩ : ∀ m n → ¬ (succ (m + n) ≤ m) 274 | ¬⟨1+m+n≤m⟩ zero n () 275 | ¬⟨1+m+n≤m⟩ (succ m) n (s≤s ¬p) = ¬⟨1+m+n≤m⟩ m n ¬p 276 | 277 | ¬1+m≤m : ∀ m → ¬ (succ m ≤ m) 278 | ¬1+m≤m zero () 279 | ¬1+m≤m (succ m) (s≤s x) = ¬1+m≤m m x 280 | 281 | x≰y⇒x≤y : ∀ {x y} → x ≰ y → y ≤ x 282 | x≰y⇒x≤y {y = zero} x≰y = z≤n 283 | x≰y⇒x≤y {zero} x≰y = exFalso (x≰y z≤n) 284 | x≰y⇒x≤y {succ _} {succ _} x≰y = s≤s (x≰y⇒x≤y (λ x≤y → x≰y (s≤s x≤y))) 285 | 286 | ¬⟨m≤n⟩ : ∀ m n → (x : succ (m + ((n ∸ 1) ∸ m)) ≤ m) → ⊥ 287 | ¬⟨m≤n⟩ m n = ¬⟨1+m+n≤m⟩ m ((n ∸ 1) ∸ m) 288 | 289 | ∃-larger-number : ∀ a → ∃ (λ b → a < b) 290 | ∃-larger-number n = (succ n , refl-≤) 291 | 292 | x-y≡0 : ∀ x y → x ≤ y → x ∸ y ≡ 0 293 | x-y≡0 zero y x₁ = refl 294 | x-y≡0 (succ x) _ (s≤s {n = n} e) = x-y≡0 x n e 295 | 296 | _≤?_ : Decidable _≤_ 297 | zero ≤? zero = yes z≤n 298 | zero ≤? succ y = yes z≤n 299 | succ x ≤? zero = no (λ ()) 300 | succ x ≤? succ y with x ≤? y 301 | … | yes x≤y = yes (s≤s x≤y) 302 | … | no x≰y = no (λ ≤succ → x≰y (pred-≤ ≤succ)) 303 | 304 | module test-≤ where 305 | test₁ : 1 ≤ 2 306 | test₁ = s≤s z≤n 307 | 308 | test₂ : 2 ≤ 7 309 | test₂ = s≤s (s≤s z≤n) 310 | 311 | test₃ : 8 ≤ 8 312 | test₃ = refl-≤ 313 | 314 | test₄ : ¬ (4 ≤ 3) 315 | test₄ (s≤s (s≤s (s≤s ()))) 316 | 317 | test₅ : ¬ (10 ≤ 7) 318 | test₅ = ¬⟨1+m+n≤m⟩ 7 ((10 ∸ 1) ∸ 7) 319 | 320 | -- Try auto-deriving this proof ;-) 321 | test₆ : 222 ≤ 228 322 | test₆ = m≤m+n 222 (228 ∸ 222) 323 | 324 | 325 | module comparizon-zoo-theorems where 326 | 327 | ¬⟨x+y auto :-) 111 | pointwiseApply : ∀ {n} {a b : Set} → Vec (a → b) n → Vec a n → Vec b n 112 | pointwiseApply [] [] = [] 113 | pointwiseApply (f ∷ fs) (x ∷ xs) = f x ∷ pointwiseApply fs xs 114 | 115 | -- Case -> Autoderive :-) 116 | map : ∀ {α β n} {a : Set α} {b : Set β} 117 | → (a → b) → Vec a n → Vec b n 118 | map f [] = [] 119 | map f (x ∷ xs) = f x ∷ map f xs 120 | 121 | -- Case split -> auto :-) 122 | zip : ∀ {n} {a b : Set} → Vec a n → Vec b n → Vec (a ∧ b) n 123 | zip [] [] = [] 124 | zip (x ∷ xs) (y ∷ ys) = (x , y) ∷ zip xs ys 125 | 126 | -- Case split -> auto :-) 127 | zipWith : ∀ {α n} {a b c : Set α} → (a → b → c) → Vec a n → Vec b n → Vec c n 128 | zipWith f [] [] = [] 129 | zipWith f (x ∷ xs) (y ∷ ys) = f x y ∷ zipWith f xs ys 130 | 131 | module zipWith-properties where 132 | zipWith[]⁇≡[] : ∀ {α} {a b c : Set α} {f : a → b → c} {xs : Vec a 0} {ys : Vec b 0} → zipWith f [] ys ≡ [] 133 | zipWith[]⁇≡[] {xs = []} {[]} = refl 134 | 135 | zipWith⁇[]≡[] : ∀ {α} {a b c : Set α} {f : a → b → c} {xs : Vec a 0} {ys : Vec b 0} → zipWith f xs [] ≡ [] 136 | zipWith⁇[]≡[] {xs = []} {[]} = refl 137 | 138 | tabulate : ∀ {n} {a : Set} → (Fin n → a) → Vec a n 139 | tabulate {zero} f = [] 140 | tabulate {succ n} f = f zero ∷ tabulate (f ∘ succ) 141 | 142 | module tabulate-test where 143 | square : ∀ {n} → Fin n → ℕ 144 | square = (λ n → n * n) ∘ toℕ 145 | 146 | test₁ : tabulate square ≡ 0 ∷ 1 ∷ 4 ∷ 9 ∷ 16 ∷ 25 ∷ 36 ∷ [] 147 | test₁ = refl 148 | 149 | transpose : ∀ {α m n} {a : Set α} → Vec (Vec a m) n → Vec (Vec a n) m 150 | transpose [] = replicate [] 151 | transpose (xs ∷ xss) = zipWith _∷_ xs (transpose xss) 152 | 153 | module matrix-test where 154 | testMatrix : Vec (Vec ℕ 3) 4 155 | testMatrix = ((11 ∷ 12 ∷ 13 ∷ []) 156 | ∷ (21 ∷ 22 ∷ 23 ∷ []) 157 | ∷ (31 ∷ 32 ∷ 33 ∷ []) 158 | ∷ (41 ∷ 42 ∷ 43 ∷ []) ∷ []) 159 | 160 | testTranspose : transpose testMatrix 161 | ≡ (11 ∷ 21 ∷ 31 ∷ 41 ∷ []) 162 | ∷ (12 ∷ 22 ∷ 32 ∷ 42 ∷ []) 163 | ∷ (13 ∷ 23 ∷ 33 ∷ 43 ∷ []) ∷ [] 164 | testTranspose = refl 165 | 166 | -- Auto-derivable :-) 167 | index-is-inverse-of-tabulate 168 | : ∀ {n} {a : Set} 169 | → (f : Fin n → a) (i : Fin n) → index i (tabulate f) ≡ f i 170 | index-is-inverse-of-tabulate f zero = refl 171 | index-is-inverse-of-tabulate f (succ i) = index-is-inverse-of-tabulate (f ∘ succ) i 172 | 173 | tabulate-is-inverse-of-index 174 | : ∀ {n} {a : Set} 175 | → (xs : Vec a n) 176 | → tabulate (λ k → index k xs) ≡ xs 177 | tabulate-is-inverse-of-index [] = refl 178 | tabulate-is-inverse-of-index (x ∷ xs) = cong (λ e → x ∷ e) (tabulate-is-inverse-of-index xs) 179 | 180 | -- The pairs are autoderivable when the first component is known 181 | filter : ∀ {α n} {A : Set α} → (A → Bool) → Vec A n → Σ ℕ (Vec A) 182 | filter _ [] = zero , [] 183 | filter p (x ∷ xs) with p x | filter p xs 184 | … | true | (n , vs) = (succ n , x ∷ vs) 185 | … | false | (n , vs) = (n , vs) 186 | 187 | fromList : ∀ {α} {A : Set α} → 𝕃.List A → Σ ℕ (Vec A) 188 | fromList 𝕃.[] = zero , [] 189 | fromList (x 𝕃.∷ xs) with fromList xs 190 | fromList (x 𝕃.∷ xs) | n , vs = succ n , x ∷ vs 191 | 192 | toList : ∀ {α n} {A : Set α} → Vec A n → 𝕃.List A 193 | toList = foldr 𝕃._∷_ 𝕃.[] 194 | 195 | module Sort where 196 | -- module first-attempt where 197 | -- Sorted' : ∀ {n} → Vec ℕ n → Set 198 | -- Sorted' [] = ⊤ 199 | -- Sorted' (x ∷ []) = ⊤ 200 | -- Sorted' (x₁ ∷ x₂ ∷ xs) with ⌊ x₁ ≤? x₂ ⌋ 201 | -- … | true = Sorted' (x₂ ∷ xs) 202 | -- … | false = ⊥ 203 | -- 204 | -- insertSorted : ∀ {n} → ℕ → Vec ℕ n → Vec ℕ (succ n) 205 | -- insertSorted n [] = [ n ] 206 | -- insertSorted n (x ∷ xs) with ⌊ n ≤? x ⌋ 207 | -- … | true = n ∷ x ∷ xs 208 | -- … | false = x ∷ insertSorted n xs 209 | -- 210 | -- module insertSorted-preserves-sorted where 211 | -- test : ∀ {n} (x : ℕ) → (xs : Vec ℕ n) → Sorted' xs → Sorted' (insertSorted x xs) 212 | -- test _ [] _ = tt 213 | -- test x (y ∷ ys) x₂ with ⌊ x ≤? y ⌋ 214 | -- test x (y ∷ ys) x₂ | true = {! !} 215 | -- test x (y ∷ ys) x₂ | false = {! !} 216 | -- 217 | -- sort : ∀ {n} → Vec ℕ n → Vec ℕ n 218 | -- sort [] = [] 219 | -- sort (x ∷ xs) = insertSorted x (sort xs) 220 | -- 221 | -- module second-attempt where 222 | -- 223 | -- data Sorted : ∀ {n} → Vec ℕ n → Set where 224 | -- [≤] : Sorted [] 225 | -- [≤_] : (x : ℕ) → Sorted [ x ] 226 | -- _∷≤_ 227 | -- : ∀ {k : ℕ} x₁ x₂ {xs : Vec ℕ k} 228 | -- → (p : x₁ ≤ x₂) 229 | -- → (ys : Sorted (_∷_ x₂ xs)) 230 | -- → Sorted (_∷_ x₁ (_∷_ x₂ xs)) 231 | -- 232 | -- insertSorted : ∀ {n} → (x : ℕ) → Vec ℕ n → Vec ℕ (succ n) 233 | -- insertSorted x [] = [ x ] 234 | -- insertSorted x (y ∷ ys) with x ≤? y 235 | -- ... | yes _ = x ∷ y ∷ ys 236 | -- ... | no _ = y ∷ insertSorted x ys 237 | -- 238 | -- insertSorted' : ∀ {n} {xs : Vec ℕ n} → (x : ℕ) → Sorted xs → Sorted (insertSorted x xs) 239 | -- insertSorted' zero [≤] = [≤ zero ] 240 | -- insertSorted' zero [≤ x ] = {! !} 241 | -- insertSorted' zero ((x₁ ∷≤ x₂) p x₃) = {! !} 242 | -- insertSorted' (succ x) [≤] = [≤ succ x ] 243 | -- insertSorted' (succ x₁) [≤ x ] = {! !} 244 | -- insertSorted' (succ x) ((x₁ ∷≤ x₂) p x₃) = {! !} 245 | -- 246 | -- sort : ∀ {n} → (xs : Vec ℕ n) → Sorted xs 247 | -- sort [] = [≤] 248 | -- sort (x ∷ xs) with insertSorted x (sort xs) 249 | -- ... | foo = ? 250 | 251 | -- module third-attempt where 252 | -- 253 | -- data Sorted : ∀ {n} → Vec ℕ n → Set where 254 | -- [≤] : Sorted [] 255 | -- [≤_] : (x : ℕ) → Sorted [ x ] 256 | -- _≤_⟨_⟩∷_ 257 | -- : ∀ {k} x₁ x₂ {xs : Vec ℕ k} 258 | -- → (p : x₁ ≤ x₂) 259 | -- → (ys : Sorted (_∷_ x₂ xs)) 260 | -- → Sorted (_∷_ x₁ (_∷_ x₂ xs)) 261 | -- 262 | -- uncons 263 | -- : ∀ {k x₁ x₂} {x₁≤x₂ : x₁ ≤ x₂} {xs : Vec ℕ k} 264 | -- → Sorted (x₁ ∷ x₂ ∷ xs) 265 | -- → ℕ × ((x₁ ≤ x₂) × Sorted (x₂ ∷ xs)) 266 | -- uncons (x ≤ _ ⟨ p ⟩∷ xs) = (x , p , xs) 267 | -- 268 | -- Vecℕₛ : (n : ℕ) → Set 269 | -- Vecℕₛ n = Σ (Vec ℕ n) Sorted 270 | -- 271 | -- merge : ∀ {nxs nys} → (xsₛ : Vecℕₛ nxs) → (ysₛ : Vecℕₛ nys) → Vecℕₛ (nxs + nys) 272 | -- 273 | -- -- Empty list 274 | -- merge ([] , _) ysₛ = ysₛ 275 | -- merge xsₛ ([] , _) = subst Vecℕₛ (comm-+ 0 _) xsₛ 276 | -- 277 | -- -- Merge of two singletons 278 | -- merge (x ∷ .[] , [≤ .x ]) (y ∷ .[] , [≤ .y ]) with x ≤? y 279 | -- … | yes x≤y = (x ∷ y ∷ [] , (x ≤ y ⟨ x≤y ⟩∷ [≤ y ])) 280 | -- … | no x≰y = (y ∷ x ∷ [] , y ≤ x ⟨ ¬⟨x≤y⟩⇒x≤y x≰y ⟩∷ [≤ x ]) 281 | -- 282 | -- -- Merge of a singleton with a list 283 | -- merge (x ∷ .[] , [≤ .x ]) (y₁ ∷ .(y₂ ∷ []) , (.y₁ ≤ y₂ ⟨ p ⟩∷ [≤ .y₂ ])) with x ≤? y₁ 284 | -- … | yes x≤y₁ = (x ∷ y₁ ∷ y₂ ∷ [] , (x ≤ y₁ ⟨ x≤y₁ ⟩∷ (y₁ ≤ y₂ ⟨ p ⟩∷ [≤ y₂ ]))) 285 | -- … | no x≰y₁ with {! !} 286 | -- … | foo = {! !} 287 | -- merge (x ∷ .[] , [≤ .x ]) (y₁ ∷ _ , (.y₁ ≤ y₂ ⟨ p ⟩∷ (.y₂ ≤ x₂ ⟨ p₁ ⟩∷ sy))) = {! !} 288 | -- -- … | yes x≤y₁ = (x ∷ y₁ ∷ y₂ ∷ ys , (x ≤ y₁ ⟨ x≤y₁ ⟩∷ (y₁ ≤ y₂ ⟨ p ⟩∷ sy))) 289 | -- -- … | no x≰y₁ with uncons sy 290 | -- -- … | foo = ? 291 | -- 292 | -- merge (x₁ ∷ _ , (.x₁ ≤ x₂ ⟨ p ⟩∷ sx)) (y ∷ .[] , [≤ .y ]) = {! !} 293 | -- merge (x₁ ∷ _ , (.x₁ ≤ x₂ ⟨ x₁≤x₂ ⟩∷ sx)) (y₁ ∷ _ , (.y₁ ≤ y₂ ⟨ y₁≤y₂ ⟩∷ sy)) = {! !} 294 | -- -- merge (x ∷ xs , sx) (y ∷ ys , sy) = {! sx sy !} 295 | 296 | -- module fourth-attempt where 297 | -- 298 | -- data Sorted : ∀ {n} → Vec ℕ n → Set where 299 | -- [≤] : Sorted [] 300 | -- [≤_] : ∀ x → Sorted [ x ] 301 | -- _≤_⟨_⟩∷_ 302 | -- : ∀ {k} x₁ x₂ {xs : Vec _ k} 303 | -- → (p : x₁ ≤ x₂) 304 | -- → (ys : Sorted (x₂ ∷ xs)) 305 | -- → Sorted (x₁ ∷ x₂ ∷ xs) 306 | -- 307 | -- insert : ∀ {n} {xs : Vec ℕ n} {ys : Vec ℕ (succ n)} x (ys≡x∷xs : ys ≡ x ∷ xs) → Sorted xs → Sorted ys 308 | -- insert x refl [≤] = [≤ x ] 309 | -- insert x refl [≤ y ] with x ≤? y 310 | -- insert x refl [≤ y ] | yes x≤y = x ≤ y ⟨ x≤y ⟩∷ [≤ y ] 311 | -- insert x refl [≤ y ] | no x≰y = {! !} ≤ {! !} ⟨ {! !} ⟩∷ {! !} 312 | -- insert x refl (y₁ ≤ y₂ ⟨ y₁≤y₂ ⟩∷ ys) with x ≤? y₁ 313 | -- insert x refl (y₁ ≤ y₂ ⟨ y₁≤y₂ ⟩∷ ys) | yes x≤y₁ = x ≤ y₁ ⟨ x≤y₁ ⟩∷ (y₁ ≤ y₂ ⟨ y₁≤y₂ ⟩∷ ys) 314 | -- insert x refl (y₁ ≤ y₂ ⟨ y₁≤y₂ ⟩∷ ys) | no x≰y₁ with insert x refl ys 315 | -- ... | foo = {! foo !} 316 | --------------------------------------------------------------------------------