├── .gitignore ├── Everything.agda ├── LICENSE ├── Makefile ├── README.md ├── linear-algebra.agda-lib └── src └── Algebra ├── Linear.agda ├── Linear ├── Category │ ├── End.agda │ └── Vect.agda ├── Construct.agda ├── Construct │ ├── Matrix.agda │ ├── Matrix │ │ └── Square.agda │ └── Vector.agda ├── Core.agda ├── Morphism.agda ├── Morphism │ ├── Bundles.agda │ ├── Bundles │ │ └── VectorSpace.agda │ ├── Definitions.agda │ ├── Properties.agda │ ├── Structures │ │ └── Vect.agda │ └── VectorSpace.agda ├── Space │ ├── FiniteDimensional │ │ ├── .#Hom.agda │ │ ├── Hom.agda │ │ └── Product.agda │ ├── Hom.agda │ └── Product.agda ├── Structures.agda └── Structures │ ├── Bundles.agda │ ├── Bundles │ └── FiniteDimensional.agda │ ├── FiniteDimensional.agda │ ├── GL.agda │ └── VectorSpace.agda └── Structures ├── Bundles └── Field.agda ├── Field.agda └── Field └── Utils.agda /.gitignore: -------------------------------------------------------------------------------- 1 | .DS_Store 2 | **/*.agdai 3 | ignore -------------------------------------------------------------------------------- /Everything.agda: -------------------------------------------------------------------------------- 1 | module Everything where 2 | 3 | import Algebra.Structures.Field 4 | import Algebra.Structures.Bundles.Field 5 | 6 | import Algebra.Linear.Core 7 | 8 | import Algebra.Linear.Space 9 | import Algebra.Linear.Space.Hom 10 | import Algebra.Linear.Space.Product 11 | import Algebra.Linear.Space.FiniteDimensional 12 | import Algebra.Linear.Space.FiniteDimensional.Hom 13 | import Algebra.Linear.Space.FiniteDimensional.Product 14 | 15 | import Algebra.Linear.Construct 16 | import Algebra.Linear.Construct.Vector 17 | import Algebra.Linear.Construct.Matrix 18 | 19 | import Algebra.Linear.Morphism 20 | import Algebra.Linear.Morphism.Definitions 21 | import Algebra.Linear.Morphism.VectorSpace 22 | import Algebra.Linear.Morphism.Bundles 23 | import Algebra.Linear.Morphism.Bundles.VectorSpace 24 | 25 | import Algebra.Linear.Structures 26 | import Algebra.Linear.Structures.VectorSpace 27 | import Algebra.Linear.Structures.FiniteDimensional 28 | import Algebra.Linear.Structures.Bundles 29 | import Algebra.Linear.Structures.Bundles.FiniteDimensional 30 | -------------------------------------------------------------------------------- /LICENSE: -------------------------------------------------------------------------------- 1 | MIT License 2 | 3 | Copyright (c) 2019 felko 4 | 5 | Permission is hereby granted, free of charge, to any person obtaining a copy 6 | of this software and associated documentation files (the "Software"), to deal 7 | in the Software without restriction, including without limitation the rights 8 | to use, copy, modify, merge, publish, distribute, sublicense, and/or sell 9 | copies of the Software, and to permit persons to whom the Software is 10 | furnished to do so, subject to the following conditions: 11 | 12 | The above copyright notice and this permission notice shall be included in all 13 | copies or substantial portions of the Software. 14 | 15 | THE SOFTWARE IS PROVIDED "AS IS", WITHOUT WARRANTY OF ANY KIND, EXPRESS OR 16 | IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF MERCHANTABILITY, 17 | FITNESS FOR A PARTICULAR PURPOSE AND NONINFRINGEMENT. IN NO EVENT SHALL THE 18 | AUTHORS OR COPYRIGHT HOLDERS BE LIABLE FOR ANY CLAIM, DAMAGES OR OTHER 19 | LIABILITY, WHETHER IN AN ACTION OF CONTRACT, TORT OR OTHERWISE, ARISING FROM, 20 | OUT OF OR IN CONNECTION WITH THE SOFTWARE OR THE USE OR OTHER DEALINGS IN THE 21 | SOFTWARE. 22 | -------------------------------------------------------------------------------- /Makefile: -------------------------------------------------------------------------------- 1 | check: 2 | for f in src/**/*.agda ; do agda --library=linear-algebra $$f ; done 3 | 4 | clean: 5 | rm src/**/*.agdai 6 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # linear-algebra 2 | Linear algebra formalization in Agda, as a learning project 3 | -------------------------------------------------------------------------------- /linear-algebra.agda-lib: -------------------------------------------------------------------------------- 1 | name: linear-algebra 2 | depend: standard-library agda-categories 3 | include: src 4 | -------------------------------------------------------------------------------- /src/Algebra/Linear.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Algebra.Linear where 4 | 5 | open import Algebra.Linear.Core public 6 | open import Algebra.Linear.Construct public 7 | open import Algebra.Linear.Morphism public 8 | open import Algebra.Linear.Structures public 9 | -------------------------------------------------------------------------------- /src/Algebra/Linear/Category/End.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Algebra.Structures.Bundles.Field 4 | 5 | module Algebra.Linear.Category.End 6 | {k ℓᵏ} (K : Field k ℓᵏ) 7 | {a ℓ} 8 | where 9 | 10 | open import Relation.Binary using (Setoid) 11 | open import Algebra.Linear.Structures.Bundles 12 | open import Algebra.Linear.Morphism.Bundles K 13 | open import Categories.Category 14 | 15 | open LinearMap 16 | -------------------------------------------------------------------------------- /src/Algebra/Linear/Category/Vect.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Algebra.Structures.Bundles.Field 4 | 5 | module Algebra.Linear.Category.Vect 6 | {k ℓᵏ} (K : Field k ℓᵏ) 7 | {a ℓ} 8 | where 9 | 10 | open import Relation.Binary using (Setoid) 11 | open import Algebra.Linear.Structures.Bundles 12 | open import Algebra.Linear.Morphism.Bundles K 13 | open import Categories.Category 14 | 15 | open LinearMap 16 | 17 | id : ∀ {V : VectorSpace K a ℓ} -> LinearMap V V 18 | id {V = V} = record 19 | { ⟦_⟧ = λ u -> u 20 | ; isLinearMap = record 21 | { isAbelianGroupMorphism = record 22 | { gp-homo = record 23 | { mn-homo = record 24 | { sm-homo = record 25 | { ⟦⟧-cong = λ r -> r 26 | ; ∙-homo = λ x y → VectorSpace.refl V } 27 | ; ε-homo = VectorSpace.refl V } } } 28 | ; ∙-homo = λ c u → VectorSpace.refl V } } 29 | 30 | _∘_ : ∀ {U V W : VectorSpace K a ℓ} -> LinearMap V W -> LinearMap U V -> LinearMap U W 31 | _∘_ {U = U} {V} {W} f g = record 32 | { ⟦_⟧ = λ u → f ⟪$⟫ g ⟪$⟫ u 33 | ; isLinearMap = record 34 | { isAbelianGroupMorphism = record 35 | { gp-homo = record 36 | { mn-homo = record 37 | { sm-homo = record 38 | { ⟦⟧-cong = λ r → ⟦⟧-cong f (⟦⟧-cong g r) 39 | ; ∙-homo = λ u v → VectorSpace.trans W (⟦⟧-cong f (+-homo g u v)) (+-homo f (g ⟪$⟫ u) (g ⟪$⟫ v)) 40 | } 41 | ; ε-homo = VectorSpace.trans W (⟦⟧-cong f (0#-homo g)) (0#-homo f) 42 | } 43 | } 44 | } 45 | ; ∙-homo = λ c u → VectorSpace.trans W (⟦⟧-cong f (∙-homo g c u)) (∙-homo f c (g ⟪$⟫ u)) 46 | } 47 | } 48 | 49 | private 50 | module Assoc 51 | {U V W X : VectorSpace K a ℓ} 52 | {f : LinearMap U V} {g : LinearMap V W} {h : LinearMap W X} 53 | where 54 | 55 | open Setoid (linearMap-setoid U X) 56 | open import Relation.Binary.EqReasoning (linearMap-setoid U X) 57 | 58 | right : ((h ∘ g) ∘ f) ≈ (h ∘ (g ∘ f)) 59 | right r = ⟦⟧-cong h (⟦⟧-cong g (⟦⟧-cong f r)) 60 | 61 | left : (h ∘ (g ∘ f)) ≈ ((h ∘ g) ∘ f) 62 | left r = ⟦⟧-cong h (⟦⟧-cong g (⟦⟧-cong f r)) 63 | 64 | Vect : Category _ _ _ 65 | Vect = record 66 | { Obj = VectorSpace K a ℓ 67 | ; _⇒_ = λ V W → LinearMap V W 68 | ; _≈_ = λ {V} {W} -> Setoid._≈_ (linearMap-setoid V W) 69 | ; id = id 70 | ; _∘_ = _∘_ 71 | ; assoc = λ {U} {V} {W} {X} {f} {g} {h} -> Assoc.right {U} {V} {W} {X} {f} {g} {h} 72 | ; sym-assoc = λ {U} {V} {W} {X} {f} {g} {h} -> Assoc.left {U} {V} {W} {X} {f} {g} {h} 73 | ; identityˡ = λ {U} {V} {f} -> ⟦⟧-cong f 74 | ; identityʳ = λ {U} {V} {f} → ⟦⟧-cong f 75 | ; equiv = λ {V} {W} -> Setoid.isEquivalence (linearMap-setoid V W) 76 | ; ∘-resp-≈ = λ {U} {V} {W} {f} {h} {g} {i} rf rg rx → 77 | VectorSpace.trans W (rf (⟦⟧-cong g rx)) (⟦⟧-cong h (rg (VectorSpace.refl U))) 78 | } 79 | -------------------------------------------------------------------------------- /src/Algebra/Linear/Construct.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Algebra.Linear.Construct where 4 | 5 | import Algebra.Linear.Construct.Vector 6 | import Algebra.Linear.Construct.Matrix 7 | -------------------------------------------------------------------------------- /src/Algebra/Linear/Construct/Matrix.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Algebra.Structures.Bundles.Field 4 | 5 | module Algebra.Linear.Construct.Matrix 6 | {k ℓ} (K : Field k ℓ) 7 | where 8 | 9 | open import Level using (_⊔_) 10 | open import Data.Product hiding (map) 11 | open import Data.Fin using (Fin; toℕ; fromℕ; _≤_) 12 | open import Data.Fin.Properties using (¬Fin0) 13 | open import Data.Nat hiding (_⊔_; _≤_) renaming (_+_ to _+ℕ_; _*_ to _*ℕ_) 14 | open import Data.Nat.Properties using (1+n≢0) 15 | 16 | open import Relation.Binary 17 | import Data.Vec.Relation.Binary.Pointwise.Inductive as PW 18 | open import Algebra.Linear.Core 19 | 20 | import Data.Vec.Properties as VP 21 | 22 | open import Algebra.Structures.Field.Utils K 23 | 24 | import Algebra.Linear.Construct.Vector K as V 25 | open V 26 | using (Vec; zipWith; replicate) 27 | renaming 28 | ( _+_ to _+v_ 29 | ; _∙_ to _∙v_ 30 | ; -_ to -v_ 31 | ) 32 | 33 | open import Relation.Binary.PropositionalEquality as P 34 | using (_≡_; subst; subst-subst-sym; _≗_) 35 | renaming 36 | ( refl to ≡-refl 37 | ; sym to ≡-sym 38 | ; trans to ≡-trans 39 | ) 40 | 41 | import Algebra.Linear.Structures.VectorSpace as VS 42 | 43 | open VS.VectorSpaceField K 44 | 45 | open import Data.Nat.Properties 46 | using 47 | ( ≤-refl 48 | ; ≤-reflexive 49 | ; ≤-antisym 50 | ; n∸n≡0 51 | ; m+[n∸m]≡n 52 | ; m≤m+n 53 | ; suc-injective 54 | ) 55 | renaming 56 | ( +-identityˡ to +ℕ-identityˡ 57 | ; +-identityʳ to +ℕ-identityʳ  58 | ) 59 | 60 | Matrix : ℕ -> ℕ -> Set k 61 | Matrix n p = Vec (Vec K' p) n 62 | 63 | private 64 | M : ℕ -> ℕ -> Set k 65 | M = Matrix 66 | 67 | _≈ʰ_ : ∀ {n p n' p'} (A : M n p) (B : M n' p') → Set (k ⊔ ℓ) 68 | _≈ʰ_ = PW.Pointwise V._≈ʰ_ 69 | 70 | module _ {n p} where 71 | setoid : Setoid k (k ⊔ ℓ) 72 | setoid = record 73 | { Carrier = M n p 74 | ; _≈_ = _≈ʰ_ {n} {p} 75 | ; isEquivalence = PW.isEquivalence (V.≈-isEquiv {p}) n 76 | } 77 | 78 | open Setoid setoid public 79 | renaming 80 | ( refl to ≈-refl 81 | ; sym to ≈-sym 82 | ; trans to ≈-trans 83 | ; reflexive to ≈-reflexive 84 | ; isEquivalence to ≈-isEquiv 85 | ) 86 | 87 | import Algebra.FunctionProperties as FP 88 | 89 | tabulate : ∀ {n p} -> (Fin n -> Fin p -> K') -> M n p 90 | tabulate f = V.tabulate λ i -> V.tabulate λ j -> f i j 91 | 92 | tabulate⁺ : ∀ {n p} {f g : Fin n -> Fin p -> K'} 93 | -> (∀ i j -> f i j ≈ᵏ g i j) 94 | -> tabulate f ≈ tabulate g 95 | tabulate⁺ {0} r = PW.[] 96 | tabulate⁺ {suc n} {p} {f} {g} r = 97 | (PW.tabulate⁺ (r Fin.zero)) PW.∷ tabulate⁺ {n} {p} λ i j → r (Fin.suc i) j 98 | 99 | fromVec : ∀ {n p} -> V.Vec K' (n *ℕ p) -> M n p 100 | fromVec {0} V.[] = V.[] 101 | fromVec {suc n} {p} xs = 102 | let (vp , vnp , _) = V.splitAt p xs 103 | in vp V.∷ fromVec {n} vnp 104 | 105 | toVec : ∀ {n p} -> M n p -> V.Vec K' (n *ℕ p) 106 | toVec {n} {p} = V.concat {m = p} {n = n} 107 | 108 | concat∘fromVec : ∀ {n p} (v : V.Vec K' (n *ℕ p)) -> toVec {n} {p} (fromVec v) V.≈ v 109 | concat∘fromVec {0} V.[] = PW.[] 110 | concat∘fromVec {suc n} {p} v = 111 | let (vn , vnp , r) = V.splitAt p v 112 | in begin 113 | toVec {suc n} {p} (fromVec v) 114 | ≡⟨⟩ 115 | vn V.++ toVec {n} {p} (fromVec vnp) 116 | ≈⟨ V.++-cong {p} {n *ℕ p} V.≈-refl (concat∘fromVec {n} {p} vnp) ⟩ 117 | vn V.++ vnp 118 | ≈⟨ V.≈-sym (V.≈-reflexive r) ⟩ 119 | v 120 | ∎ 121 | where open import Relation.Binary.EqReasoning (V.setoid (suc n *ℕ p)) 122 | 123 | _++_ : ∀ {n p q} -> M n p -> M n q -> M n (p +ℕ q) 124 | _++_ = zipWith V._++_ 125 | 126 | _‡_ : ∀ {n m p} -> M n p -> M m p -> M (n +ℕ m) p 127 | _‡_ = V._++_ 128 | 129 | lookup : ∀ {n p} -> M n p -> Fin n -> Fin p -> K' 130 | lookup A i j = V.lookup (V.lookup A i) j 131 | 132 | _⟪_,_⟫ : ∀ {n p} -> M n p -> Fin n -> Fin p -> K' 133 | _⟪_,_⟫ = lookup 134 | 135 | lookup-cong : ∀ {n p} {A B : M n p} (i : Fin n) (j : Fin p) -> A ≈ B -> (A ⟪ i , j ⟫) ≈ᵏ (B ⟪ i , j ⟫) 136 | lookup-cong i j rs = PW.lookup (PW.lookup rs i) j 137 | 138 | tabulate∘lookup : ∀ {n p} (A : M n p) → tabulate (lookup A) ≡ A 139 | tabulate∘lookup A = 140 | begin 141 | tabulate (lookup A) 142 | ≡⟨⟩ 143 | V.tabulate (λ i -> V.tabulate λ j -> V.lookup (V.lookup A i) j) 144 | ≡⟨ VP.tabulate-cong (λ i → VP.tabulate∘lookup (V.lookup A i)) ⟩ 145 | V.tabulate (λ i -> V.lookup A i) 146 | ≡⟨ VP.tabulate∘lookup A ⟩ 147 | A 148 | ∎ 149 | where open import Relation.Binary.PropositionalEquality as Eq 150 | open Eq.≡-Reasoning 151 | 152 | lookup∘tabulate : ∀ {n p} (f : Fin n -> Fin p -> K') (i : Fin n) (j : Fin p) 153 | -> lookup (tabulate f) i j ≡ f i j 154 | lookup∘tabulate {suc n} f i j = 155 | begin 156 | lookup (tabulate f) i j 157 | ≡⟨⟩ 158 | V.lookup (V.lookup (V.tabulate λ i′ -> V.tabulate λ j′ -> f i′ j′) i) j 159 | ≡⟨ cong (λ u → V.lookup u j) (VP.lookup∘tabulate (λ i′ -> V.tabulate λ j′ -> f i′ j′) i) ⟩ 160 | V.lookup (V.tabulate λ j′ -> f i j′) j 161 | ≡⟨ VP.lookup∘tabulate (λ j′ -> f i j′) j ⟩ 162 | f i j 163 | ∎ 164 | where open import Relation.Binary.PropositionalEquality as Eq 165 | open Eq.≡-Reasoning 166 | 167 | tabulate-cong-≡ : ∀ {n p} {f g : Fin n -> Fin p -> K'} -> (∀ i j -> f i j ≡ g i j) -> tabulate f ≡ tabulate g 168 | tabulate-cong-≡ {f = f} {g = g} r = VP.tabulate-cong (λ i → VP.tabulate-cong (λ j → r i j)) 169 | 170 | tabulate-cong : ∀ {n p} {f g : Fin n -> Fin p -> K'} -> (∀ i j -> f i j ≈ᵏ g i j) -> tabulate f ≈ tabulate g 171 | tabulate-cong {f = f} {g = g} r = PW.tabulate⁺ (λ i → V.tabulate-cong (λ j → r i j)) 172 | 173 | transpose : ∀ {n p} -> M n p -> M p n 174 | transpose A = tabulate λ i j -> A ⟪ j , i ⟫ 175 | 176 | _ᵀ : ∀ {n p} -> M n p -> M p n 177 | _ᵀ = transpose 178 | 179 | transpose-involutive : ∀ {n p} (A : M n p) -> ((A ᵀ) ᵀ) ≡ A 180 | transpose-involutive A = 181 | begin 182 | ((A ᵀ)ᵀ) 183 | ≡⟨⟩ 184 | (tabulate λ i j -> (tabulate λ i′ j′ -> A ⟪ j′ , i′ ⟫) ⟪ j , i ⟫) 185 | ≡⟨ tabulate-cong-≡ (λ i j → lookup∘tabulate (λ i′ j′ -> A ⟪ j′ , i′ ⟫) j i) ⟩ 186 | (tabulate λ i j -> A ⟪ i , j ⟫) 187 | ≡⟨ tabulate∘lookup A ⟩ 188 | A 189 | ∎ 190 | where open import Relation.Binary.PropositionalEquality as Eq 191 | open Eq.≡-Reasoning 192 | 193 | map : ∀ {n p} -> (K' -> K') -> M n p -> M n p 194 | map f = V.map (V.map f) 195 | 196 | mapRows : ∀ {n p q} -> (V.Vec K' p -> V.Vec K' q) -> M n p -> M n q 197 | mapRows = V.map 198 | 199 | mapCols : ∀ {n m p} -> (V.Vec K' n -> V.Vec K' m) -> M n p -> M m p 200 | mapCols f A = (V.map f (A ᵀ)) ᵀ 201 | 202 | map-cong : ∀ {n p} {f g : K' -> K'} -> f ≗ g -> map {n} {p} f ≗ map g 203 | map-cong r = VP.map-cong (VP.map-cong r) 204 | 205 | mapRows-cong : ∀ {n p q} {f g : V.Vec K' p -> V.Vec K' q} 206 | -> f ≗ g -> mapRows {n} f ≗ mapRows g 207 | mapRows-cong = VP.map-cong 208 | 209 | mapCols-cong : ∀ {n m p} {f g : V.Vec K' n -> V.Vec K' m} 210 | -> f ≗ g -> mapCols {n} {m} {p} f ≗ mapCols g 211 | mapCols-cong r A = P.cong transpose (VP.map-cong r (transpose A)) 212 | 213 | open import Data.Nat.DivMod 214 | 215 | _+_ : ∀ {n p} -> FP.Op₂ (M n p) 216 | _+_ = zipWith V._+_ 217 | 218 | _∙_ : ∀ {n p} -> ScalarMultiplication K' (M n p) 219 | _∙_ k = V.map (k V.∙_) 220 | 221 | -_ : ∀ {n p} -> FP.Op₁ (M n p) 222 | -_ = V.map V.-_ 223 | 224 | 0# : ∀ {n p} -> M n p 225 | 0# = replicate V.0# 226 | 227 | +-cong : ∀ {n p} {A B C D : M n p} -> A ≈ B -> C ≈ D -> (A + C) ≈ (B + D) 228 | +-cong PW.[] PW.[] = PW.[] 229 | +-cong (r₁ PW.∷ rs₁) (r₂ PW.∷ rs₂) = V.+-cong r₁ r₂ PW.∷ +-cong rs₁ rs₂ 230 | 231 | +-assoc : ∀ {n p} (A B C : M n p) -> ((A + B) + C) ≈ (A + (B + C)) 232 | +-assoc V.[] V.[] V.[] = PW.[] 233 | +-assoc (u V.∷ us) (v V.∷ vs) (w V.∷ ws) = 234 | V.+-assoc u v w PW.∷ +-assoc us vs ws 235 | 236 | +-identityˡ : ∀ {n p} (A : M n p) -> (0# + A) ≈ A 237 | +-identityˡ V.[] = PW.[] 238 | +-identityˡ (u V.∷ us) = V.+-identityˡ u PW.∷ +-identityˡ us 239 | 240 | +-identityʳ : ∀ {n p} (A : M n p) -> (A + 0#) ≈ A 241 | +-identityʳ V.[] = PW.[] 242 | +-identityʳ (u V.∷ us) = V.+-identityʳ u PW.∷ +-identityʳ us 243 | 244 | +-identity : ∀ {n p} -> ((∀ (A : M n p) -> ((0# + A) ≈ A)) × (∀ (A : M n p) -> ((A + 0#) ≈ A))) 245 | +-identity = +-identityˡ , +-identityʳ 246 | 247 | +-comm : ∀ {n p} (A B : M n p) -> (A + B) ≈ (B + A) 248 | +-comm V.[] V.[] = PW.[] 249 | +-comm (u V.∷ us) (v V.∷ vs) = (V.+-comm u v) PW.∷ (+-comm us vs) 250 | 251 | *ᵏ-∙-compat : ∀ {n p} (a b : K') (A : M n p) -> ((a *ᵏ b) ∙ A) ≈ (a ∙ (b ∙ A)) 252 | *ᵏ-∙-compat a b V.[] = PW.[] 253 | *ᵏ-∙-compat a b (u V.∷ us) = (V.*ᵏ-∙-compat a b u) PW.∷ (*ᵏ-∙-compat a b us) 254 | 255 | ∙-+-distrib : ∀ {n p} (a : K') (A B : M n p) -> (a ∙ (A + B)) ≈ ((a ∙ A) + (a ∙ B)) 256 | ∙-+-distrib a V.[] V.[] = PW.[] 257 | ∙-+-distrib a (u V.∷ us) (v V.∷ vs) = (V.∙-+-distrib a u v) PW.∷ (∙-+-distrib a us vs) 258 | 259 | ∙-+ᵏ-distrib : ∀ {n p} (a b : K') (A : M n p) -> ((a +ᵏ b) ∙ A) ≈ ((a ∙ A) + (b ∙ A)) 260 | ∙-+ᵏ-distrib a b V.[] = PW.[] 261 | ∙-+ᵏ-distrib a b (u V.∷ us) = (V.∙-+ᵏ-distrib a b u) PW.∷ (∙-+ᵏ-distrib a b us) 262 | 263 | ∙-cong : ∀ {n p} {a b : K'} {A B : M n p} → a ≈ᵏ b -> A ≈ B -> (a ∙ A) ≈ (b ∙ B) 264 | ∙-cong rᵏ PW.[] = PW.[] 265 | ∙-cong rᵏ (r PW.∷ rs) = (V.∙-cong rᵏ r) PW.∷ (∙-cong rᵏ rs) 266 | 267 | ∙-identity : ∀ {n p} (A : M n p) → (1ᵏ ∙ A) ≈ A 268 | ∙-identity V.[] = PW.[] 269 | ∙-identity (u V.∷ us) = (V.∙-identity u) PW.∷ (∙-identity us) 270 | 271 | ∙-absorbˡ : ∀ {n p} (A : M n p) → (0ᵏ ∙ A) ≈ 0# 272 | ∙-absorbˡ V.[] = PW.[] 273 | ∙-absorbˡ (u V.∷ us) = (V.∙-absorbˡ u) PW.∷ (∙-absorbˡ us) 274 | 275 | -‿inverseˡ : ∀ {n p} (A : M n p) -> ((- A) + A) ≈ 0# 276 | -‿inverseˡ V.[] = PW.[] 277 | -‿inverseˡ (u V.∷ us) = (V.-‿inverseˡ u) PW.∷ (-‿inverseˡ us) 278 | 279 | -‿inverseʳ : ∀ {n p} (A : M n p) -> (A + (- A)) ≈ 0# 280 | -‿inverseʳ V.[] = PW.[] 281 | -‿inverseʳ (u V.∷ us) = (V.-‿inverseʳ u) PW.∷ (-‿inverseʳ us) 282 | 283 | -‿inverse : ∀ {n p} → (∀ (A : M n p) -> ((- A) + A) ≈ 0#) × (∀ (A : M n p) -> (A + (- A)) ≈ 0#) 284 | -‿inverse = -‿inverseˡ , -‿inverseʳ 285 | 286 | -‿cong : ∀ {n p} {A B : M n p} -> A ≈ B -> (- A) ≈ (- B) 287 | -‿cong PW.[] = PW.[] 288 | -‿cong (r PW.∷ rs) = (V.-‿cong r) PW.∷ (-‿cong rs) 289 | 290 | concat-+ : ∀ {n p} (A B : M n p) -> V.concat (V.zipWith V._+_ A B) V.≈ (V.concat A) +v (V.concat B) 291 | concat-+ {0} {p} V.[] V.[] = 292 | V.≈-trans (PW.concat⁺ {m = p} {p = 0} PW.[]) (V.≈-sym (V.+-identityˡ V.[])) 293 | concat-+ {suc n} {p} (u V.∷ us) (v V.∷ vs) = 294 | begin 295 | V.concat (V.zipWith V._+_ (u V.∷ us) (v V.∷ vs)) 296 | ≡⟨⟩ 297 | (u V.+ v) V.++ V.concat (V.zipWith V._+_ us vs) 298 | ≈⟨ V.++-cong V.≈-refl (concat-+ {n} {p} us vs) ⟩ 299 | (u V.+ v) V.++ (V.concat us V.+ V.concat vs) 300 | ≈⟨ V.≈-sym (V.+-++-distrib u (V.concat us) v (V.concat vs)) ⟩ 301 | (V.concat (u V.∷ us)) V.+ (V.concat (v V.∷ vs)) 302 | ∎ 303 | where open import Relation.Binary.EqReasoning (V.setoid (p +ℕ n *ℕ p)) 304 | 305 | concat-0# : ∀ {n p} -> V.concat (0# {n} {p}) V.≈ (V.0# {n *ℕ p}) 306 | concat-0# {0} = PW.[] 307 | concat-0# {suc n} {p} = 308 | begin 309 | V.concat (0# {suc n} {p}) 310 | ≡⟨⟩ 311 | V.0# {p} V.++ V.concat (0# {n} {p}) 312 | ≈⟨ PW.++⁺ V.≈-refl (concat-0# {n} {p}) ⟩ 313 | V.0# {p} V.++ V.0# {n *ℕ p} 314 | ≈⟨ V.0++0≈0 {p} {n *ℕ p} ⟩ 315 | V.0# {p +ℕ n *ℕ p} 316 | ∎ 317 | where open import Relation.Binary.EqReasoning (V.setoid (p +ℕ n *ℕ p)) 318 | 319 | concat-∙ : ∀ {n p} (c : K') (A : M n p) -> V.concat (c ∙ A) V.≈ (c V.∙ V.concat A) 320 | concat-∙ {0} c V.[] = PW.[] 321 | concat-∙ {suc n} {p} c (u V.∷ us) = 322 | begin 323 | V.concat ((c V.∙ u) V.∷ (c ∙ us)) 324 | ≡⟨⟩ 325 | (c V.∙ u) V.++ (V.concat (c ∙ us)) 326 | ≈⟨ V.++-cong V.≈-refl (concat-∙ {n} {p} c us) ⟩ 327 | (c V.∙ u) V.++ (c V.∙ V.concat us) 328 | ≈⟨ V.≈-sym (V.∙-++-distrib c u (V.concat us)) ⟩ 329 | c V.∙ V.concat (u V.∷ us) 330 | ∎ 331 | where open import Relation.Binary.EqReasoning (V.setoid (p +ℕ n *ℕ p)) 332 | 333 | I : ∀ {n p} -> M n p 334 | I = tabulate δ 335 | 336 | _*_ : ∀ {n p q} -> M n p -> M p q -> M n q 337 | A * B = tabulate λ i j -> V.sum-tab λ k -> (A ⟪ i , k ⟫) *ᵏ (B ⟪ k , j ⟫) 338 | 339 | *-cong : ∀ {n p q} {A B : M n p} {C D : M p q} 340 | -> A ≈ B -> C ≈ D -> (A * C) ≈ (B * D) 341 | *-cong {n} {q = q} {A} {B} {C} {D} r₁ r₂ = 342 | begin 343 | A * C 344 | ≡⟨⟩ 345 | tabulate (λ i j -> V.sum-tab λ k -> (A ⟪ i , k ⟫) *ᵏ (C ⟪ k , j ⟫)) 346 | ≈⟨ tabulate⁺ (λ i j -> V.sum-tab-cong λ k → *ᵏ-cong (lookup-cong i k r₁) (lookup-cong k j r₂)) ⟩ 347 | B * D 348 | ∎ 349 | where open import Relation.Binary.EqReasoning (setoid {n} {q}) 350 | 351 | *-assoc : ∀ {n p q r} (A : M n p) (B : M p q) (C : M q r) 352 | -> ((A * B) * C) ≈ (A * (B * C)) 353 | *-assoc {n} {r = r} A B C = tabulate⁺ λ i j -> 354 | begin 355 | V.sum-tab (λ k′ -> ((A * B) ⟪ i , k′ ⟫) *ᵏ (C ⟪ k′ , j ⟫)) 356 | ≈⟨ V.sum-tab-cong (λ k′ -> *ᵏ-cong (≈ᵏ-reflexive (lookup∘tabulate 357 | (λ i′ j′ -> V.sum-tab λ k -> (A ⟪ i′ , k ⟫) *ᵏ (B ⟪ k , j′ ⟫)) i k′)) ≈ᵏ-refl) ⟩ 358 | V.sum-tab (λ k′ -> (V.sum-tab λ k -> (A ⟪ i , k ⟫) *ᵏ (B ⟪ k , k′ ⟫)) *ᵏ (C ⟪ k′ , j ⟫)) 359 | ≈⟨ V.sum-tab-cong (λ k′ -> V.*ᵏ-sum-tab-distribʳ (C ⟪ k′ , j ⟫) λ k -> (A ⟪ i , k ⟫) *ᵏ (B ⟪ k , k′ ⟫)) ⟩ 360 | V.sum-tab (λ k′ -> V.sum-tab λ k -> ((A ⟪ i , k ⟫) *ᵏ (B ⟪ k , k′ ⟫)) *ᵏ (C ⟪ k′ , j ⟫)) 361 | ≈⟨ V.sum-tab-cong (λ k′ -> V.sum-tab-cong λ k -> *ᵏ-assoc (A ⟪ i , k ⟫) (B ⟪ k , k′ ⟫) (C ⟪ k′ , j ⟫)) ⟩ 362 | V.sum-tab (λ k′ -> V.sum-tab λ k -> (A ⟪ i , k ⟫) *ᵏ ((B ⟪ k , k′ ⟫) *ᵏ (C ⟪ k′ , j ⟫))) 363 | ≈⟨ V.sum-tab-cong (λ k′ -> V.sum-tab-cong λ k -> *ᵏ-comm (A ⟪ i , k ⟫) ((B ⟪ k , k′ ⟫) *ᵏ (C ⟪ k′ , j ⟫))) ⟩ 364 | V.sum-tab (λ k′ -> V.sum-tab λ k -> ((B ⟪ k , k′ ⟫) *ᵏ (C ⟪ k′ , j ⟫)) *ᵏ (A ⟪ i , k ⟫)) 365 | ≈⟨ V.sum-tab-swap (λ k k′ -> (B ⟪ k , k′ ⟫) *ᵏ (C ⟪ k′ , j ⟫)) (λ k -> A ⟪ i , k ⟫) ⟩ 366 | V.sum-tab (λ k -> V.sum-tab λ k′ -> ((B ⟪ k , k′ ⟫) *ᵏ (C ⟪ k′ , j ⟫)) *ᵏ (A ⟪ i , k ⟫)) 367 | ≈⟨ V.sum-tab-cong (λ k -> ≈ᵏ-trans 368 | (≈ᵏ-sym (V.*ᵏ-sum-tab-distribʳ (A ⟪ i , k ⟫) λ k′ -> (B ⟪ k , k′ ⟫) *ᵏ (C ⟪ k′ , j ⟫))) 369 | (*ᵏ-comm (V.sum-tab λ k′ -> (B ⟪ k , k′ ⟫) *ᵏ (C ⟪ k′ , j ⟫)) (A ⟪ i , k ⟫))) ⟩ 370 | V.sum-tab (λ k -> (A ⟪ i , k ⟫) *ᵏ V.sum-tab (λ k′ -> (B ⟪ k , k′ ⟫) *ᵏ (C ⟪ k′ , j ⟫))) 371 | ≈⟨ V.sum-tab-cong (λ k -> *ᵏ-cong ≈ᵏ-refl 372 | (≈ᵏ-sym (≈ᵏ-reflexive (lookup∘tabulate 373 | (λ i′ j′ -> V.sum-tab λ k′ -> (B ⟪ i′ , k′ ⟫) *ᵏ (C ⟪ k′ , j′ ⟫)) k j)))) ⟩ 374 | V.sum-tab (λ k -> (A ⟪ i , k ⟫) *ᵏ ((B * C) ⟪ k , j ⟫)) 375 | ∎ 376 | where open import Relation.Binary.EqReasoning (Field.setoid K) 377 | 378 | *-identityˡ : ∀ {n p} (A : M n p) -> I * A ≈ A 379 | *-identityˡ {n} {p} A = 380 | begin 381 | I * A 382 | ≡⟨⟩ 383 | tabulate (λ i j -> V.sum-tab λ k -> (I ⟪ i , k ⟫) *ᵏ (A ⟪ k , j ⟫)) 384 | ≈⟨ tabulate-cong (λ i j -> V.sum-tab-cong {n} λ k -> *ᵏ-cong (≈ᵏ-reflexive (lookup∘tabulate δ i k)) ≈ᵏ-refl) ⟩ 385 | tabulate (λ i j -> V.sum-tab λ k -> δ i k *ᵏ (A ⟪ k , j ⟫)) 386 | ≈⟨ tabulate-cong (λ i j -> V.sum-tab-δ (λ k -> A ⟪ k , j ⟫) i) ⟩ 387 | tabulate (λ i j -> A ⟪ i , j ⟫) 388 | ≡⟨ tabulate∘lookup A ⟩ 389 | A 390 | ∎ 391 | where open import Relation.Binary.EqReasoning (setoid {n} {p}) 392 | 393 | *-identityʳ : ∀ {n p} (A : M n p) -> A * I ≈ A 394 | *-identityʳ {n} {p} A = 395 | begin 396 | A * I 397 | ≡⟨⟩ 398 | tabulate (λ i j -> V.sum-tab λ k -> (A ⟪ i , k ⟫) *ᵏ (I ⟪ k , j ⟫)) 399 | ≈⟨ tabulate-cong (λ i j -> V.sum-tab-cong {p} λ k -> *ᵏ-cong ≈ᵏ-refl (≈ᵏ-reflexive (lookup∘tabulate δ k j))) ⟩ 400 | tabulate (λ i j -> V.sum-tab λ k -> (A ⟪ i , k ⟫) *ᵏ δ k j) 401 | ≈⟨ tabulate-cong (λ i j -> V.sum-tab-cong {p} λ k -> *ᵏ-cong ≈ᵏ-refl (δ-comm k j)) ⟩ 402 | tabulate (λ i j -> V.sum-tab λ k -> (A ⟪ i , k ⟫) *ᵏ δ j k) 403 | ≈⟨ tabulate-cong (λ i j -> ≈ᵏ-trans (V.sum-tab-cong {p} (λ k -> *ᵏ-comm ((A ⟪ i , k ⟫)) (δ j k))) 404 | (V.sum-tab-δ (λ k -> A ⟪ i , k ⟫) j)) ⟩ 405 | tabulate (λ i j -> A ⟪ i , j ⟫) 406 | ≡⟨ tabulate∘lookup A ⟩ 407 | A 408 | ∎ 409 | where open import Relation.Binary.EqReasoning (setoid {n} {p}) 410 | 411 | module _ {n p} where 412 | open IsEquivalence (≈-isEquiv {n} {p}) public 413 | using () 414 | renaming 415 | ( refl to ≈-refl 416 | ; sym to ≈-sym 417 | ; trans to ≈-trans 418 | ) 419 | 420 | open FP (_≈_ {n} {p}) 421 | 422 | open import Algebra.Structures (_≈_ {n} {p}) 423 | open import Algebra.Linear.Structures.Bundles 424 | 425 | isMagma : IsMagma _+_ 426 | isMagma = record 427 | { isEquivalence = ≈-isEquiv 428 | ; ∙-cong = +-cong 429 | } 430 | 431 | isSemigroup : IsSemigroup _+_ 432 | isSemigroup = record 433 | { isMagma = isMagma 434 | ; assoc = +-assoc 435 | } 436 | 437 | isMonoid : IsMonoid _+_ 0# 438 | isMonoid = record 439 | { isSemigroup = isSemigroup 440 | ; identity = +-identity 441 | } 442 | 443 | isGroup : IsGroup _+_ 0# -_ 444 | isGroup = record 445 | { isMonoid = isMonoid 446 | ; inverse = -‿inverse 447 | ; ⁻¹-cong = -‿cong 448 | } 449 | 450 | isAbelianGroup : IsAbelianGroup _+_ 0# -_ 451 | isAbelianGroup = record 452 | { isGroup = isGroup 453 | ; comm = +-comm 454 | } 455 | 456 | open VS K 457 | 458 | isVectorSpace : VS.IsVectorSpace K (_≈_ {n}) _+_ _∙_ -_ 0# 459 | isVectorSpace = record 460 | { isAbelianGroup = isAbelianGroup 461 | ; *ᵏ-∙-compat = *ᵏ-∙-compat 462 | ; ∙-+-distrib = ∙-+-distrib 463 | ; ∙-+ᵏ-distrib = ∙-+ᵏ-distrib 464 | ; ∙-cong = ∙-cong 465 | ; ∙-identity = ∙-identity 466 | ; ∙-absorbˡ = ∙-absorbˡ 467 | } 468 | 469 | vectorSpace : VectorSpace K k (k ⊔ ℓ) 470 | vectorSpace = record { isVectorSpace = isVectorSpace } 471 | 472 | open import Algebra.Linear.Structures.FiniteDimensional K 473 | open import Algebra.Linear.Morphism.VectorSpace K 474 | open import Algebra.Linear.Morphism.Bundles K 475 | 476 | open import Algebra.Morphism.Definitions (M n p) (Vec K' (n *ℕ p)) V._≈_ 477 | open import Algebra.Linear.Morphism.Definitions K (M n p) (Vec K' (n *ℕ p)) V._≈_ 478 | import Relation.Binary.Morphism.Definitions (M n p) (Vec K' (n *ℕ p)) as R 479 | open import Function 480 | open import Relation.Binary.EqReasoning (V.setoid (n *ℕ p)) 481 | 482 | ⟦_⟧ : M n p -> Vec K' (n *ℕ p) 483 | ⟦_⟧ = V.concat 484 | 485 | ⟦⟧-cong : R.Homomorphic₂ _≈_ (V._≈_ {n *ℕ p}) ⟦_⟧ 486 | ⟦⟧-cong = PW.concat⁺ 487 | 488 | +-homo : Homomorphic₂ ⟦_⟧ _+_ _+v_ 489 | +-homo = concat-+ 490 | 491 | 0#-homo : Homomorphic₀ ⟦_⟧ 0# V.0# 492 | 0#-homo = concat-0# {n} {p} 493 | 494 | ∙-homo : ScalarHomomorphism ⟦_⟧ _∙_ _∙v_ 495 | ∙-homo = concat-∙ 496 | 497 | ⟦⟧-injective : Injective (_≈_ {n} {p}) (V._≈_ {n *ℕ p}) ⟦_⟧ 498 | ⟦⟧-injective {A} {B} r = PW.concat⁻ A B r 499 | 500 | ⟦⟧-surjective : Surjective (_≈_ {n} {p}) (V._≈_ {n *ℕ p}) ⟦_⟧ 501 | ⟦⟧-surjective v = fromVec {n} {p} v , concat∘fromVec {n} {p} v 502 | 503 | embed : LinearIsomorphism vectorSpace (V.vectorSpace {n *ℕ p}) 504 | embed = record 505 | { ⟦_⟧ = ⟦_⟧ 506 | ; isLinearIsomorphism = record 507 | { isLinearMonomorphism = record 508 | { isLinearMap = record 509 | { isAbelianGroupMorphism = record 510 | { gp-homo = record 511 | { mn-homo = record 512 | { sm-homo = record 513 | { ⟦⟧-cong = ⟦⟧-cong 514 | ; ∙-homo = +-homo 515 | } 516 | ; ε-homo = 0#-homo 517 | } 518 | } 519 | } 520 | ; ∙-homo = ∙-homo 521 | } 522 | ; injective = ⟦⟧-injective 523 | } 524 | ; surjective = ⟦⟧-surjective 525 | } 526 | } 527 | 528 | isFiniteDimensional : IsFiniteDimensional _≈_ _+_ _∙_ -_ 0# (n *ℕ p) 529 | isFiniteDimensional = record 530 | { isVectorSpace = isVectorSpace 531 | ; embed = embed 532 | } 533 | -------------------------------------------------------------------------------- /src/Algebra/Linear/Construct/Matrix/Square.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Algebra.Structures.Bundles.Field 4 | 5 | module Algebra.Linear.Construct.Matrix.Square 6 | {k ℓ} (K : Field k ℓ) 7 | where 8 | 9 | open import Level using (_⊔_) 10 | open import Data.Nat using (ℕ) 11 | open import Data.Fin 12 | open import Data.Product 13 | 14 | open import Algebra.Structures.Field.Utils K 15 | open import Algebra.Linear.Construct.Matrix K public 16 | 17 | module F = Field K 18 | 19 | private 20 | M : ℕ -> Set k 21 | M n = Matrix n n 22 | 23 | K' : Set k 24 | K' = F.Carrier 25 | 26 | UpperTriangular : ℕ -> Set (k ⊔ ℓ) 27 | UpperTriangular n = ∃ λ (A : M n) -> ∀ {i j} -> j ≤ i -> (A ⟪ i , j ⟫) F.≈ F.0# 28 | 29 | LowerTriangular : ℕ -> Set (k ⊔ ℓ) 30 | LowerTriangular n = ∃ λ (A : M n) -> ∀ {i j} -> i ≤ j -> (A ⟪ i , j ⟫) F.≈ F.0# 31 | 32 | Invertible : ℕ -> Set (k ⊔ ℓ) 33 | Invertible n = ∃₂ λ (A B : M n) -> (A * B) ≈ I n × (B * A) ≈ I n 34 | 35 | trace : ∀ {n} -> M n n -> K' 36 | trace A = vsum (V.tabulate λ i → A ⟪ i , i ⟫) 37 | 38 | det : ∀ {n} -> M n -> K' 39 | det A = {!!} 40 | -------------------------------------------------------------------------------- /src/Algebra/Linear/Construct/Vector.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Algebra.Structures.Bundles.Field 4 | 5 | module Algebra.Linear.Construct.Vector 6 | {k ℓ} (K : Field k ℓ) 7 | where 8 | 9 | open import Level using (_⊔_) 10 | open import Data.Product using (_×_; _,_) 11 | open import Data.Vec public hiding (sum) 12 | import Data.Vec.Properties as VP 13 | import Data.Vec.Relation.Binary.Pointwise.Inductive as PW 14 | open import Data.Nat hiding (_⊔_) renaming (_+_ to _+ℕ_) 15 | open import Data.Fin using (Fin) 16 | 17 | open import Relation.Binary 18 | open import Algebra.Linear.Core 19 | open import Algebra.FunctionProperties 20 | 21 | open import Relation.Binary.PropositionalEquality as P 22 | using (_≡_; subst; subst-subst-sym) 23 | renaming 24 | ( refl to ≡-refl 25 | ; sym to ≡-sym 26 | ; trans to ≡-trans 27 | ) 28 | 29 | open import Algebra.Structures.Field.Utils K 30 | import Algebra.Linear.Structures.VectorSpace as VS 31 | 32 | open VS.VectorSpaceField K 33 | 34 | open import Data.Nat.Properties 35 | using 36 | ( ≤-refl 37 | ; ≤-reflexive 38 | ; ≤-antisym 39 | ; n∸n≡0 40 | ; m+[n∸m]≡n 41 | ; m≤m+n 42 | ; suc-injective 43 | ) 44 | renaming 45 | ( +-identityˡ to +ℕ-identityˡ 46 | ; +-identityʳ to +ℕ-identityʳ  47 | ) 48 | 49 | private 50 | V : ℕ -> Set k 51 | V = Vec K' 52 | 53 | splitAt' : ∀ {n} (i : ℕ) (j : ℕ) -> i +ℕ j ≡ n -> V n → V i × V j 54 | splitAt' {0} 0 0 r [] = ([] , []) 55 | splitAt' {suc n} 0 (suc j) r u = ([] , subst V (≡-sym r) u) 56 | splitAt' {suc n} (suc i) j r (x ∷ xs) = 57 | let (xs₁ , xs₂) = splitAt' {n} i j (suc-injective r) xs 58 | in (x ∷ xs₁ , xs₂) 59 | 60 | open PW public 61 | using () 62 | renaming 63 | ( _∷_ to ≈-cons 64 | ; [] to ≈-null 65 | ; head to ≈-head 66 | ; tail to ≈-tail 67 | ; uncons to ≈-uncons 68 | ; lookup to ≈-lookup 69 | ; map to ≈-map 70 | ) 71 | 72 | _≈ʰ_ : ∀ {m n} (xs : Vec K' m) (ys : Vec K' n) → Set (k ⊔ ℓ) 73 | _≈ʰ_ = PW.Pointwise _≈ᵏ_ 74 | 75 | _≈_ : ∀ {n} (xs : Vec K' n) (ys : Vec K' n) → Set (k ⊔ ℓ) 76 | _≈_ = _≈ʰ_ 77 | 78 | ≈-isEquiv : ∀ {n} -> IsEquivalence (_≈_ {n}) 79 | ≈-isEquiv {n} = PW.isEquivalence ≈ᵏ-isEquiv n 80 | 81 | setoid : ℕ -> Setoid k (k ⊔ ℓ) 82 | setoid n = record { isEquivalence = ≈-isEquiv {n} } 83 | 84 | module _ {n} where 85 | open IsEquivalence (≈-isEquiv {n}) public 86 | using () 87 | renaming 88 | ( refl to ≈-refl 89 | ; sym to ≈-sym 90 | ; trans to ≈-trans 91 | ; reflexive to ≈-reflexive 92 | ) 93 | 94 | 0# : ∀ {n} → V n 95 | 0# {0} = [] 96 | 0# {suc n} = 0ᵏ ∷ 0# {n} 97 | 98 | -_ : ∀ {n} → V n → V n 99 | -_ = map (-ᵏ_) 100 | 101 | infixr 25 _+_ 102 | _+_ : ∀ {n} → Op₂ (V n) 103 | _+_ = zipWith _+ᵏ_ 104 | 105 | infixr 35 _∙_ 106 | _∙_ : ∀ {n} → K' → V n → V n 107 | _∙_ k = map (k *ᵏ_) 108 | 109 | {- 110 | module HeterogeneousEquivalence where 111 | import Relation.Binary.Indexed.Heterogeneous as IH 112 | 113 | data _≈ʰ_ : IH.IRel V (k ⊔ ℓ) where 114 | ≈ʰ-null : ∀ {n p} {u : V n} {v : V p} -> n ≡ 0 -> p ≡ 0 -> u ≈ʰ v 115 | ≈ʰ-cons : ∀ {n p} {x y : K'} {u : V n} {v : V p} -> n ≡ p -> x ≈ᵏ y -> u ≈ʰ v -> (x ∷ u) ≈ʰ (y ∷ v) 116 | 117 | pattern ≈ʰ-zero = ≈ʰ-null ≡-refl ≡-refl 118 | 119 | ≈ʰ-refl : IH.Reflexive V _≈ʰ_ 120 | ≈ʰ-refl {0} {[]} = ≈ʰ-zero 121 | ≈ʰ-refl {suc n} {x ∷ xs} = ≈ʰ-cons (≡-refl) ≈ᵏ-refl (≈ʰ-refl {n} {xs}) 122 | 123 | ≈ʰ-sym : IH.Symmetric V _≈ʰ_ 124 | ≈ʰ-sym {0} ≈ʰ-zero = ≈ʰ-zero 125 | ≈ʰ-sym {suc n} (≈ʰ-cons rn r rs) = ≈ʰ-cons (≡-sym rn) (≈ᵏ-sym r) (≈ʰ-sym rs) 126 | 127 | ≈ʰ-trans : IH.Transitive V _≈ʰ_ 128 | ≈ʰ-trans (≈ʰ-null rn rp) (≈ʰ-null rn' rp') = ≈ʰ-null rn rp' 129 | ≈ʰ-trans (≈ʰ-cons rn r rs) (≈ʰ-cons rn' r' rs') = ≈ʰ-cons (≡-trans rn rn') (≈ᵏ-trans r r') (≈ʰ-trans rs rs') 130 | 131 | ≈ʰ-isEquiv : IH.IsIndexedEquivalence V _≈ʰ_ 132 | ≈ʰ-isEquiv = record 133 | { refl = ≈ʰ-refl 134 | ; sym = ≈ʰ-sym 135 | ; trans = ≈ʰ-trans 136 | } 137 | 138 | ≈-to-≈ʰ : ∀ {n} {u v : V n} -> u ≈ v -> u ≈ʰ v 139 | ≈-to-≈ʰ {0} ≈-null = ≈ʰ-zero 140 | ≈-to-≈ʰ {suc n} (≈-cons r rs) = ≈ʰ-cons ≡-refl r (≈-to-≈ʰ rs) 141 | 142 | ≈ʰ-to-≈ : ∀ {n} {u v : V n} -> u ≈ʰ v -> u ≈ v 143 | ≈ʰ-to-≈ {0} {[]} {[]} ≈ʰ-zero = ≈-null 144 | ≈ʰ-to-≈ {suc n} (≈ʰ-cons _ r rs) = ≈-cons r (≈ʰ-to-≈ rs) 145 | 146 | ≈ʰ-tail : ∀ {n p} {x y : K'} {u : V n} {v : V p} -> (x ∷ u) ≈ʰ (x ∷ v) -> u ≈ʰ v 147 | ≈ʰ-tail (≈ʰ-cons _ _ r) = r 148 | 149 | indexedSetoid : IH.IndexedSetoid ℕ k (ℓ ⊔ k) 150 | indexedSetoid = record 151 | { Carrier = V 152 | ; _≈_ = _≈ʰ_ 153 | ; isEquivalence = ≈ʰ-isEquiv 154 | } 155 | setoid : ℕ -> Setoid k (k ⊔ ℓ) 156 | setoid n = record 157 | { Carrier = V n 158 | ; _≈_ = _≈_ 159 | ; isEquivalence = ≈-isEquiv 160 | } 161 | -} 162 | 163 | ++-identityˡ : ∀ {n} (u : V n) -> ([] ++ u) ≈ u 164 | ++-identityˡ _ = ≈-refl 165 | 166 | ++-cong : ∀ {n p} {u₁ v₁ : V n} {u₂ v₂ : V p} 167 | -> u₁ ≈ v₁ -> u₂ ≈ v₂ -> (u₁ ++ u₂) ≈ (v₁ ++ v₂) 168 | ++-cong ≈-null r₂ = r₂ 169 | ++-cong (≈-cons r₁ rs₁) r₂ = ≈-cons r₁ (++-cong rs₁ r₂) 170 | 171 | ++-split : ∀ {n p} {u₁ v₁ : V n} {u₂ v₂ : V p} 172 | -> (u₁ ++ u₂) ≈ (v₁ ++ v₂) -> (u₁ ≈ v₁ × u₂ ≈ v₂) 173 | ++-split {0} {p} {[]} {[]} r = ≈-null , r 174 | ++-split {suc n} {p} {x ∷ xs} {y ∷ ys} (≈-cons r rs) = 175 | let (r₁ , r₂) = ++-split {n} {p} rs 176 | in (≈-cons r r₁) , r₂ 177 | 178 | 0++0≈0 : ∀ {n p} -> (0# {n} ++ 0# {p}) ≈ 0# {n +ℕ p} 179 | 0++0≈0 {zero} = ++-identityˡ 0# 180 | 0++0≈0 {suc n} = ≈-cons ≈ᵏ-refl (0++0≈0 {n}) 181 | 182 | {- 183 | open HeterogeneousEquivalence 184 | 185 | ++-identityʳ-≈ʰ : ∀ {n} (u : V n) -> (u ++ []) ≈ʰ u 186 | ++-identityʳ-≈ʰ [] = ≈ʰ-refl 187 | ++-identityʳ-≈ʰ {suc n} (x ∷ xs) = ≈ʰ-cons (+ℕ-identityʳ n) ≈ᵏ-refl (++-identityʳ-≈ʰ {n} xs) 188 | 189 | ++-cong-≈ʰ : ∀ {n p} {u₁ v₁ : V n} {u₂ v₂ : V p} 190 | -> u₁ ≈ʰ v₁ -> u₂ ≈ʰ v₂ -> (u₁ ++ u₂) ≈ʰ (v₁ ++ v₂) 191 | ++-cong-≈ʰ r₁ r₂ = ≈-to-≈ʰ (++-cong (≈ʰ-to-≈ r₁) (≈ʰ-to-≈ r₂)) 192 | 193 | splitAt-identityˡ : ∀ {n} (u : V n) -> let (x₁ , x₂) = splitAt 0 z≤n u in (x₁ ≈ʰ []) × (x₂ ≈ʰ u) 194 | splitAt-identityˡ {0}_ = ≈ʰ-zero , ≈ʰ-zero 195 | splitAt-identityˡ {suc n} _ = ≈ʰ-zero , ≈ʰ-refl 196 | 197 | splitAt-identityʳ : ∀ {n} (u : V n) -> let (x₁ , x₂) = splitAt n ≤-refl u in (x₁ ≈ʰ u) × (x₂ ≈ʰ []) 198 | splitAt-identityʳ {0} [] = ≈ʰ-refl , ≈ʰ-refl 199 | splitAt-identityʳ {suc n} (x ∷ xs) = ≈ʰ-cons ≡-refl ≈ᵏ-refl (proj₁ (splitAt-identityʳ xs)) 200 | , ≈ʰ-null (n∸n≡0 (suc n)) ≡-refl 201 | 202 | ++-splitAt : ∀ {n} (k : ℕ) (r : k ≤ n) (u : V n) -> let (x₁ , x₂, _) = splitAt n u in (x₁ ++ x₂) ≡ u 203 | ++-splitAt {0} 0 r u = ≈ʰ-null (m+[n∸m]≡n r) ≡-refl 204 | ++-splitAt {suc n} 0 r (x ∷ xs) = ≈ʰ-refl 205 | ++-splitAt {suc n} (suc k) sk≤sn (x ∷ xs) = 206 | let k≤n = ≤-pred sk≤sn 207 | in ≈ʰ-cons (m+[n∸m]≡n k≤n) ≈ᵏ-refl (++-splitAt k k≤n xs) 208 | 209 | splitAt-++ : ∀ {n p} (u : Vec n) (v : Vec p) -> 210 | let (u' , v') = splitAt n (m≤m+n n p) (u ++ v) in (u ≈ʰ u') × (v ≈ʰ v') 211 | splitAt-++ {0} {0} [] [] = ≈ʰ-zero , ≈ʰ-zero 212 | splitAt-++ {0} {suc p} [] (y ∷ ys) = ≈ʰ-zero , ≈ʰ-refl 213 | splitAt-++ {suc n} {0} (x ∷ xs) [] = 214 | let (rxs , rv) = splitAt-++ xs [] 215 | in ≈ʰ-cons ≡-refl ≈ᵏ-refl rxs , rv 216 | splitAt-++ {suc n} {suc p} (x ∷ xs) (y ∷ ys) = 217 | let (u' , v') = splitAt (suc n) (m≤m+n (suc n) (suc p)) (x ∷ xs ++ y ∷ ys) 218 | (ru , rys) = splitAt-++ {suc n} {p} (x ∷ xs) ys 219 | (rxs , rv) = splitAt-++ {n} {suc p} xs (y ∷ ys) 220 | in ≈ʰ-cons ≡-refl ≈ᵏ-refl rxs , rv 221 | 222 | ++-splitAt' : ∀ {n p} (u : Vec (n +ℕ p)) -> let (x₁ , x₂) = splitAt' n p ≡-refl u in x₁ ++ x₂ ≈ u 223 | ++-splitAt' {0} [] = ≈-null 224 | ++-splitAt' {0} (x ∷ xs) = ≈-refl 225 | ++-splitAt' {suc n} (x ∷ xs) = ≈-cons ≈ᵏ-refl (++-splitAt' {n} xs) 226 | 227 | splitAt'-++ : ∀ {n p} (u : Vec n) (v : Vec p) -> 228 | let (u' , v') = splitAt' n p ≡-refl (u ++ v) in (u ≈ u') × (v ≈ v') 229 | splitAt'-++ {0} {0} [] [] = ≈-null , ≈-null 230 | splitAt'-++ {0} {suc p} [] (y ∷ ys) = ≈-null , ≈-refl 231 | splitAt'-++ {suc n} {0} (x ∷ xs) [] = 232 | let (rxs , rv) = splitAt'-++ xs [] 233 | in ≈-cons ≈ᵏ-refl rxs , rv 234 | splitAt'-++ {suc n} {suc p} (x ∷ xs) (y ∷ ys) = 235 | let (u' , v') = splitAt' (suc n) (suc p) ≡-refl (x ∷ xs ++ y ∷ ys) 236 | (ru , rys) = splitAt'-++ {suc n} {p} (x ∷ xs) ys 237 | (rxs , rv) = splitAt'-++ {n} {suc p} xs (y ∷ ys) 238 | in ≈-cons ≈ᵏ-refl rxs , rv 239 | -} 240 | 241 | +-cong : ∀ {n} → Congruent₂ (_≈_ {n}) _+_ 242 | +-cong ≈-null ≈-null = ≈-null 243 | +-cong (≈-cons r₁ rs₁) (≈-cons r₂ rs₂) = ≈-cons (+ᵏ-cong r₁ r₂) (+-cong rs₁ rs₂) 244 | 245 | +-assoc : ∀ {n} → Associative (_≈_ {n}) _+_ 246 | +-assoc [] [] [] = ≈-null 247 | +-assoc (x ∷ xs) (y ∷ ys) (z ∷ zs) = ≈-cons (+ᵏ-assoc x y z) (+-assoc xs ys zs) 248 | 249 | +-identityˡ : ∀ {n} → LeftIdentity (_≈_ {n}) 0# _+_ 250 | +-identityˡ [] = ≈-null 251 | +-identityˡ (x ∷ xs) = ≈-cons (+ᵏ-identityˡ x) (+-identityˡ xs) 252 | 253 | +-identityʳ : ∀ {n} → RightIdentity (_≈_ {n}) 0# _+_ 254 | +-identityʳ [] = ≈-null 255 | +-identityʳ (x ∷ xs) = ≈-cons (+ᵏ-identityʳ x) (+-identityʳ xs) 256 | 257 | +-identity : ∀ {n} → Identity (_≈_ {n}) 0# _+_ 258 | +-identity = +-identityˡ , +-identityʳ 259 | 260 | +-comm : ∀ {n} -> Commutative (_≈_ {n}) _+_ 261 | +-comm [] [] = ≈-null 262 | +-comm (x ∷ xs) (y ∷ ys) = ≈-cons (+ᵏ-comm x y) (+-comm xs ys) 263 | 264 | *ᵏ-∙-compat : ∀ {n} (a b : K') (u : V n) -> ((a *ᵏ b) ∙ u) ≈ (a ∙ (b ∙ u)) 265 | *ᵏ-∙-compat a b [] = ≈-null 266 | *ᵏ-∙-compat a b (x ∷ xs) = ≈-cons (*ᵏ-assoc a b x) (*ᵏ-∙-compat a b xs) 267 | 268 | ∙-+-distrib : ∀ {n} (a : K') (u v : V n) -> (a ∙ (u + v)) ≈ ((a ∙ u) + (a ∙ v)) 269 | ∙-+-distrib a [] [] = ≈-null 270 | ∙-+-distrib a (x ∷ xs) (y ∷ ys) = ≈-cons (*ᵏ-+ᵏ-distribˡ a x y) (∙-+-distrib a xs ys) 271 | 272 | ∙-+ᵏ-distrib : ∀ {n} (a b : K') (u : V n) -> ((a +ᵏ b) ∙ u) ≈ ((a ∙ u) + (b ∙ u)) 273 | ∙-+ᵏ-distrib a b [] = ≈-null 274 | ∙-+ᵏ-distrib a b (x ∷ u) = ≈-cons (*ᵏ-+ᵏ-distribʳ x a b) (∙-+ᵏ-distrib a b u) 275 | 276 | ∙-cong : ∀ {n} {a b : K'} {u v : V n} → a ≈ᵏ b -> u ≈ v -> (a ∙ u) ≈ (b ∙ v) 277 | ∙-cong rᵏ ≈-null = ≈-null 278 | ∙-cong rᵏ (≈-cons r rs) = ≈-cons (*ᵏ-cong rᵏ r) (∙-cong rᵏ rs) 279 | 280 | ∙-identity : ∀ {n} (x : V n) → (1ᵏ ∙ x) ≈ x 281 | ∙-identity [] = ≈-null 282 | ∙-identity (x ∷ xs) = ≈-cons (*ᵏ-identityˡ x) (∙-identity xs) 283 | 284 | ∙-absorbˡ : ∀ {n} (x : V n) → (0ᵏ ∙ x) ≈ 0# 285 | ∙-absorbˡ [] = ≈-null 286 | ∙-absorbˡ (x ∷ xs) = ≈-cons (*ᵏ-zeroˡ x) (∙-absorbˡ xs) 287 | 288 | -‿inverseˡ : ∀ {n} → LeftInverse (_≈_ {n}) 0# -_ _+_ 289 | -‿inverseˡ [] = ≈-null 290 | -‿inverseˡ (x ∷ xs) = ≈-cons (-ᵏ‿inverseˡ x) (-‿inverseˡ xs) 291 | 292 | -‿inverseʳ : ∀ {n} → RightInverse (_≈_ {n}) 0# -_ _+_ 293 | -‿inverseʳ [] = ≈-null 294 | -‿inverseʳ (x ∷ xs) = ≈-cons (-ᵏ‿inverseʳ x) (-‿inverseʳ xs) 295 | 296 | -‿inverse : ∀ {n} → Inverse (_≈_ {n}) 0# -_ _+_ 297 | -‿inverse = -‿inverseˡ , -‿inverseʳ 298 | 299 | -‿cong : ∀ {n} -> Congruent₁ (_≈_ {n}) -_ 300 | -‿cong ≈-null = ≈-null 301 | -‿cong (≈-cons r rs) = ≈-cons (-ᵏ‿cong r) (-‿cong rs) 302 | 303 | +-++-distrib : ∀ {n p} 304 | (u₁ : V n) (u₂ : V p) 305 | (v₁ : V n) (v₂ : V p) 306 | -> ((u₁ ++ u₂) + (v₁ ++ v₂)) ≈ ((u₁ + v₁) ++ (u₂ + v₂)) 307 | +-++-distrib [] [] [] [] = ≈-null 308 | +-++-distrib (x₁ ∷ xs₁) [] (y₁ ∷ ys₁) [] = ++-cong ≈-refl (+-++-distrib xs₁ [] ys₁ []) 309 | +-++-distrib {0} {suc p} [] (x₂ ∷ xs₂) [] (y₂ ∷ ys₂) = ++-identityˡ ((x₂ +ᵏ y₂) ∷ (xs₂ + ys₂)) 310 | +-++-distrib (x₁ ∷ xs₁) u₂ (y₁ ∷ ys₁) v₂ = 311 | ≈-cons ≈ᵏ-refl (++-cong ≈-refl (+-++-distrib xs₁ u₂ ys₁ v₂)) 312 | 313 | ∙-++-distrib : ∀ {n p} (a : K') (u : V n) (v : V p) -> (a ∙ (u ++ v)) ≈ ((a ∙ u) ++ (a ∙ v)) 314 | ∙-++-distrib a [] v = ≈-refl 315 | ∙-++-distrib a (x ∷ xs) v = ≈-cons ≈ᵏ-refl (∙-++-distrib a xs v) 316 | 317 | sum : ∀ {n} -> V n -> K' 318 | sum = foldr _ _+ᵏ_ 0ᵏ 319 | 320 | sum-tab : ∀ {n} -> (Fin n -> K') -> K' 321 | sum-tab f = sum (tabulate f) 322 | 323 | sum-cong : ∀ {n} {u v : V n} -> u ≈ v -> sum u ≈ᵏ sum v 324 | sum-cong {0} PW.[] = ≈ᵏ-refl 325 | sum-cong {suc n} (r PW.∷ rs) = +ᵏ-cong r (sum-cong {n} rs) 326 | 327 | tabulate-cong-≡ : ∀ {n} {f g : Fin n -> K'} -> (∀ i -> f i ≡ g i) -> tabulate f ≡ tabulate g 328 | tabulate-cong-≡ = VP.tabulate-cong 329 | 330 | tabulate-cong : ∀ {n} {f g : Fin n -> K'} -> (∀ i -> f i ≈ᵏ g i) -> tabulate f ≈ tabulate g 331 | tabulate-cong {0} r = PW.[] 332 | tabulate-cong {suc n} r = r Fin.zero PW.∷ tabulate-cong (λ i → r (Fin.suc i)) 333 | 334 | sum-tab-cong : ∀ {n} {f g : Fin n -> K'} -> (∀ i -> f i ≈ᵏ g i) -> sum-tab f ≈ᵏ sum-tab g 335 | sum-tab-cong r = sum-cong (tabulate-cong r) 336 | 337 | *ᵏ-sum-tab-distribʳ : ∀ {n} (a : K') (f : Fin n -> K') 338 | → (sum-tab f *ᵏ a) ≈ᵏ sum-tab (λ k -> f k *ᵏ a) 339 | *ᵏ-sum-tab-distribʳ {0} a f = *ᵏ-zeroˡ a 340 | *ᵏ-sum-tab-distribʳ {suc n} a f = 341 | begin 342 | sum-tab f *ᵏ a 343 | ≈⟨ *ᵏ-+ᵏ-distribʳ a (f Fin.zero) (sum-tab (λ k -> f (Fin.suc k))) ⟩ 344 | (f Fin.zero *ᵏ a) +ᵏ (sum-tab (λ k -> f (Fin.suc k)) *ᵏ a) 345 | ≈⟨ +ᵏ-cong ≈ᵏ-refl (*ᵏ-sum-tab-distribʳ a (λ k -> f (Fin.suc k))) ⟩ 346 | sum-tab (λ k -> f k *ᵏ a) 347 | ∎ 348 | where open import Relation.Binary.EqReasoning (Field.setoid K) 349 | 350 | sum-assoc : ∀ {n} (f g : Fin n -> K') 351 | -> (sum-tab f +ᵏ sum-tab g) ≈ᵏ sum-tab (λ k -> f k +ᵏ g k) 352 | sum-assoc {0} f g = +ᵏ-identityˡ 0ᵏ 353 | sum-assoc {suc n} f g = 354 | begin 355 | sum-tab f +ᵏ sum-tab g 356 | ≡⟨⟩ 357 | (f Fin.zero +ᵏ sum-tab {n} (λ k -> f (Fin.suc k))) +ᵏ (g Fin.zero +ᵏ sum-tab {n} (λ k -> g (Fin.suc k))) 358 | ≈⟨ +ᵏ-assoc (f Fin.zero) (sum-tab {n} (λ k -> f (Fin.suc k))) (g Fin.zero +ᵏ sum-tab {n} (λ k -> g (Fin.suc k))) ⟩ 359 | f Fin.zero +ᵏ (sum-tab (λ k → f (Fin.suc k)) +ᵏ (g Fin.zero +ᵏ sum-tab (λ k → g (Fin.suc k)))) 360 | ≈⟨ +ᵏ-cong ≈ᵏ-refl (+ᵏ-comm (sum-tab (λ k -> f (Fin.suc k))) ((g Fin.zero +ᵏ sum-tab (λ k → g (Fin.suc k))))) ⟩ 361 | f Fin.zero +ᵏ ((g Fin.zero +ᵏ sum-tab (λ k → g (Fin.suc k))) +ᵏ sum-tab (λ k → f (Fin.suc k))) 362 | ≈⟨ +ᵏ-cong ≈ᵏ-refl (+ᵏ-assoc (g Fin.zero) (sum-tab (λ k -> g (Fin.suc k))) (sum-tab (λ k -> f (Fin.suc k)))) ⟩ 363 | f Fin.zero +ᵏ (g Fin.zero +ᵏ (sum-tab (λ k → g (Fin.suc k)) +ᵏ sum-tab (λ k → f (Fin.suc k)))) 364 | ≈⟨ ≈ᵏ-sym (+ᵏ-assoc (f Fin.zero) (g Fin.zero) (sum-tab (λ k → g (Fin.suc k)) +ᵏ sum-tab (λ k → f (Fin.suc k)))) ⟩ 365 | (f Fin.zero +ᵏ g Fin.zero) +ᵏ (sum-tab (λ k → g (Fin.suc k)) +ᵏ sum-tab (λ k → f (Fin.suc k))) 366 | ≈⟨ +ᵏ-cong ≈ᵏ-refl (+ᵏ-comm (sum-tab λ k -> g (Fin.suc k)) (sum-tab λ k -> f (Fin.suc k))) ⟩ 367 | (f Fin.zero +ᵏ g Fin.zero) +ᵏ (sum-tab (λ k → f (Fin.suc k)) +ᵏ sum-tab (λ k → g (Fin.suc k))) 368 | ≈⟨ +ᵏ-cong ≈ᵏ-refl (sum-assoc {n} (λ k -> f (Fin.suc k)) (λ k -> g (Fin.suc k))) ⟩ 369 | sum-tab (λ k -> f k +ᵏ g k) 370 | ∎ 371 | where open import Relation.Binary.EqReasoning (Field.setoid K) 372 | 373 | sum-tab-0ᵏ : ∀ {n} -> sum-tab {n} (λ k -> 0ᵏ) ≈ᵏ 0ᵏ 374 | sum-tab-0ᵏ {0} = ≈ᵏ-refl 375 | sum-tab-0ᵏ {suc n} = ≈ᵏ-trans (+ᵏ-identityˡ (sum-tab {n} λ k -> 0ᵏ)) (sum-tab-0ᵏ {n}) 376 | 377 | sum-tab-swap : ∀ {n p} (f : Fin n -> Fin p -> K') (g : Fin n -> K') 378 | -> sum-tab (λ k′ -> sum-tab λ k -> f k k′ *ᵏ g k) 379 | ≈ᵏ sum-tab (λ k -> sum-tab λ k′ -> f k k′ *ᵏ g k) 380 | sum-tab-swap {n} {0} f g = ≈ᵏ-sym (sum-tab-0ᵏ {n}) 381 | sum-tab-swap {n} {suc p} f g = 382 | begin 383 | sum-tab (λ k′ -> sum-tab λ k -> f k k′ *ᵏ g k) 384 | ≡⟨⟩ 385 | (sum-tab λ k -> f k Fin.zero *ᵏ g k) +ᵏ (sum-tab λ k′ -> sum-tab λ k -> f k (Fin.suc k′) *ᵏ g k) 386 | ≈⟨ +ᵏ-cong ≈ᵏ-refl (sum-tab-swap {n} {p} (λ k k′ -> f k (Fin.suc k′)) g) ⟩ 387 | sum-tab (λ k → f k Fin.zero *ᵏ g k) +ᵏ sum-tab (λ k → sum-tab (λ k′ → f k (Fin.suc k′) *ᵏ g k)) 388 | ≈⟨ +ᵏ-cong ≈ᵏ-refl (sum-tab-cong λ k -> ≈ᵏ-sym (*ᵏ-sum-tab-distribʳ {p} (g k) λ k′ -> f k (Fin.suc k′))) ⟩ 389 | sum-tab (λ k → f k Fin.zero *ᵏ g k) +ᵏ sum-tab (λ k → sum-tab (λ k′ → f k (Fin.suc k′)) *ᵏ g k) 390 | ≈⟨ sum-assoc (λ k -> f k Fin.zero *ᵏ g k) (λ k -> sum-tab (λ k′ -> f k (Fin.suc k′)) *ᵏ g k) ⟩ 391 | sum-tab (λ k -> (f k Fin.zero *ᵏ g k) +ᵏ (sum-tab (λ k′ -> f k (Fin.suc k′)) *ᵏ g k)) 392 | ≈⟨ sum-tab-cong (λ k -> ≈ᵏ-sym (*ᵏ-+ᵏ-distribʳ (g k) (f k Fin.zero) (sum-tab λ k′ -> f k (Fin.suc k′)))) ⟩ 393 | sum-tab (λ k -> sum-tab (λ k′ -> f k k′) *ᵏ g k) 394 | ≈⟨ sum-tab-cong (λ k -> *ᵏ-sum-tab-distribʳ {suc p} (g k) (λ k′ -> f k k′)) ⟩ 395 | sum-tab (λ k -> sum-tab λ k′ -> f k k′ *ᵏ g k) 396 | ∎ 397 | where open import Relation.Binary.EqReasoning (Field.setoid K) 398 | 399 | sum-tab-δ : ∀ {n} (f : Fin n -> K') (i : Fin n) -> sum-tab (λ k -> δ i k *ᵏ f k) ≈ᵏ f i 400 | sum-tab-δ {suc n} f i@Fin.zero = 401 | begin 402 | sum-tab (λ k -> δ i k *ᵏ f k) 403 | ≡⟨⟩ 404 | (δ i i *ᵏ f i) +ᵏ sum-tab (λ k -> δ i (Fin.suc k) *ᵏ f (Fin.suc k)) 405 | ≈⟨ +ᵏ-cong (*ᵏ-identityˡ (f i)) 406 | (sum-tab-cong (λ k -> ≈ᵏ-trans (*ᵏ-cong (δ-cancelˡ {n} k) ≈ᵏ-refl) (*ᵏ-zeroˡ (f (Fin.suc k))))) ⟩ 407 | f i +ᵏ sum-tab {n} (λ k -> 0ᵏ) 408 | ≈⟨ ≈ᵏ-trans (+ᵏ-cong ≈ᵏ-refl (sum-tab-0ᵏ {n})) (+ᵏ-identityʳ (f i)) ⟩ 409 | f i 410 | ∎ 411 | where open import Relation.Binary.EqReasoning (Field.setoid K) 412 | sum-tab-δ {suc n} f (Fin.suc i) = 413 | begin 414 | sum-tab (λ k -> δ (Fin.suc i) k *ᵏ f k) 415 | ≡⟨⟩ 416 | (δ (Fin.suc i) (Fin.zero {suc n}) *ᵏ f Fin.zero) +ᵏ sum-tab (λ k -> δ (Fin.suc i) (Fin.suc k) *ᵏ f (Fin.suc k)) 417 | ≈⟨ +ᵏ-cong (≈ᵏ-trans (*ᵏ-cong (δ-cancelʳ {p = suc n} (Fin.suc i)) ≈ᵏ-refl) (*ᵏ-zeroˡ (f Fin.zero))) ≈ᵏ-refl ⟩ 418 | 0ᵏ +ᵏ sum-tab (λ k → δ (Fin.suc i) (Fin.suc k) *ᵏ f (Fin.suc k)) 419 | ≈⟨ +ᵏ-identityˡ (sum-tab λ k -> δ (Fin.suc i) (Fin.suc k) *ᵏ f (Fin.suc k)) ⟩ 420 | sum-tab (λ k → δ (Fin.suc i) (Fin.suc k) *ᵏ f (Fin.suc k)) 421 | ≈⟨ sum-tab-δ (λ k -> f (Fin.suc k)) i ⟩ 422 | f (Fin.suc i) 423 | ∎ 424 | where open import Relation.Binary.EqReasoning (Field.setoid K) 425 | 426 | module _ {n} where 427 | open import Algebra.Structures (_≈_ {n}) 428 | open import Algebra.Linear.Structures.Bundles 429 | 430 | isMagma : IsMagma _+_ 431 | isMagma = record 432 | { isEquivalence = ≈-isEquiv 433 | ; ∙-cong = +-cong 434 | } 435 | 436 | isSemigroup : IsSemigroup _+_ 437 | isSemigroup = record 438 | { isMagma = isMagma 439 | ; assoc = +-assoc 440 | } 441 | 442 | isMonoid : IsMonoid _+_ 0# 443 | isMonoid = record 444 | { isSemigroup = isSemigroup 445 | ; identity = +-identity 446 | } 447 | 448 | isGroup : IsGroup _+_ 0# -_ 449 | isGroup = record 450 | { isMonoid = isMonoid 451 | ; inverse = -‿inverse 452 | ; ⁻¹-cong = -‿cong 453 | } 454 | 455 | isAbelianGroup : IsAbelianGroup _+_ 0# -_ 456 | isAbelianGroup = record 457 | { isGroup = isGroup 458 | ; comm = +-comm 459 | } 460 | 461 | open VS K 462 | 463 | isVectorSpace : VS.IsVectorSpace K (_≈_ {n}) _+_ _∙_ -_ 0# 464 | isVectorSpace = record 465 | { isAbelianGroup = isAbelianGroup 466 | ; *ᵏ-∙-compat = *ᵏ-∙-compat 467 | ; ∙-+-distrib = ∙-+-distrib 468 | ; ∙-+ᵏ-distrib = ∙-+ᵏ-distrib 469 | ; ∙-cong = ∙-cong 470 | ; ∙-identity = ∙-identity 471 | ; ∙-absorbˡ = ∙-absorbˡ 472 | } 473 | 474 | vectorSpace : VectorSpace K k (k ⊔ ℓ) 475 | vectorSpace = record { isVectorSpace = isVectorSpace } 476 | 477 | open import Algebra.Linear.Morphism.VectorSpace K 478 | 479 | open import Algebra.Linear.Morphism.Bundles K 480 | 481 | open import Function 482 | 483 | embed : LinearIsomorphism vectorSpace vectorSpace 484 | embed = record 485 | { ⟦_⟧ = id 486 | ; isLinearIsomorphism = record 487 | { isLinearMonomorphism = record 488 | { isLinearMap = record 489 | { isAbelianGroupMorphism = record 490 | { gp-homo = record 491 | { mn-homo = record 492 | { sm-homo = record 493 | { ⟦⟧-cong = id 494 | ; ∙-homo = λ x y → ≈-refl 495 | } 496 | ; ε-homo = ≈-refl 497 | } 498 | } 499 | } 500 | ; ∙-homo = λ c u → ≈-refl 501 | } 502 | ; injective = id 503 | } 504 | ; surjective = λ y → y , ≈-refl 505 | } 506 | } 507 | -------------------------------------------------------------------------------- /src/Algebra/Linear/Core.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Algebra.Linear.Core where 4 | 5 | open import Relation.Binary 6 | open import Relation.Nullary using (¬_) 7 | open import Level using (_⊔_) 8 | open import Data.Fin using (Fin; suc; zero) 9 | 10 | VectorAddition : ∀ {c} -> Set c -> Set c 11 | VectorAddition V = V -> V -> V 12 | 13 | ScalarMultiplication : ∀ {c k} -> Set k -> Set c -> Set (k ⊔ c) 14 | ScalarMultiplication K V = K -> V -> V 15 | -------------------------------------------------------------------------------- /src/Algebra/Linear/Morphism.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Algebra.Structures.Bundles.Field 4 | 5 | module Algebra.Linear.Morphism 6 | {k ℓᵏ} (K : Field k ℓᵏ) 7 | where 8 | 9 | open import Algebra.Linear.Morphism.Definitions K public 10 | open import Algebra.Linear.Morphism.Bundles K public 11 | open import Algebra.Linear.Morphism.VectorSpace K public 12 | -------------------------------------------------------------------------------- /src/Algebra/Linear/Morphism/Bundles.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Algebra.Structures.Bundles.Field 4 | 5 | module Algebra.Linear.Morphism.Bundles 6 | {k ℓᵏ} (K : Field k ℓᵏ) 7 | where 8 | 9 | open import Algebra.FunctionProperties as FP 10 | open import Relation.Binary using (Rel; Setoid) 11 | open import Level 12 | 13 | import Algebra.Linear.Structures.Bundles as VS 14 | 15 | open import Algebra.Linear.Morphism.VectorSpace K 16 | 17 | open import Function.Equality as FunEq using (_⟶_; _⟨$⟩_; setoid) 18 | 19 | module _ 20 | {a₁ ℓ₁} (From : VS.VectorSpace K a₁ ℓ₁) 21 | {a₂ ℓ₂} (To : VS.VectorSpace K a₂ ℓ₂) 22 | where 23 | 24 | private 25 | module F = VS.VectorSpace From 26 | module T = VS.VectorSpace To 27 | 28 | K' : Set k 29 | K' = Field.Carrier K 30 | 31 | open F using () 32 | renaming 33 | ( Carrier to A 34 | ; _≈_ to _≈₁_ 35 | ; refl to ≈₁-refl 36 | ; sym to ≈₁-sym 37 | ; setoid to A-setoid 38 | ; _+_ to +₁ 39 | ; _∙_ to ∙₁ 40 | ) 41 | 42 | open T using () 43 | renaming 44 | ( Carrier to B 45 | ; _≈_ to _≈₂_ 46 | ; refl to ≈₂-refl 47 | ; sym to ≈₂-sym 48 | ; trans to ≈₂-trans 49 | ; setoid to B-setoid 50 | ; _+_ to +₂ 51 | ; _∙_ to ∙₂ 52 | ) 53 | 54 | hom-setoid : ∀ {Carrier} -> (Carrier -> (A-setoid ⟶ B-setoid)) 55 | -> Setoid (suc (k ⊔ a₁ ⊔ a₂ ⊔ ℓ₁ ⊔ ℓ₂)) (a₁ ⊔ ℓ₁ ⊔ ℓ₂) 56 | hom-setoid {Carrier} app = record 57 | { Carrier = Carrier 58 | ; _≈_ = λ f g → ∀ {x y} (r : x ≈₁ y) -> (app f ⟨$⟩ x) ≈₂ (app g ⟨$⟩ y) 59 | ; isEquivalence = record 60 | { refl = λ {f} -> FunEq.cong (app f) 61 | ; sym = λ r rₓ → ≈₂-sym (r (≈₁-sym rₓ)) 62 | ; trans = λ r₁ r₂ rₓ → ≈₂-trans (r₁ ≈₁-refl) (r₂ rₓ) 63 | } 64 | } 65 | 66 | open LinearDefinitions A B _≈₂_ 67 | 68 | record LinearMap : Set (suc (k ⊔ a₁ ⊔ a₂ ⊔ ℓ₁ ⊔ ℓ₂)) where 69 | field 70 | ⟦_⟧ : Morphism 71 | isLinearMap : IsLinearMap From To ⟦_⟧ 72 | 73 | infixr 25 _⟪$⟫_ 74 | _⟪$⟫_ : A -> B 75 | _⟪$⟫_ = ⟦_⟧ 76 | 77 | open IsLinearMap isLinearMap public 78 | 79 | linearMap-setoid : Setoid (suc (k ⊔ a₁ ⊔ a₂ ⊔ ℓ₁ ⊔ ℓ₂)) (a₁ ⊔ ℓ₁ ⊔ ℓ₂) 80 | linearMap-setoid = hom-setoid {LinearMap} λ f → record 81 | { _⟨$⟩_ = λ x → f LinearMap.⟪$⟫ x 82 | ; cong = LinearMap.⟦⟧-cong f 83 | } 84 | 85 | record LinearIsomorphism : Set (suc (k ⊔ a₁ ⊔ a₂ ⊔ ℓ₁ ⊔ ℓ₂)) where 86 | field 87 | ⟦_⟧ : Morphism 88 | isLinearIsomorphism : IsLinearIsomorphism From To ⟦_⟧ 89 | 90 | open IsLinearIsomorphism isLinearIsomorphism public 91 | 92 | linearMap : LinearMap 93 | linearMap = record { ⟦_⟧ = ⟦_⟧ ; isLinearMap = isLinearMap } 94 | 95 | open LinearMap linearMap public 96 | using (_⟪$⟫_) 97 | 98 | linearIsomorphism-setoid : Setoid (suc (k ⊔ a₁ ⊔ a₂ ⊔ ℓ₁ ⊔ ℓ₂)) (a₁ ⊔ ℓ₁ ⊔ ℓ₂) 99 | linearIsomorphism-setoid = hom-setoid {LinearIsomorphism} λ f → record 100 | { _⟨$⟩_ = λ x -> f LinearIsomorphism.⟪$⟫ x 101 | ; cong = LinearIsomorphism.⟦⟧-cong f 102 | } 103 | 104 | module _ {a ℓ} (V : VS.VectorSpace K a ℓ) where 105 | LinearEndomorphism : Set (suc (k ⊔ a ⊔ ℓ)) 106 | LinearEndomorphism = LinearMap V V 107 | 108 | LinearAutomorphism : Set (suc (k ⊔ a ⊔ ℓ)) 109 | LinearAutomorphism = LinearIsomorphism V V 110 | -------------------------------------------------------------------------------- /src/Algebra/Linear/Morphism/Bundles/VectorSpace.agda: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/felko/linear-algebra/d87c5a1eb5dd0569238272e67bce1899616b789a/src/Algebra/Linear/Morphism/Bundles/VectorSpace.agda -------------------------------------------------------------------------------- /src/Algebra/Linear/Morphism/Definitions.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Relation.Binary using (Rel) 4 | 5 | open import Algebra.Structures.Bundles.Field 6 | 7 | module Algebra.Linear.Morphism.Definitions 8 | {k ℓᵏ} (K : Field k ℓᵏ) 9 | {a} (A : Set a) 10 | {b} (B : Set b) 11 | {ℓ} (_≈_ : Rel B ℓ) 12 | where 13 | 14 | open import Algebra.Linear.Core 15 | open import Function 16 | 17 | import Algebra.Morphism as Morphism 18 | open Morphism.Definitions A B _≈_ 19 | 20 | private 21 | K' : Set k 22 | K' = Field.Carrier K 23 | 24 | Linear : Morphism 25 | -> VectorAddition A -> ScalarMultiplication K' A 26 | -> VectorAddition B -> ScalarMultiplication K' B 27 | -> Set _ 28 | Linear ⟦_⟧ _+₁_ _∙₁_ _+₂_ _∙₂_ = 29 | ∀ (a b : K') (u v : A) -> 30 | ⟦ (a ∙₁ u) +₁ (b ∙₁ v) ⟧ ≈ ((a ∙₂ ⟦ u ⟧) +₂ (b ∙₂ ⟦ v ⟧)) 31 | 32 | ScalarHomomorphism : Morphism 33 | -> ScalarMultiplication K' A 34 | -> ScalarMultiplication K' B 35 | -> Set _ 36 | ScalarHomomorphism ⟦_⟧ _∙_ _∘_ = ∀ (c : K') (u : A) -> ⟦ c ∙ u ⟧ ≈ (c ∘ ⟦ u ⟧) 37 | -------------------------------------------------------------------------------- /src/Algebra/Linear/Morphism/Properties.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Algebra.Structures.Bundles.Field 4 | 5 | module Algebra.Linear.Morphism.Properties 6 | {k ℓᵏ} (K : Field k ℓᵏ) 7 | where 8 | 9 | open import Level 10 | open import Algebra.FunctionProperties as FP 11 | import Algebra.Linear.Morphism.Definitions as LinearMorphismDefinitions 12 | import Algebra.Morphism as Morphism 13 | open import Relation.Binary using (Rel; Setoid) 14 | open import Relation.Binary.Morphism.Structures 15 | open import Algebra.Morphism 16 | open import Algebra.Linear.Structures.VectorSpace K 17 | open import Algebra.Linear.Structures.Bundles 18 | 19 | module _ 20 | {a₁ ℓ₁} (From : VectorSpace K a₁ ℓ₁) 21 | {a₂ ℓ₂} (To : VectorSpace K a₂ ℓ₂) 22 | where 23 | 24 | open import Algebra.Linear.Morphism.VectorSpace K 25 | 26 | private 27 | module F = VectorSpace From 28 | module T = VectorSpace To 29 | 30 | open VectorSpaceField 31 | 32 | open F 33 | using () 34 | renaming 35 | ( Carrier to A 36 | ; _≈_ to _≈₁_ 37 | ; refl to ≈₁-refl 38 | ; sym to ≈₁-sym 39 | ; trans to ≈₁-trans 40 | ; setoid to setoid₁ 41 | ; _+_ to _+₁_ 42 | ; +-cong to +₁-cong 43 | ; +-identityˡ to +₁-identityˡ 44 | ; _∙_ to _∙₁_ 45 | ; ∙-identity to ∙₁-identity 46 | ; ∙-absorbˡ to ∙₁-absorbˡ 47 | ; ∙-absorbʳ to ∙₁-absorbʳ 48 | ; 0# to 0₁ 49 | ) 50 | 51 | open T 52 | using () 53 | renaming 54 | ( Carrier to B 55 | ; _≈_ to _≈₂_ 56 | ; refl to ≈₂-refl 57 | ; sym to ≈₂-sym 58 | ; trans to ≈₂-trans 59 | ; setoid to setoid₂ 60 | ; _+_ to _+₂_ 61 | ; +-cong to +₂-cong 62 | ; +-identityˡ to +₂-identityˡ 63 | ; _∙_ to _∙₂_ 64 | ; ∙-identity to ∙₂-identity 65 | ; ∙-absorbˡ to ∙₂-absorbˡ 66 | ; ∙-absorbʳ to ∙₂-absorbʳ 67 | ; 0# to 0₂ 68 | ) 69 | 70 | open import Function 71 | open import Function.Equality 72 | 73 | open LinearDefinitions (VectorSpace.Carrier From) (VectorSpace.Carrier To) _≈₂_ 74 | 75 | linear : ∀ (f : setoid₁ ⟶ setoid₂) 76 | -> (∀ (a b : K') (u v : A) -> (f ⟨$⟩ (a ∙₁ u +₁ b ∙₁ v)) ≈₂ a ∙₂ (f ⟨$⟩ u) +₂ b ∙₂ (f ⟨$⟩ v)) 77 | -> IsLinearMap From To (f ⟨$⟩_) 78 | linear f lin = record 79 | { isAbelianGroupMorphism = record 80 | { gp-homo = record 81 | { mn-homo = record 82 | { sm-homo = record 83 | { ⟦⟧-cong = cong f 84 | ; ∙-homo = λ x y → 85 | begin 86 | f ⟨$⟩ x +₁ y 87 | ≈⟨ cong f (+₁-cong (≈₁-sym (∙₁-identity x)) (≈₁-sym (∙₁-identity y))) ⟩ 88 | f ⟨$⟩ (1ᵏ ∙₁ x +₁ 1ᵏ ∙₁ y) 89 | ≈⟨ lin 1ᵏ 1ᵏ x y ⟩ 90 | 1ᵏ ∙₂ (f ⟨$⟩ x) +₂ 1ᵏ ∙₂ (f ⟨$⟩ y) 91 | ≈⟨ +₂-cong (∙₂-identity (f ⟨$⟩ x)) (∙₂-identity (f ⟨$⟩ y)) ⟩ 92 | (f ⟨$⟩ x) +₂ (f ⟨$⟩ y) 93 | ∎ 94 | } 95 | ; ε-homo = 96 | begin 97 | f ⟨$⟩ 0₁ 98 | ≈⟨ cong f (≈₁-trans (≈₁-sym (+₁-identityˡ 0₁)) (+₁-cong (≈₁-sym (∙₁-absorbˡ 0₁)) (≈₁-sym (∙₁-absorbˡ 0₁)))) ⟩ 99 | f ⟨$⟩ 0ᵏ ∙₁ 0₁ +₁ 0ᵏ ∙₁ 0₁ 100 | ≈⟨ lin 0ᵏ 0ᵏ 0₁ 0₁ ⟩ 101 | 0ᵏ ∙₂ (f ⟨$⟩ 0₁) +₂ 0ᵏ ∙₂ (f ⟨$⟩ 0₁) 102 | ≈⟨ ≈₂-trans (+₂-cong (∙₂-absorbˡ (f ⟨$⟩ 0₁)) (∙₂-absorbˡ (f ⟨$⟩ 0₁))) (+₂-identityˡ 0₂) ⟩ 103 | 0₂ 104 | ∎ 105 | } 106 | } 107 | } 108 | ; ∙-homo = λ c u → 109 | begin 110 | f ⟨$⟩ c ∙₁ u 111 | ≈⟨ cong f (≈₁-trans (≈₁-sym (+₁-identityˡ (c ∙₁ u))) (+₁-cong (≈₁-sym (∙₁-absorbˡ 0₁)) ≈₁-refl)) ⟩ 112 | f ⟨$⟩ 0ᵏ ∙₁ 0₁ +₁ c ∙₁ u 113 | ≈⟨ lin 0ᵏ c 0₁ u ⟩ 114 | 0ᵏ ∙₂ (f ⟨$⟩ 0₁) +₂ c ∙₂ (f ⟨$⟩ u) 115 | ≈⟨ ≈₂-trans (+₂-cong (∙₂-absorbˡ (f ⟨$⟩ 0₁)) ≈₂-refl) (+₂-identityˡ (c ∙₂ (f ⟨$⟩ u))) ⟩ 116 | c ∙₂ (f ⟨$⟩ u) 117 | ∎ 118 | } 119 | where open import Relation.Binary.EqReasoning setoid₂ 120 | -------------------------------------------------------------------------------- /src/Algebra/Linear/Morphism/Structures/Vect.agda: -------------------------------------------------------------------------------- 1 | module Algebra.Linear.Category.Vect where 2 | -------------------------------------------------------------------------------- /src/Algebra/Linear/Morphism/VectorSpace.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Algebra.Structures.Bundles.Field 4 | 5 | module Algebra.Linear.Morphism.VectorSpace 6 | {k ℓᵏ} (K : Field k ℓᵏ) 7 | where 8 | 9 | open import Level 10 | open import Algebra.FunctionProperties as FP 11 | import Algebra.Linear.Morphism.Definitions as LinearMorphismDefinitions 12 | import Algebra.Morphism as Morphism 13 | open import Relation.Binary using (Rel; Setoid) 14 | open import Relation.Binary.Morphism.Structures 15 | open import Algebra.Morphism 16 | 17 | open import Algebra.Linear.Core 18 | open import Algebra.Linear.Structures.VectorSpace K 19 | open import Algebra.Linear.Structures.Bundles 20 | 21 | module LinearDefinitions {a₁ a₂ ℓ₂} (A : Set a₁) (B : Set a₂) (_≈_ : Rel B ℓ₂) where 22 | open Morphism.Definitions A B _≈_ public 23 | open LinearMorphismDefinitions K A B _≈_ public 24 | 25 | module _ 26 | {a₁ ℓ₁} (From : VectorSpace K a₁ ℓ₁) 27 | {a₂ ℓ₂} (To : VectorSpace K a₂ ℓ₂) 28 | where 29 | 30 | private 31 | module F = VectorSpace From 32 | module T = VectorSpace To 33 | 34 | K' : Set k 35 | K' = Field.Carrier K 36 | 37 | open F using () renaming (Carrier to A; _≈_ to _≈₁_; _+_ to _+₁_; _∙_ to _∙₁_) 38 | open T using () renaming (Carrier to B; _≈_ to _≈₂_; _+_ to _+₂_; _∙_ to _∙₂_) 39 | 40 | open import Function 41 | 42 | open LinearDefinitions (VectorSpace.Carrier From) (VectorSpace.Carrier To) _≈₂_ 43 | 44 | record IsLinearMap (⟦_⟧ : Morphism) : Set (k ⊔ a₁ ⊔ a₂ ⊔ ℓ₁ ⊔ ℓ₂) where 45 | field 46 | isAbelianGroupMorphism : IsAbelianGroupMorphism F.abelianGroup T.abelianGroup ⟦_⟧ 47 | ∙-homo : ScalarHomomorphism ⟦_⟧ _∙₁_ _∙₂_ 48 | 49 | open IsAbelianGroupMorphism isAbelianGroupMorphism public 50 | renaming (∙-homo to +-homo; ε-homo to 0#-homo) 51 | 52 | distrib-linear : ∀ (a b : K') (u v : A) -> ⟦ a ∙₁ u +₁ b ∙₁ v ⟧ ≈₂ a ∙₂ ⟦ u ⟧ +₂ b ∙₂ ⟦ v ⟧ 53 | distrib-linear a b u v = T.trans (+-homo (a ∙₁ u) (b ∙₁ v)) (T.+-cong (∙-homo a u) (∙-homo b v)) 54 | 55 | record IsLinearMonomorphism (⟦_⟧ : Morphism) : Set (k ⊔ a₁ ⊔ a₂ ⊔ ℓ₁ ⊔ ℓ₂) where 56 | field 57 | isLinearMap : IsLinearMap ⟦_⟧ 58 | injective : Injective _≈₁_ _≈₂_ ⟦_⟧ 59 | 60 | open IsLinearMap isLinearMap public 61 | 62 | record IsLinearIsomorphism (⟦_⟧ : Morphism) : Set (k ⊔ a₁ ⊔ a₂ ⊔ ℓ₁ ⊔ ℓ₂) where 63 | field 64 | isLinearMonomorphism : IsLinearMonomorphism ⟦_⟧ 65 | surjective : Surjective _≈₁_ _≈₂_ ⟦_⟧ 66 | 67 | open IsLinearMonomorphism isLinearMonomorphism public 68 | 69 | module _ 70 | {a ℓ} (V : VectorSpace K a ℓ) 71 | where 72 | 73 | open VectorSpace V 74 | 75 | open LinearDefinitions Carrier Carrier _≈_ 76 | 77 | IsLinearEndomorphism : Morphism -> Set (k ⊔ a ⊔ ℓ) 78 | IsLinearEndomorphism ⟦_⟧ = IsLinearMap V V ⟦_⟧ 79 | 80 | IsLinearAutomorphism : Morphism -> Set (k ⊔ a ⊔ ℓ) 81 | IsLinearAutomorphism ⟦_⟧ = IsLinearIsomorphism V V ⟦_⟧ 82 | -------------------------------------------------------------------------------- /src/Algebra/Linear/Space/FiniteDimensional/.#Hom.agda: -------------------------------------------------------------------------------- 1 | joseph@Josephs-MacBook-Air.local.964 -------------------------------------------------------------------------------- /src/Algebra/Linear/Space/FiniteDimensional/Hom.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Algebra.Structures.Bundles.Field 4 | import Algebra.Linear.Structures.Bundles.FiniteDimensional as FDB 5 | 6 | module Algebra.Linear.Space.FiniteDimensional.Hom 7 | {k ℓᵏ} (K : Field k ℓᵏ) 8 | {p a₁ ℓ₁} (V₁-space : FDB.FiniteDimensional K a₁ ℓ₁ p) 9 | {n a₂ ℓ₂} (V₂-space : FDB.FiniteDimensional K a₂ ℓ₂ n) 10 | where 11 | 12 | open import Algebra.Linear.Structures.VectorSpace K 13 | open import Algebra.Linear.Structures.Bundles as VS 14 | open import Algebra.Linear.Structures.FiniteDimensional K 15 | 16 | open import Algebra.Linear.Morphism.Bundles K 17 | 18 | open VectorSpaceField 19 | 20 | open FDB.FiniteDimensional V₁-space 21 | using () 22 | renaming 23 | ( Carrier to V₁ 24 | ; _≈_ to _≈₁_ 25 | ; isEquivalence to ≈₁-isEquiv 26 | ; refl to ≈₁-refl 27 | ; sym to ≈₁-sym 28 | ; trans to ≈₁-trans 29 | ; _+_ to _+₁_ 30 | ; _∙_ to _∙₁_ 31 | ; -_ to -₁_ 32 | ; 0# to 0₁ 33 | ; +-identityˡ to +₁-identityˡ 34 | ; +-identityʳ to +₁-identityʳ 35 | ; +-identity to +₁-identity 36 | ; +-cong to +₁-cong 37 | ; +-assoc to +₁-assoc 38 | ; +-comm to +₁-comm 39 | ; *ᵏ-∙-compat to *ᵏ-∙₁-compat 40 | ; ∙-+-distrib to ∙₁-+₁-distrib 41 | ; ∙-+ᵏ-distrib to ∙₁-+ᵏ-distrib 42 | ; ∙-cong to ∙₁-cong 43 | ; ∙-identity to ∙₁-identity 44 | ; ∙-absorbˡ to ∙₁-absorbˡ 45 | ; ∙-absorbʳ to ∙₁-absorbʳ 46 | ; -‿cong to -₁‿cong 47 | ; -‿inverseˡ to -₁‿inverseˡ 48 | ; -‿inverseʳ to -₁‿inverseʳ 49 | ; vectorSpace to vectorSpace₁ 50 | ; embed to embed₁ 51 | ) 52 | 53 | open FDB.FiniteDimensional V₂-space 54 | using () 55 | renaming 56 | ( Carrier to V₂ 57 | ; _≈_ to _≈₂_ 58 | ; isEquivalence to ≈₂-isEquiv 59 | ; refl to ≈₂-refl 60 | ; sym to ≈₂-sym 61 | ; trans to ≈₂-trans 62 | ; _+_ to _+₂_ 63 | ; _∙_ to _∙₂_ 64 | ; -_ to -₂_ 65 | ; 0# to 0₂ 66 | ; +-identityˡ to +₂-identityˡ 67 | ; +-identityʳ to +₂-identityʳ 68 | ; +-identity to +₂-identity 69 | ; +-cong to +₂-cong 70 | ; +-assoc to +₂-assoc 71 | ; +-comm to +₂-comm 72 | ; *ᵏ-∙-compat to *ᵏ-∙₂-compat 73 | ; ∙-+-distrib to ∙₂-+₂-distrib 74 | ; ∙-+ᵏ-distrib to ∙₂-+ᵏ-distrib 75 | ; ∙-cong to ∙₂-cong 76 | ; ∙-identity to ∙₂-identity 77 | ; ∙-absorbˡ to ∙₂-absorbˡ 78 | ; ∙-absorbʳ to ∙₂-absorbʳ 79 | ; -‿cong to -₂‿cong 80 | ; -‿inverseˡ to -₂‿inverseˡ 81 | ; -‿inverseʳ to -₂‿inverseʳ 82 | ; vectorSpace to vectorSpace₂ 83 | ; embed to embed₂ 84 | ) 85 | 86 | open LinearIsomorphism embed₁ 87 | using () 88 | renaming 89 | ( ⟦_⟧ to ⟦_⟧₁ 90 | ; ⟦⟧-cong to ⟦⟧₁-cong 91 | ; injective to ⟦⟧₁-injective 92 | ; surjective to ⟦⟧₁-surjective 93 | ; +-homo to +₁-homo 94 | ; ∙-homo to ∙₁-homo 95 | ; 0#-homo to 0₁-homo 96 | ) 97 | 98 | open LinearIsomorphism embed₂ 99 | using () 100 | renaming 101 | ( ⟦_⟧ to ⟦_⟧₂ 102 | ; ⟦⟧-cong to ⟦⟧₂-cong 103 | ; injective to ⟦⟧₂-injective 104 | ; surjective to ⟦⟧₂-surjective 105 | ; +-homo to +₂-homo 106 | ; ∙-homo to ∙₂-homo 107 | ; 0#-homo to 0₂-homo 108 | ) 109 | 110 | import Algebra.Linear.Construct.HomSpace K vectorSpace₁ vectorSpace₂ as PS 111 | 112 | open VS.VectorSpace PS.vectorSpace 113 | renaming 114 | ( refl to ≈-refl 115 | ; sym to ≈-sym 116 | ; trans to ≈-trans 117 | ) 118 | 119 | import Algebra.Linear.Construct.Vector as Vec 120 | open Vec K 121 | using 122 | ( ++-cong 123 | ; ++-identityˡ 124 | ; +-distrib-++ 125 | ; ∙-distrib-++ 126 | ; 0++0≈0 127 | ) 128 | renaming 129 | ( _≈_ to _≈v_ 130 | ; ≈-refl to ≈v-refl 131 | ; ≈-sym to ≈v-sym 132 | ; ≈-trans to ≈v-trans 133 | ; _+_ to _+v_ 134 | ; _∙_ to _∙v_ 135 | ; 0# to 0v 136 | ; +-cong to +v-cong 137 | ; setoid to vec-setoid 138 | ; vectorSpace to vector-vectorSpace 139 | ) 140 | 141 | import Algebra.Linear.Construct.Matrix K as M 142 | 143 | open import Data.Nat hiding (_+_) renaming (_*_ to _*ℕ_) 144 | 145 | open import Relation.Binary.PropositionalEquality as P 146 | using (_≡_; subst; subst-subst-sym) 147 | renaming 148 | ( refl to ≡-refl 149 | ; sym to ≡-sym 150 | ; trans to ≡-trans 151 | ) 152 | 153 | open import Data.Vec 154 | open import Data.Product 155 | open import Data.Fin 156 | open import Function 157 | 158 | δ : ∀ {n p} -> Fin n -> Fin p -> K' 159 | δ zero zero = 1ᵏ 160 | δ (suc n) zero = 0ᵏ 161 | δ zero (suc p) = 0ᵏ 162 | δ (suc n) (suc p) = δ n p 163 | 164 | canonicalBasis : M.Matrix n p 165 | canonicalBasis = M.tabulate δ 166 | 167 | module _ where 168 | open import Algebra.Morphism.Definitions (LinearMap vectorSpace₁ vectorSpace₂) (M.Matrix n p) M._≈_ 169 | open import Algebra.Linear.Morphism.Definitions K (LinearMap vectorSpace₁ vectorSpace₂) (M.Matrix n p) M._≈_ 170 | import Relation.Binary.Morphism.Definitions (LinearMap vectorSpace₁ vectorSpace₂) (M.Matrix n p) as R 171 | open import Relation.Binary.EqReasoning (M.setoid {n} {p}) 172 | 173 | Mat : LinearMap vectorSpace₁ vectorSpace₂ -> M.Matrix n p 174 | Mat f = M.mapCols (λ u -> embed₂ LinearIsomorphism.⟪$⟫ (f LinearMap.⟪$⟫ {! embed₁ ⁻¹ ⟪$⟫ u!})) 175 | canonicalBasis 176 | 177 | -- (λ u → embed₂ LinearIsomorphism.⟪$⟫ (proj₁ (⟦⟧₂-surjective u))) canonicalBasis 178 | 179 | Mat-cong : R.Homomorphic₂ _≈_ M._≈_ Mat 180 | Mat-cong {A} {B} f = 181 | begin 182 | Mat A 183 | ≈⟨ M.≈-reflexive (M.mapCols-cong (λ u → {!!}) canonicalBasis) ⟩ 184 | {!!} 185 | ≈⟨ {!!} ⟩ 186 | Mat B 187 | ∎ 188 | 189 | {- 190 | 191 | Mat-cong : R.Homomorphic₂ PS._≈_ _≈v_ Mat 192 | ⟦⟧-cong (r₁ , r₂) = ++-cong (⟦⟧₁-cong r₁) (⟦⟧₂-cong r₂) 193 | 194 | +-homo : Homomorphic₂ ⟦_⟧ _+_ _+v_ 195 | +-homo (x₁ , x₂) (y₁ , y₂) = 196 | begin 197 | ⟦ (x₁ , x₂) + (y₁ , y₂) ⟧ 198 | ≡⟨⟩ 199 | ⟦ x₁ +₁ y₁ ⟧₁ ++ ⟦ x₂ +₂ y₂ ⟧₂ 200 | ≈⟨ ++-cong (+₁-homo x₁ y₁) (+₂-homo x₂ y₂) ⟩ 201 | (⟦ x₁ ⟧₁ +v ⟦ y₁ ⟧₁) ++ (⟦ x₂ ⟧₂ +v ⟦ y₂ ⟧₂) 202 | ≈⟨ ≈v-sym (+-distrib-++ ⟦ x₁ ⟧₁ ⟦ x₂ ⟧₂ ⟦ y₁ ⟧₁ ⟦ y₂ ⟧₂) ⟩ 203 | (⟦ x₁ ⟧₁ ++ ⟦ x₂ ⟧₂) +v (⟦ y₁ ⟧₁ ++ ⟦ y₂ ⟧₂) 204 | ≈⟨ +v-cong ≈v-refl ≈v-refl ⟩ 205 | ⟦ x₁ , x₂ ⟧ +v ⟦ y₁ , y₂ ⟧ 206 | ∎ 207 | 208 | 0#-homo : Homomorphic₀ ⟦_⟧ 0# 0v 209 | 0#-homo = 210 | begin 211 | ⟦ 0# ⟧ 212 | ≡⟨⟩ 213 | ⟦ 0₁ ⟧₁ ++ ⟦ 0₂ ⟧₂ 214 | ≈⟨ ++-cong 0₁-homo 0₂-homo ⟩ 215 | (0v {n}) ++ (0v {p}) 216 | ≈⟨ 0++0≈0 {n} {p} ⟩ 217 | 0v {n +ℕ p} 218 | ∎ 219 | 220 | ∙-homo : ScalarHomomorphism ⟦_⟧ _∙_ _∙v_ 221 | ∙-homo c (x₁ , x₂) = 222 | begin 223 | ⟦ c ∙ (x₁ , x₂) ⟧ 224 | ≈⟨ ++-cong (∙₁-homo c x₁) (∙₂-homo c x₂) ⟩ 225 | (c ∙v ⟦ x₁ ⟧₁) ++ (c ∙v ⟦ x₂ ⟧₂) 226 | ≈⟨ ≈v-sym (∙-distrib-++ c ⟦ x₁ ⟧₁ ⟦ x₂ ⟧₂) ⟩ 227 | c ∙v ⟦ x₁ , x₂ ⟧ 228 | ∎ 229 | 230 | ⟦⟧-injective : Injective ⟦_⟧ 231 | ⟦⟧-injective {x₁ , x₂} {y₁ , y₂} r = 232 | let (r₁ , r₂) = ++-split r 233 | in ⟦⟧₁-injective r₁ , ⟦⟧₂-injective r₂ 234 | 235 | ⟦⟧-surjective : Surjective ⟦_⟧ 236 | ⟦⟧-surjective y = 237 | let (x₁ , x₂) = splitAt' n p ≡-refl y 238 | (u , r₁) = ⟦⟧₁-surjective x₁ 239 | (v , r₂) = ⟦⟧₂-surjective x₂ 240 | in (u , v) , ≈v-trans (++-cong r₁ r₂) (++-splitAt' {n} {p} y) 241 | 242 | embed : LinearIsomorphism PS.vectorSpace (vector-vectorSpace {n +ℕ p}) 243 | embed = record 244 | { ⟦_⟧ = ⟦_⟧ 245 | ; isLinearIsomorphism = record 246 | { isLinearMonomorphism = record 247 | { isLinearMap = record 248 | { isAbelianGroupMorphism = record 249 | { gp-homo = record 250 | { mn-homo = record 251 | { sm-homo = record 252 | { ⟦⟧-cong = ⟦⟧-cong 253 | ; ∙-homo = +-homo 254 | } 255 | ; ε-homo = 0#-homo 256 | } 257 | } 258 | } 259 | ; ∙-homo = ∙-homo 260 | } 261 | ; injective = ⟦⟧-injective 262 | } 263 | ; surjective = ⟦⟧-surjective 264 | } 265 | } 266 | 267 | isFiniteDimensional : IsFiniteDimensional _≈_ _+_ _∙_ -_ 0# (n +ℕ p) 268 | isFiniteDimensional = record 269 | { isVectorSpace = isVectorSpace 270 | ; embed = embed 271 | } 272 | -} 273 | -------------------------------------------------------------------------------- /src/Algebra/Linear/Space/FiniteDimensional/Product.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Algebra.Structures.Bundles.Field 4 | import Algebra.Linear.Structures.Bundles.FiniteDimensional as FDB 5 | 6 | module Algebra.Linear.Space.FiniteDimensional.Product 7 | {k ℓᵏ} (K : Field k ℓᵏ) 8 | {n a₁ ℓ₁} (V₁-space : FDB.FiniteDimensional K a₁ ℓ₁ n) 9 | {p a₂ ℓ₂} (V₂-space : FDB.FiniteDimensional K a₂ ℓ₂ p) 10 | where 11 | 12 | open import Algebra.Linear.Structures.VectorSpace K 13 | open import Algebra.Linear.Structures.Bundles as VS 14 | open import Algebra.Linear.Structures.FiniteDimensional K 15 | 16 | open import Algebra.Linear.Morphism.Bundles K 17 | 18 | open VectorSpaceField 19 | 20 | open FDB.FiniteDimensional V₁-space 21 | using () 22 | renaming 23 | ( Carrier to V₁ 24 | ; _≈_ to _≈₁_ 25 | ; isEquivalence to ≈₁-isEquiv 26 | ; refl to ≈₁-refl 27 | ; sym to ≈₁-sym 28 | ; trans to ≈₁-trans 29 | ; reflexive to ≈₁-reflexive 30 | ; _+_ to _+₁_ 31 | ; _∙_ to _∙₁_ 32 | ; -_ to -₁_ 33 | ; 0# to 0₁ 34 | ; +-identityˡ to +₁-identityˡ 35 | ; +-identityʳ to +₁-identityʳ 36 | ; +-identity to +₁-identity 37 | ; +-cong to +₁-cong 38 | ; +-assoc to +₁-assoc 39 | ; +-comm to +₁-comm 40 | ; *ᵏ-∙-compat to *ᵏ-∙₁-compat 41 | ; ∙-+-distrib to ∙₁-+₁-distrib 42 | ; ∙-+ᵏ-distrib to ∙₁-+ᵏ-distrib 43 | ; ∙-cong to ∙₁-cong 44 | ; ∙-identity to ∙₁-identity 45 | ; ∙-absorbˡ to ∙₁-absorbˡ 46 | ; ∙-absorbʳ to ∙₁-absorbʳ 47 | ; -‿cong to -₁‿cong 48 | ; -‿inverseˡ to -₁‿inverseˡ 49 | ; -‿inverseʳ to -₁‿inverseʳ 50 | ; vectorSpace to vectorSpace₁ 51 | ; embed to embed₁ 52 | ) 53 | 54 | open FDB.FiniteDimensional V₂-space 55 | using () 56 | renaming 57 | ( Carrier to V₂ 58 | ; _≈_ to _≈₂_ 59 | ; isEquivalence to ≈₂-isEquiv 60 | ; refl to ≈₂-refl 61 | ; sym to ≈₂-sym 62 | ; trans to ≈₂-trans 63 | ; reflexive to ≈₂-reflexive 64 | ; _+_ to _+₂_ 65 | ; _∙_ to _∙₂_ 66 | ; -_ to -₂_ 67 | ; 0# to 0₂ 68 | ; +-identityˡ to +₂-identityˡ 69 | ; +-identityʳ to +₂-identityʳ 70 | ; +-identity to +₂-identity 71 | ; +-cong to +₂-cong 72 | ; +-assoc to +₂-assoc 73 | ; +-comm to +₂-comm 74 | ; *ᵏ-∙-compat to *ᵏ-∙₂-compat 75 | ; ∙-+-distrib to ∙₂-+₂-distrib 76 | ; ∙-+ᵏ-distrib to ∙₂-+ᵏ-distrib 77 | ; ∙-cong to ∙₂-cong 78 | ; ∙-identity to ∙₂-identity 79 | ; ∙-absorbˡ to ∙₂-absorbˡ 80 | ; ∙-absorbʳ to ∙₂-absorbʳ 81 | ; -‿cong to -₂‿cong 82 | ; -‿inverseˡ to -₂‿inverseˡ 83 | ; -‿inverseʳ to -₂‿inverseʳ 84 | ; vectorSpace to vectorSpace₂ 85 | ; embed to embed₂ 86 | ) 87 | 88 | open LinearIsomorphism embed₁ 89 | using () 90 | renaming 91 | ( ⟦_⟧ to ⟦_⟧₁ 92 | ; ⟦⟧-cong to ⟦⟧₁-cong 93 | ; injective to ⟦⟧₁-injective 94 | ; surjective to ⟦⟧₁-surjective 95 | ; +-homo to +₁-homo 96 | ; ∙-homo to ∙₁-homo 97 | ; 0#-homo to 0₁-homo 98 | ) 99 | 100 | open LinearIsomorphism embed₂ 101 | using () 102 | renaming 103 | ( ⟦_⟧ to ⟦_⟧₂ 104 | ; ⟦⟧-cong to ⟦⟧₂-cong 105 | ; injective to ⟦⟧₂-injective 106 | ; surjective to ⟦⟧₂-surjective 107 | ; +-homo to +₂-homo 108 | ; ∙-homo to ∙₂-homo 109 | ; 0#-homo to 0₂-homo 110 | ) 111 | 112 | open import Data.Product 113 | import Algebra.Linear.Construct.ProductSpace K vectorSpace₁ vectorSpace₂ as PS 114 | open VS.VectorSpace PS.vectorSpace 115 | renaming 116 | ( refl to ≈-refl 117 | ; sym to ≈-sym 118 | ; trans to ≈-trans 119 | ) 120 | 121 | open import Algebra.Linear.Construct.Vector K 122 | using 123 | ( ++-cong 124 | ; ++-identityˡ 125 | ; ++-split 126 | ; +-distrib-++ 127 | ; ∙-distrib-++ 128 | ; 0++0≈0 129 | ) 130 | renaming 131 | ( _≈_ to _≈v_ 132 | ; ≈-refl to ≈v-refl 133 | ; ≈-sym to ≈v-sym 134 | ; ≈-trans to ≈v-trans 135 | ; ≈-reflexive to ≈v-reflexive 136 | ; _+_ to _+v_ 137 | ; _∙_ to _∙v_ 138 | ; 0# to 0v 139 | ; +-cong to +v-cong 140 | ; setoid to vec-setoid 141 | ; vectorSpace to vector-vectorSpace 142 | ) 143 | 144 | open import Data.Nat using (ℕ) renaming (_+_ to _+ℕ_) 145 | 146 | open import Relation.Binary.PropositionalEquality as P 147 | using (_≡_; subst; subst-subst-sym) 148 | renaming 149 | ( refl to ≡-refl 150 | ; sym to ≡-sym 151 | ; trans to ≡-trans 152 | ) 153 | 154 | open import Data.Vec 155 | open import Algebra.Morphism.Definitions (V₁ × V₂) (Vec K' (n +ℕ p)) _≈v_ 156 | open import Algebra.Linear.Morphism.Definitions K (V₁ × V₂) (Vec K' (n +ℕ p)) _≈v_ 157 | import Relation.Binary.Morphism.Definitions (V₁ × V₂) (Vec K' (n +ℕ p)) as R 158 | open import Function 159 | open import Relation.Binary.EqReasoning (vec-setoid (n +ℕ p)) 160 | 161 | ⟦_⟧ : V₁ × V₂ -> Vec K' (n +ℕ p) 162 | ⟦ (u , v) ⟧ = ⟦ u ⟧₁ ++ ⟦ v ⟧₂ 163 | 164 | ⟦⟧-cong : R.Homomorphic₂ PS._≈_ _≈v_ ⟦_⟧ 165 | ⟦⟧-cong (r₁ , r₂) = ++-cong (⟦⟧₁-cong r₁) (⟦⟧₂-cong r₂) 166 | 167 | +-homo : Homomorphic₂ ⟦_⟧ _+_ _+v_ 168 | +-homo (x₁ , x₂) (y₁ , y₂) = 169 | begin 170 | ⟦ (x₁ , x₂) + (y₁ , y₂) ⟧ 171 | ≡⟨⟩ 172 | ⟦ x₁ +₁ y₁ ⟧₁ ++ ⟦ x₂ +₂ y₂ ⟧₂ 173 | ≈⟨ ++-cong (+₁-homo x₁ y₁) (+₂-homo x₂ y₂) ⟩ 174 | (⟦ x₁ ⟧₁ +v ⟦ y₁ ⟧₁) ++ (⟦ x₂ ⟧₂ +v ⟦ y₂ ⟧₂) 175 | ≈⟨ ≈v-sym (+-distrib-++ ⟦ x₁ ⟧₁ ⟦ x₂ ⟧₂ ⟦ y₁ ⟧₁ ⟦ y₂ ⟧₂) ⟩ 176 | (⟦ x₁ ⟧₁ ++ ⟦ x₂ ⟧₂) +v (⟦ y₁ ⟧₁ ++ ⟦ y₂ ⟧₂) 177 | ≈⟨ +v-cong ≈v-refl ≈v-refl ⟩ 178 | ⟦ x₁ , x₂ ⟧ +v ⟦ y₁ , y₂ ⟧ 179 | ∎ 180 | 181 | 0#-homo : Homomorphic₀ ⟦_⟧ 0# 0v 182 | 0#-homo = 183 | begin 184 | ⟦ 0# ⟧ 185 | ≡⟨⟩ 186 | ⟦ 0₁ ⟧₁ ++ ⟦ 0₂ ⟧₂ 187 | ≈⟨ ++-cong 0₁-homo 0₂-homo ⟩ 188 | (0v {n}) ++ (0v {p}) 189 | ≈⟨ 0++0≈0 {n} {p} ⟩ 190 | 0v {n +ℕ p} 191 | ∎ 192 | 193 | ∙-homo : ScalarHomomorphism ⟦_⟧ _∙_ _∙v_ 194 | ∙-homo c (x₁ , x₂) = 195 | begin 196 | ⟦ c ∙ (x₁ , x₂) ⟧ 197 | ≈⟨ ++-cong (∙₁-homo c x₁) (∙₂-homo c x₂) ⟩ 198 | (c ∙v ⟦ x₁ ⟧₁) ++ (c ∙v ⟦ x₂ ⟧₂) 199 | ≈⟨ ≈v-sym (∙-distrib-++ c ⟦ x₁ ⟧₁ ⟦ x₂ ⟧₂) ⟩ 200 | c ∙v ⟦ x₁ , x₂ ⟧ 201 | ∎ 202 | 203 | ⟦⟧-injective : Injective PS._≈_ (_≈v_ {n +ℕ p}) ⟦_⟧ 204 | ⟦⟧-injective {x₁ , x₂} {y₁ , y₂} r = 205 | let (r₁ , r₂) = ++-split r 206 | in ⟦⟧₁-injective r₁ , ⟦⟧₂-injective r₂ 207 | 208 | ⟦⟧-surjective : Surjective PS._≈_ (_≈v_ {n +ℕ p}) ⟦_⟧ 209 | ⟦⟧-surjective y = 210 | let (x₁ , x₂ , r) = splitAt n y 211 | (u , r₁) = ⟦⟧₁-surjective x₁ 212 | (v , r₂) = ⟦⟧₂-surjective x₂ 213 | in (u , v) , ≈v-trans (++-cong r₁ r₂) (≈v-sym (≈v-reflexive r)) 214 | 215 | embed : LinearIsomorphism PS.vectorSpace (vector-vectorSpace {n +ℕ p}) 216 | embed = record 217 | { ⟦_⟧ = ⟦_⟧ 218 | ; isLinearIsomorphism = record 219 | { isLinearMonomorphism = record 220 | { isLinearMap = record 221 | { isAbelianGroupMorphism = record 222 | { gp-homo = record 223 | { mn-homo = record 224 | { sm-homo = record 225 | { ⟦⟧-cong = ⟦⟧-cong 226 | ; ∙-homo = +-homo 227 | } 228 | ; ε-homo = 0#-homo 229 | } 230 | } 231 | } 232 | ; ∙-homo = ∙-homo 233 | } 234 | ; injective = ⟦⟧-injective 235 | } 236 | ; surjective = ⟦⟧-surjective 237 | } 238 | } 239 | 240 | isFiniteDimensional : IsFiniteDimensional _≈_ _+_ _∙_ -_ 0# (n +ℕ p) 241 | isFiniteDimensional = record 242 | { isVectorSpace = isVectorSpace 243 | ; embed = embed 244 | } 245 | -------------------------------------------------------------------------------- /src/Algebra/Linear/Space/Hom.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Algebra.Structures.Bundles.Field 4 | open import Algebra.Linear.Structures.Bundles 5 | 6 | module Algebra.Linear.Space.Hom 7 | {k ℓ} (K : Field k ℓ) 8 | {a₁ ℓ₁} (V₁-space : VectorSpace K a₁ ℓ₁) 9 | {a₂ ℓ₂} (V₂-space : VectorSpace K a₂ ℓ₂) 10 | where 11 | 12 | open import Relation.Binary 13 | open import Level using (_⊔_; suc) 14 | open import Data.Nat using (ℕ) renaming (_+_ to _+ℕ_) 15 | open import Algebra.Linear.Morphism.Bundles K 16 | open import Algebra.Linear.Morphism.Properties K 17 | open import Algebra.Linear.Structures.VectorSpace 18 | 19 | open import Data.Product 20 | 21 | open VectorSpaceField K 22 | 23 | open VectorSpace V₁-space 24 | using () 25 | renaming 26 | ( Carrier to V₁ 27 | ; _≈_ to _≈₁_ 28 | ; isEquivalence to ≈₁-isEquiv 29 | ; refl to ≈₁-refl 30 | ; sym to ≈₁-sym 31 | ; trans to ≈₁-trans 32 | ; setoid to setoid₁ 33 | ; _+_ to _+₁_ 34 | ; _∙_ to _∙₁_ 35 | ; -_ to -₁_ 36 | ; 0# to 0₁ 37 | ; +-identityˡ to +₁-identityˡ 38 | ; +-identityʳ to +₁-identityʳ 39 | ; +-identity to +₁-identity 40 | ; +-cong to +₁-cong 41 | ; +-assoc to +₁-assoc 42 | ; +-comm to +₁-comm 43 | ; *ᵏ-∙-compat to *ᵏ-∙₁-compat 44 | ; ∙-+-distrib to ∙₁-+₁-distrib 45 | ; ∙-+ᵏ-distrib to ∙₁-+ᵏ-distrib 46 | ; ∙-cong to ∙₁-cong 47 | ; ∙-identity to ∙₁-identity 48 | ; ∙-absorbˡ to ∙₁-absorbˡ 49 | ; ∙-absorbʳ to ∙₁-absorbʳ 50 | ; -‿cong to -₁‿cong 51 | ; -‿inverseˡ to -₁‿inverseˡ 52 | ; -‿inverseʳ to -₁‿inverseʳ 53 | ) 54 | 55 | open VectorSpace V₂-space 56 | using () 57 | renaming 58 | ( Carrier to V₂ 59 | ; _≈_ to _≈₂_ 60 | ; isEquivalence to ≈₂-isEquiv 61 | ; refl to ≈₂-refl 62 | ; sym to ≈₂-sym 63 | ; trans to ≈₂-trans 64 | ; setoid to setoid₂ 65 | ; _+_ to _+₂_ 66 | ; _∙_ to _∙₂_ 67 | ; -_ to -₂_ 68 | ; 0# to 0₂ 69 | ; +-identityˡ to +₂-identityˡ 70 | ; +-identityʳ to +₂-identityʳ 71 | ; +-identity to +₂-identity 72 | ; +-cong to +₂-cong 73 | ; +-assoc to +₂-assoc 74 | ; +-comm to +₂-comm 75 | ; *ᵏ-∙-compat to *ᵏ-∙₂-compat 76 | ; ∙-+-distrib to ∙₂-+₂-distrib 77 | ; ∙-+ᵏ-distrib to ∙₂-+ᵏ-distrib 78 | ; ∙-cong to ∙₂-cong 79 | ; ∙-identity to ∙₂-identity 80 | ; ∙-absorbˡ to ∙₂-absorbˡ 81 | ; ∙-absorbʳ to ∙₂-absorbʳ 82 | ; ∙-∙-comm to ∙₂-∙₂-comm 83 | ; -‿cong to -₂‿cong 84 | ; -‿inverseˡ to -₂‿inverseˡ 85 | ; -‿inverseʳ to -₂‿inverseʳ 86 | ; -‿+-comm to -₂‿+₂-comm 87 | ; -1∙u≈-u to -1∙u≈₂-u 88 | ) 89 | 90 | import Function 91 | 92 | private 93 | V : Set (suc (k ⊔ a₁ ⊔ a₂ ⊔ ℓ₁ ⊔ ℓ₂)) 94 | V = LinearMap V₁-space V₂-space 95 | 96 | open LinearMap using (_⟪$⟫_; ⟦⟧-cong; distrib-linear) 97 | 98 | setoid : Setoid (suc (k ⊔ a₁ ⊔ a₂ ⊔ ℓ₁ ⊔ ℓ₂)) _ 99 | setoid = linearMap-setoid V₁-space V₂-space 100 | 101 | open Setoid setoid public 102 | renaming (isEquivalence to ≈-isEquiv) 103 | 104 | open import Algebra.FunctionProperties _≈_ 105 | open import Algebra.Structures _≈_ 106 | open import Function.Equality hiding (setoid; _∘_) 107 | 108 | open import Relation.Binary.EqReasoning setoid₂ 109 | 110 | 0# : V 111 | 0# = record 112 | { ⟦_⟧ = λ x → 0₂ 113 | ; isLinearMap = linear V₁-space V₂-space π 114 | λ a b u v → 115 | begin 116 | 0₂ 117 | ≈⟨ ≈₂-trans (≈₂-sym (∙₂-absorbʳ b)) (≈₂-sym (+₂-identityˡ (b ∙₂ 0₂))) ⟩ 118 | 0₂ +₂ b ∙₂ 0₂ 119 | ≈⟨ +₂-cong (≈₂-sym (∙₂-absorbʳ a)) ≈₂-refl ⟩ 120 | a ∙₂ 0₂ +₂ b ∙₂ 0₂ 121 | ∎ 122 | } 123 | where π : setoid₁ ⟶ setoid₂ 124 | π = record 125 | { _⟨$⟩_ = λ x → 0₂ 126 | ; cong = λ r → ≈₂-refl } 127 | 128 | -_ : Op₁ V 129 | - f = record 130 | { ⟦_⟧ = λ x → -₂ f ⟪$⟫ x 131 | ; isLinearMap = linear V₁-space V₂-space π 132 | let -[a∙u]≈₂a∙[-u] : ∀ (a : K') (u : V₂) -> (-₂ (a ∙₂ u)) ≈₂ a ∙₂ (-₂ u) 133 | -[a∙u]≈₂a∙[-u] a u = 134 | begin 135 | (-₂ (a ∙₂ u)) 136 | ≈⟨ ≈₂-sym (-1∙u≈₂-u (a ∙₂ u)) ⟩ 137 | (-ᵏ 1ᵏ) ∙₂ (a ∙₂ u) 138 | ≈⟨ ≈₂-sym (*ᵏ-∙₂-compat (-ᵏ 1ᵏ) a u) ⟩ 139 | ((-ᵏ 1ᵏ) *ᵏ a) ∙₂ u 140 | ≈⟨ ∙₂-cong (≈ᵏ-trans (≈ᵏ-sym (-ᵏ‿distribˡ-*ᵏ 1ᵏ a)) (-ᵏ‿cong (*ᵏ-identityˡ a))) ≈₂-refl ⟩ 141 | (-ᵏ a) ∙₂ u 142 | ≈⟨ ∙₂-cong (≈ᵏ-trans (-ᵏ‿cong (≈ᵏ-sym (*ᵏ-identityʳ a))) (-ᵏ‿distribʳ-*ᵏ a 1ᵏ)) ≈₂-refl ⟩ 143 | (a *ᵏ (-ᵏ 1ᵏ)) ∙₂ u 144 | ≈⟨ ≈₂-trans (*ᵏ-∙₂-compat a (-ᵏ 1ᵏ) u) (∙₂-cong ≈ᵏ-refl (-1∙u≈₂-u u)) ⟩ 145 | a ∙₂ (-₂ u) 146 | ∎ 147 | in λ a b u v → 148 | begin 149 | π ⟨$⟩ a ∙₁ u +₁ b ∙₁ v 150 | ≈⟨ -₂‿cong (distrib-linear f a b u v) ⟩ 151 | -₂ (a ∙₂ (f ⟪$⟫ u) +₂ b ∙₂ (f ⟪$⟫ v)) 152 | ≈⟨ ≈₂-sym (-₂‿+₂-comm (a ∙₂ (f ⟪$⟫ u)) ( b ∙₂ (f ⟪$⟫ v))) ⟩ 153 | ((-₂ (a ∙₂ (f ⟪$⟫ u))) +₂ (-₂ (b ∙₂ (f ⟪$⟫ v)))) 154 | ≈⟨ +₂-cong (-[a∙u]≈₂a∙[-u] a (f ⟪$⟫ u)) (-[a∙u]≈₂a∙[-u] b (f ⟪$⟫ v)) ⟩ 155 | a ∙₂ (π ⟨$⟩ u) +₂ b ∙₂ (π ⟨$⟩ v) 156 | ∎ 157 | } where π : setoid₁ ⟶ setoid₂ 158 | π = record 159 | { _⟨$⟩_ = λ x → -₂ f ⟪$⟫ x 160 | ; cong = λ r → -₂‿cong (⟦⟧-cong f r) 161 | } 162 | 163 | infixr 25 _+_ 164 | _+_ : Op₂ V 165 | f + g = record 166 | { ⟦_⟧ = λ x -> (f ⟪$⟫ x) +₂ (g ⟪$⟫ x) 167 | ; isLinearMap = linear V₁-space V₂-space π 168 | let reassoc : ∀ (u v u' v' : V₂) -> (u +₂ v) +₂ (u' +₂ v') ≈₂ (u +₂ u') +₂ (v +₂ v') 169 | reassoc u v u' v' = 170 | begin 171 | (u +₂ v) +₂ (u' +₂ v') 172 | ≈⟨ +₂-assoc u v (u' +₂ v') ⟩ 173 | u +₂ (v +₂ (u' +₂ v')) 174 | ≈⟨ +₂-cong ≈₂-refl (≈₂-trans (≈₂-sym (+₂-assoc v u' v')) (+₂-cong (+₂-comm v u') ≈₂-refl)) ⟩ 175 | u +₂ ((u' +₂ v) +₂ v') 176 | ≈⟨ ≈₂-trans (+₂-cong ≈₂-refl (+₂-assoc u' v v')) (≈₂-sym (+₂-assoc u u' (v +₂ v'))) ⟩ 177 | (u +₂ u') +₂ (v +₂ v') 178 | ∎ 179 | distrib-f-g k w = ≈₂-sym (∙₂-+₂-distrib k (f ⟪$⟫ w) (g ⟪$⟫ w)) 180 | in λ a b u v → 181 | begin 182 | π ⟨$⟩ a ∙₁ u +₁ b ∙₁ v 183 | ≡⟨⟩ 184 | (f ⟪$⟫ (a ∙₁ u +₁ b ∙₁ v)) +₂ (g ⟪$⟫ (a ∙₁ u +₁ b ∙₁ v)) 185 | ≈⟨ +₂-cong (distrib-linear f a b u v) (distrib-linear g a b u v) ⟩ 186 | (a ∙₂ (f ⟪$⟫ u) +₂ b ∙₂ (f ⟪$⟫ v)) +₂ ((a ∙₂ (g ⟪$⟫ u) +₂ b ∙₂ (g ⟪$⟫ v))) 187 | ≈⟨ reassoc (a ∙₂ (f ⟪$⟫ u)) (b ∙₂ (f ⟪$⟫ v)) (a ∙₂ (g ⟪$⟫ u)) (b ∙₂ (g ⟪$⟫ v)) ⟩ 188 | (a ∙₂ (f ⟪$⟫ u) +₂ a ∙₂ (g ⟪$⟫ u)) +₂ ((b ∙₂ (f ⟪$⟫ v)) +₂ b ∙₂ (g ⟪$⟫ v)) 189 | ≈⟨ +₂-cong (distrib-f-g a u) (distrib-f-g b v) ⟩ 190 | a ∙₂ (π ⟨$⟩ u) +₂ b ∙₂ (π ⟨$⟩ v) 191 | ∎ 192 | } where π : setoid₁ ⟶ setoid₂ 193 | π = record 194 | { _⟨$⟩_ = λ x → (f ⟪$⟫ x) +₂ (g ⟪$⟫ x) 195 | ; cong = λ r → +₂-cong (⟦⟧-cong f r) (⟦⟧-cong g r) 196 | } 197 | 198 | infixr 30 _∙_ 199 | _∙_ : K' -> V -> V 200 | c ∙ f = record 201 | { ⟦_⟧ = λ x → c ∙₂ (f ⟪$⟫ x) 202 | ; isLinearMap = linear V₁-space V₂-space π 203 | λ a b u v → 204 | begin 205 | π ⟨$⟩ a ∙₁ u +₁ b ∙₁ v 206 | ≈⟨ ∙₂-cong ≈ᵏ-refl (distrib-linear f a b u v) ⟩ 207 | c ∙₂ (a ∙₂ (f ⟪$⟫ u) +₂ b ∙₂ (f ⟪$⟫ v)) 208 | ≈⟨ ∙₂-+₂-distrib c (a ∙₂ (f ⟪$⟫ u)) ( b ∙₂ (f ⟪$⟫ v)) ⟩ 209 | (c ∙₂ (a ∙₂ (f ⟪$⟫ u))) +₂ (c ∙₂ (b ∙₂ (f ⟪$⟫ v))) 210 | ≈⟨ +₂-cong (∙₂-∙₂-comm c a (f ⟪$⟫ u)) (∙₂-∙₂-comm c b (f ⟪$⟫ v)) ⟩ 211 | a ∙₂ (π ⟨$⟩ u) +₂ b ∙₂ (π ⟨$⟩ v) 212 | ∎ 213 | } where π : setoid₁ ⟶ setoid₂ 214 | π = record 215 | { _⟨$⟩_ = λ x → c ∙₂ (f ⟪$⟫ x) 216 | ; cong = λ r → ∙₂-cong ≈ᵏ-refl (⟦⟧-cong f r) 217 | } 218 | 219 | +-cong : Congruent₂ _+_ 220 | +-cong r₁ r₂ = λ rₓ -> +₂-cong (r₁ rₓ) (r₂ rₓ) 221 | 222 | +-assoc : Associative _+_ 223 | +-assoc f g h {y = y} rₓ = ≈₂-trans 224 | (+₂-cong (+₂-cong (⟦⟧-cong f rₓ) (⟦⟧-cong g rₓ)) (⟦⟧-cong h rₓ)) 225 | (+₂-assoc (f ⟪$⟫ y) (g ⟪$⟫ y) (h ⟪$⟫ y)) 226 | 227 | +-identityˡ : LeftIdentity 0# _+_ 228 | +-identityˡ f {x} rₓ = ≈₂-trans 229 | (+₂-identityˡ (f ⟪$⟫ x)) 230 | (⟦⟧-cong f rₓ) 231 | 232 | +-identityʳ : RightIdentity 0# _+_ 233 | +-identityʳ f {x} rₓ = ≈₂-trans 234 | (+₂-identityʳ (f ⟪$⟫ x)) 235 | (⟦⟧-cong f rₓ) 236 | 237 | +-identity : Identity 0# _+_ 238 | +-identity = +-identityˡ , +-identityʳ 239 | 240 | +-comm : Commutative _+_ 241 | +-comm f g {x} rₓ = ≈₂-trans 242 | (+₂-comm (f ⟪$⟫ x) (g ⟪$⟫ x)) 243 | (+₂-cong (⟦⟧-cong g rₓ) (⟦⟧-cong f rₓ)) 244 | 245 | *ᵏ-∙-compat : ∀ (a b : K') (f : V) -> ((a *ᵏ b) ∙ f) ≈ (a ∙ (b ∙ f)) 246 | *ᵏ-∙-compat a b f {x} rₓ = ≈₂-trans 247 | (*ᵏ-∙₂-compat a b (f ⟪$⟫ x)) 248 | (∙₂-cong ≈ᵏ-refl (∙₂-cong ≈ᵏ-refl (⟦⟧-cong f rₓ))) 249 | 250 | ∙-+-distrib : ∀ (a : K') (f g : V) -> (a ∙ (f + g)) ≈ ((a ∙ f) + (a ∙ g)) 251 | ∙-+-distrib a f g {x} rₓ = ≈₂-trans 252 | (∙₂-+₂-distrib a (f ⟪$⟫ x) (g ⟪$⟫ x)) 253 | (+₂-cong (∙₂-cong ≈ᵏ-refl (⟦⟧-cong f rₓ)) (∙₂-cong ≈ᵏ-refl (⟦⟧-cong g rₓ))) 254 | 255 | ∙-+ᵏ-distrib : ∀ (a b : K') (f : V) -> ((a +ᵏ b) ∙ f) ≈ ((a ∙ f) + (b ∙ f)) 256 | ∙-+ᵏ-distrib a b f {x} rₓ = ≈₂-trans 257 | (∙₂-+ᵏ-distrib a b (f ⟪$⟫ x)) 258 | (+₂-cong cf cf) 259 | where cf : ∀ {c : K'} -> _ 260 | cf {c} = ∙₂-cong (≈ᵏ-refl {c}) (⟦⟧-cong f rₓ) 261 | 262 | ∙-cong : ∀ {a b : K'} {f g : V} -> a ≈ᵏ b -> f ≈ g -> (a ∙ f) ≈ (b ∙ g) 263 | ∙-cong {f = f} {g = g} rk rf rₓ = ≈₂-trans 264 | (∙₂-cong rk (rf rₓ)) 265 | (∙₂-cong ≈ᵏ-refl (⟦⟧-cong g ≈₁-refl)) 266 | 267 | ∙-identity : ∀ (f : V) → (1ᵏ ∙ f) ≈ f 268 | ∙-identity f {x} rₓ = ≈₂-trans (∙₂-identity (f ⟪$⟫ x)) (⟦⟧-cong f rₓ) 269 | 270 | ∙-absorbˡ : ∀ (x : V) → (0ᵏ ∙ x) ≈ 0# 271 | ∙-absorbˡ f {x} _ = ≈₂-trans (∙₂-absorbˡ (f ⟪$⟫ x)) ≈₂-refl 272 | 273 | -‿inverseˡ : LeftInverse 0# -_ _+_ 274 | -‿inverseˡ f {x} _ = ≈₂-trans (-₂‿inverseˡ (f ⟪$⟫ x)) ≈₂-refl 275 | 276 | -‿inverseʳ : RightInverse 0# -_ _+_ 277 | -‿inverseʳ f {x} _ = -₂‿inverseʳ (f ⟪$⟫ x) 278 | 279 | -‿inverse : Inverse 0# -_ _+_ 280 | -‿inverse = -‿inverseˡ , -‿inverseʳ 281 | 282 | -‿cong : Congruent₁ -_ 283 | -‿cong r rₓ = -₂‿cong (r rₓ) 284 | 285 | isMagma : IsMagma _+_ 286 | isMagma = record 287 | { isEquivalence = ≈-isEquiv 288 | -- I don't know how to solve this: I'd like to write simply ∙-cong = +-cong 289 | ; ∙-cong = λ {f} {g} {u} {v} -> +-cong {f} {g} {u} {v} 290 | } 291 | 292 | isSemigroup : IsSemigroup _+_ 293 | isSemigroup = record 294 | { isMagma = isMagma 295 | ; assoc = +-assoc 296 | } 297 | 298 | isMonoid : IsMonoid _+_ 0# 299 | isMonoid = record 300 | { isSemigroup = isSemigroup 301 | ; identity = +-identity 302 | } 303 | 304 | isGroup : IsGroup _+_ 0# -_ 305 | isGroup = record 306 | { isMonoid = isMonoid 307 | ; inverse = -‿inverse 308 | ; ⁻¹-cong = λ {f} {g} -> -‿cong {f} {g} 309 | } 310 | 311 | isAbelianGroup : IsAbelianGroup _+_ 0# -_ 312 | isAbelianGroup = record 313 | { isGroup = isGroup 314 | ; comm = +-comm 315 | } 316 | 317 | isVectorSpace : IsVectorSpace K _≈_ _+_ _∙_ -_ 0# 318 | isVectorSpace = record 319 | { isAbelianGroup = isAbelianGroup 320 | ; *ᵏ-∙-compat = *ᵏ-∙-compat 321 | ; ∙-+-distrib = ∙-+-distrib 322 | ; ∙-+ᵏ-distrib = ∙-+ᵏ-distrib 323 | ; ∙-cong = λ {a} {b} {f} {g} -> ∙-cong {a} {b} {f} {g} 324 | ; ∙-identity = ∙-identity 325 | ; ∙-absorbˡ = ∙-absorbˡ 326 | } 327 | 328 | vectorSpace : VectorSpace K (suc (k ⊔ a₁ ⊔ a₂ ⊔ ℓ₁ ⊔ ℓ₂)) (a₁ ⊔ ℓ₁ ⊔ ℓ₂) 329 | vectorSpace = record { isVectorSpace = isVectorSpace } 330 | -------------------------------------------------------------------------------- /src/Algebra/Linear/Space/Product.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Algebra.Structures.Bundles.Field 4 | open import Algebra.Linear.Structures.Bundles 5 | 6 | module Algebra.Linear.Space.Product 7 | {k ℓ} (K : Field k ℓ) 8 | {a₁ ℓ₁} (V₁-space : VectorSpace K a₁ ℓ₁) 9 | {a₂ ℓ₂} (V₂-space : VectorSpace K a₂ ℓ₂) 10 | where 11 | 12 | open import Relation.Binary 13 | open import Level using (_⊔_) 14 | open import Data.Nat using (ℕ; zero; suc) renaming (_+_ to _+ℕ_) 15 | open import Algebra.Linear.Structures.VectorSpace 16 | import Algebra.FunctionProperties as FP 17 | 18 | open VectorSpaceField K 19 | 20 | open VectorSpace V₁-space 21 | using () 22 | renaming 23 | ( Carrier to V₁ 24 | ; _≈_ to _≈₁_ 25 | ; isEquivalence to ≈₁-isEquiv 26 | ; refl to ≈₁-refl 27 | ; sym to ≈₁-sym 28 | ; trans to ≈₁-trans 29 | ; _+_ to _+₁_ 30 | ; _∙_ to _∙₁_ 31 | ; -_ to -₁_ 32 | ; 0# to 0₁ 33 | ; +-identityˡ to +₁-identityˡ 34 | ; +-identityʳ to +₁-identityʳ 35 | ; +-identity to +₁-identity 36 | ; +-cong to +₁-cong 37 | ; +-assoc to +₁-assoc 38 | ; +-comm to +₁-comm 39 | ; *ᵏ-∙-compat to *ᵏ-∙₁-compat 40 | ; ∙-+-distrib to ∙₁-+₁-distrib 41 | ; ∙-+ᵏ-distrib to ∙₁-+ᵏ-distrib 42 | ; ∙-cong to ∙₁-cong 43 | ; ∙-identity to ∙₁-identity 44 | ; ∙-absorbˡ to ∙₁-absorbˡ 45 | ; ∙-absorbʳ to ∙₁-absorbʳ 46 | ; -‿cong to -₁‿cong 47 | ; -‿inverseˡ to -₁‿inverseˡ 48 | ; -‿inverseʳ to -₁‿inverseʳ 49 | ) 50 | 51 | open VectorSpace V₂-space 52 | using () 53 | renaming 54 | ( Carrier to V₂ 55 | ; _≈_ to _≈₂_ 56 | ; isEquivalence to ≈₂-isEquiv 57 | ; refl to ≈₂-refl 58 | ; sym to ≈₂-sym 59 | ; trans to ≈₂-trans 60 | ; _+_ to _+₂_ 61 | ; _∙_ to _∙₂_ 62 | ; -_ to -₂_ 63 | ; 0# to 0₂ 64 | ; +-identityˡ to +₂-identityˡ 65 | ; +-identityʳ to +₂-identityʳ 66 | ; +-identity to +₂-identity 67 | ; +-cong to +₂-cong 68 | ; +-assoc to +₂-assoc 69 | ; +-comm to +₂-comm 70 | ; *ᵏ-∙-compat to *ᵏ-∙₂-compat 71 | ; ∙-+-distrib to ∙₂-+₂-distrib 72 | ; ∙-+ᵏ-distrib to ∙₂-+ᵏ-distrib 73 | ; ∙-cong to ∙₂-cong 74 | ; ∙-identity to ∙₂-identity 75 | ; ∙-absorbˡ to ∙₂-absorbˡ 76 | ; ∙-absorbʳ to ∙₂-absorbʳ 77 | ; -‿cong to -₂‿cong 78 | ; -‿inverseˡ to -₂‿inverseˡ 79 | ; -‿inverseʳ to -₂‿inverseʳ 80 | ) 81 | 82 | open import Data.Product 83 | open import Data.Product.Relation.Binary.Pointwise.NonDependent 84 | 85 | private 86 | V : Set (a₁ ⊔ a₂) 87 | V = V₁ × V₂ 88 | 89 | _≈_ : Rel V (ℓ₁ ⊔ ℓ₂) 90 | _≈_ = Pointwise _≈₁_ _≈₂_ 91 | 92 | ≈-isEquiv : IsEquivalence _≈_ 93 | ≈-isEquiv = ×-isEquivalence ≈₁-isEquiv ≈₂-isEquiv 94 | 95 | prod-setoid : Setoid (a₁ ⊔ a₂) (ℓ₁ ⊔ ℓ₂) 96 | prod-setoid = record 97 | { Carrier = V 98 | ; _≈_ = _≈_ 99 | ; isEquivalence = ≈-isEquiv 100 | } 101 | 102 | open IsEquivalence ≈-isEquiv renaming 103 | ( refl to ≈-refl 104 | ; sym to ≈-sym 105 | ; trans to ≈-trans 106 | ) 107 | 108 | open import Algebra.FunctionProperties _≈_ 109 | open import Algebra.Structures _≈_ 110 | 111 | 0# : V 112 | 0# = (0₁ , 0₂) 113 | 114 | -_ : Op₁ V 115 | - (x₁ , x₂) = (-₁ x₁ , -₂ x₂) 116 | 117 | infixr 25 _+_ 118 | _+_ : Op₂ V 119 | (x₁ , x₂) + (y₁ , y₂) = (x₁ +₁ y₁ , x₂ +₂ y₂) 120 | 121 | infixr 30 _∙_ 122 | _∙_ : K' -> V -> V 123 | k ∙ (x₁ , x₂) = (k ∙₁ x₁ , k ∙₂ x₂) 124 | 125 | +-cong : Congruent₂ _+_ 126 | +-cong (r₁ , r₂) (s₁ , s₂) = ( +₁-cong r₁ s₁ , +₂-cong r₂ s₂ ) 127 | 128 | +-assoc : Associative _+_ 129 | +-assoc (x₁ , x₂) (y₁ , y₂) (z₁ , z₂) = ( +₁-assoc x₁ y₁ z₁ , +₂-assoc x₂ y₂ z₂ ) 130 | 131 | +-identityˡ : LeftIdentity 0# _+_ 132 | +-identityˡ (x₁ , x₂) = ( +₁-identityˡ x₁ , +₂-identityˡ x₂ ) 133 | 134 | +-identityʳ : RightIdentity 0# _+_ 135 | +-identityʳ (x₁ , x₂) = ( +₁-identityʳ x₁ , +₂-identityʳ x₂ ) 136 | 137 | +-identity : Identity 0# _+_ 138 | +-identity = +-identityˡ , +-identityʳ 139 | 140 | +-comm : Commutative _+_ 141 | +-comm (x₁ , x₂) (y₁ , y₂) = ( +₁-comm x₁ y₁ , +₂-comm x₂ y₂ ) 142 | 143 | *ᵏ-∙-compat : ∀ (a b : K') (u : V) -> ((a *ᵏ b) ∙ u) ≈ (a ∙ (b ∙ u)) 144 | *ᵏ-∙-compat a b (x₁ , x₂) = ( *ᵏ-∙₁-compat a b x₁ , *ᵏ-∙₂-compat a b x₂ ) 145 | 146 | ∙-+-distrib : ∀ (a : K') (u v : V) -> (a ∙ (u + v)) ≈ ((a ∙ u) + (a ∙ v)) 147 | ∙-+-distrib a (x₁ , x₂) (y₁ , y₂) = ( ∙₁-+₁-distrib a x₁ y₁ , ∙₂-+₂-distrib a x₂ y₂ ) 148 | 149 | ∙-+ᵏ-distrib : ∀ (a b : K') (u : V) -> ((a +ᵏ b) ∙ u) ≈ ((a ∙ u) + (b ∙ u)) 150 | ∙-+ᵏ-distrib a b (x₁ , x₂) = ( ∙₁-+ᵏ-distrib a b x₁ , ∙₂-+ᵏ-distrib a b x₂ ) 151 | 152 | ∙-cong : ∀ {a b : K'} {u v : V} -> a ≈ᵏ b -> u ≈ v -> (a ∙ u) ≈ (b ∙ v) 153 | ∙-cong rₓ (r₁ , r₂) = ( ∙₁-cong rₓ r₁ , ∙₂-cong rₓ r₂ ) 154 | 155 | ∙-identity : ∀ (x : V) → (1ᵏ ∙ x) ≈ x 156 | ∙-identity (x₁ , x₂) = ( ∙₁-identity x₁ , ∙₂-identity x₂ ) 157 | 158 | ∙-absorbˡ : ∀ (x : V) → (0ᵏ ∙ x) ≈ 0# 159 | ∙-absorbˡ (x₁ , x₂) = ( ∙₁-absorbˡ x₁ , ∙₂-absorbˡ x₂ ) 160 | 161 | -‿inverseˡ : LeftInverse 0# -_ _+_ 162 | -‿inverseˡ (x₁ , x₂) = ( -₁‿inverseˡ x₁ , -₂‿inverseˡ x₂ ) 163 | 164 | -‿inverseʳ : RightInverse 0# -_ _+_ 165 | -‿inverseʳ (x₁ , x₂) = ( -₁‿inverseʳ x₁ , -₂‿inverseʳ x₂ ) 166 | 167 | -‿inverse : Inverse 0# -_ _+_ 168 | -‿inverse = -‿inverseˡ , -‿inverseʳ 169 | 170 | -‿cong : Congruent₁ -_ 171 | -‿cong (x₁ , x₂) = ( -₁‿cong x₁ , -₂‿cong x₂ ) 172 | 173 | isMagma : IsMagma _+_ 174 | isMagma = record 175 | { isEquivalence = ≈-isEquiv 176 | ; ∙-cong = +-cong 177 | } 178 | 179 | isSemigroup : IsSemigroup _+_ 180 | isSemigroup = record 181 | { isMagma = isMagma 182 | ; assoc = +-assoc 183 | } 184 | 185 | isMonoid : IsMonoid _+_ 0# 186 | isMonoid = record 187 | { isSemigroup = isSemigroup 188 | ; identity = +-identity 189 | } 190 | 191 | isGroup : IsGroup _+_ 0# -_ 192 | isGroup = record 193 | { isMonoid = isMonoid 194 | ; inverse = -‿inverse 195 | ; ⁻¹-cong = -‿cong 196 | } 197 | 198 | isAbelianGroup : IsAbelianGroup _+_ 0# -_ 199 | isAbelianGroup = record 200 | { isGroup = isGroup 201 | ; comm = +-comm 202 | } 203 | 204 | isVectorSpace : IsVectorSpace K _≈_ _+_ _∙_ -_ 0# 205 | isVectorSpace = record 206 | { isAbelianGroup = isAbelianGroup 207 | ; *ᵏ-∙-compat = *ᵏ-∙-compat 208 | ; ∙-+-distrib = ∙-+-distrib 209 | ; ∙-+ᵏ-distrib = ∙-+ᵏ-distrib 210 | ; ∙-cong = ∙-cong 211 | ; ∙-identity = ∙-identity 212 | ; ∙-absorbˡ = ∙-absorbˡ 213 | } 214 | 215 | vectorSpace : VectorSpace K (a₁ ⊔ a₂) (ℓ₁ ⊔ ℓ₂) 216 | vectorSpace = record { isVectorSpace = isVectorSpace } 217 | -------------------------------------------------------------------------------- /src/Algebra/Linear/Structures.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Algebra.Linear.Structures where 4 | 5 | open import Algebra.Structures.Field public 6 | open import Algebra.Linear.Structures.VectorSpace public 7 | -------------------------------------------------------------------------------- /src/Algebra/Linear/Structures/Bundles.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Algebra.Linear.Structures.Bundles where 4 | 5 | open import Algebra 6 | open import Level using (Level; suc; _⊔_) 7 | open import Relation.Binary using (Rel; IsEquivalence) 8 | 9 | open import Algebra.FunctionProperties 10 | open import Algebra.Structures.Field 11 | open import Algebra.Structures.Bundles.Field 12 | open import Algebra.Linear.Structures.VectorSpace 13 | 14 | record RawVectorSpace {k ℓᵏ} (K : Field k ℓᵏ) (c ℓ : Level) : Set (suc (c ⊔ k ⊔ ℓ)) where 15 | infix 20 _≈_ 16 | infixl 25 _+_ 17 | infixr 30 _∙_ 18 | 19 | field 20 | Carrier : Set c 21 | _≈_ : Rel Carrier ℓ 22 | _+_ : Carrier -> Carrier -> Carrier 23 | _∙_ : Field.Carrier K -> Carrier -> Carrier 24 | -_ : Carrier -> Carrier 25 | 0# : Carrier 26 | 27 | rawGroup : RawGroup c ℓ 28 | rawGroup = record { _≈_ = _≈_; _∙_ = _+_; ε = 0#; _⁻¹ = -_ } 29 | 30 | open RawGroup rawGroup public 31 | using (rawMagma; rawMonoid) 32 | 33 | record VectorSpace {k ℓᵏ} (K : Field k ℓᵏ) (c ℓ : Level) : Set (suc (c ⊔ k ⊔ ℓ ⊔ ℓᵏ)) where 34 | infix 20 _≈_ 35 | infixl 25 _+_ 36 | infixr 30 _∙_ 37 | 38 | field 39 | Carrier : Set c 40 | _≈_ : Rel Carrier ℓ 41 | _+_ : Carrier -> Carrier -> Carrier 42 | _∙_ : Field.Carrier K -> Carrier -> Carrier 43 | -_ : Carrier -> Carrier 44 | 0# : Carrier 45 | isVectorSpace : IsVectorSpace K _≈_ _+_ _∙_ -_ 0# 46 | 47 | open IsVectorSpace isVectorSpace public 48 | 49 | rawVectorSpace : RawVectorSpace K c ℓ 50 | rawVectorSpace = record 51 | { _≈_ = _≈_; _+_ = _+_; _∙_ = _∙_; -_ = -_; 0# = 0# } 52 | 53 | abelianGroup : AbelianGroup c ℓ 54 | abelianGroup = record { isAbelianGroup = isAbelianGroup } 55 | 56 | open AbelianGroup abelianGroup public 57 | using (rawMagma; magma; rawMonoid; monoid; rawGroup; group) 58 | -------------------------------------------------------------------------------- /src/Algebra/Linear/Structures/Bundles/FiniteDimensional.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Algebra.Linear.Structures.Bundles.FiniteDimensional where 4 | 5 | open import Algebra 6 | open import Algebra.FunctionProperties 7 | open import Relation.Binary using (Rel) 8 | open import Level using (Level; suc; _⊔_) 9 | 10 | open import Algebra.Structures.Bundles.Field 11 | open import Algebra.Linear.Structures.Bundles 12 | 13 | open import Algebra.Linear.Structures.VectorSpace 14 | open import Algebra.Linear.Structures.FiniteDimensional 15 | 16 | open import Data.Nat using (ℕ) 17 | 18 | record FiniteDimensional {k ℓᵏ} (K : Field k ℓᵏ) (c ℓ : Level) (n : ℕ) : Set (suc (c ⊔ k ⊔ ℓ ⊔ ℓᵏ)) where 19 | field 20 | Carrier : Set c 21 | _≈_ : Rel Carrier ℓ 22 | _+_ : Carrier -> Carrier -> Carrier 23 | _∙_ : Field.Carrier K -> Carrier -> Carrier 24 | -_ : Carrier -> Carrier 25 | 0# : Carrier 26 | isFiniteDimensional : IsFiniteDimensional K _≈_ _+_ _∙_ -_ 0# n 27 | 28 | open IsFiniteDimensional isFiniteDimensional public 29 | 30 | vectorSpace : VectorSpace K c ℓ 31 | vectorSpace = record { isVectorSpace = isVectorSpace } 32 | -------------------------------------------------------------------------------- /src/Algebra/Linear/Structures/FiniteDimensional.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Algebra.Structures.Bundles.Field 4 | 5 | module Algebra.Linear.Structures.FiniteDimensional 6 | {k ℓᵏ} (K : Field k ℓᵏ) 7 | where 8 | 9 | open import Algebra.Linear.Core 10 | open import Algebra.FunctionProperties 11 | open import Relation.Binary using (Rel) 12 | 13 | open import Algebra.Linear.Morphism.Bundles K 14 | 15 | open import Algebra.Linear.Structures.VectorSpace K 16 | 17 | open import Function.Equality 18 | 19 | open import Level using (_⊔_; suc) 20 | open import Data.Nat using (ℕ) 21 | import Algebra.Linear.Construct.Vector K as Vec 22 | 23 | private 24 | K' : Set k 25 | K' = Field.Carrier K 26 | 27 | record IsFiniteDimensional 28 | {v ℓ} {V : Set v} 29 | (_≈_ : Rel V ℓ) 30 | (_+_ : Op₂ V) (_∙_ : K' → V → V) (-_ : Op₁ V) (0# : V) 31 | (n : ℕ) : Set (suc (v ⊔ k ⊔ ℓ ⊔ ℓᵏ)) 32 | where 33 | field 34 | isVectorSpace : IsVectorSpace _≈_ _+_ _∙_ -_ 0# 35 | embed : LinearIsomorphism (record { isVectorSpace = isVectorSpace }) 36 | (record { isVectorSpace = Vec.isVectorSpace {n} }) 37 | 38 | open IsVectorSpace isVectorSpace public 39 | 40 | module Vector {n} where 41 | open Vec 42 | 43 | isFiniteDimensional : IsFiniteDimensional (Vec._≈_) (Vec._+_) (Vec._∙_) (Vec.-_) (Vec.0#) n 44 | isFiniteDimensional = record 45 | { isVectorSpace = Vec.isVectorSpace 46 | ; embed = Vec.embed 47 | } 48 | -------------------------------------------------------------------------------- /src/Algebra/Linear/Structures/GL.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Algebra.Structures.Bundles.Field 4 | 5 | module Algebra.Linear.Structures.GL 6 | {k ℓᵏ} (K : Field k ℓᵏ) 7 | {a ℓ} 8 | where 9 | 10 | open import Relation.Binary using (Setoid) 11 | open import Algebra.Linear.Structures.Bundles 12 | open import Algebra.Linear.Morphism.Bundles K 13 | open import Categories.Category 14 | 15 | open LinearMap 16 | -------------------------------------------------------------------------------- /src/Algebra/Linear/Structures/VectorSpace.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Algebra.Structures.Bundles.Field 4 | 5 | module Algebra.Linear.Structures.VectorSpace 6 | {k ℓᵏ} (K : Field k ℓᵏ) 7 | where 8 | 9 | open import Algebra.FunctionProperties 10 | open import Relation.Binary using (Rel) 11 | open import Level using (_⊔_) 12 | open import Data.Nat using (ℕ) 13 | 14 | module VectorSpaceField where 15 | open Field K public 16 | using () 17 | renaming 18 | ( Carrier to K' 19 | ; _≈_ to _≈ᵏ_ 20 | ; refl to ≈ᵏ-refl 21 | ; sym to ≈ᵏ-sym 22 | ; trans to ≈ᵏ-trans 23 | ; reflexive to ≈ᵏ-reflexive 24 | ; isEquivalence to ≈ᵏ-isEquiv 25 | ; setoid to K-setoid 26 | ; _+_ to _+ᵏ_ 27 | ; _*_ to _*ᵏ_ 28 | ; _-_ to _-ᵏ_ 29 | ; 0# to 0ᵏ 30 | ; 1# to 1ᵏ 31 | ; -_ to -ᵏ_ 32 | ; _⁻¹ to _⁻¹ᵏ 33 | ; +-cong to +ᵏ-cong 34 | ; +-identity to +ᵏ-identity 35 | ; +-identityˡ to +ᵏ-identityˡ 36 | ; +-identityʳ to +ᵏ-identityʳ 37 | ; +-assoc to +ᵏ-assoc 38 | ; +-comm to +ᵏ-comm 39 | ; *-cong to *ᵏ-cong 40 | ; *-identity to *ᵏ-identity 41 | ; *-identityˡ to *ᵏ-identityˡ 42 | ; *-identityʳ to *ᵏ-identityʳ 43 | ; *-assoc to *ᵏ-assoc 44 | ; *-comm to *ᵏ-comm 45 | ; zero to *ᵏ-zero 46 | ; zeroˡ to *ᵏ-zeroˡ 47 | ; zeroʳ to *ᵏ-zeroʳ 48 | ; distrib to *ᵏ-+ᵏ-distrib 49 | ; distribˡ to *ᵏ-+ᵏ-distribˡ 50 | ; distribʳ to *ᵏ-+ᵏ-distribʳ 51 | ; -‿distribˡ-* to -ᵏ‿distribˡ-*ᵏ 52 | ; -‿distribʳ-* to -ᵏ‿distribʳ-*ᵏ 53 | ; -‿cong to -ᵏ‿cong 54 | ; -‿inverse to -ᵏ‿inverse 55 | ; -‿inverseˡ to -ᵏ‿inverseˡ 56 | ; -‿inverseʳ to -ᵏ‿inverseʳ 57 | ; uniqueˡ-⁻¹ to uniqueˡ-⁻¹ᵏ 58 | ; uniqueʳ-⁻¹ to uniqueʳ-⁻¹ᵏ 59 | ; _⁻¹-inverse to _⁻¹ᵏ-inverse 60 | ; _⁻¹-involutive to _⁻¹ᵏ-involutive 61 | ; 0#-not-1# to 0ᵏ-not-1ᵏ 62 | ) 63 | 64 | module _ 65 | {v ℓ} {V : Set v} 66 | (_≈_ : Rel V ℓ) 67 | where 68 | 69 | open VectorSpaceField 70 | 71 | open import Algebra.Structures _≈_ 72 | 73 | record IsVectorSpace (_+_ : Op₂ V) (_∙_ : K' → V → V) (-_ : Op₁ V) (0# : V) : Set (v ⊔ k ⊔ ℓ ⊔ ℓᵏ) where 74 | field 75 | isAbelianGroup : IsAbelianGroup _+_ 0# -_ 76 | *ᵏ-∙-compat : ∀ (a b : K') (u : V) -> ((a *ᵏ b) ∙ u) ≈ (a ∙ (b ∙ u)) 77 | ∙-+-distrib : ∀ (a : K') (u v : V) -> (a ∙ (u + v)) ≈ ((a ∙ u) + (a ∙ v)) 78 | ∙-+ᵏ-distrib : ∀ (a b : K') (u : V) -> ((a +ᵏ b) ∙ u) ≈ ((a ∙ u) + (b ∙ u)) 79 | ∙-cong : ∀ {a b : K'} {u v : V} -> a ≈ᵏ b -> u ≈ v -> (a ∙ u) ≈ (b ∙ v) 80 | ∙-identity : ∀ (x : V) → (1ᵏ ∙ x) ≈ x 81 | ∙-absorbˡ : ∀ (x : V) → (0ᵏ ∙ x) ≈ 0# 82 | 83 | open IsAbelianGroup isAbelianGroup public 84 | renaming 85 | ( assoc to +-assoc 86 | ; ∙-cong to +-cong 87 | ; identity to +-identity 88 | ; identityˡ to +-identityˡ 89 | ; identityʳ to +-identityʳ 90 | ; comm to +-comm 91 | ; inverse to -‿inverse 92 | ; inverseˡ to -‿inverseˡ 93 | ; inverseʳ to -‿inverseʳ 94 | ; ⁻¹-cong to -‿cong 95 | ) 96 | 97 | open import Algebra.Properties.AbelianGroup (record { isAbelianGroup = isAbelianGroup }) public 98 | renaming 99 | ( ε⁻¹≈ε to -0≈0 100 | ; ∙-cancelˡ to +-cancelˡ 101 | ; ∙-cancelʳ to +-cancelʳ 102 | ; ∙-cancel to +-cancel 103 | ; identityˡ-unique to +-identityˡ-unique 104 | ; identityʳ-unique to +-identityʳ-unique 105 | ; identity-unique to +-identity-unique 106 | ; inverseˡ-unique to +-inverseˡ-unique 107 | ; inverseʳ-unique to +-inverseʳ-unique 108 | ; xyx⁻¹≈y to x+y-x≈y 109 | ; ⁻¹-∙-comm to -‿+-comm 110 | ) 111 | 112 | open import Relation.Binary.EqReasoning setoid 113 | 114 | ∙-absorbʳ : ∀ (a : K') -> (a ∙ 0#) ≈ 0# 115 | ∙-absorbʳ a = 116 | begin 117 | a ∙ 0# 118 | ≈⟨ trans (∙-cong ≈ᵏ-refl (sym (∙-absorbˡ 0#))) (sym (*ᵏ-∙-compat a 0ᵏ 0#)) ⟩ 119 | (a *ᵏ 0ᵏ) ∙ 0# 120 | ≈⟨ ∙-cong (*ᵏ-zeroʳ a) refl ⟩ 121 | 0ᵏ ∙ 0# 122 | ≈⟨ ∙-absorbˡ 0# ⟩ 123 | 0# 124 | ∎ 125 | 126 | ∙-∙-comm : ∀ (a b : K') (u : V) -> (a ∙ (b ∙ u)) ≈ (b ∙ (a ∙ u)) 127 | ∙-∙-comm a b u = 128 | begin 129 | a ∙ (b ∙ u) 130 | ≈⟨ sym (*ᵏ-∙-compat a b u) ⟩ 131 | (a *ᵏ b) ∙ u 132 | ≈⟨ ∙-cong (*ᵏ-comm a b) refl ⟩ 133 | (b *ᵏ a) ∙ u 134 | ≈⟨ *ᵏ-∙-compat b a u ⟩ 135 | b ∙ (a ∙ u) 136 | ∎ 137 | 138 | private 139 | lemma-inverseʳ : ∀ (u : V) -> (u + ((-ᵏ 1ᵏ) ∙ u)) ≈ 0# 140 | lemma-inverseʳ u = 141 | begin 142 | u + ((-ᵏ 1ᵏ) ∙ u) 143 | ≈⟨ +-cong (sym (∙-identity u)) refl ⟩ 144 | (1ᵏ ∙ u) + ((-ᵏ 1ᵏ) ∙ u) 145 | ≈⟨ sym (∙-+ᵏ-distrib 1ᵏ (-ᵏ 1ᵏ) u) ⟩ 146 | (1ᵏ +ᵏ (-ᵏ 1ᵏ)) ∙ u 147 | ≈⟨ ∙-cong (-ᵏ‿inverseʳ 1ᵏ) refl ⟩ 148 | 0ᵏ ∙ u 149 | ≈⟨ ∙-absorbˡ u ⟩ 150 | 0# 151 | ∎ 152 | 153 | -1∙u≈-u : ∀ (u : V) -> ((-ᵏ 1ᵏ) ∙ u) ≈ (- u) 154 | -1∙u≈-u u = 155 | begin 156 | ((-ᵏ 1ᵏ) ∙ u) 157 | ≈⟨ +-inverseʳ-unique u ((-ᵏ 1ᵏ) ∙ u) (lemma-inverseʳ u) ⟩ 158 | - u 159 | ∎ 160 | -------------------------------------------------------------------------------- /src/Algebra/Structures/Bundles/Field.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | module Algebra.Structures.Bundles.Field where 4 | 5 | open import Algebra 6 | open import Level using (suc; _⊔_) 7 | open import Relation.Binary using (Rel) 8 | 9 | open import Algebra.Structures.Field 10 | 11 | record Field c ℓ : Set (suc (c ⊔ ℓ)) where 12 | field 13 | Carrier : Set c 14 | _≈_ : Rel Carrier ℓ 15 | _+_ : Carrier -> Carrier -> Carrier 16 | _*_ : Carrier -> Carrier -> Carrier 17 | 0# : Carrier 18 | 1# : Carrier 19 | -_ : Carrier -> Carrier 20 | _⁻¹ : MultiplicativeInverse _≈_ 0# 21 | isField : IsField _≈_ _+_ _*_ 0# 1# -_ _⁻¹ 22 | 23 | open IsField isField public 24 | -------------------------------------------------------------------------------- /src/Algebra/Structures/Field.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Relation.Binary using (Rel) 4 | 5 | module Algebra.Structures.Field 6 | {a ℓ} {A : Set a} 7 | (_≈_ : Rel A ℓ) 8 | where 9 | 10 | open import Algebra.Linear.Core 11 | 12 | open import Relation.Nullary using (¬_) 13 | 14 | open import Algebra.Structures _≈_ 15 | open import Algebra.FunctionProperties _≈_ 16 | 17 | open import Level using (_⊔_) 18 | 19 | record NonZero (0# : A) : Set (a ⊔ ℓ) where 20 | field 21 | value : A 22 | non-zero : ¬ (value ≈ 0#) 23 | 24 | MultiplicativeInverse : ∀ (0# : A) -> Set (a ⊔ ℓ) 25 | MultiplicativeInverse 0# = NonZero 0# → NonZero 0# 26 | 27 | record IsField (_+_ _*_ : Op₂ A) (0# 1# : A) (-_ : Op₁ A) (_⁻¹ : MultiplicativeInverse 0#) : Set (a ⊔ ℓ) where 28 | field 29 | isCommutativeRing : IsCommutativeRing _+_ _*_ -_ 0# 1# 30 | _⁻¹-involutive : ∀ (x : NonZero 0#) → NonZero.value ((x ⁻¹) ⁻¹) ≈ NonZero.value x 31 | _⁻¹-inverse : ∀ (x : NonZero 0#) → ((NonZero.value x) * (NonZero.value (x ⁻¹))) ≈ 1# 32 | 0#-not-1# : ¬ (0# ≈ 1#) 33 | 34 | open IsCommutativeRing isCommutativeRing public 35 | 36 | open import Algebra.Properties.Ring (record { isRing = isRing }) public 37 | -------------------------------------------------------------------------------- /src/Algebra/Structures/Field/Utils.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --without-K --safe #-} 2 | 3 | open import Algebra.Structures.Bundles.Field 4 | 5 | module Algebra.Structures.Field.Utils 6 | {k ℓ} (K : Field k ℓ) 7 | where 8 | 9 | open Field K hiding (zero) 10 | 11 | open import Data.Fin 12 | import Data.Nat as Nat 13 | 14 | private 15 | K' : Set k 16 | K' = Carrier 17 | 18 | δ : ∀ {n p} -> Fin n -> Fin p -> K' 19 | δ zero zero = 1# 20 | δ (suc i) zero = 0# 21 | δ zero (suc j) = 0# 22 | δ (suc i) (suc j) = δ i j 23 | 24 | δ-comm : ∀ {n p} (i : Fin n) (j : Fin p) -> δ i j ≈ δ j i 25 | δ-comm zero zero = refl 26 | δ-comm (suc i) zero = refl 27 | δ-comm zero (suc j) = refl 28 | δ-comm (suc i) (suc j) = δ-comm i j 29 | 30 | δ-cancelˡ : ∀ {n p} (j : Fin p) -> δ {Nat.suc n} zero (suc j) ≈ 0# 31 | δ-cancelˡ zero = refl 32 | δ-cancelˡ (suc j) = refl 33 | 34 | δ-cancelʳ : ∀ {n p} (i : Fin n) -> δ {p = Nat.suc p} (suc i) zero ≈ 0# 35 | δ-cancelʳ {n} {p} i = trans (δ-comm (suc i) (zero {Nat.suc p})) (δ-cancelˡ {p} {n} i) 36 | --------------------------------------------------------------------------------