├── .gitignore ├── quantitative.agda-lib ├── src ├── Lib │ ├── Category.agda │ ├── Category │ │ └── Examples.agda │ ├── Dec.agda │ ├── Dec │ │ └── Properties.agda │ ├── Equality.agda │ ├── Function.agda │ ├── FunctionProperties.agda │ ├── Level.agda │ ├── List.agda │ ├── Matrix.agda │ ├── Matrix │ │ ├── Addition.agda │ │ ├── Addition │ │ │ └── Order.agda │ │ ├── Multiplication.agda │ │ ├── Multiplication │ │ │ ├── Basis.agda │ │ │ ├── Block.agda │ │ │ └── Order.agda │ │ ├── Poset.agda │ │ ├── Preorder.agda │ │ ├── Scaling.agda │ │ ├── Scaling │ │ │ └── Right.agda │ │ ├── Setoid.agda │ │ └── VecCompat.agda │ ├── Maybe.agda │ ├── Module.agda │ ├── Nat.agda │ ├── One.agda │ ├── Product.agda │ ├── Relation.agda │ ├── Relation │ │ └── Propositional.agda │ ├── Setoid.agda │ ├── Structure.agda │ ├── Structure │ │ ├── CommutativeMonoid.agda │ │ ├── Flip.agda │ │ ├── Morphism.agda │ │ ├── Semiring.agda │ │ └── Sugar.agda │ ├── Sum.agda │ ├── Sum │ │ └── Pointwise.agda │ ├── Thinning.agda │ ├── Two.agda │ ├── TypeAlgebra.agda │ ├── VZip.agda │ ├── Vec.agda │ ├── Vec │ │ └── Thinning.agda │ └── Zero.agda ├── Quantitative │ ├── Instantiation │ │ ├── Counting.agda │ │ ├── Monotonicity.agda │ │ └── Monotonicity │ │ │ └── Worlds.agda │ ├── Models │ │ ├── RelationTransformer.agda │ │ └── RelationTransformer │ │ │ └── Action.agda │ ├── Resources.agda │ ├── Resources │ │ ├── Checker.agda │ │ ├── Context.agda │ │ ├── Context │ │ │ ├── Semilattice.agda │ │ │ └── ToppedSemilattice.agda │ │ ├── Reduction.agda │ │ └── Substitution.agda │ ├── Semantics │ │ ├── Setoid.agda │ │ ├── Sets.agda │ │ ├── Sets │ │ │ └── Term.agda │ │ ├── WRel.agda │ │ └── WRel │ │ │ ├── Bang.agda │ │ │ ├── Core.agda │ │ │ └── Term.agda │ ├── Syntax.agda │ ├── Syntax │ │ ├── Direction.agda │ │ ├── Reduction.agda │ │ └── Substitution.agda │ ├── Types.agda │ └── Types │ │ ├── Checker.agda │ │ ├── Formers.agda │ │ ├── Formers │ │ └── Dec.agda │ │ ├── Reduction.agda │ │ └── Substitution.agda └── index.agda └── tex ├── abstract ├── ACM-Reference-Format.bst ├── acmart.cls ├── bib.bib ├── examples.tex ├── intro.tex ├── quantitative.bib ├── quantitative.tex ├── semantics.tex └── syntax.tex ├── lin-tlla-abstract ├── agda.sty ├── breakurl.sty ├── default.nix ├── eptcs.bst ├── eptcs.cls ├── eptcsalpha.bst ├── generic.bib ├── pkg.nix ├── quant-subst.tex └── shell.nix ├── lin-tlla-paper ├── CHANGES ├── agda.sty ├── breakurl.sty ├── default.nix ├── eptcs.bst ├── eptcs.cls ├── eptcsalpha.bst ├── generic.bib ├── pkg.nix ├── quant-subst.tex └── shell.nix ├── lin-tlla-presentation ├── agda.sty ├── default.nix ├── pkg.nix └── quant-subst.tex ├── macros.tex ├── paper-lncs ├── future-work.tex ├── instantiations.tex ├── introduction.tex ├── llncs.cls ├── motivation.tex ├── quantitative.bib ├── quantitative.tex ├── semantics.tex ├── splncs04.bst └── syntax.tex ├── paper ├── ACM-Reference-Format.bst ├── acmart.cls ├── bib.bib ├── future-work.tex ├── instantiations.tex ├── introduction.tex ├── motivation.tex ├── quantitative.bib ├── quantitative.tex ├── semantics.tex └── syntax.tex ├── presentation ├── default.nix ├── icfp-presentation-final.pdf ├── pkg.nix ├── quantitative.tex └── spls-presentation-final.pdf ├── quantitative.bib ├── subst-note ├── acmart.cls └── subst-note.tex ├── types-abstract ├── quantitative-2page.tex └── quantitative.tex └── types-presentation ├── bread.jpeg ├── default.nix ├── pkg.nix ├── quantitative.tex ├── spls-presentation-final.pdf └── types-presentation-final.pdf /.gitignore: -------------------------------------------------------------------------------- 1 | *.agdai 2 | .#* 3 | comment.cut 4 | *.aux 5 | *.log 6 | *.out 7 | /paper/quantitative.pdf 8 | *~ 9 | paper/auto 10 | presentation/auto 11 | *.bbl 12 | *.blg 13 | *.fdb_latexmk 14 | *.fls 15 | *.dvi 16 | *.nav 17 | *.pdf 18 | *.snm 19 | *.toc 20 | *.prv 21 | *.pag 22 | .auctex-auto 23 | *.envrc 24 | -------------------------------------------------------------------------------- /quantitative.agda-lib: -------------------------------------------------------------------------------- 1 | name: quantitative 2 | depend: standard-library 3 | include: src 4 | -------------------------------------------------------------------------------- /src/Lib/Category/Examples.agda: -------------------------------------------------------------------------------- 1 | module Lib.Category.Examples where 2 | 3 | open import Lib.Category 4 | open import Lib.Function as Fun using () 5 | open import Lib.Level 6 | open import Lib.One 7 | open import Lib.Product 8 | open import Lib.Relation 9 | open import Lib.Setoid 10 | open import Lib.Structure 11 | open import Lib.Structure.Morphism 12 | open import Lib.Structure.Sugar 13 | 14 | REL : ∀ {a e} (A : Setoid a e) l → Category (a ⊔ lsuc l) _ _ 15 | REL A l = record 16 | { Obj = Rel A l 17 | ; Arr = λ R S → ⊤-Setoid ([ A ] R ⇒ S) 18 | ; isCategory = record 19 | { id = λ R x y r → r 20 | ; _>>_ = λ r s x y → s x y Fun.o r x y 21 | ; id->> = λ _ → <> 22 | ; >>-id = λ _ → <> 23 | ; >>->> = λ _ _ _ → <> 24 | ; _>>-cong_ = λ _ _ → <> 25 | } 26 | } 27 | 28 | diag : ∀ {a e A l} → Functor (REL (A ×S A) l) (REL {a} {e} A l) 29 | diag = record 30 | { obj = λ R x y → R (x , x) (y , y) 31 | ; arr = →E-⊤ _ λ rr x y r → rr (x , x) (y , y) r 32 | ; isFunctor = record { arr-id = λ _ → <> ; arr->> = <> } 33 | } 34 | 35 | module _ {a b l m} (A : Setoid a l) (B : Setoid b m) where 36 | private 37 | module A = Setoid A ; module B = Setoid B 38 | 39 | RelF : ∀ {n} → (A.C → B.C) → Functor (REL B n) (REL A n) 40 | RelF f = record 41 | { obj = λ R a b → R (f a) (f b) 42 | ; arr = λ {R} {S} → →E-⊤ _ λ rs a b r → rs (f a) (f b) r 43 | ; isFunctor = record { arr-id = λ _ → <> ; arr->> = <> } 44 | } 45 | 46 | RelF′ : ∀ {a b l} {A : Set a} {B : Set b} → 47 | (A → B) → Functor (REL (≡-Setoid B) l) (REL (≡-Setoid A) l) 48 | RelF′ {A = A} {B} f = RelF (≡-Setoid A) (≡-Setoid B) f 49 | 50 | CMON : ∀ c l → Category (lsuc (c ⊔ l)) (c ⊔ l) (c ⊔ l) 51 | CMON c l = record 52 | { Obj = ΣCommutativeMonoid c l 53 | ; Arr = CommutativeMonoidArr 54 | ; isCategory = record 55 | { id = λ M → let open ΣCommutativeMonoid M in 56 | idE Carrier , (record { e = refl ; _·_ = λ _ _ → refl }) 57 | ; _>>_ = λ {M} {N} {O} h i → 58 | let open ΣCommutativeMonoid O in 59 | let module h = CommutativeMonoidArr M N h in 60 | let module i = CommutativeMonoidArr N O i in 61 | h.el >>E i.el 62 | , (record { e = trans (i.el $E= h.e) i.e 63 | ; _·_ = λ x y → trans (i.el $E= (x h.· y)) 64 | ((h.el $E x) i.· (h.el $E y)) 65 | }) 66 | ; id->> = λ {M} {N} h → let open CommutativeMonoidArr M N h in 67 | (λ xy → el $E= xy) , <> 68 | ; >>-id = λ {M} {N} h → let open CommutativeMonoidArr M N h in 69 | (λ xy → el $E= xy) , <> 70 | ; >>->> = λ {M} {N} {O} {P} h i j → 71 | let module h = CommutativeMonoidArr M N h in 72 | let module i = CommutativeMonoidArr N O i in 73 | let module j = CommutativeMonoidArr O P j in 74 | (λ xy → j.el $E= (i.el $E= (h.el $E= xy))) , <> 75 | ; _>>-cong_ = λ ff gg → (λ xx → fst gg (fst ff xx)) , <> 76 | } 77 | } 78 | 79 | -- CMON-MonoidalCategory : ∀ c l → MonoidalCategory (lsuc (c ⊔ l)) (c ⊔ l) (c ⊔ l) 80 | -- CMON-MonoidalCategory c l = record 81 | -- { C = CMON c l 82 | -- ; I = {!!} 83 | -- ; ⊗ = {!!} 84 | -- ; isMonoidal = record 85 | -- { I⊗ = {!!} 86 | -- ; ⊗I = {!!} 87 | -- ; ⊗⊗ = {!!} 88 | -- ; triangle = {!!} 89 | -- ; pentagon = {!!} 90 | -- } 91 | -- } 92 | 93 | -- Semiringoid : (o c l : Level) → Set (lsuc (o ⊔ c ⊔ l)) 94 | -- Semiringoid o c l = EnrichedCategory o (CMON-MonoidalCategory c l) 95 | -------------------------------------------------------------------------------- /src/Lib/Dec.agda: -------------------------------------------------------------------------------- 1 | module Lib.Dec where 2 | open import Lib.Function 3 | open import Lib.One 4 | open import Lib.Sum 5 | open import Lib.Two 6 | open import Lib.Zero 7 | 8 | Dec : ∀ {x} (X : Set x) → Set x 9 | Dec X = X ⊎ Not X 10 | 11 | pattern yes p = inl p 12 | pattern no np = inr np 13 | 14 | mapDec : ∀ {x y X Y} → (X → Y) → (Y → X) → Dec {x} X → Dec {y} Y 15 | mapDec f g = map⊎ f (λ ¬x y → ¬x (g y)) 16 | 17 | Not? : ∀ {x X} → Dec {x} X → Dec (Not X) 18 | Not? (yes p) = no λ np → np p 19 | Not? (no np) = yes np 20 | 21 | floor : ∀ {x X} → Dec {x} X → Two 22 | floor (yes p) = ttt 23 | floor (no np) = fff 24 | 25 | Auto : ∀ {x X} → Dec {x} X → Set 26 | Auto = T o floor 27 | 28 | fromWitness : ∀ {x X} {X? : Dec {x} X} → X → Auto X? 29 | fromWitness {X? = yes p} x = <> 30 | fromWitness {X? = no np} x = Zero-elim (np x) 31 | 32 | toWitness : ∀ {x X} {X? : Dec {x} X} → Auto X? → X 33 | toWitness {X? = yes x} auto = x 34 | toWitness {X? = no nx} () 35 | 36 | byDec : ∀ {x X} (X? : Dec {x} X) {auto : Auto X?} → X 37 | byDec X? {auto} = toWitness auto 38 | 39 | infixr 4 _>>=[_]_ _<&>[_]_ 40 | _>>=[_]_ : ∀ {a b} {A : Set a} {B : Set b} → 41 | Dec A → (B → A) → (A → Dec B) → Dec B 42 | yes a >>=[ B→A ] A→B? = A→B? a 43 | no na >>=[ B→A ] A→B? = no (na o B→A) 44 | 45 | _<&>[_]_ : ∀ {a b} {A : Set a} {B : Set b} → 46 | Dec A → (B → A) → (A → B) → Dec B 47 | A? <&>[ g ] f = mapDec f g A? 48 | -------------------------------------------------------------------------------- /src/Lib/Dec/Properties.agda: -------------------------------------------------------------------------------- 1 | module Lib.Dec.Properties where 2 | 3 | open import Lib.Dec 4 | open import Lib.Equality 5 | open import Lib.Function 6 | open import Lib.One 7 | open import Lib.Product 8 | open import Lib.Sum 9 | open import Lib.Sum.Pointwise 10 | open import Lib.Two 11 | open import Lib.Zero 12 | 13 | floor-mapDec : ∀ {x y X Y} (f : X → Y) (g : Y → X) X? → 14 | floor (mapDec {x} {y} f g X?) ≡ floor X? 15 | floor-mapDec f g (yes p) = refl 16 | floor-mapDec f g (no np) = refl 17 | 18 | floor-true : ∀ {x X} (X? : Dec {x} X) → X → floor X? ≡ ttt 19 | floor-true (yes _) p = refl 20 | floor-true (no np) p = Zero-elim (np p) 21 | 22 | true→≡yes : ∀ {x X} (X? : Dec {x} X) (p : X) → ∃ λ p′ → X? ≡ yes p′ 23 | true→≡yes (yes p′) p = p′ , refl 24 | true→≡yes (no np) p = Zero-elim (np p) 25 | 26 | false→≡no : ∀ {x X} (X? : Dec {x} X) (np : Not X) → ∃ λ np′ → X? ≡ no np′ 27 | false→≡no (yes p) np = Zero-elim (np p) 28 | false→≡no (no np′) np = np′ , refl 29 | 30 | dec-agree : ∀ {x y X Y} → (X → Y) → (Y → X) → 31 | (X? : Dec {x} X) (Y? : Dec {y} Y) → Agree X? Y? 32 | dec-agree f g (yes x) Y? rewrite true→≡yes Y? (f x) .snd = inl <> 33 | dec-agree f g (no nx) Y? rewrite false→≡no Y? (nx o g) .snd = inr <> 34 | -------------------------------------------------------------------------------- /src/Lib/Equality.agda: -------------------------------------------------------------------------------- 1 | module Lib.Equality where 2 | open import Lib.Dec 3 | open import Lib.Level 4 | open import Lib.Product 5 | open import Lib.Zero 6 | 7 | open import Relation.Binary.PropositionalEquality public using (_≡_; refl) 8 | 9 | -- Basic lemmas 10 | 11 | cong : ∀ {a b} {A : Set a} {B : Set b} {x y : A} (f : A → B) → x ≡ y → f x ≡ f y 12 | cong f refl = refl 13 | 14 | cong2 : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} 15 | {a a′ b b′} (f : A → B → C) → a ≡ a′ → b ≡ b′ → f a b ≡ f a′ b′ 16 | cong2 f refl refl = refl 17 | 18 | sym : ∀ {a} {A : Set a} {x y : A} → x ≡ y → y ≡ x 19 | sym refl = refl 20 | 21 | trans : ∀ {a} {A : Set a} {x y z : A} → x ≡ y → y ≡ z → x ≡ z 22 | trans refl q = q 23 | 24 | subst : ∀ {a p} {A : Set a} (P : A → Set p) {x y : A} → x ≡ y → P x → P y 25 | subst P refl px = px 26 | 27 | subst2 : ∀ {a b p} {A : Set a} {B : Set b} (P : A → B → Set p) 28 | {x x′ : A} {y y′ : B} → x ≡ x′ → y ≡ y′ → P x y → P x′ y′ 29 | subst2 P refl refl pxy = pxy 30 | 31 | coerce : ∀ {l} {A B : Set l} → A ≡ B → A → B 32 | coerce = subst (λ X → X) 33 | 34 | -- Propositionality 35 | 36 | IsProp : ∀ {a} → Set a → Set a 37 | IsProp A = (x y : A) → x ≡ y 38 | 39 | ≡IsProp : ∀ {a A x y} → IsProp (_≡_ {a} {A} x y) 40 | ≡IsProp refl refl = refl 41 | 42 | -- Reasoning syntax 43 | 44 | infixr 1 _=[_]=_ 45 | infixr 2 _QED 46 | 47 | _=[_]=_ : ∀ {a} {A : Set a} (x : A) {y z} → x ≡ y → y ≡ z → x ≡ z 48 | x =[ xy ]= yz = trans xy yz 49 | 50 | _QED : ∀ {a} {A : Set a} (x : A) → x ≡ x 51 | x QED = refl 52 | 53 | infix 0 _/=_ 54 | _/=_ : ∀ {a} {A : Set a} → A → A → Set a 55 | x /= y = Not (x ≡ y) 56 | 57 | record Graph {a b} {A : Set a} {B : A → Set b} (f : (x : A) → B x) (x : A) (y : B x) : Set (a ⊔ b) where 58 | constructor ingraph 59 | field 60 | eq : f x ≡ y 61 | 62 | inspect : ∀ {a b} {A : Set a} {B : A → Set b} (f : (x : A) → B x) (x : A) → Graph f x (f x) 63 | inspect f x = ingraph refl 64 | 65 | ≡× : ∀ {a b A B} {p q : _×_ {a} {b} A B} → p ≡ q → (fst p ≡ fst q) × (snd p ≡ snd q) 66 | ≡× refl = refl , refl 67 | 68 | DecEq : ∀ {x} → Set x → Set x 69 | DecEq X = (x y : X) → Dec (x ≡ y) 70 | 71 | ≡⇒refl : ∀ {a l} {A : Set a} (R : (x y : A) → Set l) → (∀ x → R x x) → 72 | {x y : A} → x ≡ y → R x y 73 | ≡⇒refl R r refl = r _ 74 | -------------------------------------------------------------------------------- /src/Lib/Function.agda: -------------------------------------------------------------------------------- 1 | module Lib.Function where 2 | id : ∀ {l} {A : Set l} → A → A 3 | id x = x 4 | 5 | infixr 5 _o_ _>>_ 6 | _o_ : ∀ {a b c} {A : Set a} {B : A → Set b} {C : ∀ {a} → B a → Set c} 7 | (f : ∀ {a} (b : B a) → C b) (g : ∀ a → B a) a → C (g a) 8 | (f o g) x = f (g x) 9 | 10 | case_return_of_ : 11 | ∀ {a b} {A : Set a} (x : A) (B : A → Set b) → (∀ x → B x) → B x 12 | case x return B of f = f x 13 | 14 | case_of_ : ∀ {a b} {A : Set a} {B : Set b} (x : A) → (A → B) → B 15 | case x of f = f x 16 | 17 | const : ∀ {a b} {A : Set a} {B : Set b} → A → B → A 18 | const x y = x 19 | 20 | infixl 2 __ 21 | __ : ∀ {a b c} {A : Set a} {B : A → Set b} {C : (x : A) → B x → Set c} → 22 | (f : (x : A) (y : B x) → C x y) → (g : (x : A) → B x) → ((x : A) → C x (g x)) 23 | (f g) x = f x (g x) 24 | 25 | flip : ∀ {a b c} {A : Set a} {B : Set b} {C : A → B → Set c} → 26 | (∀ x y → C x y) → (∀ y x → C x y) 27 | flip f y x = f x y 28 | 29 | _>>_ : ∀ {a b c} {A : Set a} {B : A → Set b} {C : ∀ {a} → B a → Set c} 30 | (g : ∀ a → B a) (f : ∀ {a} (b : B a) → C b) a → C (g a) 31 | _>>_ = flip _o_ 32 | 33 | _⟨_⟩_ : ∀ {a b c} {A : Set a} {B : A → Set b} {C : ∀ x → B x → Set c} 34 | (x : A) → ((x : A) (y : B x) → C x y) → (y : B x) → C x y 35 | x ⟨ f ⟩ y = f x y 36 | 37 | infixl 4 _on_ 38 | _on_ : ∀ {a b c} {A : Set a} {B : Set b} {C : (x y : A) → Set c} → 39 | ((x y : A) → C x y) → (f : B → A) → ((x y : B) → C (f x) (f y)) 40 | (f on g) x y = f (g x) (g y) 41 | 42 | infixl 0 _:∋_ 43 | _:∋_ : ∀ {a} (A : Set a) → A → A 44 | A :∋ x = x 45 | -------------------------------------------------------------------------------- /src/Lib/FunctionProperties.agda: -------------------------------------------------------------------------------- 1 | open import Lib.Setoid 2 | 3 | module Lib.FunctionProperties {c l} (S : Setoid c l) where 4 | 5 | open import Lib.Level 6 | open import Lib.Product 7 | 8 | open Setoid S 9 | 10 | Op2 : Set _ 11 | Op2 = C → C → C 12 | 13 | Rel : (l′ : Level) → Set _ 14 | Rel l′ = C → C → Set l′ 15 | 16 | -------------------------------------------------------------------------------- 17 | -- Operation properties 18 | 19 | Cong2 : Op2 → Set _ 20 | Cong2 _·_ = ∀ {x x′ y y′} → x ≈ x′ → y ≈ y′ → (x · y) ≈ (x′ · y′) 21 | 22 | IdentityLeft : Op2 → C → Set _ 23 | IdentityLeft _·_ e = ∀ x → (e · x) ≈ x 24 | 25 | IdentityRight : Op2 → C → Set _ 26 | IdentityRight _·_ e = ∀ x → (x · e) ≈ x 27 | 28 | Identity : Op2 → C → Set _ 29 | Identity _·_ e = IdentityLeft _·_ e × IdentityRight _·_ e 30 | 31 | Assoc : Op2 → Set _ 32 | Assoc _·_ = ∀ x y z → ((x · y) · z) ≈ (x · (y · z)) 33 | 34 | Comm : Op2 → Set _ 35 | Comm _·_ = ∀ x y → (x · y) ≈ (y · x) 36 | 37 | AnnihilLeft : Op2 → C → Set _ 38 | AnnihilLeft _·_ z = ∀ x → (x · z) ≈ z 39 | 40 | AnnihilRight : Op2 → C → Set _ 41 | AnnihilRight _·_ z = ∀ x → (z · x) ≈ z 42 | 43 | Annihil : Op2 → C → Set _ 44 | Annihil _·_ z = AnnihilLeft _·_ z × AnnihilRight _·_ z 45 | 46 | DistribLeft : Op2 → Op2 → Set _ 47 | DistribLeft _*_ _+_ = ∀ x y z → (x * (y + z)) ≈ ((x * y) + (x * z)) 48 | 49 | DistribRight : Op2 → Op2 → Set _ 50 | DistribRight _*_ _+_ = ∀ x y z → ((y + z) * x) ≈ ((y * x) + (z * x)) 51 | 52 | Distrib : Op2 → Op2 → Set _ 53 | Distrib _*_ _+_ = DistribLeft _*_ _+_ × DistribRight _*_ _+_ 54 | 55 | Mono : ∀ {l′} → Rel l′ → Op2 → Set _ 56 | Mono _≤_ _·_ = ∀ {x x′ y y′} → x ≤ x′ → y ≤ y′ → (x · y) ≤ (x′ · y′) 57 | -------------------------------------------------------------------------------- /src/Lib/Level.agda: -------------------------------------------------------------------------------- 1 | module Lib.Level where 2 | open import Agda.Primitive using (Level; lzero; lsuc; _⊔_) public 3 | 4 | record Lift {a} l (A : Set a) : Set (a ⊔ l) where 5 | constructor lift 6 | field 7 | lower : A 8 | open Lift public 9 | 10 | mapLift : ∀ {a b l} {A : Set a} {B : Set b} → (A → B) → Lift l A → Lift l B 11 | mapLift f (lift x) = lift (f x) 12 | -------------------------------------------------------------------------------- /src/Lib/List.agda: -------------------------------------------------------------------------------- 1 | module Lib.List where 2 | 3 | open import Data.List public using (List; []; _∷_) 4 | 5 | -- Bidirectional order of arguments 6 | fold : ∀ {a b} {A : Set a} → List A → (B : Set b) (n : B) (c : A → B → B) → B 7 | fold [] B n c = n 8 | fold (x ∷ xs) B n c = c x (fold xs B n c) 9 | 10 | data ListR {a b r} {A : Set a} {B : Set b} (R : A → B → Set r) 11 | : List A → List B → Set r where 12 | [] : ListR R [] [] 13 | _∷_ : ∀ {x y xs ys} (r : R x y) (rs : ListR R xs ys) → 14 | ListR R (x ∷ xs) (y ∷ ys) 15 | -------------------------------------------------------------------------------- /src/Lib/Matrix.agda: -------------------------------------------------------------------------------- 1 | module Lib.Matrix {c} (C : Set c) where 2 | 3 | open import Lib.Nat 4 | open import Lib.Product 5 | open import Lib.Sum 6 | open import Lib.Thinning 7 | 8 | record Mat (mn : Nat × Nat) : Set c where 9 | constructor mk 10 | field 11 | get : Fin (mn .fst) × Fin (mn .snd) → C 12 | open Mat public 13 | 14 | infixl 5 _+↓_ _+→_ 15 | _+↓_ : ∀ {m m′ n} → Mat (m , n) → Mat (m′ , n) → Mat (m′ +N m , n) 16 | _+↓_ {m} {m′} {n} M N .get (i , j) = 17 | [ (λ i′ → N .get (i′ , j)) , (λ i′ → M .get (i′ , j)) ] (part m′ i) 18 | 19 | _+→_ : ∀ {m n n′} → Mat (m , n) → Mat (m , n′) → Mat (m , n′ +N n) 20 | _+→_ {m} {n} {n′} M N .get (i , j) = 21 | [ (λ j′ → N .get (i , j′)) , (λ j′ → M .get (i , j′)) ] (part n′ j) 22 | 23 | -- These two help size inference, and are nice visually 24 | -- Mnemonic: || for a (narrow) column; -- for a (flat) row 25 | [-_-] : ∀ {n} → C → Mat (1 , n) 26 | [- v -] .get _ = v 27 | 28 | [|_|] : ∀ {m} → C → Mat (m , 1) 29 | [| v |] .get _ = v 30 | -------------------------------------------------------------------------------- /src/Lib/Matrix/Addition.agda: -------------------------------------------------------------------------------- 1 | open import Lib.Structure 2 | open import Lib.Structure.Sugar 3 | 4 | module Lib.Matrix.Addition {c l} (M : ΣCommutativeMonoid c l) where 5 | 6 | open ΣCommutativeMonoid M 7 | 8 | open import Lib.Matrix.Setoid Carrier 9 | 10 | open import Lib.Equality as ≡ 11 | open import Lib.FunctionProperties 12 | open import Lib.Level 13 | open import Lib.Nat 14 | open import Lib.Product using (_×_; fst; snd; _,_) 15 | open import Lib.Setoid 16 | open import Lib.Thinning 17 | 18 | MatCM : Nat × Nat → ΣCommutativeMonoid c l 19 | MatCM mn@(m , n) = record 20 | { Carrier = MatS mn 21 | ; commutativeMonoid = record 22 | { e = mk λ _ → e 23 | ; _·_ = _+_ 24 | ; isCommutativeMonoid = record 25 | { comm = λ where M N .get ij → comm (M .get ij) (N .get ij) 26 | ; isMonoid = record 27 | { identity = λ where 28 | .fst N .get ij → fst identity (N .get ij) 29 | .snd M .get ij → snd identity (M .get ij) 30 | ; assoc = λ where M N O .get ij → assoc _ _ _ 31 | ; _·-cong_ = λ where MM NN .get ij → MM .get ij ·-cong NN .get ij 32 | } 33 | } 34 | } 35 | } 36 | where 37 | _+_ : (M N : Mat mn) → Mat mn 38 | (M + N) .get ij = M .get ij · N .get ij 39 | 40 | private 41 | module Explicit (mn : Nat × Nat) = ΣCommutativeMonoid (MatCM mn) 42 | module Implicit {mn : Nat × Nat} = Explicit mn 43 | 44 | open Explicit public using () 45 | renaming (commutativeMonoid to Mat-commutativeMonoid; monoid to Mat-monoid) 46 | open Implicit public using () 47 | renaming ( e to 0M; _·_ to _+M_; identity to +M-identity; assoc to +M-assoc 48 | ; _·-cong_ to _+M-cong_; comm to +M-comm 49 | ; isCommutativeMonoid to isCommutativeMonoidM 50 | ; isMonoid to isMonoidM 51 | ) 52 | -------------------------------------------------------------------------------- /src/Lib/Matrix/Addition/Order.agda: -------------------------------------------------------------------------------- 1 | open import Lib.Structure 2 | open import Lib.Structure.Sugar 3 | 4 | module Lib.Matrix.Addition.Order {c l l′} 5 | (PM : ΣCommutativePomonoid c l l′) where 6 | 7 | open ΣCommutativePomonoid PM 8 | 9 | open import Lib.Matrix.Setoid Carrier 10 | open import Lib.Matrix.Addition σCommutativeMonoid 11 | open import Lib.Matrix.Poset σPoset 12 | 13 | open import Lib.Level 14 | open import Lib.Nat 15 | open import Lib.Product 16 | open import Lib.Setoid 17 | 18 | MatPCM : Nat × Nat → ΣCommutativePomonoid c l l′ 19 | MatPCM mn@(m , n) = record 20 | { Carrier = MatS mn 21 | ; commutativePomonoid = record 22 | { _≤_ = _≤M_ 23 | ; e = 0M 24 | ; _·_ = _+M_ 25 | ; isCommutativePomonoid = record 26 | { _·-mono_ = λ where MM NN .get ij → MM .get ij ·-mono NN .get ij 27 | ; isPoset = isPosetM 28 | ; isCommutativeMonoid = isCommutativeMonoidM 29 | } 30 | } 31 | } 32 | 33 | private 34 | module Explicit (mn : Nat × Nat) = ΣCommutativePomonoid (MatPCM mn) 35 | module Implicit {mn : Nat × Nat} = Explicit mn 36 | 37 | open Explicit public using () 38 | renaming (commutativePomonoid to Mat-commutativeMonoid 39 | ; pomonoid to Mat-pomonoid) 40 | open Implicit public using () 41 | renaming (_·-mono_ to _+M-mono_) 42 | -------------------------------------------------------------------------------- /src/Lib/Matrix/Multiplication/Basis.agda: -------------------------------------------------------------------------------- 1 | open import Lib.Structure 2 | open import Lib.Structure.Sugar 3 | 4 | module Lib.Matrix.Multiplication.Basis {c l} (R : ΣSemiring c l) where 5 | 6 | open ΣSemiring R 7 | open import Lib.Structure.Semiring R 8 | 9 | open import Lib.Matrix.Setoid Carrier 10 | open import Lib.Matrix.Multiplication R 11 | 12 | open import Lib.Dec 13 | open import Lib.Dec.Properties 14 | open import Lib.Equality as ≡ using (_≡_; ≡⇒refl) 15 | open import Lib.Function 16 | open import Lib.Nat 17 | open import Lib.Product 18 | open import Lib.Setoid 19 | open import Lib.Sum 20 | open import Lib.Thinning 21 | 22 | choose-col : ∀ {m n} (j : Fin n) (M : Mat (m , n)) → 23 | M *M basis-col j ≈M thin oi j $E M 24 | choose-col {m} {succ n} (os j) M .get (i , k) = 25 | (M *M basis-col (os j)) .get (i , k) ≈[ refl ]≈ 26 | (sum λ j′ → M .get (i , j′) * basis-col (os j) .get (j′ , k)) 27 | ≈[ refl ]≈ 28 | M .get (i , zeroth) * basis-col (os j) .get (zeroth , k) 29 | + (sum λ j′ → M .get (i , o′ j′) * basis-col (os j) .get (o′ j′ , k)) 30 | ≈[ 31 | (M .get (i , zeroth) * basis-col (os j) .get (zeroth , k) 32 | ≈[ refl *-cong lemma2 ]≈ 33 | M .get (i , zeroth) * e1 34 | ≈[ *-identity .snd _ ]≈ 35 | M .get (i , zeroth) 36 | ≈[ ≡⇒≈ (≡.sym (≡.cong (M .get) (≡.cong2 _,_ (comp-oi i) lemma))) ]≈ 37 | M .get (i ≤-comp oi , k ≤-comp os j) QED) 38 | +-cong 39 | ((sum λ j′ → M .get (i , o′ j′) * basis-col (os j) .get (o′ j′ , k)) 40 | ≈[ (sum-cong {n} λ j′ → refl *-cong lemma3 j′) ]≈ 41 | (sum λ j′ → M .get (i , o′ j′) * e0) 42 | ≈[ (sum-cong {n} λ j′ → annihil .fst _) ]≈ 43 | (sum {n} λ j′ → e0) ≈[ sum-e0 n ]≈ 44 | e0 QED) 45 | ]≈ 46 | M .get (i ≤-comp oi , k ≤-comp os j) + e0 ≈[ +-identity .snd _ ]≈ 47 | M .get (i ≤-comp oi , k ≤-comp os j) ≈[ refl ]≈ 48 | (thin oi (os j) $E M) .get (i , k) QED 49 | where 50 | open SetoidReasoning Carrier 51 | 52 | lemma : k ≤-comp os j ≡ zeroth 53 | lemma rewrite oi-unique k zeroth | oe-unique (oz ≤-comp j) oe = ≡.refl 54 | 55 | lemma2 : basis-col (os j) .get (zeroth , k) ≈ e1 56 | lemma2 rewrite true→≡yes (oe ⊆? j) (empty-⊆ oe j) .snd 57 | | true→≡yes (k ⊆? zeroth) 58 | (≡⇒refl _⊆_ ⊆-refl (oi-unique k zeroth)) 59 | .snd 60 | = refl 61 | 62 | lemma3 : ∀ j′ → basis-col (os j) .get (o′ j′ , k) ≈ e0 63 | lemma3 j′ rewrite false→≡no (j′ ⊆? j) ((λ ()) o ⊆⇒≤) .snd = refl 64 | 65 | choose-col {m} {succ n} (o′ j) M .get (i , k) = 66 | (M *M basis-col (o′ j)) .get (i , k) ≈[ refl ]≈ 67 | (sum λ j′ → M .get (i , j′) * basis-col (o′ j) .get (j′ , k)) ≈[ refl ]≈ 68 | M .get (i , zeroth) * basis-col (o′ j) .get (zeroth , k) 69 | + (sum λ j′ → M .get (i , o′ j′) * basis-col (o′ j) .get (o′ j′ , k)) 70 | ≈[ 71 | (M .get (i , zeroth) * basis-col (o′ j) .get (zeroth , k) ≈[ refl ]≈ 72 | M .get (i , zeroth) * e0 ≈[ annihil .fst _ ]≈ 73 | e0 QED) 74 | +-cong 75 | ((sum λ j′ → M .get (i , o′ j′) * basis-col (o′ j) .get (o′ j′ , k)) 76 | ≈[ (sum-cong {n} λ j′ → refl *-cong lemma j′) ]≈ 77 | (sum λ j′ → M .get (i , o′ j′) * basis-col j .get (j′ , k)) 78 | ≈[ choose-col {m} {n} j (remove-col $E M) .get (i , k) ]≈ 79 | M .get (i ≤-comp oi , o′ (k ≤-comp j)) QED) 80 | ]≈ 81 | e0 + M .get (i ≤-comp oi , o′ (k ≤-comp j)) ≈[ +-identity .fst _ ]≈ 82 | M .get (i ≤-comp oi , o′ (k ≤-comp j)) ≈[ refl ]≈ 83 | (thin oi (o′ j) $E M) .get (i , k) QED 84 | where 85 | open SetoidReasoning Carrier 86 | 87 | lemma : ∀ j′ → basis-col (o′ j) .get (o′ j′ , k) ≈ basis-col j .get (j′ , k) 88 | lemma j′ with j′ ⊆? j 89 | lemma j′ | yes sub rewrite true→≡yes (k ⊆? zeroth) 90 | (≡⇒refl _⊆_ ⊆-refl (oi-unique k zeroth)) 91 | .snd 92 | = refl 93 | lemma j′ | no nsub = refl 94 | -------------------------------------------------------------------------------- /src/Lib/Matrix/Multiplication/Block.agda: -------------------------------------------------------------------------------- 1 | open import Lib.Structure 2 | open import Lib.Structure.Sugar 3 | 4 | module Lib.Matrix.Multiplication.Block {c l} (R : ΣSemiring c l) where 5 | 6 | open ΣSemiring R 7 | open import Lib.Structure.Semiring R 8 | 9 | open import Lib.Matrix.Setoid Carrier 10 | open import Lib.Matrix.Addition +-σCommutativeMonoid 11 | open import Lib.Matrix.Multiplication R 12 | 13 | open import Lib.Equality as ≡ using (_≡_; ≡⇒refl) 14 | open import Lib.Function 15 | open import Lib.Nat 16 | open import Lib.Product 17 | open import Lib.Setoid 18 | open import Lib.Sum 19 | open import Lib.Thinning 20 | 21 | insert-blocks : ∀ {m n n′ o} (M : Mat (m , n)) (M′ : Mat (m , n′)) 22 | (N : Mat (n , o)) (N′ : Mat (n′ , o)) → 23 | (M +→ M′) *M (N +↓ N′) ≈M M *M N +M M′ *M N′ 24 | insert-blocks {n′ = zero} M M′ N N′ = symM (+M-identity .snd (M *M N)) 25 | insert-blocks {m} {n} {n′ = succ n′} {o} M M′ N N′ .get (i , k) = 26 | ((M +→ M′) *M (N +↓ N′)) .get (i , k) ≈[ refl ]≈ 27 | (M +→ M′) .get (i , zeroth) * (N +↓ N′) .get (zeroth , k) 28 | + (remove-col $E (M +→ M′) *M remove-row $E (N +↓ N′)) .get (i , k) 29 | ≈[ (lemma3 *-cong lemma4) +-cong (lemma0 *M-cong lemma1) .get (i , k) ]≈ 30 | M′ .get (i , zeroth) * N′ .get (zeroth , k) 31 | + ((M +→ remove-col $E M′) *M (N +↓ remove-row $E N′)) .get (i , k) 32 | ≈[ refl +-cong insert-blocks {n′ = n′} M (remove-col $E M′) 33 | N (remove-row $E N′) 34 | .get (i , k) ]≈ 35 | M′ .get (i , zeroth) * N′ .get (zeroth , k) + (M *M N) .get (i , k) 36 | + (remove-col $E M′ *M remove-row $E N′) .get (i , k) ≈[ lemma2 ]≈ 37 | (M *M N) .get (i , k) + M′ .get (i , zeroth) * N′ .get (zeroth , k) 38 | + (remove-col $E M′ *M remove-row $E N′) .get (i , k) ≈[ refl ]≈ 39 | (M *M N +M M′ *M N′) .get (i , k) QED 40 | where 41 | open SetoidReasoning Carrier 42 | 43 | lemma0 : remove-col $E (M +→ M′) ≈M M +→ remove-col $E M′ 44 | lemma0 .get (i , j) with ≤-+ n′ j 45 | ... | 0 , .1 , _ , j′ , ≡.refl = refl 46 | ... | 1 , .0 , j′ , _ , ≡.refl = refl 47 | ... | succ (succ _) , _ , _ , _ , () 48 | 49 | lemma1 : remove-row $E (N +↓ N′) ≈M N +↓ remove-row $E N′ 50 | lemma1 .get (j , k) with ≤-+ n′ j 51 | ... | 0 , .1 , _ , j′ , ≡.refl = refl 52 | ... | 1 , .0 , j′ , _ , ≡.refl = refl 53 | ... | succ (succ _) , _ , _ , _ , () 54 | 55 | lemma2 : ∀ {x y z} → x + y + z ≈ y + x + z 56 | lemma2 {x} {y} {z} = 57 | x + y + z ≈[ sym (+-assoc x y z) ]≈ 58 | (x + y) + z ≈[ +-comm x y +-cong refl ]≈ 59 | (y + x) + z ≈[ +-assoc y x z ]≈ 60 | y + x + z QED 61 | 62 | lemma3 : (M +→ M′) .get (i , zeroth) ≈ M′ .get (i , zeroth) 63 | lemma3 with ≤-+ n′ (oe {n′ +N n}) 64 | lemma3 | 0 , .0 , e , _ , ≡.refl rewrite oe-unique e oe = refl 65 | lemma3 | succ _ , _ , _ , _ , () 66 | 67 | lemma4 : (N +↓ N′) .get (zeroth , k) ≈ N′ .get (zeroth , k) 68 | lemma4 with ≤-+ n′ (oe {n′ +N n}) 69 | lemma4 | 0 , .0 , e , _ , ≡.refl rewrite oe-unique e oe = refl 70 | lemma4 | succ _ , _ , _ , _ , () 71 | -------------------------------------------------------------------------------- /src/Lib/Matrix/Multiplication/Order.agda: -------------------------------------------------------------------------------- 1 | open import Lib.Structure 2 | open import Lib.Structure.Sugar 3 | 4 | module Lib.Matrix.Multiplication.Order {c l l′} 5 | (PR : ΣPosemiring c l l′) where 6 | 7 | open ΣPosemiring PR 8 | 9 | open import Lib.Matrix.Setoid Carrier 10 | open import Lib.Matrix.Addition +-σCommutativeMonoid 11 | open import Lib.Matrix.Multiplication σSemiring 12 | open import Lib.Matrix.Poset σPoset 13 | 14 | open import Lib.Function 15 | open import Lib.Level 16 | open import Lib.Nat 17 | open import Lib.Product 18 | open import Lib.Setoid 19 | open import Lib.Thinning using (Fin; oz; os; o′; oe) renaming (_≤_ to _≤th_) 20 | 21 | sum-mono : ∀ {n u v} → (∀ i → u i ≤ v i) → sum u ≤ sum {n} v 22 | sum-mono {zero} uv = ≤-refl 23 | sum-mono {succ n} uv = uv _ +-mono sum-mono {n} (uv o o′) 24 | 25 | _*M-mono_ : ∀ {m n o} {M M′ : Mat (m , n)} {N N′ : Mat (n , o)} → 26 | M ≤M M′ → N ≤M N′ → M *M N ≤M M′ *M N′ 27 | (MM *M-mono NN) .get (i , k) = 28 | sum-mono (λ j → MM .get (i , j) *-mono NN .get (j , k)) 29 | -------------------------------------------------------------------------------- /src/Lib/Matrix/Poset.agda: -------------------------------------------------------------------------------- 1 | open import Lib.Structure 2 | open import Lib.Structure.Sugar 3 | 4 | module Lib.Matrix.Poset {c l l′} (Po : ΣPoset c l l′) where 5 | 6 | open ΣPoset Po 7 | 8 | open import Lib.Matrix.Setoid Carrier 9 | open import Lib.Matrix.Preorder σPreorder public 10 | 11 | open import Lib.Equality as ≡ using (_≡_) 12 | open import Lib.Level 13 | open import Lib.Nat 14 | open import Lib.Product 15 | open import Lib.Setoid 16 | 17 | MatPo : Nat × Nat → ΣPoset c l l′ 18 | MatPo mn@(m , n) = record 19 | { Carrier = MatS mn 20 | ; poset = record 21 | { _≤_ = _≤M_ 22 | ; isPoset = record 23 | { antisym = λ where MN NM .get ij → antisym (MN .get ij) (NM .get ij) 24 | ; isPreorder = isPreorderM 25 | } 26 | } 27 | } 28 | 29 | private 30 | module Explicit (mn : Nat × Nat) = ΣPoset (MatPo mn) 31 | module Implicit {mn : Nat × Nat} = Explicit mn 32 | 33 | open Explicit public using () 34 | renaming (poset to Mat-poset) 35 | open Implicit public using () 36 | renaming (antisym to antisymM; isPoset to isPosetM) 37 | -------------------------------------------------------------------------------- /src/Lib/Matrix/Preorder.agda: -------------------------------------------------------------------------------- 1 | open import Lib.Structure 2 | open import Lib.Structure.Sugar 3 | 4 | module Lib.Matrix.Preorder {c l l′} (PreO : ΣPreorder c l l′) where 5 | 6 | open ΣPreorder PreO 7 | 8 | open import Lib.Matrix.Setoid Carrier 9 | 10 | open import Lib.Equality as ≡ using (_≡_) 11 | open import Lib.Level 12 | open import Lib.Nat 13 | open import Lib.Product 14 | open import Lib.Setoid 15 | open import Lib.Sum 16 | open import Lib.Thinning renaming (_≤_ to _≤th_) using (Fin; part) 17 | 18 | infix 4 _≤M_ 19 | record _≤M_ {mn} (M N : Mat mn) : Set l′ where 20 | constructor mk 21 | field 22 | get : ∀ ij → M .get ij ≤ N .get ij 23 | open _≤M_ public 24 | 25 | MatPreO : Nat × Nat → ΣPreorder c l l′ 26 | MatPreO mn@(m , n) = record 27 | { Carrier = MatS mn 28 | ; preorder = record 29 | { _≤_ = _≤M_ 30 | ; isPreorder = record 31 | { ≤-reflexive = λ where MM .get ij → ≤-reflexive (MM .get ij) 32 | ; ≤-trans = λ where MN NO .get ij → ≤-trans (MN .get ij) (NO .get ij) 33 | } 34 | } 35 | } 36 | 37 | private 38 | module Explicit (mn : Nat × Nat) = ΣPreorder (MatPreO mn) 39 | module Implicit {mn : Nat × Nat} = Explicit mn 40 | 41 | open Explicit public using () 42 | renaming (preorder to Mat-preorder) 43 | open Implicit public using () 44 | renaming ( ≤-refl to ≤M-refl; ≤-reflexive to ≤M-reflexive 45 | ; ≤-trans to ≤M-trans; isPreorder to isPreorderM 46 | ) 47 | 48 | infixl 5 _+↓-mono_ 49 | _+↓-mono_ : ∀ {m m′ n} {M M′ : Mat (m , n)} {N N′ : Mat (m′ , n)} → 50 | M ≤M M′ → N ≤M N′ → M +↓ N ≤M M′ +↓ N′ 51 | _+↓-mono_ {m} {m′} {n} MM NN .get (i , j) with part m′ i 52 | ... | inl i′ = NN .get (i′ , j) 53 | ... | inr i′ = MM .get (i′ , j) 54 | 55 | thin-≤M : ∀ {m n m′ n′} (mm : m′ ≤th m) (nn : n′ ≤th n) {M N : Mat (m , n)} → 56 | M ≤M N → thin mm nn $E M ≤M thin mm nn $E N 57 | thin-≤M mm nn MN .get ij = MN .get _ 58 | -------------------------------------------------------------------------------- /src/Lib/Matrix/Scaling.agda: -------------------------------------------------------------------------------- 1 | open import Lib.Structure 2 | open import Lib.Structure.Sugar 3 | 4 | module Lib.Matrix.Scaling {c l} (R : ΣSemiring c l) where 5 | 6 | open ΣSemiring R 7 | open import Lib.Structure.Semiring R 8 | 9 | open import Lib.Matrix.Setoid Carrier 10 | 11 | open import Lib.Equality as ≡ using (_≡_) 12 | open import Lib.Level 13 | open import Lib.Matrix.Addition +-σCommutativeMonoid 14 | open import Lib.Matrix.Multiplication R 15 | open import Lib.Nat 16 | open import Lib.Product using (Σ; _×_; fst; snd; _,_; uncurry) 17 | open import Lib.Setoid 18 | open import Lib.Thinning 19 | 20 | private 21 | module Size (mn : Nat × Nat) where 22 | open import Lib.Module Carrier (MatS mn) 23 | 24 | scaleMl : Carrier →E MatS mn →S MatS mn 25 | scaleMl ._$E_ x ._$E_ M .get ij = x * M .get ij 26 | scaleMl ._$E_ x ._$E=_ MM .get ij = refl *-cong MM .get ij 27 | scaleMl ._$E=_ xx MM .get ij = xx *-cong MM .get ij 28 | 29 | infixr 7 _*l_ 30 | 31 | _*l_ : C → Mat mn → Mat mn 32 | x *l M = scaleMl $E x $E M 33 | 34 | Mat-semimodule : Semimodule 35 | Mat-semimodule = record 36 | { 0s = e0 37 | ; 1s = e1 38 | ; _+s_ = _+_ 39 | ; _*s_ = _*_ 40 | ; 0f = 0M 41 | ; _+f_ = _+M_ 42 | ; _*f_ = _*l_ 43 | ; isSemimodule = record 44 | { +*s-isSemiring = isSemiring 45 | ; +f-isCommutativeMonoid = isCommutativeMonoidM 46 | ; _*f-cong_ = λ where 47 | xx MM .get ii → xx *-cong MM .get ii 48 | ; annihil = λ where 49 | .fst x .get ij → annihil .fst x 50 | .snd M .get ij → annihil .snd (M .get ij) 51 | ; distrib = λ where 52 | .fst x M N .get ij → distrib .fst x (M .get ij) (N .get ij) 53 | .snd x y M .get ij → distrib .snd (M .get ij) x y 54 | ; assoc = λ where 55 | x y M .get ij → *-assoc x y (M .get ij) 56 | ; identity = λ where 57 | M .get ij → *-identity .fst (M .get ij) 58 | } 59 | } 60 | open Semimodule Mat-semimodule public 61 | 62 | module Size-implicit {mn : Nat × Nat} = Size mn 63 | 64 | open Size public using (Mat-semimodule) 65 | open Size-implicit public using (scaleMl; _*l_) 66 | renaming ( isSemimodule to isSemimoduleM; _*f-cong_ to _*l-cong_ 67 | ; annihil to *l-annihil; distrib to *l-distrib; assoc to *l-assoc 68 | ; identity to *l-identity) 69 | 70 | open SetoidReasoning Carrier 71 | 72 | *l-linear : ∀ {m n o} x (M : Mat (m , n)) (N : Mat (n , o)) → 73 | (x *l M) *M N ≈M x *l (M *M N) 74 | *l-linear {n = n} x M N .get (i , k) = 75 | ((x *l M) *M N) .get (i , k) ≈[ refl ]≈ 76 | (sum λ j → (x * M .get (i , j)) * N .get (j , k)) 77 | ≈[ (sum-cong {n} λ j → *-assoc _ _ _) ]≈ 78 | (sum λ j → x * (M .get (i , j) * N .get (j , k))) 79 | ≈[ sym (*-sum x λ j → M .get (i , j) * N .get (j , k)) ]≈ 80 | x * (sum λ j → M .get (i , j) * N .get (j , k)) ≈[ refl ]≈ 81 | (x *l (M *M N)) .get (i , k) QED 82 | -------------------------------------------------------------------------------- /src/Lib/Matrix/Scaling/Right.agda: -------------------------------------------------------------------------------- 1 | open import Lib.Structure 2 | open import Lib.Structure.Sugar 3 | open import Lib.Structure.Flip 4 | 5 | module Lib.Matrix.Scaling.Right {c l} (R : ΣSemiring c l) where 6 | 7 | open ΣSemiring R 8 | 9 | open import Lib.Matrix.Setoid Carrier 10 | open import Lib.Matrix.Multiplication R 11 | open import Lib.Matrix.Scaling (flip-σSemiring R) as Sc 12 | using (_*l_; _*l-cong_) 13 | 14 | open import Lib.Function 15 | open import Lib.Product 16 | open import Lib.Setoid 17 | 18 | open Sc public using () 19 | renaming ( scaleMl to scaleMr; Mat-semimodule to Mat-rSemimodule 20 | ; isSemimoduleM to isRSemimoduleM; *l-annihil to *r-annihil 21 | ; *l-distrib to *r-distrib; *l-assoc to *r-assoc 22 | ; *l-identity to *r-identity 23 | ) 24 | 25 | infixl 7 _*r_ _*r-cong_ 26 | 27 | _*r_ : ∀ {mn} → Mat mn → C → Mat mn 28 | _*r_ = flip _*l_ 29 | 30 | _*r-cong_ : ∀ {mn} {M M′ : Mat mn} {x x′ : C} → 31 | M ≈M M′ → x ≈ x′ → M *r x ≈M M′ *r x′ 32 | _*r-cong_ = flip _*l-cong_ 33 | 34 | open SetoidReasoning Carrier 35 | 36 | *r-linear : ∀ {m n o} (M : Mat (m , n)) (N : Mat (n , o)) x → 37 | M *M (N *r x) ≈M (M *M N) *r x 38 | *r-linear {n = n} M N x .get (i , k) = 39 | (M *M (N *r x)) .get (i , k) ≈[ refl ]≈ 40 | (sum λ j → M .get (i , j) * N .get (j , k) * x) 41 | ≈[ (sum-cong {n} λ j → sym (*-assoc _ _ _)) ]≈ 42 | (sum λ j → (M .get (i , j) * N .get (j , k)) * x) 43 | ≈[ sym (sum-* (λ j → M .get (i , j) * N .get (j , k)) x) ]≈ 44 | (sum λ j → M .get (i , j) * N .get (j , k)) * x ≈[ refl ]≈ 45 | ((M *M N) *r x) .get (i , k) QED 46 | -------------------------------------------------------------------------------- /src/Lib/Matrix/Setoid.agda: -------------------------------------------------------------------------------- 1 | open import Lib.Setoid 2 | 3 | module Lib.Matrix.Setoid {c l} (A : Setoid c l) where 4 | 5 | open Setoid A 6 | 7 | open import Lib.Matrix C public 8 | 9 | open import Lib.Dec 10 | open import Lib.Equality as ≡ using (_≡_) 11 | open import Lib.Function using (id) 12 | open import Lib.Level 13 | open import Lib.Nat 14 | open import Lib.Product 15 | open import Lib.Sum 16 | open import Lib.Thinning 17 | open import Lib.Two 18 | 19 | infix 4 _≈M_ 20 | record _≈M_ {mn} (M N : Mat mn) : Set l where 21 | constructor mk 22 | field 23 | get : ∀ ij → M .get ij ≈ N .get ij 24 | open _≈M_ public 25 | 26 | -- TODO: this could be abstracted out as a generic pointwise setoid 27 | -- (assuming propositional equality on the indices) at no intensional cost 28 | MatS : Nat × Nat → Setoid c l 29 | MatS mn@(m , n) = record 30 | { C = Mat mn 31 | ; setoidOver = record 32 | { _≈_ = _≈M_ 33 | ; isSetoid = record 34 | { refl = λ where .get ij → refl 35 | ; sym = λ where MN .get ij → sym (MN .get ij) 36 | ; trans = λ where MN NO .get ij → trans (MN .get ij) (NO .get ij) 37 | } 38 | } 39 | } 40 | 41 | private 42 | module DummyIm {mn : Nat × Nat} = Setoid (MatS mn) 43 | using () 44 | renaming (refl to reflM; sym to symM; trans to transM) 45 | 46 | open DummyIm public 47 | 48 | transpose : ∀ {m n} → MatS (m , n) →E MatS (n , m) 49 | transpose {m} {n} ._$E_ M .get (i , j) = M .get (j , i) 50 | transpose {m} {n} ._$E=_ MM .get (i , j) = MM .get (j , i) 51 | 52 | thin : ∀ {m n m′ n′} → m′ ≤ m → n′ ≤ n → MatS (m , n) →E MatS (m′ , n′) 53 | thin mm nn ._$E_ M .get (i , j) = M .get (i ≤-comp mm , j ≤-comp nn) 54 | thin mm nn ._$E=_ MM .get _ = MM .get _ 55 | 56 | -- Leave special cases of thin for intensional reasons 57 | near-col : ∀ {m n} → MatS (m , succ n) →E MatS (m , 1) 58 | (near-col $E M) .get (i , _) = M .get (i , zeroth) 59 | (near-col $E= MM) .get (i , _) = MM .get (i , zeroth) 60 | 61 | near-row : ∀ {m n} → MatS (succ m , n) →E MatS (1 , n) 62 | (near-row $E M) .get (_ , j) = M .get (zeroth , j) 63 | (near-row $E= MM) .get (_ , j) = MM .get (zeroth , j) 64 | 65 | remove-col : ∀ {m n} → MatS (m , succ n) →E MatS (m , n) 66 | remove-col ._$E_ M .get (i , j) = M .get (i , o′ j) 67 | remove-col ._$E=_ MM .get _ = MM .get _ 68 | 69 | remove-row : ∀ {m n} → MatS (succ m , n) →E MatS (m , n) 70 | remove-row ._$E_ M .get (i , j) = M .get (o′ i , j) 71 | remove-row ._$E=_ MM .get _ = MM .get _ 72 | 73 | set : ∀ {m n m′ n′} → m ≤ m′ → n ≤ n′ → 74 | (MatS (m , n) →E MatS (m , n)) → (MatS (m′ , n′) →E MatS (m′ , n′)) 75 | set mm nn f ._$E_ M .get (i , j) = 76 | [ (λ stuff → let imm , jnn = stuff in 77 | let M′ = thin mm nn $E M in 78 | (f $E M′) .get (⊆⇒≤ imm , ⊆⇒≤ jnn)) 79 | , (λ _ → M .get (i , j)) 80 | ] (i ⊆? mm ×? j ⊆? nn) 81 | set mm nn f ._$E=_ MM .get (i , j) with i ⊆? mm ×? j ⊆? nn 82 | ... | yes _ = (f $E= (thin _ _ $E= MM)) .get _ 83 | ... | no _ = MM .get _ 84 | 85 | set′ : ∀ {m n m′ n′} → m ≤ m′ → n ≤ n′ → 86 | Mat (m , n) → (MatS (m′ , n′) →E MatS (m′ , n′)) 87 | set′ mm nn N = set mm nn (constE $E N) 88 | 89 | infixl 5 _+↓-cong_ 90 | _+↓-cong_ : ∀ {m m′ n} {M M′ : Mat (m , n)} {N N′ : Mat (m′ , n)} → 91 | M ≈M M′ → N ≈M N′ → M +↓ N ≈M M′ +↓ N′ 92 | _+↓-cong_ {m} {m′} {n} MM NN .get (i , j) with part m′ i 93 | ... | inl i′ = NN .get (i′ , j) 94 | ... | inr i′ = MM .get (i′ , j) 95 | 96 | view-near-row : ∀ {m n} (Mv : Mat (succ m , n)) → 97 | ∃ \ M → ∃ \ (v : Mat (1 , n)) → Mv ≈M M +↓ v 98 | view-near-row Mv = remove-row $E Mv , near-row $E Mv , λ where 99 | .get (os e , j) → ≡⇒≈ (≡.cong (λ z → Mv .get (os z , j)) (oe-unique e oe)) 100 | .get (o′ i , j) → refl 101 | 102 | view-near-col : ∀ {m n} (Mv : Mat (m , succ n)) → 103 | ∃ \ M → ∃ \ (v : Mat (m , 1)) → Mv ≈M M +→ v 104 | view-near-col Mv = remove-col $E Mv , near-col $E Mv , λ where 105 | .get (i , os e) → ≡⇒≈ (≡.cong (λ z → Mv .get (i , os z)) (oe-unique e oe)) 106 | .get (i , o′ j) → refl 107 | 108 | view-scalar : (M : Mat (1 , 1)) → ∃ \ x → M ≈M mk λ _ → x 109 | view-scalar M = M .get (zeroth , zeroth) , λ where 110 | .get (os oz , os oz) → refl 111 | -------------------------------------------------------------------------------- /src/Lib/Matrix/VecCompat.agda: -------------------------------------------------------------------------------- 1 | open import Lib.Setoid 2 | 3 | module Lib.Matrix.VecCompat {c l} (A : Setoid c l) where 4 | 5 | open Setoid A 6 | 7 | open import Lib.Function 8 | open import Lib.Matrix.Setoid A 9 | open import Lib.Nat 10 | open import Lib.Product 11 | open import Lib.Thinning 12 | open import Lib.Vec 13 | 14 | from-col-vec : ∀ {m} → Vec C m → Mat (m , 1) 15 | from-col-vec xs .get (i , _) = lookup i xs 16 | 17 | to-col-vec : ∀ {m} → Mat (m , 1) → Vec C m 18 | to-col-vec M = tabulate (λ i → M .get (i , zeroth)) 19 | 20 | from-row-vec : ∀ {n} → Vec C n → Mat (1 , n) 21 | from-row-vec = transpose $E_ o from-col-vec 22 | 23 | to-row-vec : ∀ {n} → Mat (1 , n) → Vec C n 24 | to-row-vec = to-col-vec o transpose $E_ 25 | -------------------------------------------------------------------------------- /src/Lib/Maybe.agda: -------------------------------------------------------------------------------- 1 | module Lib.Maybe where 2 | open import Lib.Dec 3 | open import Lib.Equality 4 | open import Lib.Product 5 | open import Lib.Zero 6 | 7 | data Maybe {a} (A : Set a) : Set a where 8 | just : A → Maybe A 9 | nothing : Maybe A 10 | 11 | mapMaybe : ∀ {a b} {A : Set a} {B : Set b} → (A → B) → Maybe A → Maybe B 12 | mapMaybe f (just x) = just (f x) 13 | mapMaybe f nothing = nothing 14 | 15 | infixr 4 _>>=_ _×M_ 16 | _>>=_ : ∀ {a b} {A : Set a} {B : Set b} → Maybe A → (A → Maybe B) → Maybe B 17 | just a >>= amb = amb a 18 | nothing >>= amb = nothing 19 | 20 | Dec→Maybe : ∀ {a A} → Dec {a} A → Maybe A 21 | Dec→Maybe (yes p) = just p 22 | Dec→Maybe (no np) = nothing 23 | 24 | _×M_ : ∀ {a b A B} → Maybe {a} A → Maybe {b} B → Maybe (A × B) 25 | just x ×M just y = just (x , y) 26 | just x ×M nothing = nothing 27 | nothing ×M mb = nothing 28 | 29 | nothing/=just : ∀ {a A x} → nothing {a} {A} ≡ just x → Zero 30 | nothing/=just () 31 | -------------------------------------------------------------------------------- /src/Lib/Nat.agda: -------------------------------------------------------------------------------- 1 | module Lib.Nat where 2 | open import Lib.Dec 3 | open import Lib.Equality 4 | 5 | open import Data.Nat public using (zero) renaming (ℕ to Nat; suc to succ) 6 | 7 | infixr 6 _+N_ 8 | _+N_ : Nat → Nat → Nat 9 | zero +N n = n 10 | succ m +N n = succ (m +N n) 11 | 12 | +N-zero : ∀ m → m +N zero ≡ m 13 | +N-zero zero = refl 14 | +N-zero (succ m) = cong succ (+N-zero m) 15 | 16 | +N-succ : ∀ m n → m +N succ n ≡ succ (m +N n) 17 | +N-succ zero n = refl 18 | +N-succ (succ m) n = cong succ (+N-succ m n) 19 | 20 | succInj : ∀ {m n} → succ m ≡ succ n → m ≡ n 21 | succInj refl = refl 22 | 23 | _≟Nat_ : DecEq Nat 24 | zero ≟Nat zero = yes refl 25 | zero ≟Nat succ n = no (λ ()) 26 | succ m ≟Nat zero = no (λ ()) 27 | succ m ≟Nat succ n = mapDec (cong succ) succInj (m ≟Nat n) 28 | 29 | +N-assoc : ∀ m n o → (m +N n) +N o ≡ m +N (n +N o) 30 | +N-assoc zero n o = refl 31 | +N-assoc (succ x) n o = cong succ (+N-assoc x n o) 32 | 33 | +N-comm : ∀ m n → m +N n ≡ n +N m 34 | +N-comm m zero = +N-zero m 35 | +N-comm m (succ n) = trans (+N-succ m n) (cong succ (+N-comm m n)) 36 | -------------------------------------------------------------------------------- /src/Lib/One.agda: -------------------------------------------------------------------------------- 1 | module Lib.One where 2 | 3 | record One : Set where 4 | constructor <> 5 | open One public 6 | -------------------------------------------------------------------------------- /src/Lib/Product.agda: -------------------------------------------------------------------------------- 1 | module Lib.Product where 2 | open import Lib.Dec 3 | open import Lib.Level 4 | 5 | open import Data.Product public 6 | using (Σ; module Σ; _,_) 7 | renaming (proj₁ to fst; proj₂ to snd) 8 | 9 | infixr 2 _×_ _×?_ 10 | 11 | _×_ : ∀ {a b} → Set a → Set b → Set (a ⊔ b) 12 | A × B = Σ A λ _ → B 13 | 14 | ∃ : ∀ {a b} {A : Set a} (B : A → Set b) → Set (a ⊔ b) 15 | ∃ = Σ _ 16 | 17 | ∃2 : ∀ {a b c} {A : Set a} {B : Set b} (C : A → B → Set c) → Set (a ⊔ b ⊔ c) 18 | ∃2 C = Σ _ (λ a → Σ _ (λ b → C a b)) 19 | 20 | mapΣ : ∀ {a a′ b b′} {A : Set a} {A′ : Set a′} {B : A → Set b} {B′ : A′ → Set b′} 21 | (fa : A → A′) → (∀ {a} → B a → B′ (fa a)) → Σ A B → Σ A′ B′ 22 | mapΣ fa fb (a , b) = fa a , fb b 23 | 24 | uncurry : ∀ {a b c} {A : Set a} {B : A → Set b} {C : (a : A) → B a → Set c} → 25 | ((a : A) (b : B a) → C a b) → (ab : Σ A B) → C (fst ab) (snd ab) 26 | uncurry f (a , b) = f a b 27 | 28 | swap : ∀ {a b} {A : Set a} {B : Set b} → A × B → B × A 29 | swap (x , y) = y , x 30 | 31 | assoc : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} → 32 | A × (B × C) → (A × B) × C 33 | assoc (x , (y , z)) = ((x , y) , z) 34 | 35 | unassoc : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} → 36 | (A × B) × C → A × (B × C) 37 | unassoc ((x , y) , z) = (x , (y , z)) 38 | 39 | map× : ∀ {a a′ b b′} {A : Set a} {A′ : Set a′} {B : Set b} {B′ : Set b′} → 40 | (A → A′) → (B → B′) → A × B → A′ × B′ 41 | map× fa fb = mapΣ fa fb 42 | 43 | <_,_> : ∀ {a b c} {A : Set a} {B : A → Set b} 44 | {C : ∀ {x} → B x → Set c} 45 | (f : ∀ x → B x) (g : ∀ x → C (f x)) 46 | (x : A) → Σ (B x) C 47 | < f , g > x = f x , g x 48 | 49 | pure : ∀ {a} {A : Set a} → A → A × A 50 | pure x = x , x 51 | 52 | infixl 4 _<*>_ 53 | _<*>_ : ∀ {a a′ b b′} {A : Set a} {A′ : Set a′} {B : A → Set b} {B′ : A′ → Set b′} → 54 | (Σ (A → A′) λ fa → ∀ {a} → B a → B′ (fa a)) → Σ A B → Σ A′ B′ 55 | _<*>_ (fa , fb) = mapΣ fa fb 56 | 57 | _×?_ : ∀ {a b A B} → Dec {a} A → Dec {b} B → Dec (A × B) 58 | _×?_ (yes a) (yes b) = yes (a , b) 59 | _×?_ (yes a) (no nb) = no (λ { (_ , b) → nb b }) 60 | _×?_ (no na) B? = no (λ { (a , b) → na a }) 61 | -------------------------------------------------------------------------------- /src/Lib/Relation.agda: -------------------------------------------------------------------------------- 1 | module Lib.Relation where 2 | 3 | import Lib.FunctionProperties as FP 4 | open FP using (Rel) public 5 | open FP hiding (Rel) 6 | 7 | open import Lib.Function 8 | open import Lib.Level 9 | open import Lib.One 10 | open import Lib.Product 11 | open import Lib.Setoid 12 | open import Lib.Sum 13 | open import Lib.Sum.Pointwise 14 | 15 | module _ {a l} (A : Setoid a l) where 16 | open Setoid A 17 | 18 | liftR0 : ∀ {z} → Set z → Rel A z 19 | liftR0 X _ _ = X 20 | 21 | liftR : ∀ {x y z} → (Set x → Set y → Set z) → Rel A x → Rel A y → Rel A z 22 | liftR P R S x y = P (R x y) (S x y) 23 | 24 | --_⟨_⟩R_ : ∀ {x y z} → Rel A x → (Set x → Set y → Set z) → Rel A y → Rel A z 25 | --(R ⟨ P ⟩R S) x y = P (R x y) (S x y) 26 | 27 | SymClosure : ∀ {m} → Rel A m → Rel A m 28 | SymClosure R = liftR _⊎_ R (flip R) 29 | 30 | AntisymClosure : ∀ {m} → Rel A m → Rel A m 31 | AntisymClosure R = liftR _×_ R (flip R) 32 | 33 | [_]_⇒_ : ∀ {x y} → Rel A x → Rel A y → Set _ 34 | [_]_⇒_ R S = ∀ x y → liftR (λ X Y → X → Y) R S x y 35 | 36 | [_]_⇔_ : ∀ {x y} → Rel A x → Rel A y → Set _ 37 | [_]_⇔_ R S = ∀ x y → liftR (λ X Y → (X → Y) × (Y → X)) R S x y 38 | 39 | module _ {a b l m} (A : Setoid a l) (B : Setoid b m) where 40 | private 41 | module A = Setoid A ; module B = Setoid B 42 | 43 | [_,_]_×R_ : ∀ {x y} → Rel A x → Rel B y → Rel (A ×S B) _ 44 | [_,_]_×R_ R S (xa , xb) (ya , yb) = R xa ya × S xb yb 45 | 46 | [_,_]_⊎R_ : ∀ {x y} → Rel A x → Rel B y → Rel (A ⊎S B) _ 47 | [_,_]_⊎R_ = _⊎R_ 48 | 49 | 1R : Rel OneS _ 50 | 1R _ _ = One 51 | -------------------------------------------------------------------------------- /src/Lib/Relation/Propositional.agda: -------------------------------------------------------------------------------- 1 | module Lib.Relation.Propositional where 2 | 3 | import Lib.Relation as R 4 | 5 | open import Lib.Equality 6 | open import Lib.Function 7 | open import Lib.Level 8 | open import Lib.One 9 | open import Lib.Product 10 | open import Lib.Setoid 11 | 12 | Rel : ∀ {a} → Set a → (l : Level) → Set _ 13 | Rel A = R.Rel (≡-Setoid A) 14 | 15 | liftR0 : ∀ {a z} {A : Set a} → Set z → Rel A z 16 | liftR0 {A = A} = R.liftR0 (≡-Setoid A) 17 | 18 | liftR : ∀ {a x y z} {A : Set a} → 19 | (Set x → Set y → Set z) → Rel A x → Rel A y → Rel A z 20 | liftR {A = A} = R.liftR (≡-Setoid A) 21 | 22 | SymClosure : ∀ {a l} {A : Set a} → Rel A l → Rel A l 23 | SymClosure {A = A} = R.SymClosure (≡-Setoid A) 24 | 25 | AntisymClosure : ∀ {a l} {A : Set a} → Rel A l → Rel A l 26 | AntisymClosure {A = A} = R.AntisymClosure (≡-Setoid A) 27 | 28 | _⟨_⟩R_ : ∀ {a x y z} {A : Set a} → 29 | Rel A x → (Set x → Set y → Set z) → Rel A y → Rel A z 30 | X ⟨ R ⟩R Y = liftR R X Y 31 | 32 | _⇒_ : ∀ {a x y} {A : Set a} → Rel A x → Rel A y → Set _ 33 | _⇒_ {A = A} = R.[ ≡-Setoid A ]_⇒_ 34 | 35 | _⇔_ : ∀ {a x y} {A : Set a} → Rel A x → Rel A y → Set _ 36 | _⇔_ {A = A} = R.[ ≡-Setoid A ]_⇔_ 37 | 38 | _×R_ : ∀ {a b x y} {A : Set a} {B : Set b} → Rel A x → Rel B y → Rel (A × B) _ 39 | _×R_ {A = A} {B} = R.[ ≡-Setoid A , ≡-Setoid B ]_×R_ 40 | 41 | 1R : Rel One _ 42 | 1R = R.1R 43 | -------------------------------------------------------------------------------- /src/Lib/Structure/CommutativeMonoid.agda: -------------------------------------------------------------------------------- 1 | open import Lib.Structure 2 | open import Lib.Structure.Sugar 3 | 4 | module Lib.Structure.CommutativeMonoid {c l} (M : ΣCommutativeMonoid c l) where 5 | 6 | open ΣCommutativeMonoid M 7 | 8 | open import Lib.Setoid 9 | 10 | open SetoidReasoning Carrier 11 | 12 | -- Middle-four exchange: 13 | -- a · b = ab 14 | -- · · · 15 | -- c · d = cd 16 | -- ∥ ∥ ∥ 17 | -- ac · bd = r 18 | 19 | exchange : ∀ a b c d → (a · b) · (c · d) ≈ (a · c) · (b · d) 20 | exchange a b c d = 21 | (a · b) · (c · d) ≈[ sym (assoc (a · b) c d) ]≈ 22 | ((a · b) · c) · d ≈[ 23 | ((a · b) · c ≈[ assoc a b c ]≈ 24 | a · (b · c) ≈[ refl ·-cong comm b c ]≈ 25 | a · (c · b) ≈[ sym (assoc a c b) ]≈ 26 | (a · c) · b QED) 27 | ·-cong refl 28 | ]≈ 29 | ((a · c) · b) · d ≈[ assoc (a · c) b d ]≈ 30 | (a · c) · (b · d) QED 31 | -------------------------------------------------------------------------------- /src/Lib/Structure/Flip.agda: -------------------------------------------------------------------------------- 1 | module Lib.Structure.Flip where 2 | 3 | open import Lib.Function 4 | open import Lib.Product using (Σ; _×_; _,_; fst; snd; swap) 5 | open import Lib.Setoid 6 | open import Lib.Structure 7 | open import Lib.Structure.Sugar 8 | 9 | module _ {c l} {S : Setoid c l} where 10 | open Setoid S 11 | 12 | flip-isMonoid : ∀ {e _·_} → IsMonoid S e _·_ → IsMonoid S e (flip _·_) 13 | flip-isMonoid M = record 14 | { identity = swap identity 15 | ; assoc = λ x y z → sym (assoc z y x) 16 | ; _·-cong_ = flip _·-cong_ 17 | } where open IsMonoid M 18 | 19 | flip-isSemiring : 20 | ∀ {e0 e1 _+_ _*_} → 21 | IsSemiring S e0 e1 _+_ _*_ → IsSemiring S e0 e1 _+_ (flip _*_) 22 | flip-isSemiring R = record 23 | { +-isCommutativeMonoid = +-isCommutativeMonoid 24 | ; *-isMonoid = flip-isMonoid *-isMonoid 25 | ; annihil = swap annihil 26 | ; distrib = swap distrib 27 | } where open IsSemiring R 28 | 29 | -- 30 | 31 | flip-monoid : Monoid S → Monoid S 32 | flip-monoid M = record { isMonoid = flip-isMonoid isMonoid } 33 | where open Monoid M 34 | 35 | flip-semiring : Semiring S → Semiring S 36 | flip-semiring M = record { isSemiring = flip-isSemiring isSemiring } 37 | where open Semiring M 38 | 39 | -- 40 | 41 | flip-σMonoid : ∀ {c l} → ΣMonoid c l → ΣMonoid c l 42 | flip-σMonoid M = record { monoid = flip-monoid monoid } 43 | where open ΣMonoid M 44 | 45 | flip-σSemiring : ∀ {c l} → ΣSemiring c l → ΣSemiring c l 46 | flip-σSemiring M = record { semiring = flip-semiring semiring } 47 | where open ΣSemiring M 48 | -------------------------------------------------------------------------------- /src/Lib/Structure/Morphism.agda: -------------------------------------------------------------------------------- 1 | -- Algebra homomorphisms 2 | 3 | module Lib.Structure.Morphism where 4 | 5 | open import Lib.Level 6 | open import Lib.Product 7 | open import Lib.Setoid 8 | open import Lib.Structure 9 | open import Lib.Structure.Sugar 10 | 11 | module _ {c l} (M N : ΣMonoid c l) where 12 | private 13 | module M = ΣMonoid M ; module N = ΣMonoid N 14 | 15 | -- record Monoid=> : Set (c ⊔ l) where 16 | -- field 17 | -- el : M.Carrier →E N.Carrier 18 | -- e : el $E M.e N.≈ N.e 19 | -- _·_ : ∀ x y → el $E (x M.· y) N.≈ el $E x N.· el $E y 20 | 21 | record IsMonoidArr (el : M.Carrier →E N.Carrier) : Set (c ⊔ l) where 22 | field 23 | e : el $E M.e N.≈ N.e 24 | _·_ : ∀ x y → el $E (x M.· y) N.≈ el $E x N.· el $E y 25 | 26 | MonoidArr : Setoid (c ⊔ l) (c ⊔ l) 27 | MonoidArr = let S = (M.Carrier →S N.Carrier) in 28 | let open Setoid S in 29 | Subsetoid setoidOver IsMonoidArr 30 | 31 | module MonoidArr (h : Setoid.C MonoidArr) where 32 | 33 | el : M.Carrier →E N.Carrier 34 | el = fst h 35 | 36 | isMonoidArr : IsMonoidArr el 37 | isMonoidArr = snd h 38 | open IsMonoidArr isMonoidArr public 39 | 40 | module _ {c l} (M N : ΣCommutativeMonoid c l) where 41 | private 42 | module M = ΣCommutativeMonoid M ; module N = ΣCommutativeMonoid N 43 | 44 | CommutativeMonoidArr : Setoid (c ⊔ l) (c ⊔ l) 45 | CommutativeMonoidArr = MonoidArr M.σMonoid N.σMonoid 46 | 47 | module CommutativeMonoidArr (h : Setoid.C CommutativeMonoidArr) = 48 | MonoidArr M.σMonoid N.σMonoid h 49 | -------------------------------------------------------------------------------- /src/Lib/Structure/Semiring.agda: -------------------------------------------------------------------------------- 1 | open import Lib.Structure 2 | open import Lib.Structure.Sugar 3 | 4 | module Lib.Structure.Semiring {c l} (R : ΣSemiring c l) where 5 | 6 | open ΣSemiring R 7 | 8 | open import Lib.Structure.CommutativeMonoid +-σCommutativeMonoid public using () 9 | renaming (exchange to +-exchange) 10 | -------------------------------------------------------------------------------- /src/Lib/Structure/Sugar.agda: -------------------------------------------------------------------------------- 1 | -------------------------------------------------------------------------------- 2 | -- TODO: 3 | -- Do renaming: 4 | -- - from Lib.Structure: 5 | -- - - Monoid ↦ MonoidOver 6 | -- - - CommutativeMonoid ↦ CommutativeMonoidOver 7 | -- - - &c 8 | -- - from Lib.Structure.Sugar: 9 | -- - - ΣMonoid ↦ Monoid 10 | -- - - ΣCommutativeMonoid ↦ CommutativeMonoid 11 | -- - - &c 12 | -- Maybe move Lib.Structure ↦ Lib.Structure.Over 13 | -- and Lib.Structure.Sugar ↦ Lib.Structure 14 | 15 | module Lib.Structure.Sugar where 16 | 17 | open import Lib.Level 18 | open import Lib.Setoid 19 | open import Lib.Structure 20 | 21 | -- Order 22 | 23 | record ΣPreorder c l l′ : Set (lsuc (c ⊔ l ⊔ l′)) where 24 | field 25 | Carrier : Setoid c l 26 | preorder : Preorder Carrier l′ 27 | open Setoid Carrier public 28 | open Preorder preorder public 29 | 30 | record ΣPoset c l l′ : Set (lsuc (c ⊔ l ⊔ l′)) where 31 | field 32 | Carrier : Setoid c l 33 | poset : Poset Carrier l′ 34 | open Setoid Carrier public 35 | open Poset poset public 36 | 37 | σPreorder : ΣPreorder c l l′ 38 | σPreorder = record { Carrier = Carrier ; preorder = preorder } 39 | 40 | -- Monoid-like 41 | 42 | record ΣMonoid c l : Set (lsuc (c ⊔ l)) where 43 | field 44 | Carrier : Setoid c l 45 | monoid : Monoid Carrier 46 | open Setoid Carrier public 47 | open Monoid monoid public 48 | 49 | record ΣCommutativeMonoid c l : Set (lsuc (c ⊔ l)) where 50 | field 51 | Carrier : Setoid c l 52 | commutativeMonoid : CommutativeMonoid Carrier 53 | open Setoid Carrier public 54 | open CommutativeMonoid commutativeMonoid public 55 | 56 | σMonoid : ΣMonoid c l 57 | σMonoid = record { Carrier = Carrier ; monoid = monoid } 58 | 59 | -- Semiring-like 60 | 61 | record ΣSemiring c l : Set (lsuc (c ⊔ l)) where 62 | field 63 | Carrier : Setoid c l 64 | semiring : Semiring Carrier 65 | open Setoid Carrier public 66 | open Semiring semiring public 67 | 68 | +-σCommutativeMonoid : ΣCommutativeMonoid c l 69 | +-σCommutativeMonoid = record { commutativeMonoid = +-commutativeMonoid } 70 | open ΣCommutativeMonoid +-σCommutativeMonoid public 71 | using () 72 | renaming (σMonoid to +-σMonoid) 73 | 74 | *-σMonoid : ΣMonoid c l 75 | *-σMonoid = record { monoid = *-monoid } 76 | 77 | -- Order and structure 78 | 79 | record ΣPomonoid c l l′ : Set (lsuc (c ⊔ l ⊔ l′)) where 80 | field 81 | Carrier : Setoid c l 82 | pomonoid : Pomonoid Carrier l′ 83 | open Setoid Carrier public 84 | open Pomonoid pomonoid public 85 | 86 | σMonoid : ΣMonoid c l 87 | σMonoid = record { Carrier = Carrier ; monoid = monoid } 88 | 89 | σPoset : ΣPoset c l l′ 90 | σPoset = record { Carrier = Carrier ; poset = poset } 91 | open ΣPoset σPoset public using (σPreorder) 92 | 93 | record ΣCommutativePomonoid c l l′ : Set (lsuc (c ⊔ l ⊔ l′)) where 94 | field 95 | Carrier : Setoid c l 96 | commutativePomonoid : CommutativePomonoid Carrier l′ 97 | open Setoid Carrier public 98 | open CommutativePomonoid commutativePomonoid public 99 | 100 | σCommutativeMonoid : ΣCommutativeMonoid c l 101 | σCommutativeMonoid = record 102 | { Carrier = Carrier 103 | ; commutativeMonoid = commutativeMonoid 104 | } 105 | 106 | σPomonoid : ΣPomonoid c l l′ 107 | σPomonoid = record { Carrier = Carrier ; pomonoid = pomonoid } 108 | open ΣPomonoid σPomonoid public using (σMonoid; σPoset; σPreorder) 109 | 110 | record ΣPosemiring c l l′ : Set (lsuc (c ⊔ l ⊔ l′)) where 111 | field 112 | Carrier : Setoid c l 113 | posemiring : Posemiring Carrier l′ 114 | open Setoid Carrier public 115 | open Posemiring posemiring public 116 | 117 | σSemiring : ΣSemiring c l 118 | σSemiring = record { semiring = semiring } 119 | open ΣSemiring σSemiring public 120 | using (+-σCommutativeMonoid; +-σMonoid; *-σMonoid) 121 | 122 | +-σCommutativePomonoid : ΣCommutativePomonoid c l l′ 123 | +-σCommutativePomonoid = 124 | record { commutativePomonoid = +-commutativePomonoid } 125 | open ΣCommutativePomonoid +-σCommutativePomonoid public 126 | using (σPoset; σPreorder) 127 | renaming (σPomonoid to +-σPomonoid) 128 | 129 | *-σPomonoid : ΣPomonoid c l l′ 130 | *-σPomonoid = record { pomonoid = *-pomonoid } 131 | -------------------------------------------------------------------------------- /src/Lib/Sum.agda: -------------------------------------------------------------------------------- 1 | module Lib.Sum where 2 | open import Lib.Level 3 | 4 | infixr 1 _⊎_ 5 | data _⊎_ {a b} (A : Set a) (B : Set b) : Set (a ⊔ b) where 6 | inl : (a : A) → A ⊎ B 7 | inr : (b : B) → A ⊎ B 8 | 9 | map⊎ : ∀ {a a′ b b′} {A : Set a} {A′ : Set a′} {B : Set b} {B′ : Set b′} → 10 | (A → A′) → (B → B′) → A ⊎ B → A′ ⊎ B′ 11 | map⊎ f g (inl a) = inl (f a) 12 | map⊎ f g (inr b) = inr (g b) 13 | 14 | [_,_] : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} → 15 | (A → C) → (B → C) → (A ⊎ B → C) 16 | [ f , g ] (inl a) = f a 17 | [ f , g ] (inr b) = g b 18 | 19 | swap⊎ : ∀ {a b} {A : Set a} {B : Set b} → A ⊎ B → B ⊎ A 20 | swap⊎ (inl a) = inr a 21 | swap⊎ (inr b) = inl b 22 | -------------------------------------------------------------------------------- /src/Lib/Sum/Pointwise.agda: -------------------------------------------------------------------------------- 1 | module Lib.Sum.Pointwise where 2 | open import Lib.Equality 3 | open import Lib.Function 4 | open import Lib.Level 5 | open import Lib.One 6 | open import Lib.Sum 7 | 8 | data _⊎R_ {a a′ b b′ r s} {A : Set a} {A′ : Set a′} {B : Set b} {B′ : Set b′} 9 | (R : A → A′ → Set r) (S : B → B′ → Set s) : 10 | A ⊎ B → A′ ⊎ B′ → Set (r ⊔ s) where 11 | inl : ∀ {x y} (r : R x y) → (R ⊎R S) (inl x) (inl y) 12 | inr : ∀ {x y} (s : S x y) → (R ⊎R S) (inr x) (inr y) 13 | 14 | map⊎R : ∀ {a a′ b b′ c c′ d d′ r r′ s s′} 15 | {A : Set a} {A′ : Set a′} {B : Set b} {B′ : Set b′} 16 | {C : Set c} {C′ : Set c′} {D : Set d} {D′ : Set d′} 17 | {fx : A → A′} {fy : B → B′} {gx : C → C′} {gy : D → D′} 18 | {R : A → B → Set r} {S : C → D → Set s} 19 | {R′ : A′ → B′ → Set r′} {S′ : C′ → D′ → Set s′} → 20 | (∀ {x y} → R x y → R′ (fx x) (fy y)) → 21 | (∀ {x y} → S x y → S′ (gx x) (gy y)) → 22 | ∀ {x y} → (R ⊎R S) x y → (R′ ⊎R S′) (map⊎ fx gx x) (map⊎ fy gy y) 23 | map⊎R fr gs (inl r) = inl (fr r) 24 | map⊎R fr gs (inr s) = inr (gs s) 25 | 26 | map⊎R′ : ∀ {a b c d r r′ s s′} 27 | {A : Set a} {B : Set b} {C : Set c} {D : Set d} 28 | {R : A → B → Set r} {S : C → D → Set s} 29 | {R′ : A → B → Set r′} {S′ : C → D → Set s′} → 30 | (∀ {x y} → R x y → R′ x y) → (∀ {x y} → S x y → S′ x y) → 31 | ∀ {x y} → (R ⊎R S) x y → (R′ ⊎R S′) x y 32 | map⊎R′ f g (inl r) = inl (f r) 33 | map⊎R′ f g (inr s) = inr (g s) 34 | 35 | Agree : ∀ {a a′ b b′} {A : Set a} {A′ : Set a′} {B : Set b} {B′ : Set b′} → 36 | A ⊎ B → A′ ⊎ B′ → Set 37 | Agree = (λ _ _ → One) ⊎R (λ _ _ → One) 38 | 39 | ⇒Agree : ∀ {a a′ b b′ r s} {A : Set a} {A′ : Set a′} {B : Set b} {B′ : Set b′} 40 | {R : A → A′ → Set r} {S : B → B′ → Set s} {x y} → 41 | (R ⊎R S) x y → Agree x y 42 | ⇒Agree = map⊎R′ (λ _ → <>) (λ _ → <>) 43 | 44 | map⊎-rel : ∀ {a a′ b b′} {A : Set a} {A′ : Set a′} {B : Set b} {B′ : Set b′} 45 | (f : A → A′) (g : B → B′) (z : A ⊎ B) → 46 | ((λ x fx → f x ≡ fx) ⊎R (λ y gy → g y ≡ gy)) z (map⊎ f g z) 47 | map⊎-rel f g (inl x) = inl refl 48 | map⊎-rel f g (inr y) = inr refl 49 | -------------------------------------------------------------------------------- /src/Lib/Two.agda: -------------------------------------------------------------------------------- 1 | module Lib.Two where 2 | open import Lib.One 3 | open import Lib.Zero 4 | 5 | data Two : Set where 6 | ttt fff : Two 7 | 8 | T : Two → Set 9 | T ttt = One 10 | T fff = Zero 11 | 12 | Two-ind : ∀ {p} (P : Two → Set p) → P ttt → P fff → ∀ x → P x 13 | Two-ind P t f ttt = t 14 | Two-ind P t f fff = f 15 | 16 | Two-rec : ∀ {a} {A : Set a} → A → A → Two → A 17 | Two-rec = Two-ind (λ _ → _) 18 | -------------------------------------------------------------------------------- /src/Lib/TypeAlgebra.agda: -------------------------------------------------------------------------------- 1 | module Lib.TypeAlgebra where 2 | 3 | open import Lib.Product 4 | open import Lib.Sum 5 | 6 | ×-⊎-distrib-l : ∀ {a b c} {A : Set a} {B : Set b} {C : Set c} → 7 | (A ⊎ B) × C → (A × C) ⊎ (B × C) 8 | ×-⊎-distrib-l (inl a , c) = inl (a , c) 9 | ×-⊎-distrib-l (inr b , c) = inr (b , c) 10 | -------------------------------------------------------------------------------- /src/Lib/Vec/Thinning.agda: -------------------------------------------------------------------------------- 1 | module Lib.Vec.Thinning where 2 | 3 | open import Lib.Equality 4 | open import Lib.Nat 5 | open import Lib.Thinning 6 | open import Lib.Vec 7 | open import Lib.VZip 8 | 9 | module _ {a} {A : Set a} where 10 | 11 | thin : ∀ {m n} → m ≤ n → Vec A n → Vec A m 12 | thin oz nil = nil 13 | thin (os th) (x :: xs) = x :: thin th xs 14 | thin (o′ th) (x :: xs) = thin th xs 15 | 16 | thin-oi : ∀ {n} (xs : Vec A n) → VZip _≡_ (thin oi xs) xs 17 | thin-oi nil = nil 18 | thin-oi (x :: xs) = refl :: (thin-oi xs) 19 | 20 | thin-comp : ∀ {m n o} (th : m ≤ n) (ph : n ≤ o) xs → 21 | VZip _≡_ (thin (th ≤-comp ph) xs) (thin th (thin ph xs)) 22 | thin-comp th oz nil = ≡VZip refl 23 | thin-comp (os th) (os ph) (x :: xs) = refl :: thin-comp th ph xs 24 | thin-comp (o′ th) (os ph) (x :: xs) = thin-comp th ph xs 25 | thin-comp th (o′ ph) (x :: xs) = thin-comp th ph xs 26 | 27 | lookup-is-thin : ∀ {n} (i : 1 ≤ n) xs → lookup i xs ≡ headVec (thin i xs) 28 | lookup-is-thin (os i) (x :: xs) = refl 29 | lookup-is-thin (o′ i) (x :: xs) = lookup-is-thin i xs 30 | 31 | lookup′ : ∀ {n} → 1 ≤ n → Vec A n → A 32 | lookup′ th xs = headVec (thin th xs) 33 | -------------------------------------------------------------------------------- /src/Lib/Zero.agda: -------------------------------------------------------------------------------- 1 | module Lib.Zero where 2 | 3 | data Zero : Set where 4 | 5 | Zero-elim : ∀ {l} {A : Set l} → Zero → A 6 | Zero-elim () 7 | 8 | Not : ∀ {a} → Set a → Set a 9 | Not A = A → Zero 10 | 11 | aboutZero : ∀ {p} (P : Zero → Set p) {x} → P x 12 | aboutZero P {()} 13 | -------------------------------------------------------------------------------- /src/Quantitative/Models/RelationTransformer/Action.agda: -------------------------------------------------------------------------------- 1 | module Quantitative.Models.RelationTransformer.Action where 2 | 3 | open import Lib.Category 4 | open import Lib.Equality 5 | open import Lib.Function 6 | open import Lib.Level 7 | open import Lib.One 8 | open import Lib.Product 9 | open import Lib.Relation.Propositional 10 | open import Lib.Structure using (DecToppedMeetSemilatticeSemiring) 11 | open import Lib.Setoid 12 | open import Lib.Sum 13 | 14 | open import Quantitative.Models.RelationTransformer hiding (_≤_) 15 | open DecToppedMeetSemilatticeSemiring decToppedMeetSemilatticeSemiring 16 | 17 | act : ∀ {a} {A : Set a} → RelAct → Rel A a → Rel A a 18 | act zer R = λ _ _ → Lift _ One 19 | act cov R = R 20 | act con R = flip R 21 | act inv R = AntisymClosure R 22 | 23 | act-* : ∀ {a} {A : Set a} x y (R : Rel A a) → 24 | act (x * y) R ⇔ (act x o act y) R 25 | act-* zer y R a b = id , id 26 | act-* cov y R a b = id , id 27 | act-* con zer R a b = id , id 28 | act-* con cov R a b = id , id 29 | act-* con con R a b = id , id 30 | act-* con inv R a b = swap , swap 31 | act-* inv zer R a b = < id , id > , fst 32 | act-* inv cov R a b = id , id 33 | act-* inv con R a b = swap , swap 34 | act-* inv inv R a b = < id , swap > , fst 35 | 36 | act-≤ : ∀ {a} {A : Set a} {x y} → x ≤ y → (R : Rel A a) → act x R ⇒ act y R 37 | act-≤ (inv≤ zer) R a b r = lift <> 38 | act-≤ (inv≤ cov) R a b (r , s) = r 39 | act-≤ (inv≤ con) R a b (r , s) = s 40 | act-≤ (inv≤ inv) R a b r = r 41 | act-≤ (≤zer x) R a b r = lift <> 42 | act-≤ cov-refl R a b r = r 43 | act-≤ con-refl R a b r = r 44 | 45 | act-+ : ∀ {a} {A : Set a} x y (R : Rel A a) → 46 | act (x + y) R ⇔ (act x R ⟨ _×_ ⟩R act y R) 47 | act-+ zer y R a b = < const (lift <>) , id > , snd 48 | act-+ cov zer R a b = < id , const (lift <>) > , fst 49 | act-+ cov cov R a b = < id , id > , fst 50 | act-+ cov con R a b = id , id 51 | act-+ cov inv R a b = < fst , id > , map× id snd 52 | act-+ con zer R a b = < id , const (lift <>) > , fst 53 | act-+ con cov R a b = swap , swap 54 | act-+ con con R a b = < id , id > , fst 55 | act-+ con inv R a b = < snd , id > , < fst o snd , fst > 56 | act-+ inv y R a b = < id , act-≤ (inv≤ y) R a b > , fst 57 | 58 | open import Quantitative.Semantics.Relational ONE {!!} {!!} 59 | (record { obj = λ _ → OneS 60 | ; arr = record { _$E_ = λ _ → →E-⊤ _ λ _ → <> ; _$E=_ = λ _ _ → <> } 61 | ; isFunctor = record { arr-id = λ _ _ → <> ; arr->> = λ _ → <> } 62 | }) 63 | (record { obj = λ _ → OneS 64 | ; arr = record { _$E_ = λ _ → →E-⊤ _ λ _ → <> ; _$E=_ = λ _ _ → <> } 65 | ; isFunctor = record { arr-id = λ _ _ → <> ; arr->> = λ _ → <> } 66 | }) 67 | (record { isPromonoidal = record 68 | { JP = record 69 | { to = →E-⊤ _ λ _ → <> 70 | ; from = record { _$E_ = λ _ → <> , <> , <> 71 | ; _$E=_ = λ _ → refl , refl , <> , <> 72 | } 73 | ; to-from = λ _ → refl , refl , <> , <> 74 | ; from-to = λ _ → <> 75 | } 76 | ; PJ = record 77 | { to = →E-⊤ _ λ _ → <> 78 | ; from = record { _$E_ = λ _ → <> , <> , <> 79 | ; _$E=_ = λ _ → refl , refl , <> , <> 80 | } 81 | ; to-from = λ _ → refl , refl , <> , <> 82 | ; from-to = λ _ → <> 83 | } 84 | ; PP = record 85 | { to = record { _$E_ = λ _ → <> , <> , <> 86 | ; _$E=_ = λ _ → refl , refl , <> , <> 87 | } 88 | ; from = record { _$E_ = λ _ → <> , <> , <> 89 | ; _$E=_ = λ _ → refl , refl , <> , <> 90 | } 91 | ; to-from = λ _ → refl , refl , <> , <> 92 | ; from-to = λ _ → refl , refl , <> , <> 93 | } 94 | } 95 | ; comm = →E-⊤ _ λ _ → <> 96 | }) 97 | RelAct posemiring (λ x → record 98 | { obj = λ R → record 99 | { obj = λ u → act x (Functor.obj R u) 100 | ; arr = record { _$E_ = λ _ _ _ S → S ; _$E=_ = λ _ → <> } 101 | ; isFunctor = record { arr-id = λ _ → <> ; arr->> = <> } 102 | } 103 | ; arr = record 104 | { _$E_ = λ α → record 105 | { η = λ X x₁ y x₂ → {!NatTrans.η α X x₁ y!} 106 | ; square = {!!} 107 | } 108 | ; _$E=_ = {!!} 109 | } 110 | ; isFunctor = record { arr-id = {!!} ; arr->> = {!!} } 111 | }) 112 | -------------------------------------------------------------------------------- /src/Quantitative/Resources.agda: -------------------------------------------------------------------------------- 1 | open import Lib.Equality 2 | open import Lib.Setoid 3 | open import Lib.Structure 4 | open import Lib.Structure.Sugar 5 | 6 | import Quantitative.Types.Formers as Form 7 | 8 | module Quantitative.Resources 9 | {c k l′} (PrimTy : Set c) (C : Set c) (open Form PrimTy C) 10 | (Const : Set k) (constTy : Const → Ty) 11 | (POS : Posemiring (≡-Setoid C) l′) where 12 | 13 | open Posemiring POS using (poset; semiring; +-commutativeMonoid) 14 | 15 | open import Quantitative.Syntax Ty Const 16 | open import Quantitative.Types PrimTy C Const constTy 17 | open import Quantitative.Resources.Context C Const POS 18 | open import Lib.Module 19 | 20 | open import Lib.Matrix.Setoid (≡-Setoid C) 21 | open import Lib.Matrix.Poset (record { poset = poset }) 22 | open import Lib.Matrix.Addition (record { commutativeMonoid = +-commutativeMonoid }) 23 | open import Lib.Matrix.Multiplication (record { semiring = semiring }) 24 | open import Lib.Matrix.Multiplication.Basis (record { semiring = semiring }) 25 | open import Lib.Matrix.Scaling.Right (record { semiring = semiring }) 26 | 27 | open import Lib.Level 28 | open import Lib.Two 29 | open import Lib.Vec 30 | open import Lib.Vec.Thinning 31 | 32 | infix 3 _⊢r_ 33 | 34 | data _⊢r_ {n} {Γ : TCtx n} (Δ : RCtx n) 35 | : ∀ {d S} {t : Term n d} (tt : Γ ⊢t t :-: S) → Set (l′ ⊔ c) where 36 | var : ∀ {th T} {eq : T ≡ lookup′ th Γ} 37 | (sub : Δ ≤M basis-col th) 38 | → 39 | Δ ⊢r var {th = th} eq 40 | const : ∀ {l} 41 | (split : Δ ≤M 0M) 42 | → 43 | Δ ⊢r const {l = l} 44 | app : ∀ {Δe Δs S T e s} {et : Γ ⊢t e ∈ S ⊸ T} {st : Γ ⊢t S ∋ s} 45 | (split : Δ ≤M Δe +M Δs) 46 | (er : Δe ⊢r et) (sr : Δs ⊢r st) 47 | → 48 | Δ ⊢r app et st 49 | bm : ∀ {Δe Δs S T ρ e s} {et : Γ ⊢t e ∈ ! ρ S} {st : S :: Γ ⊢t T ∋ s} 50 | (split : Δ ≤M Δe +M Δs) 51 | (er : Δe ⊢r et) (sr : Δs +↓ [- ρ -] ⊢r st) 52 | → 53 | Δ ⊢r bm et st 54 | del : ∀ {Δe Δs T e s} {et : Γ ⊢t e ∈ ⊗1} {st : Γ ⊢t T ∋ s} 55 | (split : Δ ≤M Δe +M Δs) 56 | (er : Δe ⊢r et) (sr : Δs ⊢r st) 57 | → 58 | Δ ⊢r del et st 59 | pm : ∀ {Δe Δs S0 S1 T e s} 60 | {et : Γ ⊢t e ∈ S0 ⊗ S1} {st : S0 :: S1 :: Γ ⊢t T ∋ s} 61 | (split : Δ ≤M Δe +M Δs) 62 | (er : Δe ⊢r et) (sr : Δs +↓ [- R.e1 -] +↓ [- R.e1 -] ⊢r st) 63 | → 64 | Δ ⊢r pm et st 65 | proj : ∀ {i S0 S1 e} {et : Γ ⊢t e ∈ S0 & S1} 66 | (er : Δ ⊢r et) 67 | → 68 | Δ ⊢r proj {i = i} et 69 | exf : ∀ {Δe Δs T e} {et : Γ ⊢t e ∈ ⊕0} 70 | (split : Δ ≤M Δe +M Δs) 71 | (er : Δe ⊢r et) 72 | → 73 | Δ ⊢r exf {T = T} et 74 | cse : ∀ {Δe Δs S0 S1 T e s0 s1} {et : Γ ⊢t e ∈ S0 ⊕ S1} 75 | {s0t : S0 :: Γ ⊢t T ∋ s0} {s1t : S1 :: Γ ⊢t T ∋ s1} 76 | (split : Δ ≤M Δe +M Δs) 77 | (er : Δe ⊢r et) (s0r : Δs +↓ [- R.e1 -] ⊢r s0t) 78 | (s1r : Δs +↓ [- R.e1 -] ⊢r s1t) 79 | → 80 | Δ ⊢r cse et s0t s1t 81 | fold : ∀ {S T e sn sc} {et : Γ ⊢t e ∈ LIST S} 82 | {snt : Γ ⊢t T ∋ sn} {sct : S :: T :: Γ ⊢t T ∋ sc} 83 | (er : Δ ⊢r et) (snr : 0M ⊢r snt) 84 | (scr : 0M +↓ [- R.e1 -] +↓ [- R.e1 -] ⊢r sct) 85 | → 86 | Δ ⊢r fold et snt sct 87 | the : ∀ {S s} {st : Γ ⊢t S ∋ s} 88 | (sr : Δ ⊢r st) 89 | → 90 | Δ ⊢r the st 91 | 92 | lam : ∀ {S T s} {st : S :: Γ ⊢t T ∋ s} 93 | (sr : Δ +↓ [- R.e1 -] ⊢r st) 94 | → 95 | Δ ⊢r lam st 96 | bang : ∀ {Δs S ρ s} {st : Γ ⊢t S ∋ s} 97 | (split : Δ ≤M Δs *r ρ) 98 | (sr : Δs ⊢r st) 99 | → 100 | Δ ⊢r bang {ρ = ρ} st 101 | unit : (split : Δ ≤M 0M) 102 | → 103 | Δ ⊢r unit 104 | ten : ∀ {Δs0 Δs1 S0 S1 s0 s1} {s0t : Γ ⊢t S0 ∋ s0} {s1t : Γ ⊢t S1 ∋ s1} 105 | (split : Δ ≤M Δs0 +M Δs1) 106 | (s0r : Δs0 ⊢r s0t) (s1r : Δs1 ⊢r s1t) 107 | → 108 | Δ ⊢r ten s0t s1t 109 | eat : Δ ⊢r eat 110 | wth : ∀ {S0 S1 s0 s1} {s0t : Γ ⊢t S0 ∋ s0} {s1t : Γ ⊢t S1 ∋ s1} 111 | (s0r : Δ ⊢r s0t) (s1r : Δ ⊢r s1t) 112 | → 113 | Δ ⊢r wth s0t s1t 114 | inj : ∀ {i S0 S1 s} {st : Γ ⊢t Two-rec S0 S1 i ∋ s} 115 | (sr : Δ ⊢r st) 116 | → 117 | Δ ⊢r inj {i = i} st 118 | nil : ∀ {S} (split : Δ ≤M 0M) 119 | → 120 | Δ ⊢r nil {S = S} 121 | cons : ∀ {Δs0 Δs1 S s0 s1} {s0t : Γ ⊢t S ∋ s0} {s1t : Γ ⊢t LIST S ∋ s1} 122 | (split : Δ ≤M Δs0 +M Δs1) 123 | (s0r : Δs0 ⊢r s0t) (s1r : Δs1 ⊢r s1t) 124 | → 125 | Δ ⊢r cons s0t s1t 126 | [_] : ∀ {S e} {et : Γ ⊢t e ∈ S} 127 | (er : Δ ⊢r et) 128 | → 129 | Δ ⊢r [ et ] 130 | -------------------------------------------------------------------------------- /src/Quantitative/Resources/Context.agda: -------------------------------------------------------------------------------- 1 | open import Lib.Dec 2 | open import Lib.Equality hiding (_QED) 3 | open import Lib.Setoid 4 | open import Lib.Structure 5 | 6 | module Quantitative.Resources.Context 7 | {c k l′} (C : Set c) (Const : Set k) (POS : Posemiring (≡-Setoid C) l′) where 8 | 9 | open import Quantitative.Syntax C Const 10 | 11 | open import Lib.Module 12 | open import Lib.Nat 13 | open import Lib.Product hiding (assoc) 14 | open import Lib.Structure.Sugar 15 | open import Lib.Thinning as Θ hiding (_≤_; ≤-refl) 16 | open import Lib.Vec 17 | open import Lib.VZip 18 | 19 | module R where 20 | open Posemiring POS public 21 | posemiring = POS 22 | 23 | infixr 1 _≤[_]_ 24 | infixr 2 _≤-QED 25 | 26 | _≤[_]_ : ∀ ρ {ρ′ ρ″} → ρ R.≤ ρ′ → ρ′ R.≤ ρ″ → ρ R.≤ ρ″ 27 | _≤[_]_ ρ {ρ′} {ρ″} xy yz = R.≤-trans {x = ρ} {y = ρ′} {z = ρ″} xy yz 28 | 29 | _≤-QED : ∀ ρ → ρ R.≤ ρ 30 | ρ ≤-QED = R.≤-refl 31 | 32 | open import Lib.Matrix.Setoid (≡-Setoid C) 33 | open import Lib.Matrix.Poset (record { poset = R.poset }) 34 | open import Lib.Matrix.Addition 35 | (record { commutativeMonoid = R.+-commutativeMonoid }) 36 | open import Lib.Matrix.Addition.Order 37 | (record { commutativePomonoid = R.+-commutativePomonoid }) 38 | open import Lib.Matrix.Multiplication (record { semiring = R.semiring }) 39 | open import Lib.Matrix.VecCompat (≡-Setoid C) 40 | 41 | -- Resource contexts 42 | 43 | RCtx : Nat → Set c 44 | RCtx n = Mat (n , 1) 45 | 46 | -- TODO: RCtx ↦ Mat, maybe in a different module 47 | 48 | -- Reasoning syntax for _≈M_ 49 | infixr 1 _≈[_]M_ 50 | infixr 2 _≈M-QED 51 | 52 | _≈[_]M_ : ∀ {n} Δ {Δ′ Δ″ : RCtx n} → Δ ≈M Δ′ → Δ′ ≈M Δ″ → Δ ≈M Δ″ 53 | _≈[_]M_ Δ {Δ′} {Δ″} xy yz = transM {x = Δ} {y = Δ′} {z = Δ″} xy yz 54 | 55 | _≈M-QED : ∀ {n} (Δ : RCtx n) → Δ ≈M Δ 56 | Δ ≈M-QED = reflM {x = Δ} 57 | 58 | -- Reasoning syntax for _≤M_ 59 | infixr 1 _≤[_]M_ 60 | infixr 2 _≤M-QED 61 | 62 | _≤[_]M_ : ∀ {n} Δ {Δ′ Δ″ : RCtx n} → Δ ≤M Δ′ → Δ′ ≤M Δ″ → Δ ≤M Δ″ 63 | _≤[_]M_ Δ {Δ′} {Δ″} xy yz = ≤M-trans {x = Δ} {y = Δ′} {z = Δ″} xy yz 64 | 65 | _≤M-QED : ∀ {n} (Δ : RCtx n) → Δ ≤M Δ 66 | Δ ≤M-QED = ≤M-refl {x = Δ} 67 | 68 | {- 69 | RCtx-setoid : Nat → Setoid _ _ 70 | RCtx-setoid n = record 71 | { C = RCtx n 72 | ; setoidOver = record 73 | { _≈_ = _≈_ 74 | ; isSetoid = record 75 | { refl = ≈-refl _ 76 | ; sym = ≈-sym 77 | ; trans = ≈-trans 78 | } 79 | } 80 | } 81 | 82 | commutativePomonoid : ∀ n → CommutativePomonoid (RCtx-setoid n) _ 83 | commutativePomonoid n = record 84 | { _≤_ = _≤_ 85 | ; e = 0Δ 86 | ; _·_ = _+Δ_ 87 | ; isCommutativePomonoid = record 88 | { _·-mono_ = _+Δ-mono_ 89 | ; isPoset = record 90 | { antisym = antisym 91 | ; isPreorder = record 92 | { ≤-reflexive = ≤-reflexive 93 | ; ≤-trans = ≤-trans 94 | } 95 | } 96 | ; isCommutativeMonoid = record 97 | { comm = +Δ-comm 98 | ; isMonoid = record 99 | { identity = fst +Δ-identity , snd +Δ-identity 100 | ; assoc = +Δ-assoc 101 | ; _·-cong_ = _+Δ-cong_ 102 | } 103 | } 104 | } 105 | } 106 | where 107 | antisym : ∀ {n} → Antisym (RCtx-setoid n) _≤_ 108 | antisym nil nil = nil 109 | antisym (≤R :: ≤Δ) (≥R :: ≥Δ) = R.antisym ≤R ≥R :: antisym ≤Δ ≥Δ 110 | 111 | posemimodule : ∀ n → Posemimodule (≡-Setoid C) (RCtx-setoid n) _ _ 112 | posemimodule n = record 113 | { _≤s_ = R._≤_ 114 | ; _≤f_ = _≤_ 115 | ; 0s = R.e0 116 | ; 1s = R.e1 117 | ; _+s_ = R._+_ 118 | ; _*s_ = R._*_ 119 | ; 0f = 0Δ 120 | ; 1f = 1Δ 121 | ; _+f_ = _+Δ_ 122 | ; _*f_ = _*Δ_ 123 | ; isPosemimodule = record 124 | { _+s-mono_ = R._+-mono_ 125 | ; _*s-mono_ = R._*-mono_ 126 | ; _+f-mono_ = _+Δ-mono_ 127 | ; _*f-mono_ = _*Δ-mono_ 128 | ; ≤s-isPoset = R.isPoset 129 | ; ≤f-isPoset = isPoset 130 | ; isSemimodule = record 131 | { +*s-isSemiring = R.isSemiring 132 | ; +f-isCommutativeMonoid = isCommutativeMonoid 133 | ; _*f-cong_ = _*Δ-cong_ 134 | ; annihil = *Δempty , e0*Δ 135 | ; distrib = *Δ-distrib-+Δ , (λ x y z → *Δ-distrib-+ z x y) 136 | ; assoc = assoc 137 | ; identity = fst *Δ-identity 138 | } 139 | } 140 | } 141 | where 142 | open CommutativePomonoid (commutativePomonoid n) 143 | using (isPoset; isCommutativeMonoid) 144 | 145 | assoc : ∀ {n} x y (zs : RCtx n) → (x R.* y) *Δ zs ≈ x *Δ (y *Δ zs) 146 | assoc x y nil = nil 147 | assoc x y (z :: zs) = R.*-assoc x y z :: assoc x y zs 148 | 149 | antisym : ∀ {n} → Antisym (RCtx-setoid n) _≤_ 150 | antisym nil nil = nil 151 | antisym (≤R :: ≤Δ) (≥R :: ≥Δ) = R.antisym ≤R ≥R :: antisym ≤Δ ≥Δ 152 | 153 | -- Some stuff depends on n, and some doesn′t. We really want n to be an 154 | -- implicit argument to all of the things that depend on it, and not be there 155 | -- in all of the things that don′t. 156 | 157 | -- Δ complements R, defined at the top of the file. 158 | 159 | module Δ {n : Nat} where 160 | open Posemimodule (posemimodule n) public 161 | using (annihil; distrib; assoc; identity) 162 | renaming (1f to e1; _*f_ to _*_; _*f-cong_ to _*-cong_; 163 | _*f-mono_ to _*-mono_) 164 | open CommutativePomonoid (commutativePomonoid n) public 165 | renaming (e to e0; _·_ to _+_; 166 | _·-cong_ to _+-cong_; _·-mono_ to _+-mono_; 167 | identity to +-identity; assoc to +-assoc; comm to +-comm) 168 | 169 | setoid : Setoid _ _ 170 | setoid = RCtx-setoid n 171 | 172 | open Setoid setoid public 173 | 174 | infixr 1 _≈[_]_ _≤[_]_ 175 | infixr 2 _≈-QED _≤-QED 176 | 177 | _≈[_]_ = _≈[_]Δ_ {n} 178 | _≈-QED = _≈Δ-QED {n} 179 | 180 | _≤[_]_ = _≤[_]Δ_ {n} 181 | _≤-QED = _≤Δ-QED {n} 182 | -} 183 | 184 | --module ≈Δ-Reasoning where 185 | 186 | --module F {n} = Posemimodule (posemimodule n) 187 | -- renaming (_≤f_ to _≤_) 188 | -- hiding (_≤s_, 0s, 1s, _+s_, _*s_, _+s-mono_, _*s-mono_, 189 | -- ≤s-reflexive, ≤s-trans, ≤s-refl, ≤s-antisym, 190 | -- ≤s-preorder, ≤s-poset, ≤s-isPreorder, ≤s-isPoset, 191 | -- +s-cong, +s-identity, +s-assoc, +s-comm, 192 | -- +s-monoid, +s-commutativeMonoid, 193 | -- +s-isMonoid, +s-isCommutativeMonoid, 194 | -- *s-cong, *s-identity, *s-assoc, *s-monoid, *s-isMonoid, 195 | -- +*s-semiring, +*s-isSemiring, ≤+*-posemiring, ≤+*-isPosemiring) 196 | -------------------------------------------------------------------------------- /src/Quantitative/Resources/Context/Semilattice.agda: -------------------------------------------------------------------------------- 1 | open import Lib.Equality 2 | open import Lib.Setoid 3 | open import Lib.Structure 4 | 5 | module Quantitative.Resources.Context.Semilattice 6 | {c l′} (C : Set c) (MSS : MeetSemilatticeSemiring (≡-Setoid C) l′) where 7 | open MeetSemilatticeSemiring MSS --using (posemiring) 8 | 9 | open import Lib.Module 10 | open import Lib.Nat 11 | open import Lib.Product 12 | open import Lib.Vec 13 | open import Lib.VZip 14 | 15 | import Quantitative.Resources.Context as Base 16 | open Base C posemiring using (RCtx) public 17 | open Base C posemiring hiding (RCtx; module R) renaming (module Δ to Δ′) 18 | 19 | module R = MeetSemilatticeSemiring MSS 20 | 21 | meetSemilatticeSemimodule : 22 | ∀ n → MeetSemilatticeSemimodule (≡-Setoid C) (RCtx-setoid n) _ _ 23 | meetSemilatticeSemimodule n = record 24 | { _∧s_ = _∧_ 25 | ; _∧f_ = vzip _∧_ 26 | ; isMeetSemilatticeSemimodule = record 27 | { ∧s-isMeetSemilattice = isMeetSemilattice 28 | ; ∧f-isMeetSemilattice = record 29 | { lowerBound = lowerBoundL , lowerBoundR 30 | ; greatest = greatest′ 31 | ; isPoset = ≤f-isPoset 32 | } 33 | ; isPosemimodule = isPosemimodule 34 | } 35 | } 36 | where 37 | open Posemimodule (posemimodule n) 38 | 39 | lowerBoundL : ∀ {n} (ρs πs : Vec C n) → vzip _∧_ ρs πs Δ′.≤ ρs 40 | lowerBoundL nil nil = nil 41 | lowerBoundL (ρ :: ρs) (π :: πs) = fst lowerBound ρ π :: lowerBoundL ρs πs 42 | 43 | lowerBoundR : ∀ {n} (ρs πs : Vec C n) → vzip _∧_ ρs πs Δ′.≤ πs 44 | lowerBoundR nil nil = nil 45 | lowerBoundR (ρ :: ρs) (π :: πs) = snd lowerBound ρ π :: lowerBoundR ρs πs 46 | 47 | greatest′ : ∀ {n} {ρs πs μs : Vec C n} → 48 | μs Δ′.≤ ρs → μs Δ′.≤ πs → μs Δ′.≤ vzip _∧_ ρs πs 49 | greatest′ nil nil = nil 50 | greatest′ (leρ :: subρ) (leπ :: subπ) = 51 | greatest leρ leπ :: greatest′ subρ subπ 52 | 53 | module Δ {n : Nat} where 54 | open Δ′ {n} public 55 | 56 | open MeetSemilatticeSemimodule (meetSemilatticeSemimodule n) public 57 | using () 58 | renaming (_∧f_ to _∧_; ∧f-greatest to greatest; 59 | ∧f-lowerBound to lowerBound) 60 | -------------------------------------------------------------------------------- /src/Quantitative/Resources/Context/ToppedSemilattice.agda: -------------------------------------------------------------------------------- 1 | open import Lib.Equality 2 | open import Lib.Setoid 3 | open import Lib.Structure 4 | 5 | module Quantitative.Resources.Context.ToppedSemilattice 6 | {c l′} (C : Set c) (TMSS : ToppedMeetSemilatticeSemiring (≡-Setoid C) l′) where 7 | open ToppedMeetSemilatticeSemiring TMSS --using (posemiring) 8 | 9 | open import Lib.Module 10 | open import Lib.Nat 11 | open import Lib.Product 12 | open import Lib.Vec 13 | open import Lib.VZip 14 | 15 | import Quantitative.Resources.Context as Base 16 | open Base C posemiring using (RCtx) public 17 | open Base C posemiring hiding (RCtx; module R) renaming (module Δ to Δ′) 18 | 19 | module R = ToppedMeetSemilatticeSemiring TMSS 20 | 21 | toppedMeetSemilatticeSemimodule : 22 | ∀ n → ToppedMeetSemilatticeSemimodule (≡-Setoid C) (RCtx-setoid n) _ _ 23 | toppedMeetSemilatticeSemimodule n = record 24 | { _∧s_ = _∧_ 25 | ; _∧f_ = vzip _∧_ 26 | ; ⊤s = ⊤ 27 | ; ⊤f = replicateVec n ⊤ 28 | ; isToppedMeetSemilatticeSemimodule = record 29 | { ∧s-isToppedMeetSemilattice = record 30 | { top = top 31 | ; isMeetSemilattice = isMeetSemilattice 32 | } 33 | ; ∧f-isToppedMeetSemilattice = record 34 | { top = top′ 35 | ; isMeetSemilattice = record 36 | { lowerBound = lowerBoundL , lowerBoundR 37 | ; greatest = greatest′ 38 | ; isPoset = ≤f-isPoset 39 | } 40 | } 41 | ; isPosemimodule = isPosemimodule 42 | } 43 | } 44 | where 45 | open Posemimodule (posemimodule n) 46 | 47 | top′ : ∀ {n} (ρs : Vec C n) → ρs Δ′.≤ replicateVec n ⊤ 48 | top′ nil = nil 49 | top′ (ρ :: ρs) = top ρ :: top′ ρs 50 | 51 | lowerBoundL : ∀ {n} (ρs πs : Vec C n) → vzip _∧_ ρs πs Δ′.≤ ρs 52 | lowerBoundL nil nil = nil 53 | lowerBoundL (ρ :: ρs) (π :: πs) = fst lowerBound ρ π :: lowerBoundL ρs πs 54 | 55 | lowerBoundR : ∀ {n} (ρs πs : Vec C n) → vzip _∧_ ρs πs Δ′.≤ πs 56 | lowerBoundR nil nil = nil 57 | lowerBoundR (ρ :: ρs) (π :: πs) = snd lowerBound ρ π :: lowerBoundR ρs πs 58 | 59 | greatest′ : ∀ {n} {ρs πs μs : Vec C n} → 60 | μs Δ′.≤ ρs → μs Δ′.≤ πs → μs Δ′.≤ vzip _∧_ ρs πs 61 | greatest′ nil nil = nil 62 | greatest′ (leρ :: subρ) (leπ :: subπ) = 63 | greatest leρ leπ :: greatest′ subρ subπ 64 | 65 | meetSemilatticeSemimodule : 66 | ∀ n → MeetSemilatticeSemimodule (≡-Setoid C) (RCtx-setoid n) _ _ 67 | meetSemilatticeSemimodule n = record 68 | { _∧s_ = _∧_ 69 | ; _∧f_ = vzip _∧_ 70 | ; isMeetSemilatticeSemimodule = record 71 | { ∧s-isMeetSemilattice = isMeetSemilattice 72 | ; ∧f-isMeetSemilattice = record 73 | { lowerBound = lowerBoundL , lowerBoundR 74 | ; greatest = greatest′ 75 | ; isPoset = ≤f-isPoset 76 | } 77 | ; isPosemimodule = isPosemimodule 78 | } 79 | } 80 | where 81 | open Posemimodule (posemimodule n) 82 | 83 | lowerBoundL : ∀ {n} (ρs πs : Vec C n) → vzip _∧_ ρs πs Δ′.≤ ρs 84 | lowerBoundL nil nil = nil 85 | lowerBoundL (ρ :: ρs) (π :: πs) = fst lowerBound ρ π :: lowerBoundL ρs πs 86 | 87 | lowerBoundR : ∀ {n} (ρs πs : Vec C n) → vzip _∧_ ρs πs Δ′.≤ πs 88 | lowerBoundR nil nil = nil 89 | lowerBoundR (ρ :: ρs) (π :: πs) = snd lowerBound ρ π :: lowerBoundR ρs πs 90 | 91 | greatest′ : ∀ {n} {ρs πs μs : Vec C n} → 92 | μs Δ′.≤ ρs → μs Δ′.≤ πs → μs Δ′.≤ vzip _∧_ ρs πs 93 | greatest′ nil nil = nil 94 | greatest′ (leρ :: subρ) (leπ :: subπ) = 95 | greatest leρ leπ :: greatest′ subρ subπ 96 | 97 | module Δ {n : Nat} where 98 | open Δ′ {n} public 99 | 100 | open ToppedMeetSemilatticeSemimodule (toppedMeetSemilatticeSemimodule n) public 101 | using () 102 | renaming (_∧f_ to _∧_; ∧f-greatest to greatest; 103 | ∧f-lowerBound to lowerBound; ⊤f to ⊤; ∧f-top to top) 104 | -------------------------------------------------------------------------------- /src/Quantitative/Resources/Reduction.agda: -------------------------------------------------------------------------------- 1 | open import Lib.Equality 2 | open import Lib.Setoid 3 | open import Lib.Structure 4 | 5 | module Quantitative.Resources.Reduction 6 | {c l′} (C : Set c) (POS : Posemiring (≡-Setoid C) l′) where 7 | 8 | open import Quantitative.Types.Formers C 9 | open import Quantitative.Syntax Ty 10 | open import Quantitative.Syntax.Substitution Ty 11 | open import Quantitative.Syntax.Reduction C 12 | open import Quantitative.Types C 13 | open import Quantitative.Types.Substitution C 14 | open import Quantitative.Types.Reduction C 15 | open import Quantitative.Resources C POS 16 | open import Quantitative.Resources.Context C POS 17 | open import Quantitative.Resources.Substitution C POS 18 | 19 | open import Lib.Nat 20 | open import Lib.Product 21 | open import Lib.Sum 22 | open import Lib.Two 23 | open import Lib.Vec 24 | open import Lib.VZip 25 | 26 | ~~>-preservesRes : ∀ {n d Γ S Δ} {t u : Term n d} 27 | {tt : Γ ⊢t t :-: S} (tr : Δ ⊢r tt) → 28 | (red : t ~~> u) → Δ ⊢r ~~>-preservesTy tt red 29 | ~~>-preservesRes [ the sr ] (upsilon S s) = sr 30 | ~~>-preservesRes {Δ = Δ} 31 | (app {Δe = Δe} {Δs} split (the (lam tr)) sr) 32 | (⊸-beta S T s t) = 33 | the (substituteRes tr (singleSubstRes R.e1 (the {S = S} sr) (split′ sr))) 34 | where 35 | split-eqs : Δe Δ.+ Δs Δ.≈ R.e1 Δ.* Δs Δ.+ Δe 36 | split-eqs = 37 | Δe Δ.+ Δs Δ.≈[ Δ.+-comm Δe Δs ] 38 | Δs Δ.+ Δe Δ.≈[ Δ.sym (Δ.identity Δs) Δ.+-cong Δ.refl ] 39 | R.e1 Δ.* Δs Δ.+ Δe Δ.≈-QED 40 | 41 | split′ : ∀ {s} → Δs ⊢r s → Δ Δ.≤ R.e1 Δ.* Δs Δ.+ Δe 42 | split′ sr = Δ.≤-trans split (Δ.≤-reflexive split-eqs) 43 | ~~>-preservesRes (lam sr) (lam-cong s s′ red) = 44 | lam (~~>-preservesRes sr red) 45 | ~~>-preservesRes (app split er sr) (app0-cong e e′ s red) = 46 | app split (~~>-preservesRes er red) sr 47 | ~~>-preservesRes (app split er sr) (app1-cong e s s′ red) = 48 | app split er (~~>-preservesRes sr red) 49 | ~~>-preservesRes {Δ = Δ} 50 | (bm {Δe = ρΔ!} {Δs} split+ (the (bang {Δs = Δ!} split* sr)) tr) 51 | (!-beta S T ρ s t) = 52 | the (substituteRes tr (singleSubstRes ρ (the {S = S} sr) split′)) 53 | where 54 | split′ : Δ Δ.≤ ρ Δ.* Δ! Δ.+ Δs 55 | split′ = 56 | Δ Δ.≤[ split+ ] 57 | ρΔ! Δ.+ Δs Δ.≤[ split* Δ.+-mono Δ.≤-refl ] 58 | ρ Δ.* Δ! Δ.+ Δs Δ.≤-QED 59 | ~~>-preservesRes (bang split sr) (bang-cong s s′ red) = 60 | bang split (~~>-preservesRes sr red) 61 | ~~>-preservesRes (bm split er sr) (bm0-cong S e e′ s red) = 62 | bm split (~~>-preservesRes er red) sr 63 | ~~>-preservesRes (bm split er sr) (bm1-cong S e s s′ red) = 64 | bm split er (~~>-preservesRes sr red) 65 | ~~>-preservesRes {Δ = Δ} (del {Δe = Δe} {Δs} split+ (the (unit split0)) tr) 66 | (⊗1-beta T t) = the (weakenRes split′ tr) 67 | where 68 | split′ : Δ Δ.≤ Δs 69 | split′ = 70 | Δ Δ.≤[ split+ ] 71 | Δe Δ.+ Δs Δ.≤[ split0 Δ.+-mono Δ.≤-refl ] 72 | Δ.e0 Δ.+ Δs Δ.≤[ Δ.≤-reflexive (fst Δ.+-identity Δs) ] 73 | Δs Δ.≤-QED 74 | ~~>-preservesRes (del split er sr) (del0-cong S e e′ s red) = 75 | del split (~~>-preservesRes er red) sr 76 | ~~>-preservesRes (del split er sr) (del1-cong S e s s′ red) = 77 | del split er (~~>-preservesRes sr red) 78 | ~~>-preservesRes {Δ = Δ} 79 | (pm {Δe = Δe} {Δs} split 80 | (the (ten {Δs0 = Δs0} {Δs1} split01 s0r s1r)) tr) 81 | (⊗-beta S0 S1 T s0 s1 t) = 82 | the (substituteRes tr (multiSubstRes vfr split)) 83 | where 84 | split01′ = Δe 85 | Δ.≤[ split01 ] 86 | Δs0 Δ.+ Δs1 87 | Δ.≤[ Δ.≤-reflexive (Δ.sym (Δ.identity _ Δ.+-cong Δ.refl)) ] 88 | R.e1 Δ.* Δs0 Δ.+ Δs1 Δ.≤-QED 89 | 90 | splits1 = Δ.≤-reflexive (Δ.sym ( 91 | R.e1 Δ.* Δs1 Δ.+ Δ.e0 Δ.≈[ snd Δ.+-identity _ ] 92 | R.e1 Δ.* Δs1 Δ.≈[ Δ.identity _ ] 93 | Δs1 Δ.≈-QED)) 94 | 95 | vfr = cons split01′ (the s0r) 96 | (cons splits1 (the s1r) 97 | (nil Δ.≤-refl)) 98 | ~~>-preservesRes (ten split s0r s1r) (ten0-cong s0 s0′ s1 red) = 99 | ten split (~~>-preservesRes s0r red) s1r 100 | ~~>-preservesRes (ten split s0r s1r) (ten1-cong s0 s1 s1′ red) = 101 | ten split s0r (~~>-preservesRes s1r red) 102 | ~~>-preservesRes (pm split er sr) (pm0-cong S e e′ s red) = 103 | pm split (~~>-preservesRes er red) sr 104 | ~~>-preservesRes (pm split er sr) (pm1-cong S e s s′ red) = 105 | pm split er (~~>-preservesRes sr red) 106 | ~~>-preservesRes (proj (the (wth s0r s1r))) (&-beta ttt S0 S1 s0 s1) = 107 | the s0r 108 | ~~>-preservesRes (proj (the (wth s0r s1r))) (&-beta fff S0 S1 s0 s1) = 109 | the s1r 110 | ~~>-preservesRes (wth s0r s1r) (wth0-cong s0 s0′ s1 red) = 111 | wth (~~>-preservesRes s0r red) s1r 112 | ~~>-preservesRes (wth s0r s1r) (wth1-cong s0 s1 s1′ red) = 113 | wth s0r (~~>-preservesRes s1r red) 114 | ~~>-preservesRes (proj er) (proj-cong i e e′ red) = 115 | proj (~~>-preservesRes er red) 116 | ~~>-preservesRes (exf split er) (exf-cong T e e′ red) = 117 | exf split (~~>-preservesRes er red) 118 | ~~>-preservesRes (cse split (the (inj s0r)) t0r t1r) 119 | (⊕-beta0 S0 S1 T s0 t0 t1) = 120 | the (substituteRes t0r 121 | (singleSubstRes R.e1 (the s0r) 122 | (Δ.≤-trans split (Δ.≤-reflexive (Δ.sym (Δ.identity _)) 123 | Δ.+-mono Δ.≤-refl)))) 124 | ~~>-preservesRes (cse split (the (inj s1r)) t0r t1r) 125 | (⊕-beta1 S0 S1 T s1 t0 t1) = 126 | the (substituteRes t1r 127 | (singleSubstRes R.e1 (the s1r) 128 | (Δ.≤-trans split (Δ.≤-reflexive (Δ.sym (Δ.identity _)) 129 | Δ.+-mono Δ.≤-refl)))) 130 | ~~>-preservesRes (inj sr) (inj-cong i s s′ red) = 131 | inj (~~>-preservesRes sr red) 132 | ~~>-preservesRes (cse split er s0r s1r) (cse0-cong T e e′ s0 s1 red) = 133 | cse split (~~>-preservesRes er red) s0r s1r 134 | ~~>-preservesRes (cse split er s0r s1r) (cse1-cong T e s0 s0′ s1 red) = 135 | cse split er (~~>-preservesRes s0r red) s1r 136 | ~~>-preservesRes (cse split er s0r s1r) (cse2-cong T e s0 s1 s1′ red) = 137 | cse split er s0r (~~>-preservesRes s1r red) 138 | -------------------------------------------------------------------------------- /src/Quantitative/Semantics/Setoid.agda: -------------------------------------------------------------------------------- 1 | open import Lib.Setoid 2 | 3 | module Quantitative.Semantics.Setoid 4 | {r c l} (R : Set r) (Base : Setoid c l) where 5 | 6 | open import Quantitative.Types.Formers R 7 | open import Quantitative.Syntax Ty 8 | open import Quantitative.Types R 9 | 10 | open import Lib.Equality 11 | open import Lib.Function 12 | open import Lib.One 13 | open import Lib.Product 14 | open import Lib.Sum 15 | open import Lib.Thinning 16 | open import Lib.Two 17 | open import Lib.Vec 18 | open import Lib.Zero 19 | -------------------------------------------------------------------------------- /src/Quantitative/Semantics/Sets.agda: -------------------------------------------------------------------------------- 1 | import Quantitative.Types.Formers as Form 2 | 3 | module Quantitative.Semantics.Sets 4 | {c k} (PrimTy : Set c) (C : Set c) (open Form PrimTy C) 5 | (Const : Set k) (constTy : Const → Ty) (Base : PrimTy -> Set) where 6 | 7 | open import Quantitative.Syntax Ty Const 8 | open import Quantitative.Types PrimTy C Const constTy 9 | 10 | open import Lib.Equality 11 | open import Lib.List 12 | open import Lib.One 13 | open import Lib.Product 14 | open import Lib.Sum 15 | open import Lib.Thinning 16 | open import Lib.Two 17 | open import Lib.Vec 18 | open import Lib.Vec.Thinning 19 | open import Lib.Zero 20 | 21 | ⟦_⟧T : Ty → Set 22 | ⟦ BASE b ⟧T = Base b 23 | ⟦ ⊗1 ⟧T = One 24 | ⟦ &1 ⟧T = One 25 | ⟦ ⊕0 ⟧T = Zero 26 | ⟦ S ⊸ T ⟧T = ⟦ S ⟧T → ⟦ T ⟧T 27 | ⟦ S ⊗ T ⟧T = ⟦ S ⟧T × ⟦ T ⟧T 28 | ⟦ S & T ⟧T = ⟦ S ⟧T × ⟦ T ⟧T 29 | ⟦ S ⊕ T ⟧T = ⟦ S ⟧T ⊎ ⟦ T ⟧T 30 | ⟦ ! ρ S ⟧T = ⟦ S ⟧T 31 | ⟦ LIST S ⟧T = List ⟦ S ⟧T 32 | 33 | ⟦_⟧Γ : ∀ {n} → TCtx n → Set 34 | ⟦ Γ ⟧Γ = Vec-rec One _×_ (vmap ⟦_⟧T Γ) 35 | 36 | ⟦lookup⟧ : ∀ {n} i {Γ : TCtx n} → ⟦ Γ ⟧Γ → ⟦ lookup′ i Γ ⟧T 37 | ⟦lookup⟧ (os i) {T :: Γ} (t , η) = t 38 | ⟦lookup⟧ (o′ i) {T :: Γ} (t , η) = ⟦lookup⟧ i η 39 | -------------------------------------------------------------------------------- /src/Quantitative/Semantics/Sets/Term.agda: -------------------------------------------------------------------------------- 1 | import Quantitative.Types.Formers as Form 2 | import Quantitative.Semantics.Sets as Sem 3 | 4 | module Quantitative.Semantics.Sets.Term 5 | {c k} (PrimTy : Set c) (C : Set c) (open Form PrimTy C) 6 | (Const : Set k) (constTy : Const → Ty) (Base : PrimTy → Set) 7 | (open Sem PrimTy C Const constTy Base) (⟦const⟧ : ∀ l → ⟦ constTy l ⟧T) where 8 | 9 | open import Quantitative.Syntax Ty Const 10 | open import Quantitative.Types PrimTy C Const constTy 11 | 12 | open import Lib.Equality 13 | open import Lib.Function as F hiding (const) 14 | open import Lib.List as L 15 | open import Lib.One 16 | open import Lib.Product 17 | open import Lib.Sum 18 | open import Lib.Thinning 19 | open import Lib.Two 20 | open import Lib.Vec 21 | open import Lib.Vec.Thinning 22 | open import Lib.Zero 23 | 24 | ⟦_⟧t : ∀ {n d Γ T} {t : Term n d} → Γ ⊢t t :-: T → ⟦ Γ ⟧Γ → ⟦ T ⟧T 25 | ⟦ var {th = i} refl ⟧t = ⟦lookup⟧ i 26 | ⟦ const {l = l} ⟧t η = ⟦const⟧ l 27 | ⟦ app et st ⟧t = ⟦ et ⟧t ⟦ st ⟧t 28 | ⟦ bm et st ⟧t η = ⟦ st ⟧t (⟦ et ⟧t η , η) 29 | ⟦ del et st ⟧t = ⟦ st ⟧t 30 | ⟦ pm et st ⟧t η = 31 | let e , f = ⟦ et ⟧t η in 32 | ⟦ st ⟧t (e , f , η) 33 | ⟦ proj {ttt} et ⟧t = fst o ⟦ et ⟧t 34 | ⟦ proj {fff} et ⟧t = snd o ⟦ et ⟧t 35 | ⟦ exf et ⟧t = Zero-elim o ⟦ et ⟧t 36 | ⟦ cse et s0t s1t ⟧t η with ⟦ et ⟧t η 37 | ... | inl s0 = ⟦ s0t ⟧t (s0 , η) 38 | ... | inr s1 = ⟦ s1t ⟧t (s1 , η) 39 | ⟦ fold et snt sct ⟧t η = 40 | L.fold (⟦ et ⟧t η) _ (⟦ snt ⟧t η) λ h acc → ⟦ sct ⟧t (h , acc , η) 41 | ⟦ the st ⟧t = ⟦ st ⟧t 42 | ⟦ lam st ⟧t η = λ s → ⟦ st ⟧t (s , η) 43 | ⟦ bang st ⟧t = ⟦ st ⟧t 44 | ⟦ unit ⟧t = λ _ → <> 45 | ⟦ ten s0t s1t ⟧t = F.const _,_ ⟦ s0t ⟧t ⟦ s1t ⟧t 46 | ⟦ eat ⟧t = λ _ → <> 47 | ⟦ wth s0t s1t ⟧t = F.const _,_ ⟦ s0t ⟧t ⟦ s1t ⟧t 48 | ⟦ inj {i = ttt} st ⟧t = inl o ⟦ st ⟧t 49 | ⟦ inj {i = fff} st ⟧t = inr o ⟦ st ⟧t 50 | ⟦ nil ⟧t = λ _ → [] 51 | ⟦ cons s0t s1t ⟧t = F.const _∷_ ⟦ s0t ⟧t ⟦ s1t ⟧t 52 | ⟦ [ et ] ⟧t = ⟦ et ⟧t 53 | -------------------------------------------------------------------------------- /src/Quantitative/Semantics/WRel.agda: -------------------------------------------------------------------------------- 1 | import Quantitative.Types.Formers as Form 2 | import Quantitative.Semantics.WRel.Core as Core 3 | 4 | open import Lib.Category 5 | open import Lib.Category.Examples 6 | open import Lib.Level 7 | open import Lib.Relation.Propositional 8 | open import Lib.Setoid 9 | open import Lib.Structure 10 | 11 | module Quantitative.Semantics.WRel 12 | {c k l} (PrimTy : Set c) (C : Set c) (open Form PrimTy C) 13 | (Const : Set k) (constTy : Const → Ty) 14 | (posemiring : Posemiring (≡-Setoid C) l) 15 | (symMonCat : SymmetricMonoidalCategory lzero lzero lzero) 16 | (open SymmetricMonoidalCategory symMonCat renaming (C to W)) 17 | (open Core symMonCat) 18 | (Base : PrimTy -> Set) (BaseR : (b : PrimTy) -> WREL.Obj (Base b)) 19 | (!W : ∀ {A} → C → EndoFunctor (WREL A)) where 20 | 21 | open import Quantitative.Syntax Ty Const renaming ([_] to emb) 22 | open import Quantitative.Types PrimTy C Const constTy renaming ([_] to emb) 23 | open import Quantitative.Resources PrimTy C Const constTy posemiring 24 | renaming ([_] to emb) 25 | open import Quantitative.Resources.Context C Const posemiring 26 | open import Quantitative.Semantics.Sets PrimTy C Const constTy Base 27 | 28 | open import Lib.Equality as ≡ using (_≡_; subst2) 29 | open import Lib.Function as F using (_on_; __) 30 | open import Lib.List as L 31 | open import Lib.Matrix.Setoid (≡-Setoid C) 32 | open import Lib.One 33 | open import Lib.Product 34 | open import Lib.Sum 35 | open import Lib.Sum.Pointwise 36 | open import Lib.Thinning 37 | open import Lib.Two 38 | open import Lib.TypeAlgebra 39 | open import Lib.Vec 40 | 41 | -- Lemmas 42 | 43 | ⊤×⊤-⇒W-⊤ : ∀ {A B} → ⊗W {A} {B} .obj (⊤W , ⊤W) ⇒W ⊤W 44 | ⊤×⊤-⇒W-⊤ .η w (a , b) (a′ , b′) (x , y , wxy , xI , yI) = 45 | wxy >> ⊗ .arr $E (xI , yI) >> I⊗ .to .η I 46 | ⊤×⊤-⇒W-⊤ .square _ = <> 47 | 48 | R-⇒W-⊤∧R : ∀ {A} (R : WREL.Obj A) → R ⇒W ∧W .obj (⊤W , R) 49 | R-⇒W-⊤∧R R .η w a b r = I , w , I⊗ .to-iso w .from , id I , r 50 | R-⇒W-⊤∧R R .square _ = <> 51 | 52 | 1-⇒W-1∧1 : 1W ⇒W ∧W .obj (1W , 1W) 53 | 1-⇒W-1∧1 = R-⇒W-⊤∧R 1W 54 | 55 | ∧⊗∧-⇒W-⊗∧⊗ : ∀ {A B} R S T U → 56 | ⊗W {A} {B} .obj (∧W .obj (R , S) , ∧W .obj (T , U)) ⇒W 57 | ∧W .obj (⊗W .obj (R , T) , ⊗W .obj (S , U)) 58 | ∧⊗∧-⇒W-⊗∧⊗ R S T U .η w (a , b) (a′ , b′) (w↑ , w↓ , w↕ 59 | , (w↖ , w↗ , w↑↔ , r , s) 60 | , (w↙ , w↘ , w↓↔ , t , u)) = 61 | ⊗ .obj (w↖ , w↙) , ⊗ .obj (w↗ , w↘) 62 | , w↕ >> ⊗ .arr $E (w↑↔ , w↓↔) 63 | >> ⊗⊗ .to .η _ 64 | >> ⊗ .arr $E (id _ , ⊗⊗ .to-iso _ .from) 65 | >> ⊗ .arr $E (id _ , ⊗ .arr $E (braid .to .η _ , id _)) 66 | >> ⊗ .arr $E (id _ , ⊗⊗ .to .η _) 67 | >> ⊗⊗ .to-iso _ .from 68 | , (w↖ , w↙ , id _ , r , t) 69 | , (w↗ , w↘ , id _ , s , u) 70 | ∧⊗∧-⇒W-⊗∧⊗ R S T U .square _ = <> 71 | 72 | evalW : ∀ {A B C} (f : A → B → C) (g : A → B) 73 | (R : WREL.Obj B) (S : WREL.Obj C) → 74 | ∧W .obj (mapW f (→W .obj (R , S)) , mapW g R) ⇒W mapW (f g) S 75 | evalW f g R S .η w a b (x , y , wxy , rs , r) = rs w y wxy (g a) (g b) r 76 | evalW f g R S .square _ = <> 77 | 78 | curryW : ∀ {A B C} (R : WREL.Obj A) (S : WREL.Obj B) (T : WREL.Obj C) 79 | (f : A → B → C) → 80 | ⊗W .obj (R , S) [ uncurry f ]⇒W T → R [ f ]⇒W →W .obj (S , T) 81 | curryW R S T f α .η x g g′ r w y wxy a a′ s = 82 | α .η w (g , a) (g′ , a′) (x , y , wxy , r , s) 83 | curryW R S T f α .square _ = <> 84 | 85 | ⊤-⇒W-1 : ∀ A → ⊤W {A} [ (λ _ → <>) ]⇒W 1W 86 | ⊤-⇒W-1 A .η w a a′ wI = wI 87 | ⊤-⇒W-1 A .square _ = <> 88 | 89 | 1⊗-⇒W : ∀ {B} (S : WREL.Obj B) → ⊗W .obj (1W , S) [ snd ]⇒W S 90 | 1⊗-⇒W S .η w (<> , b) (<> , b′) (x , y , wxy , xI , bb) = 91 | (S .arr $E (wxy >> ⊗ .arr $E (xI , id y) >> I⊗ .to .η y)) b b′ bb 92 | 1⊗-⇒W S .square _ = <> 93 | 94 | ⊗1-⇒W : ∀ {A} (R : WREL.Obj A) → ⊗W .obj (R , 1W) [ fst ]⇒W R 95 | ⊗1-⇒W R .η w (a , <>) (a′ , <>) (x , y , wxy , aa , yI) = 96 | (R .arr $E (wxy >> ⊗ .arr $E (id x , yI) >> ⊗I .to .η x)) a a′ aa 97 | ⊗1-⇒W R .square _ = <> 98 | 99 | ⊗⊗-⇒W : ∀ {A B C} (R : WREL.Obj A) (S : WREL.Obj B) (T : WREL.Obj C) → 100 | ⊗W .obj (⊗W .obj (R , S) , T) 101 | [ unassoc ]⇒W ⊗W .obj (R , ⊗W .obj (S , T)) 102 | ⊗⊗-⇒W R S T .η w ((a , b) , c) ((a′ , b′) , c′) 103 | (x , y , wxy , (xx , xy , xxxxy , r , s) , t) = 104 | xx , ⊗ .obj (xy , y) , wxy >> ⊗ .arr $E (xxxxy , id y) >> ⊗⊗ .to .η _ , r 105 | , (xy , y , id _ , s , t) 106 | ⊗⊗-⇒W R S T .square _ = <> 107 | 108 | ⊗W-swap : ∀ {A B} (R : WREL.Obj A) (S : WREL.Obj B) → 109 | ⊗W .obj (R , S) [ swap ]⇒W ⊗W .obj (S , R) 110 | ⊗W-swap R S .η w (a , b) (a′ , b′) (x , y , wxy , r , s) = 111 | y , x , wxy >> braid .to .η _ , s , r 112 | ⊗W-swap R S .square _ = <> 113 | 114 | ∧-⇒W-⊗ : ∀ {A B C} (f : A → B) (g : A → C) (R : WREL.Obj B) (S : WREL.Obj C) → 115 | ∧W .obj (mapW f R , mapW g S) [ < f , g > ]⇒W ⊗W .obj (R , S) 116 | ∧-⇒W-⊗ f g R S = idN (∧W .obj (mapW f R , mapW g S)) 117 | 118 | -- ⊗-⇒W-∧ : ∀ {A B C} (f : C → A × B) (g : A × B → C) (R : WREL.Obj A) (S : WREL.Obj B) → 119 | -- ⊗W .obj (R , S) [ g ]⇒W ∧W .obj (mapW (f F.>> fst) R , mapW (f F.>> snd) S) 120 | -- ⊗-⇒W-∧ f g R S = {!mapW->>!} 121 | 122 | caseW : ∀ {A B C} (R : WREL.Obj A) (S : WREL.Obj B) (T : WREL.Obj C) 123 | (f : A → C) (g : B → C) → 124 | R [ f ]⇒W T → S [ g ]⇒W T → ⊕W .obj (R , S) [ [ f , g ] ]⇒W T 125 | caseW R S T f g ρ σ .η w (inl a) (inl a′) (inl r) = ρ .η w a a′ r 126 | caseW R S T f g ρ σ .η w (inr b) (inr b′) (inr s) = σ .η w b b′ s 127 | caseW R S T f g ρ σ .square _ = <> 128 | 129 | projW : ∀ {A B C} (f : A → B × C) (R : WREL.Obj B) (S : WREL.Obj C) i → 130 | mapW f (&W .obj (R , S)) ⇒W Two-rec (mapW (f F.>> fst) R) 131 | (mapW (f F.>> snd) S) 132 | i 133 | projW f R S ttt = record { η = λ w a b → fst ; square = λ _ → <> } 134 | projW f R S fff = record { η = λ w a b → snd ; square = λ _ → <> } 135 | 136 | ⊗-⊕W-distrib-l : 137 | ∀ {A B C} (R : WREL.Obj A) (S : WREL.Obj B) (T : WREL.Obj C) → 138 | ⊗W .obj (⊕W .obj (R , S) , T) [ ×-⊎-distrib-l ]⇒W 139 | ⊕W .obj (⊗W .obj (R , T) , ⊗W .obj (S , T)) 140 | ⊗-⊕W-distrib-l R S T .η w (a , b) (a′ , b′) (x , y , wxy , inl r , t) = 141 | inl (x , y , wxy , r , t) 142 | ⊗-⊕W-distrib-l R S T .η w (a , b) (a′ , b′) (x , y , wxy , inr s , t) = 143 | inr (x , y , wxy , s , t) 144 | ⊗-⊕W-distrib-l R S T .square _ = <> 145 | 146 | foldW : ∀ {A B} (R : WREL.Obj A) (S : WREL.Obj B) 147 | (fn : B) (fc : A → B → B) → 148 | 1W [ F.const fn ]⇒W S → ⊗W .obj (R , S) [ uncurry fc ]⇒W S → 149 | ListW .obj R [ (λ xs → L.fold xs _ fn fc) ]⇒W S 150 | foldW R S fn fc σn σc .η w [] [] (nil wI) = σn .η w <> <> wI 151 | foldW R S fn fc σn σc .η w (x ∷ xs) (y ∷ ys) (cons a b wab r rs) = 152 | σc .η w _ _ (a , b , wab , r , foldW R S fn fc σn σc .η b xs ys rs) 153 | foldW R S fn fc σn σc .square _ = <> 154 | 155 | -- Semantics of types 156 | 157 | R⟦_⟧T : (T : Ty) → WREL.Obj ⟦ T ⟧T 158 | R⟦_,_⟧ρ : ∀ T ρ → WREL.Obj ⟦ T ⟧T 159 | 160 | R⟦ BASE b ⟧T = BaseR b 161 | R⟦ ⊗1 ⟧T = 1W 162 | R⟦ &1 ⟧T = &1W 163 | R⟦ ⊕0 ⟧T = record 164 | { obj = λ _ () 165 | ; arr = →E-⊤ _ λ _ () 166 | ; isFunctor = record { arr-id = λ _ → <> ; arr->> = <> } 167 | } 168 | R⟦ S ⊸ T ⟧T = →W .obj (R⟦ S ⟧T , R⟦ T ⟧T) 169 | R⟦ S ⊗ T ⟧T = ⊗W .obj (R⟦ S ⟧T , R⟦ T ⟧T) 170 | R⟦ S & T ⟧T = &W .obj (R⟦ S ⟧T , R⟦ T ⟧T) 171 | R⟦ S ⊕ T ⟧T = ⊕W .obj (R⟦ S ⟧T , R⟦ T ⟧T) 172 | R⟦ ! ρ S ⟧T = R⟦ S , ρ ⟧ρ 173 | R⟦ LIST S ⟧T = ListW .obj R⟦ S ⟧T 174 | 175 | R⟦ T , ρ ⟧ρ = !W ρ .obj R⟦ T ⟧T 176 | 177 | R⟦_,_⟧Δ : ∀ {n} (Γ : TCtx n) (Δ : RCtx n) → WREL.Obj ⟦ Γ ⟧Γ 178 | R⟦ nil , _ ⟧Δ = 1W 179 | R⟦ T :: Γ , Δρ ⟧Δ = 180 | let ρ = Δρ .get (zeroth , zeroth) in 181 | let Δ = remove-row $E Δρ in 182 | ⊗W .obj (R⟦ T , ρ ⟧ρ , R⟦ Γ , Δ ⟧Δ) 183 | -------------------------------------------------------------------------------- /src/Quantitative/Semantics/WRel/Bang.agda: -------------------------------------------------------------------------------- 1 | open import Lib.Category 2 | open import Lib.Category.Examples 3 | open import Lib.Level 4 | open import Lib.Setoid 5 | open import Lib.Structure 6 | 7 | module Quantitative.Semantics.WRel.Bang 8 | {c l} (C : Set c) (posemiring : Posemiring (≡-Setoid C) l) 9 | (symMonCat : SymmetricMonoidalCategory lzero lzero lzero) where 10 | 11 | open import Lib.Product 12 | 13 | open SymmetricMonoidalCategory symMonCat renaming (C to W) 14 | open import Quantitative.Semantics.WRel.Core symMonCat 15 | 16 | open Posemiring posemiring 17 | 18 | record IsAct (act : {A : Set} → C → WREL.Obj A → WREL.Obj A) 19 | : Set (lsuc lzero ⊔ c ⊔ l) where 20 | field 21 | act-≤ : ∀ {A π ρ} → π ≤ ρ → ∀ C → act {A} π C ⇒W act ρ C 22 | act-0 : ∀ {A} C → act {A} e0 C ⇒W ⊤W 23 | act-+ : ∀ {A} π ρ C → act {A} (π + ρ) C ⇒W ∧W .obj (act π C , act ρ C) 24 | act-1 : ∀ {A} C → act {A} e1 C ⇔W C 25 | act-* : ∀ {A} π ρ C → act {A} (ρ * π) C ⇒W act π (act ρ C) 26 | act-1W : ∀ ρ → 1W ⇒W act ρ 1W 27 | act-⊗W : ∀ {A B} ρ C S → ⊗W .obj (act {A} ρ C , act {B} ρ S) ⇒W 28 | act ρ (⊗W .obj (C , S)) 29 | act-mapW : ∀ {A B} ρ (f : A → B) (C : WREL.Obj B) → 30 | NatTrans (act ρ (mapW f C)) (mapW f (act ρ C)) 31 | -------------------------------------------------------------------------------- /src/Quantitative/Semantics/WRel/Core.agda: -------------------------------------------------------------------------------- 1 | open import Lib.Category 2 | open import Lib.Category.Examples 3 | open import Lib.Level 4 | open import Lib.Relation.Propositional 5 | open import Lib.Setoid 6 | 7 | module Quantitative.Semantics.WRel.Core 8 | (symMonCat : SymmetricMonoidalCategory lzero lzero lzero) where 9 | 10 | open import Lib.Equality as ≡ 11 | open import Lib.Function as F using (_on_) 12 | open import Lib.List as L 13 | open import Lib.One 14 | open import Lib.Product 15 | open import Lib.Sum 16 | open import Lib.Sum.Pointwise 17 | 18 | open SymmetricMonoidalCategory symMonCat renaming (C to W) 19 | module W = Category W 20 | WREL = λ (A : Set) → FUNCTOR (OP W) (REL (≡-Setoid A) lzero) 21 | module WREL A = Category (WREL A) 22 | 23 | infixr 3 _⇒W_ _⇔W_ 24 | 25 | _⇒W_ : ∀ {A} (R S : WREL.Obj A) → Set _ 26 | _⇒W_ {A} R S = WREL._=>_ A R S 27 | 28 | _⇔W_ : ∀ {A} (R S : WREL.Obj A) → Set _ 29 | R ⇔W S = R ⇒W S × S ⇒W R 30 | 31 | mapW : ∀ {A B} → (A → B) → WREL.Obj B → WREL.Obj A 32 | mapW f R = record 33 | { obj = λ w → R .obj w on f 34 | ; arr = →E-⊤ _ λ ww → (R .arr $E ww) on f 35 | ; isFunctor = record { arr-id = λ _ → <> ; arr->> = <> } 36 | } 37 | 38 | mapW-func : ∀ {A B R S} (f : A → B) → R ⇒W S → mapW f R ⇒W mapW f S 39 | mapW-func f rs = record 40 | { η = λ w a b fr → rs .η w (f a) (f b) fr 41 | ; square = λ _ → <> 42 | } 43 | 44 | mapW-subst : ∀ {A B} R {f g : A → B} → f ≡E g → mapW f R ⇒W mapW g R 45 | mapW-subst R fg = record 46 | { η = λ w a b r → 47 | subst2 (R .obj w) (fg (≡.refl {x = a})) (fg (≡.refl {x = b})) r 48 | ; square = λ _ → <> 49 | } 50 | 51 | mapW-id : ∀ {A} {R : WREL.Obj A} → mapW F.id R ⇒W R 52 | mapW-id = idN _ 53 | 54 | mapW->> : ∀ {A B C R} (f : A → B) (g : B → C) → 55 | mapW (f F.>> g) R ⇒W mapW f (mapW g R) 56 | mapW->> f g = idN _ 57 | 58 | _[_]⇒W_ : ∀ {A B} (R : WREL.Obj A) (f : A → B) (S : WREL.Obj B) → Set _ 59 | R [ f ]⇒W S = R ⇒W mapW f S 60 | 61 | _[_]⇒W′_ : ∀ {A B} (R : WREL.Obj A) (f : B → A) (S : WREL.Obj B) → Set _ 62 | R [ f ]⇒W′ S = WREL._=>_ _ (mapW f R) S 63 | 64 | infixr 5 _>>W_ _>>W′_ 65 | 66 | _>>W_ : ∀ {A B C} {R : WREL.Obj A} {S : WREL.Obj B} {T : WREL.Obj C} 67 | {f : A → B} {g : B → C} → 68 | R [ f ]⇒W S → S [ g ]⇒W T → R [ f F.>> g ]⇒W T 69 | _>>W_ {f = f} {g} rs st = rs >>N mapW-func f st 70 | 71 | _>>W′_ : ∀ {A B} {R : WREL.Obj A} {S T : WREL.Obj B} 72 | {f : A → B} → R [ f ]⇒W S → S ⇒W T → R [ f ]⇒W T 73 | rs >>W′ st = rs >>N mapW-func _ st 74 | 75 | 1W : WREL.Obj One 76 | 1W .obj w <> <> = w => I 77 | 1W .arr = →E-⊤ _ λ f _ _ aI → f >> aI 78 | 1W .isFunctor .arr-id _ = <> 79 | 1W .isFunctor .arr->> = <> 80 | 81 | ⊗W : ∀ {A B} → Functor (WREL A ×C WREL B) (WREL (A × B)) 82 | ⊗W .obj (R , S) .obj w (a , b) (a′ , b′) = 83 | ∃2 λ x y → w => ⊗ .obj (x , y) × R .obj x a a′ × S .obj y b b′ 84 | ⊗W .obj (R , S) .arr = →E-⊤ _ λ where 85 | ww (a , b) (a′ , b′) (x , y , wxy , r , s) → x , y , ww >> wxy , r , s 86 | ⊗W .obj (R , S) .isFunctor .arr-id _ = <> 87 | ⊗W .obj (R , S) .isFunctor .arr->> = <> 88 | ⊗W .arr ._$E_ (ρ , σ) .η w (a , b) (a′ , b′) (x , y , wxy , r , s) = 89 | x , y , wxy , ρ .η x a a′ r , σ .η y b b′ s 90 | ⊗W .arr ._$E_ (ηr , ηs) .square _ = <> 91 | ⊗W .arr ._$E=_ _ _ = <> 92 | ⊗W .isFunctor .arr-id _ _ = <> 93 | ⊗W .isFunctor .arr->> _ = <> 94 | 95 | ⊤W : ∀ {A} → WREL.Obj A 96 | ⊤W .obj w _ _ = 1W .obj w <> <> 97 | ⊤W .arr = →E-⊤ _ λ yx _ _ xI → yx >> xI 98 | ⊤W .isFunctor .arr-id _ = <> 99 | ⊤W .isFunctor .arr->> = <> 100 | 101 | diagW : ∀ {A} → Functor (WREL (A × A)) (WREL A) 102 | diagW {A} .obj R = R >>F diag {A = ≡-Setoid A} 103 | diagW .arr ._$E_ ρ .η w a b ab = ρ .η w (a , a) (b , b) ab 104 | diagW .arr ._$E_ ρ .square _ = <> 105 | diagW .arr ._$E=_ _ _ = <> 106 | diagW .isFunctor .arr-id _ _ = <> 107 | diagW .isFunctor .arr->> _ = <> 108 | 109 | ∧W : ∀ {A} → Functor (WREL A ×C WREL A) (WREL A) 110 | ∧W = ⊗W >>F diagW 111 | 112 | →W : ∀ {A B} → Functor (OP (WREL A) ×C WREL B) (WREL (A → B)) 113 | →W .obj (R , S) .obj w f f′ = 114 | ∀ x y → x => ⊗ .obj (w , y) → 115 | ∀ a a′ → R .obj y a a′ → S .obj x (f a) (f′ a′) 116 | →W .obj (R , S) .arr = 117 | →E-⊤ _ λ ww f f′ rs x y xwy a a′ r → 118 | rs x y (xwy >> ⊗ .arr $E (ww , id y)) a a′ r 119 | →W .obj (R , S) .isFunctor .arr-id _ = <> 120 | →W .obj (R , S) .isFunctor .arr->> = <> 121 | →W .arr ._$E_ (ρ , σ) .η w f f′ rs x y xwy a a′ r = 122 | σ .η x (f a) (f′ a′) (rs x y xwy a a′ (ρ .η y a a′ r)) 123 | →W .arr ._$E_ (ρ , σ) .square _ = <> 124 | →W .arr ._$E=_ _ _ = <> 125 | →W .isFunctor .arr-id _ _ = <> 126 | →W .isFunctor .arr->> _ = <> 127 | 128 | &1W : WREL.Obj One 129 | &1W = constF (liftR0 One) 130 | 131 | &W : ∀ {A B} → Functor (WREL A ×C WREL B) (WREL (A × B)) 132 | &W .obj (R , S) .obj w = R .obj w ×R S .obj w 133 | &W .obj (R , S) .arr = →E-⊤ _ λ where 134 | ww (a , b) (a′ , b′) (r , s) → (R .arr $E ww) a a′ r , (S .arr $E ww) b b′ s 135 | &W .obj (R , S) .isFunctor .arr-id _ = <> 136 | &W .obj (R , S) .isFunctor .arr->> = <> 137 | &W .arr ._$E_ (ρ , σ) .η w (a , b) (a′ , b′) (r , s) = 138 | ρ .η w a a′ r , σ .η w b b′ s 139 | &W .arr ._$E_ (ρ , σ) .square _ = <> 140 | &W .arr ._$E=_ _ _ = <> 141 | &W .isFunctor .arr-id _ _ = <> 142 | &W .isFunctor .arr->> _ = <> 143 | 144 | ⊕W : ∀ {A B} → Functor (WREL A ×C WREL B) (WREL (A ⊎ B)) 145 | ⊕W .obj (R , S) .obj w = R .obj w ⊎R S .obj w 146 | ⊕W .obj (R , S) .arr = →E-⊤ _ λ where 147 | ww (inl a) (inl a′) (inl r) → inl ((R .arr $E ww) a a′ r) 148 | ww (inr b) (inr b′) (inr s) → inr ((S .arr $E ww) b b′ s) 149 | ⊕W .obj (R , S) .isFunctor .arr-id _ = <> 150 | ⊕W .obj (R , S) .isFunctor .arr->> = <> 151 | ⊕W .arr ._$E_ (ρ , σ) .η w (inl a) (inl a′) (inl r) = inl (ρ .η w a a′ r) 152 | ⊕W .arr ._$E_ (ρ , σ) .η w (inr b) (inr b′) (inr s) = inr (σ .η w b b′ s) 153 | ⊕W .arr ._$E_ (ρ , σ) .square _ = <> 154 | ⊕W .arr ._$E=_ _ _ = <> 155 | ⊕W .isFunctor .arr-id _ _ = <> 156 | ⊕W .isFunctor .arr->> _ = <> 157 | 158 | module ListW-Data {A} (R : WREL.Obj A) where 159 | 160 | data R*-obj (w : W.Obj) : Rel (List A) lzero where 161 | nil : w => I → R*-obj w [] [] 162 | cons : ∀ {x y xs ys} a b → w => ⊗ .obj (a , b) → 163 | R .obj a x y → R*-obj b xs ys → 164 | R*-obj w (x ∷ xs) (y ∷ ys) 165 | 166 | R*-arr : ∀ {u v} (vu : v W.=> u) (xs ys : List A) 167 | (rs : R*-obj u xs ys) → R*-obj v xs ys 168 | R*-arr vu [] [] (nil sp) = nil (vu >> sp) 169 | R*-arr vu (x ∷ xs) (y ∷ ys) (cons a b uab r rs) = 170 | cons a b (vu >> uab) r rs 171 | 172 | R* : WREL.Obj (List A) 173 | R* .obj = R*-obj 174 | R* .arr = →E-⊤ _ R*-arr 175 | R* .isFunctor .arr-id _ = <> 176 | R* .isFunctor .arr->> = <> 177 | 178 | private 179 | open module Implicit {A} {R} = ListW-Data {A} R using (nil; cons) public 180 | 181 | ListW : ∀ {A} → Functor (WREL A) (WREL (List A)) 182 | ListW .obj = ListW-Data.R* 183 | ListW {A} .arr ._$E_ α .η = arr′ α 184 | where 185 | open ListW-Data using (R*; R*-obj) 186 | 187 | arr′ : ∀ {R S : WREL.Obj A} (α : NatTrans R S) w xs ys → 188 | R*-obj R w xs ys → R*-obj S w xs ys 189 | arr′ α w [] [] (nil sp) = nil sp 190 | arr′ α w (x ∷ xs) (y ∷ ys) (cons a b abw r rs) = 191 | cons a b abw (α .η a x y r) (arr′ α b xs ys rs) 192 | ListW .arr ._$E_ α .square _ = <> 193 | ListW .arr ._$E=_ _ _ = <> 194 | ListW .isFunctor .arr-id _ _ = <> 195 | ListW .isFunctor .arr->> _ = <> 196 | -------------------------------------------------------------------------------- /src/Quantitative/Syntax.agda: -------------------------------------------------------------------------------- 1 | module Quantitative.Syntax {c k} (Ty : Set c) (Const : Set k) where 2 | 3 | open import Quantitative.Syntax.Direction 4 | 5 | open import Lib.Dec 6 | open import Lib.Level 7 | open import Lib.Nat 8 | open import Lib.Product 9 | open import Lib.Thinning 10 | open import Lib.Two 11 | 12 | data Term (n : Nat) : Dir → Set (c ⊔ k) where 13 | var : (th : Fin n) → Term n syn 14 | const : (l : Const) → Term n syn 15 | app : (e : Term n syn) (s : Term n chk) → Term n syn 16 | bm : (S : Ty) (e : Term n syn) (s : Term (succ n) chk) → Term n syn 17 | del : (S : Ty) (e : Term n syn) (s : Term n chk) → Term n syn 18 | pm : (S : Ty) (e : Term n syn) (s : Term (2 +N n) chk) → Term n syn 19 | -- no elimination for &1 20 | proj : (i : Two) (e : Term n syn) → Term n syn 21 | exf : (S : Ty) (e : Term n syn) → Term n syn 22 | cse : (S : Ty) (e : Term n syn) (s0 s1 : Term (succ n) chk) → Term n syn 23 | fold : (S : Ty) (e : Term n syn) 24 | (sn : Term n chk) (sc : Term (succ (succ n)) chk) → Term n syn 25 | the : (S : Ty) (s : Term n chk) → Term n syn 26 | 27 | lam : (s : Term (succ n) chk) → Term n chk 28 | bang : (s : Term n chk) → Term n chk 29 | unit : Term n chk 30 | ten : (s0 s1 : Term n chk) → Term n chk 31 | eat : Term n chk 32 | wth : (s0 s1 : Term n chk) → Term n chk 33 | -- no introduction for ⊕0 34 | inj : (i : Two) (s : Term n chk) → Term n chk 35 | nil : Term n chk 36 | cons : (s0 s1 : Term n chk) → Term n chk 37 | [_] : (e : Term n syn) → Term n chk 38 | 39 | var# : ∀ {n} m {less : Auto (m _ 14 | data _~~>_ {n} : ∀ {d} (t u : Term n d) → Set where 15 | upsilon : ∀ S s → [ the S s ] ~~> s 16 | 17 | ⊸-beta : ∀ S T s t → app (the (S ⊸ T) (lam t)) s 18 | ~~> the T (substitute t (singleSubst (the S s))) 19 | lam-cong : ∀ s s′ → s ~~> s′ → lam s ~~> lam s′ 20 | app0-cong : ∀ e e′ s → e ~~> e′ → app e s ~~> app e′ s 21 | app1-cong : ∀ e s s′ → s ~~> s′ → app e s ~~> app e s′ 22 | 23 | !-beta : ∀ S T ρ s t → bm T (the (! ρ S) (bang s)) t 24 | ~~> the T (substitute t (singleSubst (the S s))) 25 | bang-cong : ∀ s s′ → s ~~> s′ → bang s ~~> bang s′ 26 | bm0-cong : ∀ S e e′ s → e ~~> e′ → bm S e s ~~> bm S e′ s 27 | bm1-cong : ∀ S e s s′ → s ~~> s′ → bm S e s ~~> bm S e s′ 28 | 29 | ⊗1-beta : ∀ T t → del T (the ⊗1 unit) t ~~> the T t 30 | del0-cong : ∀ S e e′ s → e ~~> e′ → del S e s ~~> del S e′ s 31 | del1-cong : ∀ S e s s′ → s ~~> s′ → del S e s ~~> del S e s′ 32 | 33 | ⊗-beta : ∀ S0 S1 T s0 s1 t → 34 | pm T (the (S0 ⊗ S1) (ten s0 s1)) t 35 | ~~> the T (substitute t (multiSubst (the S0 s0 :: the S1 s1 :: nil))) 36 | ten0-cong : ∀ s0 s0′ s1 → s0 ~~> s0′ → ten s0 s1 ~~> ten s0′ s1 37 | ten1-cong : ∀ s0 s1 s1′ → s1 ~~> s1′ → ten s0 s1 ~~> ten s0 s1′ 38 | pm0-cong : ∀ S e e′ s → e ~~> e′ → pm S e s ~~> pm S e′ s 39 | pm1-cong : ∀ S e s s′ → s ~~> s′ → pm S e s ~~> pm S e s′ 40 | 41 | -- no &1-beta 42 | 43 | &-beta : ∀ i S0 S1 s0 s1 → 44 | proj i (the (S0 & S1) (wth s0 s1)) 45 | ~~> Two-rec (the S0 s0) (the S1 s1) i 46 | wth0-cong : ∀ s0 s0′ s1 → s0 ~~> s0′ → wth s0 s1 ~~> wth s0′ s1 47 | wth1-cong : ∀ s0 s1 s1′ → s1 ~~> s1′ → wth s0 s1 ~~> wth s0 s1′ 48 | proj-cong : ∀ i e e′ → e ~~> e′ → proj i e ~~> proj i e′ 49 | 50 | -- no ⊕1-beta 51 | exf-cong : ∀ T e e′ → e ~~> e′ → exf T e ~~> exf T e′ 52 | 53 | ⊕-beta0 : ∀ S0 S1 T s0 t0 t1 → 54 | cse T (the (S0 ⊕ S1) (inj ttt s0)) t0 t1 55 | ~~> the T (substitute t0 (singleSubst (the S0 s0))) 56 | ⊕-beta1 : ∀ S0 S1 T s1 t0 t1 → 57 | cse T (the (S0 ⊕ S1) (inj fff s1)) t0 t1 58 | ~~> the T (substitute t1 (singleSubst (the S1 s1))) 59 | inj-cong : ∀ i s s′ → s ~~> s′ → inj i s ~~> inj i s′ 60 | cse0-cong : ∀ T e e′ s0 s1 → e ~~> e′ → cse T e s0 s1 ~~> cse T e′ s0 s1 61 | cse1-cong : ∀ T e s0 s0′ s1 → s0 ~~> s0′ → cse T e s0 s1 ~~> cse T e s0′ s1 62 | cse2-cong : ∀ T e s0 s1 s1′ → s1 ~~> s1′ → cse T e s0 s1 ~~> cse T e s0 s1′ 63 | -------------------------------------------------------------------------------- /src/Quantitative/Syntax/Substitution.agda: -------------------------------------------------------------------------------- 1 | module Quantitative.Syntax.Substitution {c k} (Ty : Set c) (Const : Set k) where 2 | 3 | open import Quantitative.Syntax.Direction 4 | open import Quantitative.Syntax Ty Const 5 | 6 | open import Lib.Function hiding (const) 7 | open import Lib.Level 8 | open import Lib.Nat 9 | open import Lib.Thinning 10 | open import Lib.Two 11 | open import Lib.Vec 12 | 13 | rename : ∀ {m n d} → m ≤ n → Term m d → Term n d 14 | rename th (var i) = var (i ≤-comp th) 15 | rename th (const l) = const l 16 | rename th (app e s) = app (rename th e) (rename th s) 17 | rename th (bm S e s) = bm S (rename th e) (rename (os th) s) 18 | rename th (del S e s) = del S (rename th e) (rename th s) 19 | rename th (pm S e s) = pm S (rename th e) (rename (os (os th)) s) 20 | rename th (proj i e) = proj i (rename th e) 21 | rename th (exf S e) = exf S (rename th e) 22 | rename th (cse S e s0 s1) = 23 | cse S (rename th e) (rename (os th) s0) (rename (os th) s1) 24 | rename th (fold T e sn sc) = 25 | fold T (rename th e) (rename th sn) (rename (os (os th)) sc) 26 | rename th (the S s) = the S (rename th s) 27 | rename th (lam s) = lam (rename (os th) s) 28 | rename th (bang s) = bang (rename th s) 29 | rename th unit = unit 30 | rename th (ten s0 s1) = ten (rename th s0) (rename th s1) 31 | rename th eat = eat 32 | rename th (wth s0 s1) = wth (rename th s0) (rename th s1) 33 | rename th (inj i s) = inj i (rename th s) 34 | rename th nil = nil 35 | rename th (cons s0 s1) = cons (rename th s0) (rename th s1) 36 | rename th [ e ] = [ rename th e ] 37 | 38 | Subst : Nat → Nat → Set (c ⊔ k) 39 | Subst m n = Fin m → Term n syn 40 | 41 | liftSubst : ∀ {m n} → Subst m n → Subst (succ m) (succ n) 42 | liftSubst vf (os i) = var zeroth 43 | liftSubst vf (o′ i) = rename (o′ oi) (vf i) 44 | 45 | liftSubstN : ∀ {m n} l → Subst m n → Subst (l +N m) (l +N n) 46 | liftSubstN zero vf = vf 47 | liftSubstN {m} {n} (succ l) vf = liftSubst (liftSubstN l vf) 48 | 49 | substitute : ∀ {m n d} → Subst m n → Term m d → Term n d 50 | substitute vf (var th) = vf th 51 | substitute vf (const l) = const l 52 | substitute vf (app e s) = app (substitute vf e) (substitute vf s) 53 | substitute vf (bm S e s) = 54 | bm S (substitute vf e) (substitute (liftSubst vf) s) 55 | substitute vf (del S e s) = del S (substitute vf e) (substitute vf s) 56 | substitute vf (pm S e s) = 57 | pm S (substitute vf e) (substitute (liftSubstN 2 vf) s) 58 | substitute vf (proj i e) = proj i (substitute vf e) 59 | substitute vf (exf S e) = exf S (substitute vf e) 60 | substitute vf (cse S e s0 s1) = 61 | cse S (substitute vf e) (substitute (liftSubst vf) s0) 62 | (substitute (liftSubst vf) s1) 63 | substitute vf (fold T e sn sc) = 64 | fold T (substitute vf e) (substitute vf sn) 65 | (substitute (liftSubstN 2 vf) sc) 66 | substitute vf (the S s) = the S (substitute vf s) 67 | substitute vf (lam s) = lam (substitute (liftSubst vf) s) 68 | substitute vf (bang s) = bang (substitute vf s) 69 | substitute vf unit = unit 70 | substitute vf (ten s0 s1) = ten (substitute vf s0) (substitute vf s1) 71 | substitute vf eat = eat 72 | substitute vf (wth s0 s1) = wth (substitute vf s0) (substitute vf s1) 73 | substitute vf (inj i s) = inj i (substitute vf s) 74 | substitute vf nil = nil 75 | substitute vf (cons s0 s1) = cons (substitute vf s0) (substitute vf s1) 76 | substitute vf [ e ] = [ substitute vf e ] 77 | 78 | singleSubst : ∀ {m} → Term m syn → Subst (succ m) m 79 | singleSubst e (os th) = e 80 | singleSubst e (o′ th) = var th 81 | 82 | multiSubst : ∀ {m n} (es : Vec (Term m syn) n) → Subst (n +N m) m 83 | multiSubst nil i = var i 84 | multiSubst (e :: es) (os i) = e 85 | multiSubst (e :: es) (o′ i) = multiSubst es i 86 | -------------------------------------------------------------------------------- /src/Quantitative/Types.agda: -------------------------------------------------------------------------------- 1 | import Quantitative.Types.Formers as Form 2 | 3 | module Quantitative.Types {c k} (PrimTy : Set c) (C : Set c) (open Form PrimTy C) 4 | (Const : Set k) (constTy : Const → Ty) where 5 | 6 | open import Quantitative.Syntax.Direction 7 | open import Quantitative.Syntax Ty Const 8 | 9 | open import Lib.Equality 10 | open import Lib.Function hiding (const) 11 | open import Lib.Nat 12 | open import Lib.Product 13 | open import Lib.Two 14 | open import Lib.Vec 15 | open import Lib.Vec.Thinning 16 | 17 | infix 4 _∈_ _∋_ _:-:_ 18 | infix 3 _⊢t_ 19 | 20 | record Consequent {n d} (t : Term n d) (T : Ty) : Set c where 21 | 22 | _∈_ : ∀ {n} (e : Term n syn) (T : Ty) → Consequent e T 23 | e ∈ T = _ 24 | 25 | _∋_ : ∀ {n} (T : Ty) (s : Term n chk) → Consequent s T 26 | T ∋ s = _ 27 | 28 | _:-:_ : ∀ {n d} (t : Term n d) (T : Ty) → Consequent t T 29 | t :-: T = _ 30 | 31 | TCtx : Nat → Set c 32 | TCtx = Vec Ty 33 | 34 | -- type correctness 35 | data _⊢t_ {n} (Γ : TCtx n) 36 | : ∀ {d} {t : Term n d} {T} → Consequent t T → Set c where 37 | var : ∀ {th T} → 38 | T ≡ lookup′ th Γ 39 | → 40 | Γ ⊢t var th ∈ T 41 | const : ∀ {l} → Γ ⊢t const l ∈ constTy l 42 | app : ∀ {e s S T} 43 | (et : Γ ⊢t e ∈ S ⊸ T) (st : Γ ⊢t S ∋ s) 44 | → 45 | Γ ⊢t app e s ∈ T 46 | bm : ∀ {e s S T ρ} 47 | (et : Γ ⊢t e ∈ ! ρ S) (st : S :: Γ ⊢t T ∋ s) 48 | → 49 | Γ ⊢t bm T e s ∈ T 50 | del : ∀ {e s T} 51 | (et : Γ ⊢t e ∈ ⊗1) (st : Γ ⊢t T ∋ s) 52 | → 53 | Γ ⊢t del T e s ∈ T 54 | pm : ∀ {e s S0 S1 T} 55 | (et : Γ ⊢t e ∈ S0 ⊗ S1) (st : S0 :: S1 :: Γ ⊢t T ∋ s) 56 | → 57 | Γ ⊢t pm T e s ∈ T 58 | proj : ∀ {i e S0 S1} 59 | (et : Γ ⊢t e ∈ S0 & S1) 60 | → 61 | Γ ⊢t proj i e ∈ Two-rec S0 S1 i 62 | exf : ∀ {e T} 63 | (et : Γ ⊢t e ∈ ⊕0) 64 | → 65 | Γ ⊢t exf T e ∈ T 66 | cse : ∀ {e s0 s1 S0 S1 T} 67 | (et : Γ ⊢t e ∈ S0 ⊕ S1) 68 | (s0t : S0 :: Γ ⊢t T ∋ s0) (s1t : S1 :: Γ ⊢t T ∋ s1) 69 | → 70 | Γ ⊢t cse T e s0 s1 ∈ T 71 | fold : ∀ {e sn sc S T} 72 | (et : Γ ⊢t e ∈ LIST S) 73 | (snt : Γ ⊢t T ∋ sn) (sct : S :: T :: Γ ⊢t T ∋ sc) 74 | → 75 | Γ ⊢t fold T e sn sc ∈ T 76 | the : ∀ {S s} 77 | (st : Γ ⊢t S ∋ s) 78 | → 79 | Γ ⊢t the S s ∈ S 80 | 81 | lam : ∀ {s S T} 82 | (st : S :: Γ ⊢t T ∋ s) 83 | → 84 | Γ ⊢t S ⊸ T ∋ lam s 85 | bang : ∀ {s S ρ} 86 | (st : Γ ⊢t S ∋ s) 87 | → 88 | Γ ⊢t ! ρ S ∋ bang s 89 | unit : Γ ⊢t ⊗1 ∋ unit 90 | ten : ∀ {s0 s1 S0 S1} 91 | (s0t : Γ ⊢t S0 ∋ s0) (s1t : Γ ⊢t S1 ∋ s1) 92 | → 93 | Γ ⊢t S0 ⊗ S1 ∋ ten s0 s1 94 | eat : Γ ⊢t &1 ∋ eat 95 | wth : ∀ {s0 s1 S0 S1} 96 | (s0t : Γ ⊢t S0 ∋ s0) (s1t : Γ ⊢t S1 ∋ s1) 97 | → 98 | Γ ⊢t S0 & S1 ∋ wth s0 s1 99 | inj : ∀ {i s S0 S1} 100 | (st : Γ ⊢t Two-rec S0 S1 i ∋ s) 101 | → 102 | Γ ⊢t S0 ⊕ S1 ∋ inj i s 103 | nil : ∀ {S} → Γ ⊢t LIST S ∋ nil 104 | cons : ∀ {s0 s1 S} 105 | (s0t : Γ ⊢t S ∋ s0) (s1t : Γ ⊢t LIST S ∋ s1) 106 | → 107 | Γ ⊢t LIST S ∋ cons s0 s1 108 | [_] : ∀ {e S} 109 | (et : Γ ⊢t e ∈ S) 110 | → 111 | Γ ⊢t S ∋ [ e ] 112 | 113 | TypedTerm : ∀ {n} → Dir → TCtx n → Set _ 114 | TypedTerm {n} d Γ = ∃ λ T → ∃ λ (t : Term n d) → Γ ⊢t t :-: T 115 | -------------------------------------------------------------------------------- /src/Quantitative/Types/Checker.agda: -------------------------------------------------------------------------------- 1 | open import Lib.Dec 2 | open import Lib.Equality 3 | 4 | module Quantitative.Types.Checker 5 | {c} (C : Set c) (_≟_ : (π ρ : C) → Dec (π ≡ ρ)) where 6 | 7 | open import Quantitative.Syntax.Direction 8 | open import Quantitative.Types.Formers C 9 | open import Quantitative.Types.Formers.Dec C _≟_ 10 | open import Quantitative.Syntax Ty 11 | open import Quantitative.Types C 12 | 13 | open import Lib.Function 14 | open import Lib.Product 15 | open import Lib.Two 16 | open import Lib.Vec 17 | 18 | synthUnique : ∀ {n} {Γ : TCtx n} {e : Term n syn} {S S′ : Ty} → 19 | Γ ⊢t e ∈ S → Γ ⊢t e ∈ S′ → S′ ≡ S 20 | synthUnique (var refl) (var refl) = refl 21 | synthUnique (app et st) (app et′ st′) with synthUnique et et′ 22 | ... | refl = refl 23 | synthUnique (bm et st) (bm et′ st′) with synthUnique et et′ 24 | ... | refl = refl 25 | synthUnique (del et st) (del et′ st′) with synthUnique et et′ 26 | ... | refl = refl 27 | synthUnique (pm et st) (pm et′ st′) with synthUnique et et′ 28 | ... | refl = refl 29 | synthUnique (proj {ttt} et) (proj et′) with synthUnique et et′ 30 | ... | refl = refl 31 | synthUnique (proj {fff} et) (proj et′) with synthUnique et et′ 32 | ... | refl = refl 33 | synthUnique (exf et) (exf et′) with synthUnique et et′ 34 | ... | refl = refl 35 | synthUnique (cse et s0t s1t) (cse et′ s0t′ s1t′) with synthUnique et et′ 36 | ... | refl = refl 37 | synthUnique (the st) (the st′) = refl 38 | 39 | synthType : ∀ {n} (Γ : TCtx n) (e : Term n syn) → 40 | Dec (∃ λ S → Γ ⊢t e ∈ S) 41 | checkType : ∀ {n} (Γ : TCtx n) (S : Ty) (s : Term n chk) → 42 | Dec (Γ ⊢t S ∋ s) 43 | 44 | synthType Γ (var th) = yes (lookup th Γ , var refl) 45 | synthType Γ (app e s) with synthType Γ e 46 | ... | no np = no λ { (_ , app et st) → np (_ , et) } 47 | ... | yes (S⊸T , et) with Is⊸? S⊸T 48 | ... | no np = no λ { (_ , app et′ st′) → np (_ , _ , synthUnique et et′) } 49 | ... | yes (S0 , S1 , refl) = 50 | mapDec (λ st → S1 , app et st) inv (checkType Γ S0 s) 51 | where 52 | inv : (∃ λ T′ → Γ ⊢t app e s ∈ T′) → Γ ⊢t S0 ∋ s 53 | inv (T′ , app et′ st′) with synthUnique et et′ 54 | ... | refl = st′ 55 | synthType Γ (bm T e s) with synthType Γ e 56 | ... | no np = no λ { (_ , bm et st) → np (_ , et) } 57 | ... | yes (!ρS , et) with Is!? !ρS 58 | ... | no np = no λ { (_ , bm et′ st′) → np (_ , _ , synthUnique et et′) } 59 | ... | yes (ρ , S , refl) = 60 | mapDec (λ st → _ , bm et st) inv (checkType (S :: Γ) T s) 61 | where 62 | inv : (∃ λ T′ → Γ ⊢t bm T e s ∈ T′) → S :: Γ ⊢t T ∋ s 63 | inv (T′ , bm et′ st′) with synthUnique et et′ 64 | ... | refl = st′ 65 | synthType Γ (del T e s) with synthType Γ e 66 | ... | no np = no λ { (_ , del et st) → np (_ , et) } 67 | ... | yes (S , et) with Is⊗1? S 68 | ... | no np = no λ { (_ , del et′ st′) → np (synthUnique et et′) } 69 | ... | yes refl = 70 | mapDec (λ st → _ , del et st) inv (checkType Γ T s) 71 | where 72 | inv : (∃ λ T′ → Γ ⊢t del T e s ∈ T′) → Γ ⊢t T ∋ s 73 | inv (T′ , del et′ st′) with synthUnique et et′ 74 | ... | refl = st′ 75 | synthType Γ (pm T e s) with synthType Γ e 76 | ... | no np = no λ { (_ , pm et st) → np (_ , et) } 77 | ... | yes (S0⊗S1 , et) with Is⊗? S0⊗S1 78 | ... | no np = no λ { (_ , pm et′ st′) → np (_ , _ , synthUnique et et′) } 79 | ... | yes (S0 , S1 , refl) = 80 | mapDec (λ st → _ , pm et st) inv (checkType (S0 :: S1 :: Γ) T s) 81 | where 82 | inv : (∃ λ T′ → Γ ⊢t pm T e s ∈ T′) → S0 :: S1 :: Γ ⊢t T ∋ s 83 | inv (T′ , pm et′ st′) with synthUnique et et′ 84 | ... | refl = st′ 85 | synthType Γ (proj i e) with synthType Γ e 86 | ... | no np = no λ { (_ , proj et) → np (_ , et) } 87 | ... | yes (S&T , et) = 88 | mapDec (λ { (_ , _ , refl) → _ , proj et }) inv (Is&? S&T) 89 | where 90 | inv : (∃ λ S → Γ ⊢t proj i e ∈ S) → (∃ λ S → ∃ λ T → (S & T) ≡ S&T) 91 | inv (_ , proj et′) = _ , _ , synthUnique et et′ 92 | synthType Γ (exf T e) with synthType Γ e 93 | ... | no np = no λ { (_ , exf et) → np (_ , et) } 94 | ... | yes (S , et) with Is⊕0? S 95 | ... | no np = no λ { (_ , exf et′) → np (synthUnique et et′) } 96 | ... | yes refl = yes (_ , exf et) 97 | synthType Γ (cse T e s0 s1) with synthType Γ e 98 | ... | no np = no λ { (_ , cse et s0t s1t) → np (_ , et) } 99 | ... | yes (S0⊕S1 , et) with Is⊕? S0⊕S1 100 | ... | no np = no λ { (_ , cse et′ s0t′ s1t′) → 101 | np (_ , _ , synthUnique et et′) } 102 | ... | yes (S0 , S1 , refl) = 103 | mapDec (λ { (s0t , s1t) → T , cse et s0t s1t }) inv 104 | (checkType (S0 :: Γ) T s0 ×? checkType (S1 :: Γ) T s1) 105 | where 106 | inv : (∃ λ T′ → Γ ⊢t cse T e s0 s1 ∈ T′) → 107 | (S0 :: Γ ⊢t T ∋ s0) × (S1 :: Γ ⊢t T ∋ s1) 108 | inv (T′ , cse et′ s0t′ s1t′) with synthUnique et et′ 109 | ... | refl = s0t′ , s1t′ 110 | synthType Γ (the T s) = 111 | mapDec (λ st → T , the st) (λ { (_ , the st) → st }) (checkType Γ T s) 112 | 113 | checkType Γ S⊸T (lam s) with Is⊸? S⊸T 114 | ... | no np = no λ { (lam st) → np (_ , _ , refl) } 115 | ... | yes (S , T , refl) = 116 | mapDec lam (λ { (lam st) → st }) (checkType (S :: Γ) T s) 117 | checkType Γ !ρS (bang s) with Is!? !ρS 118 | ... | no np = no λ { (bang st) → np (_ , _ , refl) } 119 | ... | yes (ρ , S , refl) = 120 | mapDec bang (λ { (bang st) → st }) (checkType Γ S s) 121 | checkType Γ S unit with Is⊗1? S 122 | ... | no np = no λ { unit → np refl } 123 | ... | yes refl = yes unit 124 | checkType Γ S⊗T (ten s0 s1) with Is⊗? S⊗T 125 | ... | no np = no λ { (ten s0t s1t) → np (_ , _ , refl) } 126 | ... | yes (S , T , refl) = 127 | mapDec (uncurry ten) (λ { (ten s0t s1t) → s0t , s1t }) 128 | (checkType Γ S s0 ×? checkType Γ T s1) 129 | checkType Γ S eat with Is&1? S 130 | ... | no np = no λ { eat → np refl } 131 | ... | yes refl = yes eat 132 | checkType Γ S&T (wth s0 s1) with Is&? S&T 133 | ... | no np = no λ { (wth s0t s1t) → np (_ , _ , refl) } 134 | ... | yes (S , T , refl) = 135 | mapDec (uncurry wth) (λ { (wth s0t s1t) → s0t , s1t }) 136 | (checkType Γ S s0 ×? checkType Γ T s1) 137 | checkType Γ S⊕T (inj i s) with Is⊕? S⊕T 138 | ... | no np = no λ { (inj st) → np (_ , _ , refl) } 139 | ... | yes (S , T , refl) = 140 | mapDec inj (λ { (inj st) → st }) (checkType Γ (Two-rec S T i) s) 141 | checkType Γ S [ e ] with synthType Γ e 142 | ... | no np = no λ { [ et ] → np (S , et) } 143 | ... | yes (S′ , et) = 144 | mapDec (λ { refl → [ et ] }) (λ { [ et′ ] → synthUnique et et′ }) (S ≟Ty S′) 145 | -------------------------------------------------------------------------------- /src/Quantitative/Types/Formers.agda: -------------------------------------------------------------------------------- 1 | open import Lib.Nat 2 | 3 | module Quantitative.Types.Formers {c} (B : Set c) (C : Set c) where 4 | 5 | open import Lib.Vec 6 | 7 | infixr 30 _⊸_ 8 | infixr 40 _⊕_ 9 | infixr 50 _⊗_ _&_ 10 | data Ty : Set c where 11 | BASE : B -> Ty 12 | ⊗1 &1 ⊕0 : Ty 13 | _⊸_ _⊗_ _&_ _⊕_ : (S T : Ty) → Ty 14 | ! : (ρ : C) (S : Ty) → Ty 15 | LIST : (S : Ty) → Ty 16 | -------------------------------------------------------------------------------- /src/Quantitative/Types/Formers/Dec.agda: -------------------------------------------------------------------------------- 1 | open import Lib.Dec 2 | open import Lib.Equality 3 | 4 | module Quantitative.Types.Formers.Dec 5 | {c} (PrimTy : Set c) (_≟P_ : (p q : PrimTy) → Dec (p ≡ q)) (C : Set c) (_≟_ : (π ρ : C) → Dec (π ≡ ρ)) where 6 | 7 | open import Quantitative.Types.Formers PrimTy C 8 | 9 | open import Lib.Product 10 | 11 | _≟Ty_ : (S S′ : Ty) → Dec (S ≡ S′) 12 | BASE p ≟Ty BASE q with p ≟P q 13 | ... | yes refl = yes refl 14 | ... | no b = no λ { refl → b refl } 15 | BASE _ ≟Ty ⊗1 = no λ () 16 | BASE _ ≟Ty &1 = no λ () 17 | BASE _ ≟Ty ⊕0 = no λ () 18 | BASE _ ≟Ty (S′ ⊸ T′) = no λ () 19 | BASE _ ≟Ty ! ρ′ S′ = no λ () 20 | BASE _ ≟Ty (S′ ⊗ T′) = no λ () 21 | BASE _ ≟Ty (S′ & T′) = no λ () 22 | BASE _ ≟Ty (S′ ⊕ T′) = no λ () 23 | ⊗1 ≟Ty BASE _ = no λ () 24 | ⊗1 ≟Ty ⊗1 = yes refl 25 | ⊗1 ≟Ty &1 = no λ () 26 | ⊗1 ≟Ty ⊕0 = no λ () 27 | ⊗1 ≟Ty (S′ ⊸ T′) = no λ () 28 | ⊗1 ≟Ty ! ρ′ S′ = no λ () 29 | ⊗1 ≟Ty (S′ ⊗ T′) = no λ () 30 | ⊗1 ≟Ty (S′ & T′) = no λ () 31 | ⊗1 ≟Ty (S′ ⊕ T′) = no λ () 32 | &1 ≟Ty BASE _ = no λ () 33 | &1 ≟Ty ⊗1 = no λ () 34 | &1 ≟Ty &1 = yes refl 35 | &1 ≟Ty ⊕0 = no λ () 36 | &1 ≟Ty (S′ ⊸ T′) = no λ () 37 | &1 ≟Ty ! ρ′ S′ = no λ () 38 | &1 ≟Ty (S′ ⊗ T′) = no λ () 39 | &1 ≟Ty (S′ & T′) = no λ () 40 | &1 ≟Ty (S′ ⊕ T′) = no λ () 41 | ⊕0 ≟Ty BASE _ = no λ () 42 | ⊕0 ≟Ty ⊗1 = no λ () 43 | ⊕0 ≟Ty &1 = no λ () 44 | ⊕0 ≟Ty ⊕0 = yes refl 45 | ⊕0 ≟Ty (S′ ⊸ T′) = no λ () 46 | ⊕0 ≟Ty ! ρ′ S′ = no λ () 47 | ⊕0 ≟Ty (S′ ⊗ T′) = no λ () 48 | ⊕0 ≟Ty (S′ & T′) = no λ () 49 | ⊕0 ≟Ty (S′ ⊕ T′) = no λ () 50 | (S ⊸ T) ≟Ty BASE _ = no λ () 51 | (S ⊸ T) ≟Ty ⊗1 = no λ () 52 | (S ⊸ T) ≟Ty &1 = no λ () 53 | (S ⊸ T) ≟Ty ⊕0 = no λ () 54 | (S ⊸ T) ≟Ty (S′ ⊸ T′) = 55 | mapDec (λ { (refl , refl) → refl }) 56 | (λ { refl → (refl , refl) }) 57 | ((S ≟Ty S′) ×? (T ≟Ty T′)) 58 | (S ⊸ T) ≟Ty ! ρ′ S′ = no λ () 59 | (S ⊸ T) ≟Ty (S′ ⊗ T′) = no λ () 60 | (S ⊸ T) ≟Ty (S′ & T′) = no λ () 61 | (S ⊸ T) ≟Ty (S′ ⊕ T′) = no λ () 62 | ! ρ S ≟Ty BASE _ = no λ () 63 | ! ρ S ≟Ty ⊗1 = no λ () 64 | ! ρ S ≟Ty &1 = no λ () 65 | ! ρ S ≟Ty ⊕0 = no λ () 66 | ! ρ S ≟Ty (S′ ⊸ T′) = no λ () 67 | ! ρ S ≟Ty ! ρ′ S′ = 68 | mapDec (λ { (refl , refl) → refl }) 69 | (λ { refl → refl , refl }) 70 | ((ρ ≟ ρ′) ×? (S ≟Ty S′)) 71 | ! ρ S ≟Ty (S′ ⊗ T′) = no λ () 72 | ! ρ S ≟Ty (S′ & T′) = no λ () 73 | ! ρ S ≟Ty (S′ ⊕ T′) = no λ () 74 | (S ⊗ T) ≟Ty BASE _ = no λ () 75 | (S ⊗ T) ≟Ty ⊗1 = no λ () 76 | (S ⊗ T) ≟Ty &1 = no λ () 77 | (S ⊗ T) ≟Ty ⊕0 = no λ () 78 | (S ⊗ T) ≟Ty (S′ ⊸ T′) = no λ () 79 | (S ⊗ T) ≟Ty (S′ ⊗ T′) = 80 | mapDec (λ { (refl , refl) → refl }) 81 | (λ { refl → (refl , refl) }) 82 | ((S ≟Ty S′) ×? (T ≟Ty T′)) 83 | (S ⊗ T) ≟Ty ! ρ S′ = no λ () 84 | (S ⊗ T) ≟Ty (S′ & T′) = no λ () 85 | (S ⊗ T) ≟Ty (S′ ⊕ T′) = no λ () 86 | (S & T) ≟Ty BASE _ = no λ () 87 | (S & T) ≟Ty ⊗1 = no λ () 88 | (S & T) ≟Ty &1 = no λ () 89 | (S & T) ≟Ty ⊕0 = no λ () 90 | (S & T) ≟Ty (S′ ⊸ T′) = no λ () 91 | (S & T) ≟Ty (S′ ⊗ T′) = no λ () 92 | (S & T) ≟Ty (S′ & T′) = 93 | mapDec (λ { (refl , refl) → refl }) 94 | (λ { refl → (refl , refl) }) 95 | ((S ≟Ty S′) ×? (T ≟Ty T′)) 96 | (S & T) ≟Ty ! ρ S′ = no λ () 97 | (S & T) ≟Ty (S′ ⊕ T′) = no λ () 98 | (S ⊕ T) ≟Ty BASE _ = no λ () 99 | (S ⊕ T) ≟Ty ⊗1 = no λ () 100 | (S ⊕ T) ≟Ty &1 = no λ () 101 | (S ⊕ T) ≟Ty ⊕0 = no λ () 102 | (S ⊕ T) ≟Ty (S′ ⊸ T′) = no λ () 103 | (S ⊕ T) ≟Ty (S′ ⊗ T′) = no λ () 104 | (S ⊕ T) ≟Ty (S′ & T′) = no λ () 105 | (S ⊕ T) ≟Ty (S′ ⊕ T′) = 106 | mapDec (λ { (refl , refl) → refl }) 107 | (λ { refl → (refl , refl) }) 108 | ((S ≟Ty S′) ×? (T ≟Ty T′)) 109 | (S ⊕ T) ≟Ty ! ρ S′ = no λ () 110 | 111 | Is⊗1? = ⊗1 ≟Ty_ 112 | Is&1? = &1 ≟Ty_ 113 | Is⊕0? = ⊕0 ≟Ty_ 114 | 115 | Is⊸? : ∀ S → Dec (∃ λ S0 → ∃ λ S1 → S0 ⊸ S1 ≡ S) 116 | Is⊸? (BASE _) = no λ { (_ , _ , ()) } 117 | Is⊸? ⊗1 = no λ { (_ , _ , ()) } 118 | Is⊸? &1 = no λ { (_ , _ , ()) } 119 | Is⊸? ⊕0 = no λ { (_ , _ , ()) } 120 | Is⊸? (S ⊸ T) = yes (S , T , refl) 121 | Is⊸? (S ⊗ T) = no λ { (_ , _ , ()) } 122 | Is⊸? (S & T) = no λ { (_ , _ , ()) } 123 | Is⊸? (S ⊕ T) = no λ { (_ , _ , ()) } 124 | Is⊸? (! ρ S) = no λ { (_ , _ , ()) } 125 | 126 | Is⊗? : ∀ S → Dec (∃ λ S0 → ∃ λ S1 → S0 ⊗ S1 ≡ S) 127 | Is⊗? (BASE _) = no λ { (_ , _ , ()) } 128 | Is⊗? ⊗1 = no λ { (_ , _ , ()) } 129 | Is⊗? &1 = no λ { (_ , _ , ()) } 130 | Is⊗? ⊕0 = no λ { (_ , _ , ()) } 131 | Is⊗? (S ⊸ T) = no λ { (_ , _ , ()) } 132 | Is⊗? (S ⊗ T) = yes (S , T , refl) 133 | Is⊗? (S & T) = no λ { (_ , _ , ()) } 134 | Is⊗? (S ⊕ T) = no λ { (_ , _ , ()) } 135 | Is⊗? (! ρ S) = no λ { (_ , _ , ()) } 136 | 137 | Is&? : ∀ S → Dec (∃ λ S0 → ∃ λ S1 → S0 & S1 ≡ S) 138 | Is&? (BASE _) = no λ { (_ , _ , ()) } 139 | Is&? ⊗1 = no λ { (_ , _ , ()) } 140 | Is&? &1 = no λ { (_ , _ , ()) } 141 | Is&? ⊕0 = no λ { (_ , _ , ()) } 142 | Is&? (S ⊸ T) = no λ { (_ , _ , ()) } 143 | Is&? (S ⊗ T) = no λ { (_ , _ , ()) } 144 | Is&? (S & T) = yes (S , T , refl) 145 | Is&? (S ⊕ T) = no λ { (_ , _ , ()) } 146 | Is&? (! ρ S) = no λ { (_ , _ , ()) } 147 | 148 | Is!? : ∀ S → Dec (∃ λ ρ → ∃ λ S0 → ! ρ S0 ≡ S) 149 | Is!? (BASE _) = no λ { (_ , _ , ()) } 150 | Is!? ⊗1 = no λ { (_ , _ , ()) } 151 | Is!? &1 = no λ { (_ , _ , ()) } 152 | Is!? ⊕0 = no λ { (_ , _ , ()) } 153 | Is!? (S ⊸ T) = no λ { (_ , _ , ()) } 154 | Is!? (S ⊗ T) = no λ { (_ , _ , ()) } 155 | Is!? (S & T) = no λ { (_ , _ , ()) } 156 | Is!? (S ⊕ T) = no λ { (_ , _ , ()) } 157 | Is!? (! ρ S) = yes (ρ , S , refl) 158 | 159 | Is⊕? : ∀ S → Dec (∃ λ S0 → ∃ λ S1 → S0 ⊕ S1 ≡ S) 160 | Is⊕? (BASE _) = no λ { (_ , _ , ()) } 161 | Is⊕? ⊗1 = no λ { (_ , _ , ()) } 162 | Is⊕? &1 = no λ { (_ , _ , ()) } 163 | Is⊕? ⊕0 = no λ { (_ , _ , ()) } 164 | Is⊕? (S ⊸ T) = no λ { (_ , _ , ()) } 165 | Is⊕? (S ⊗ T) = no λ { (_ , _ , ()) } 166 | Is⊕? (S & T) = no λ { (_ , _ , ()) } 167 | Is⊕? (S ⊕ T) = yes (S , T , refl) 168 | Is⊕? (! ρ S) = no λ { (_ , _ , ()) } 169 | -------------------------------------------------------------------------------- /src/Quantitative/Types/Reduction.agda: -------------------------------------------------------------------------------- 1 | module Quantitative.Types.Reduction {c} (C : Set c) where 2 | 3 | open import Quantitative.Syntax.Direction 4 | open import Quantitative.Types.Formers C 5 | open import Quantitative.Syntax Ty 6 | open import Quantitative.Syntax.Substitution Ty 7 | open import Quantitative.Syntax.Reduction C 8 | open import Quantitative.Types C 9 | open import Quantitative.Types.Substitution C 10 | 11 | open import Lib.Two 12 | open import Lib.Vec 13 | open import Lib.VZip 14 | 15 | ~~>-preservesTy : ∀ {n Γ d T} {t u : Term n d} (tt : Γ ⊢t t :-: T) → 16 | t ~~> u → Γ ⊢t u :-: T 17 | ~~>-preservesTy [ the st ] (upsilon S s) = st 18 | ~~>-preservesTy (app (the (lam tt)) st) (⊸-beta S T s t) = 19 | the (substituteTy tt (singleSubstTy (the st))) 20 | ~~>-preservesTy (lam s0t) (lam-cong s0 s1 red) = 21 | lam (~~>-preservesTy s0t red) 22 | ~~>-preservesTy (app et st) (app0-cong e2 e3 s red) = 23 | app (~~>-preservesTy et red) st 24 | ~~>-preservesTy (app et st) (app1-cong e s0 s1 red) = 25 | app et (~~>-preservesTy st red) 26 | ~~>-preservesTy (bm (the (bang st)) tt) (!-beta S T ρ s t) = 27 | the (substituteTy tt (singleSubstTy (the st))) 28 | ~~>-preservesTy (bang st) (bang-cong s s′ red) = 29 | bang (~~>-preservesTy st red) 30 | ~~>-preservesTy (bm et st) (bm0-cong S e e′ s red) = 31 | bm (~~>-preservesTy et red) st 32 | ~~>-preservesTy (bm et st) (bm1-cong S e s s′ red) = 33 | bm et (~~>-preservesTy st red) 34 | ~~>-preservesTy (del (the unit) tt) (⊗1-beta T t) = the tt 35 | ~~>-preservesTy (del et st) (del0-cong T e e′ s red) = 36 | del (~~>-preservesTy et red) st 37 | ~~>-preservesTy (del et st) (del1-cong T e s s′ red) = 38 | del et (~~>-preservesTy st red) 39 | ~~>-preservesTy (pm (the (ten s0t s1t)) tt) (⊗-beta S0 S1 T s0 s1 t) = 40 | the (substituteTy tt (multiSubstTy (the s0t :: the s1t :: nil))) 41 | ~~>-preservesTy (ten s0t s1t) (ten0-cong s0 s0′ s1 red) = 42 | ten (~~>-preservesTy s0t red) s1t 43 | ~~>-preservesTy (ten s0t s1t) (ten1-cong s0 s1 s1′ red) = 44 | ten s0t (~~>-preservesTy s1t red) 45 | ~~>-preservesTy (pm et st) (pm0-cong S e e′ s red) = 46 | pm (~~>-preservesTy et red) st 47 | ~~>-preservesTy (pm et st) (pm1-cong S e s s′ red) = 48 | pm et (~~>-preservesTy st red) 49 | ~~>-preservesTy (proj (the (wth s0t s1t))) (&-beta ttt S0 S1 s0 s1) = the s0t 50 | ~~>-preservesTy (proj (the (wth s0t s1t))) (&-beta fff S0 S1 s0 s1) = the s1t 51 | ~~>-preservesTy (wth s0t s1t) (wth0-cong s0 s0′ s1 red) = 52 | wth (~~>-preservesTy s0t red) s1t 53 | ~~>-preservesTy (wth s0t s1t) (wth1-cong s0 s1 s1′ red) = 54 | wth s0t (~~>-preservesTy s1t red) 55 | ~~>-preservesTy (proj et) (proj-cong i e e′ red) = 56 | proj (~~>-preservesTy et red) 57 | ~~>-preservesTy (exf et) (exf-cong T e e′ red) = exf (~~>-preservesTy et red) 58 | ~~>-preservesTy (cse (the (inj s0t)) t0t t1t) (⊕-beta0 S0 S1 T s0 t0 t1) = 59 | the (substituteTy t0t (singleSubstTy (the s0t))) 60 | ~~>-preservesTy (cse (the (inj s1t)) t0t t1t) (⊕-beta1 S0 S1 T s1 t0 t1) = 61 | the (substituteTy t1t (singleSubstTy (the s1t))) 62 | ~~>-preservesTy (inj st) (inj-cong i s s′ red) = inj (~~>-preservesTy st red) 63 | ~~>-preservesTy (cse et s0t s1t) (cse0-cong T e e′ s0 s1 red) = 64 | cse (~~>-preservesTy et red) s0t s1t 65 | ~~>-preservesTy (cse et s0t s1t) (cse1-cong T e s0 s0′ s1 red) = 66 | cse et (~~>-preservesTy s0t red) s1t 67 | ~~>-preservesTy (cse et s0t s1t) (cse2-cong T e s0 s1 s1′ red) = 68 | cse et s0t (~~>-preservesTy s1t red) 69 | -------------------------------------------------------------------------------- /src/Quantitative/Types/Substitution.agda: -------------------------------------------------------------------------------- 1 | import Quantitative.Types.Formers as Form 2 | 3 | module Quantitative.Types.Substitution 4 | {c k} (PrimTy : Set c) (C : Set c) (open Form PrimTy C) 5 | (Const : Set k) (constTy : Const → Ty) where 6 | 7 | open import Quantitative.Syntax.Direction 8 | open import Quantitative.Syntax Ty Const 9 | open import Quantitative.Syntax.Substitution Ty Const 10 | open import Quantitative.Types PrimTy C Const constTy 11 | 12 | open import Lib.Equality 13 | open import Lib.Function hiding (const) 14 | open import Lib.Thinning hiding (_∈_) 15 | open import Lib.Two 16 | open import Lib.Vec 17 | open import Lib.Vec.Thinning 18 | open import Lib.VZip 19 | 20 | RenTy : ∀ {m n} → m ≤ n → TCtx m → TCtx n → Set c 21 | RenTy {m} {n} th Γm Γn = VZip _≡_ (thin th Γn) Γm 22 | 23 | renameTy : ∀ {m n d T} {t : Term m d} {th : m ≤ n} {Γm Γn} → 24 | RenTy th Γm Γn → Γm ⊢t t :-: T → Γn ⊢t rename th t :-: T 25 | renameTy {th = th} {Γm} {Γn} tht (var {th = i} refl) = var (sym q′) 26 | where 27 | q′ : lookup′ (i ≤-comp th) Γn ≡ lookup′ i Γm 28 | q′ = cong headVec (trans (VZip≡ (thin-comp i th Γn)) 29 | (cong (thin i) (VZip≡ tht))) 30 | renameTy tht const = const 31 | renameTy tht (app et st) = app (renameTy tht et) (renameTy tht st) 32 | renameTy tht (bm et st) = bm (renameTy tht et) (renameTy (refl :: tht) st) 33 | renameTy tht (del et st) = del (renameTy tht et) (renameTy tht st) 34 | renameTy tht (pm et st) = 35 | pm (renameTy tht et) (renameTy (refl :: refl :: tht) st) 36 | renameTy tht (proj et) = proj (renameTy tht et) 37 | renameTy tht (exf et) = exf (renameTy tht et) 38 | renameTy tht (cse et s0t s1t) = 39 | cse (renameTy tht et) (renameTy (refl :: tht) s0t) 40 | (renameTy (refl :: tht) s1t) 41 | renameTy tht (fold et snt sct) = 42 | fold (renameTy tht et) (renameTy tht snt) 43 | (renameTy (refl :: refl :: tht) sct) 44 | renameTy tht (the st) = the (renameTy tht st) 45 | renameTy tht (lam st) = lam (renameTy (refl :: tht) st) 46 | renameTy tht (bang st) = bang (renameTy tht st) 47 | renameTy tht unit = unit 48 | renameTy tht (ten s0t s1t) = ten (renameTy tht s0t) (renameTy tht s1t) 49 | renameTy tht eat = eat 50 | renameTy tht (wth s0t s1t) = wth (renameTy tht s0t) (renameTy tht s1t) 51 | renameTy tht (inj st) = inj (renameTy tht st) 52 | renameTy tht nil = nil 53 | renameTy tht (cons s0t s1t) = cons (renameTy tht s0t) (renameTy tht s1t) 54 | renameTy tht [ et ] = [ renameTy tht et ] 55 | 56 | SubstTy : ∀ {m n} → Subst m n → TCtx m → TCtx n → Set c 57 | SubstTy {m} {n} vf Γm Γn = (th : Fin m) → Γn ⊢t vf th ∈ lookup′ th Γm 58 | 59 | liftSubstTy : ∀ {m n} {vf : Subst m n} {Γm Γn T} → 60 | SubstTy vf Γm Γn → SubstTy (liftSubst vf) (T :: Γm) (T :: Γn) 61 | liftSubstTy vft (os th) = var refl 62 | liftSubstTy vft (o′ th) = renameTy (thin-oi _) (vft th) 63 | 64 | liftSubstNTy : ∀ {m n l} {vf : Subst m n} {Γm Γn} {Γl : TCtx l} → 65 | SubstTy vf Γm Γn → 66 | SubstTy (liftSubstN l vf) (Γl +V Γm) (Γl +V Γn) 67 | liftSubstNTy {Γl = nil} vft = vft 68 | liftSubstNTy {Γl = S :: Γl} vft = liftSubstTy (liftSubstNTy vft) 69 | 70 | substituteTy : 71 | ∀ {m n d} {t : Term m d} {vf : Subst m n} 72 | {Γm Γn T} → Γm ⊢t t :-: T → SubstTy vf Γm Γn → 73 | Γn ⊢t substitute vf t :-: T 74 | substituteTy (var {th = th} refl) vft = vft th 75 | substituteTy const vft = const 76 | substituteTy (app et st) vft = 77 | app (substituteTy et vft) (substituteTy st vft) 78 | substituteTy (bm et st) vft = 79 | bm (substituteTy et vft) 80 | (substituteTy st (liftSubstTy vft)) 81 | substituteTy (del et st) vft = 82 | del (substituteTy et vft) (substituteTy st vft) 83 | substituteTy (pm et st) vft = 84 | pm (substituteTy et vft) 85 | (substituteTy st (liftSubstNTy vft)) 86 | substituteTy (proj et) vft = proj (substituteTy et vft) 87 | substituteTy (exf st) vft = exf (substituteTy st vft) 88 | substituteTy (cse et s0t s1t) vft = 89 | cse (substituteTy et vft) (substituteTy s0t (liftSubstTy vft)) 90 | (substituteTy s1t (liftSubstTy vft)) 91 | substituteTy (fold et snt sct) vft = 92 | fold (substituteTy et vft) (substituteTy snt vft) 93 | (substituteTy sct (liftSubstNTy vft)) 94 | substituteTy (the st) vft = the (substituteTy st vft) 95 | substituteTy (lam st) vft = 96 | lam (substituteTy st (liftSubstTy vft)) 97 | substituteTy (bang st) vft = 98 | bang (substituteTy st vft) 99 | substituteTy unit vft = unit 100 | substituteTy (ten s0 s1) vft = 101 | ten (substituteTy s0 vft) (substituteTy s1 vft) 102 | substituteTy eat vft = eat 103 | substituteTy (wth s0 s1) vft = 104 | wth (substituteTy s0 vft) (substituteTy s1 vft) 105 | substituteTy (inj st) vft = 106 | inj (substituteTy st vft) 107 | substituteTy nil vft = nil 108 | substituteTy (cons s0t s1t) vft = 109 | cons (substituteTy s0t vft) (substituteTy s1t vft) 110 | substituteTy [ et ] vft = [ substituteTy et vft ] 111 | 112 | singleSubstTy : ∀ {m e Γ S} → Γ ⊢t e ∈ S → 113 | SubstTy (singleSubst {m} e) (S :: Γ) Γ 114 | singleSubstTy et (os th) = et 115 | singleSubstTy et (o′ th) = var refl 116 | 117 | multiSubstTy : ∀ {m n} {Γm : TCtx m} {Γn : TCtx n} 118 | {es : Vec (Term m syn) n} 119 | (ets : VZip (λ S e → Γm ⊢t e ∈ S) Γn es) → 120 | SubstTy (multiSubst es) (Γn +V Γm) Γm 121 | multiSubstTy nil i = var refl 122 | multiSubstTy (et :: ets) (os i) = et 123 | multiSubstTy (et :: ets) (o′ i) = multiSubstTy ets i 124 | -------------------------------------------------------------------------------- /src/index.agda: -------------------------------------------------------------------------------- 1 | module index where 2 | import Lib.Level 3 | import Lib.One 4 | import Lib.Product 5 | import Lib.Function 6 | import Lib.Level 7 | import Lib.Zero 8 | import Lib.Equality 9 | import Lib.Sum 10 | import Lib.Two 11 | import Lib.Dec 12 | import Lib.Module 13 | import Lib.Nat 14 | import Lib.Thinning 15 | import Lib.Vec 16 | import Lib.VZip 17 | import Lib.Maybe 18 | import Lib.FunctionProperties 19 | import Lib.Setoid 20 | import Lib.Structure 21 | import Lib.Relation 22 | import Lib.Relation.Propositional 23 | import Lib.Category 24 | import Lib.Category.Examples 25 | 26 | import Quantitative.Syntax 27 | import Quantitative.Syntax.Direction 28 | -- import Quantitative.Syntax.Reduction 29 | import Quantitative.Syntax.Substitution 30 | import Quantitative.Types 31 | -- import Quantitative.Types.Checker 32 | import Quantitative.Types.Formers 33 | import Quantitative.Types.Formers.Dec 34 | -- import Quantitative.Types.Reduction 35 | import Quantitative.Types.Substitution 36 | import Quantitative.Resources 37 | -- import Quantitative.Resources.Checker 38 | import Quantitative.Resources.Context 39 | -- import Quantitative.Resources.Context.Semilattice 40 | -- import Quantitative.Resources.Context.ToppedSemilattice 41 | -- import Quantitative.Resources.Reduction 42 | import Quantitative.Resources.Substitution 43 | -- import Quantitative.Semantics.Relational 44 | -- import Quantitative.Semantics.Setoid 45 | import Quantitative.Semantics.Sets 46 | import Quantitative.Semantics.Sets.Term 47 | import Quantitative.Semantics.WRel 48 | import Quantitative.Semantics.WRel.Term 49 | import Quantitative.Models.RelationTransformer 50 | -- import Quantitative.Models.RelationTransformer.Action 51 | -------------------------------------------------------------------------------- /tex/abstract/bib.bib: -------------------------------------------------------------------------------- 1 | @INPROCEEDINGS{Barendregt92lambdacalculi, 2 | author = {Henk Barendregt}, 3 | title = {Lambda Calculi with Types}, 4 | booktitle = {HANDBOOK OF LOGIC IN COMPUTER SCIENCE}, 5 | year = {1992}, 6 | pages = {117--309}, 7 | publisher = {Oxford University Press} 8 | } 9 | 10 | @Article{deBruijn:dummies, 11 | author = {Nicolas G. de~Bruijn}, 12 | title = "{Lambda Calculus notation with nameless dummies: a tool for automatic formula manipulation}", 13 | journal = {Indagationes Mathematic{\ae}}, 14 | year = {1972}, 15 | volume = {34}, 16 | pages = {381--392}, 17 | } 18 | 19 | @article{DBLP:journals/toplas/PierceT00, 20 | author = {Benjamin C. Pierce and David N. Turner}, 21 | title = {Local type inference}, 22 | journal = {{ACM} Trans. Program. Lang. Syst.}, 23 | volume = {22}, 24 | number = {1}, 25 | pages = {1--44}, 26 | year = {2000}, 27 | url = {http://doi.acm.org/10.1145/345099.345100}, 28 | doi = {10.1145/345099.345100}, 29 | timestamp = {Wed, 26 Nov 2003 14:26:46 +0100}, 30 | biburl = {https://dblp.org/rec/bib/journals/toplas/PierceT00}, 31 | bibsource = {dblp computer science bibliography, https://dblp.org} 32 | } -------------------------------------------------------------------------------- /tex/abstract/examples.tex: -------------------------------------------------------------------------------- 1 | The ingredients of our fundamental lemma are perhaps well known 2 | (relational interpretations, Kripke-indexing), but the value of our 3 | framework lies in the generality of being able to choose $\mathcal{W}$ 4 | and its promonoidal structure, and the interpretion the $\oc_\rho$ 5 | modality as a relation transformer. Examples include: 6 | 7 | \vspace{-0.6em} 8 | 9 | \paragraph{Permutation Types} With the $\{0,1,\omega\}$ semiring, we 10 | take the category $\mathcal{W}$ to consist of lists of some type of 11 | keys, and permutations between them. The relation transformer is 12 | defined as: $\oc_0 R~l = \top$, where $\top$ is the total relation, 13 | $\oc_1 R~l = R~l$ and $R_\omega~R~l = (l = []) \land R~l$. With 14 | suitable types of keys and lists of keys, the fundamental lemma states 15 | that all programs are permutations. This result has already been 16 | formalised in a one-off type system at 17 | \url{https://github.com/bobatkey/sorting-types}. 18 | 19 | \vspace{-0.6em} 20 | 21 | \paragraph{Monotonicity Types} With $R$ the partially ordered semiring 22 | with carrier $\{0,\uparrow,\downarrow,\equiv\}$ ordered 23 | ${\equiv} \leq {\uparrow},{\downarrow}$ and ${\uparrow}, {\downarrow} \leq 0$, 24 | we take $\mathcal{W}$ to be the one-object, one-arrow category, and 25 | define the relation transformer $\oc$ to be: 26 | \begin{mathpar} 27 | \oc_0~R = \top 28 | 29 | \oc_\uparrow~R = R 30 | 31 | \oc_\downarrow~R = R^{op} 32 | 33 | \oc_\equiv~R = R \cap R^{op} 34 | \end{mathpar} 35 | If we let our base type be natural numbers with the relational 36 | interpretation $R_{\mathrm{nat}}(n,n') \Leftrightarrow n \leq n'$, 37 | then the fundamental lemma states that a program of type 38 | $\ctxvar{x}{\mathrm{nat}}{\uparrow} \vdash t : \mathrm{nat}$ is 39 | covariant (and similarly for contravariant and invariant). 40 | 41 | \vspace{-0.6em} 42 | 43 | \paragraph{Sensitivity Analysis} With the 44 | $R = \mathbb{R} \cup \{\infty\}$ semiring, we let $\mathcal{W}$ be $R$ 45 | as well. The relation transformer is given by scaling. With a base 46 | type of real numbers with relational intepretation 47 | $R_{\mathrm{real}}~k~(x,x') \Leftrightarrow |x-x'| \leq k$, then the 48 | fundamental lemma states that the usage annotations on the input 49 | variables tracks the extent to which the program is sensitive to 50 | changes in those variables. 51 | 52 | \vspace{-0.6em} 53 | 54 | \paragraph{Information Flow} With the $R = \mathcal{P}(L)$ semiring, 55 | we again take $\mathcal{W} = R$, and let the relation transformer to 56 | be 57 | $\oc_l R~l' = \{\top~\textrm{when }l \geq l'; R~\textrm{otherwise}\}$. 58 | Then the fundamental lemma yields the same non-interference properties 59 | as stated by Abadi et al. for the DCC \cite{abadi99core}. 60 | 61 | 62 | % Local Variables: 63 | % TeX-master: "quantitative" 64 | % End: 65 | -------------------------------------------------------------------------------- /tex/abstract/intro.tex: -------------------------------------------------------------------------------- 1 | In normal typed $\lambda$-calculi, variables may be used multiple 2 | times, in multiple contexts, for multiple reasons, as long as the 3 | types agree. The disciplines of linear types \cite{girard87linear} and 4 | coeffects \cite{PetricekOM14,BrunelGMZ14,GhicaS14} refine this by 5 | tracking variable usage. We might track how many times a variable is 6 | used, or if it is used co-, contra-, or invariantly. Such a discipline 7 | yields a general framework of ``context constrained computing'', where 8 | constraints on variables in the context tell us something interesting 9 | about the computation being performed. 10 | % Thus we put the 11 | % type information to work to tell us facts about programs that might 12 | % not otherwise be apparent. 13 | 14 | We will present work in progress on capturing the ``intensional'' 15 | properties of programs via a family of Kripke indexed relational 16 | semantics that refines a simple set-theoretic semantics of 17 | programs. The value of our approach lies in its generality. We can 18 | accommodate the following examples: 19 | \begin{enumerate} 20 | \item Linear types that capture properties like ``all list 21 | manipulating programs are permutations''. This example uses the 22 | Kripke-indexing to track the collection of datums currently being 23 | manipulated by the program. 24 | \item Monotonicity coeffects that track whether a program uses inputs 25 | co-, contra-, or in-variantly (or not at all). 26 | \item Sensitivity typing, tracking the sensitivity of programs in 27 | terms of input changes. This forms the core of systems for 28 | differential privacy \cite{reed10distance}. 29 | \item Information flow typing, in the style of the Dependency Core 30 | Calculus \cite{abadi99core}. 31 | \end{enumerate} 32 | Through discussion at the workshop, we hope to discover more 33 | applications of our framework. In future work, we plan to extend our 34 | framework with type dependency, and to explore the space of inductive 35 | data types and elimination principles possible in the presence of 36 | usage information. 37 | 38 | The syntax and semantics we present here have been formalised in Agda: 39 | \url{https://github.com/laMudri/quantitative/}. Formalisation of the 40 | examples is in progress. 41 | 42 | % Local Variables: 43 | % TeX-master: "quantitative" 44 | % End: 45 | -------------------------------------------------------------------------------- /tex/abstract/quantitative.bib: -------------------------------------------------------------------------------- 1 | @inproceedings{BrunelGMZ14, 2 | author = {A. Brunel and 3 | M. Gaboardi and 4 | D. Mazza and 5 | S. Zdancewic}, 6 | title = {{A Core Quantitative Coeffect Calculus}}, 7 | booktitle = {{ESOP} 2014}, 8 | pages = {351--370}, 9 | year = 2014, 10 | timestamp = {Sun, 23 Mar 2014 10:48:25 +0100}, 11 | biburl = {http://dblp.uni-trier.de/rec/bib/conf/esop/BrunelGMZ14}, 12 | bibsource = {dblp computer science bibliography, http://dblp.org} 13 | } 14 | 15 | @inproceedings{GhicaS14, 16 | author = {Dan R. Ghica and 17 | Alex I. Smith}, 18 | title = {Bounded Linear Types in a Resource Semiring}, 19 | booktitle = {{ESOP} 2014}, 20 | pages = {331--350}, 21 | year = {2014}, 22 | timestamp = {Sun, 23 Mar 2014 10:48:25 +0100}, 23 | biburl = {http://dblp.uni-trier.de/rec/bib/conf/esop/GhicaS14}, 24 | bibsource = {dblp computer science bibliography, http://dblp.org} 25 | } 26 | 27 | @inproceedings{PetricekOM14, 28 | author = {Tomas Petricek and 29 | Dominic A. Orchard and 30 | Alan Mycroft}, 31 | title = {Coeffects: a calculus of context-dependent computation}, 32 | booktitle = {{ICFP} 2014}, 33 | pages = {123--135}, 34 | year = {2014}, 35 | timestamp = {Sun, 04 Jun 2017 10:05:10 +0200}, 36 | biburl = {http://dblp.uni-trier.de/rec/bib/conf/icfp/PetricekOM14}, 37 | bibsource = {dblp computer science bibliography, http://dblp.org} 38 | } 39 | 40 | @article{ girard87linear, 41 | author = "Jean-Yves Girard", 42 | title = "Linear Logic", 43 | journal = "Theor. Comp. Sci.", 44 | volume = 50, 45 | pages = "1--101", 46 | year = 1987 } 47 | 48 | @InProceedings{reed10distance, 49 | author = {J. Reed and B. C. Pierce}, 50 | title = {Distance Makes the Types Grow Stronger}, 51 | booktitle = {ICFP 2010}, 52 | year = 2010, 53 | editor = {P. Hudak and S. Weirich}, 54 | pages = {157-168}} 55 | 56 | @inproceedings{abadi99core, 57 | author = {M. Abadi and 58 | A. Banerjee and 59 | N. Heintze and 60 | J. G. Riecke}, 61 | title = {{A Core Calculus of Dependency}}, 62 | booktitle = {POPL '99}, 63 | year = 1999, 64 | pages = {147-160}, 65 | } 66 | 67 | @article{DBLP:journals/toplas/PierceT00, 68 | author = {Benjamin C. Pierce and David N. Turner}, 69 | title = {Local type inference}, 70 | journal = {{ACM} {TOPLAS}}, 71 | volume = {22}, 72 | number = {1}, 73 | pages = {1--44}, 74 | year = {2000}, 75 | timestamp = {Wed, 26 Nov 2003 14:26:46 +0100}, 76 | biburl = {https://dblp.org/rec/bib/journals/toplas/PierceT00}, 77 | bibsource = {dblp computer science bibliography, https://dblp.org} 78 | } -------------------------------------------------------------------------------- /tex/abstract/semantics.tex: -------------------------------------------------------------------------------- 1 | \paragraph{Underlying Semantics} We give a standard semantics of types 2 | and well typed terms into sets and functions. This semantics ignores 3 | the usage information. For types, we have: 4 | \begin{displaymath} 5 | \begin{array}{ll} 6 | \llbracket \base \rrbracket = A_\base \\ 7 | \llbracket \fun{S}{T} \rrbracket = \llbracket S \rrbracket \rightarrow \llbracket T \rrbracket & 8 | \llbracket \excl{\rho}{S} \rrbracket = \llbracket S \rrbracket \\ 9 | \llbracket \tensorOne \rrbracket = \llbracket \withTOne \rrbracket = \{*\} & 10 | \llbracket \tensor{S}{T} \rrbracket = \llbracket \withT{S}{T} \rrbracket = \llbracket S \rrbracket \times \llbracket T \rrbracket \\ 11 | \llbracket \sumTZero \rrbracket = \{\} & 12 | \llbracket \sumT{S}{T} \rrbracket = \llbracket S \rrbracket \uplus \llbracket T \rrbracket \\ 13 | \end{array} 14 | \end{displaymath} 15 | Contexts are interpreted as left-nested products. Terms are assigned 16 | the usual semantics as functions $\sem{t} : \sem{\Gamma} \to \sem{S}$. 17 | 18 | \paragraph{Usage-tracking semantics} To derive interesting properties 19 | from our type system, we refine the set-theoretic semantics by 20 | Kripke-indexed binary relations. This gives a fundamental lemma for 21 | our system, that when instantitated in different ways captures the 22 | examples in the introduction. 23 | 24 | Our framework is parameterised by a category $\mathcal{W}$ of 25 | \emph{possible worlds} that track how resources are distributed by 26 | programs. To interpret resource separation, we assume that 27 | $\mathcal{W}$ has \emph{symmetric promonoidal structure}: profunctors 28 | $J : 1 \tobar \mathcal{W}$ and 29 | $P : \mathcal{W} \times \mathcal{W} \tobar \mathcal{W}$ such 30 | that $P \odot (J \times 1) \cong 1$, $P \odot (1 \times J) \cong 1$, 31 | $P \odot (1 \times P) \cong P \odot (P \times 1)$, and 32 | $P \cong P \odot (\pi_2 \times \pi_1)$, and the triangle, pentagon, 33 | and hexagon laws hold\footnote{We don't need the laws to hold to prove 34 | the fundamental lemma.}. 35 | 36 | We now assign to each type $T$ a functor 37 | $\sem{T}^R : \mathcal{W}^{op} \to \mathrm{Rel}~\sem{T}$ that captures 38 | a notion of $\mathcal{W}$-indexed ``indistinguishability''. To 39 | interpret $\oc_\rho S$, we assume we are given a relation transformer 40 | $\oc_A : R^{op} \to \mathrm{Rel}(A)^{\mathcal{W}^{op}} \to 41 | \mathrm{Rel}(A)^{\mathcal{W}^{op}}$ 42 | that satisfies the axioms of a monoidal exponential comonad. The 43 | interesting cases are for functions, $\otimes$-products and the 44 | $\oc_\rho$ modality: 45 | \begin{displaymath} 46 | \begin{array}{l} 47 | \llbracket \fun{S}{T} \rrbracket^R~w~(f,f') = \\ 48 | \quad \forall x,y.~P(y,w)x \Rightarrow \forall s,s'.~\llbracket S \rrbracket^R y (s, s') \Rightarrow \llbracket 49 | T \rrbracket^R x (f~s, f'~s') 50 | \\ 51 | \llbracket \tensor{S}{T} \rrbracket^R~w~((s, t), (s', t')) =\\ 52 | \quad 53 | \exists x,y.~P(x,y)w \wedge \llbracket S \rrbracket^R x (s, s') \wedge 54 | \llbracket T \rrbracket^R y (t, t') 55 | \\ 56 | \llbracket \excl{\rho}{S} \rrbracket^R~w~(s,s')= 57 | \oc_\rho \llbracket S \rrbracket^R~w~(s,s')\\ 58 | \end{array} 59 | \end{displaymath} 60 | Contexts 61 | $\ctxvar{x_1}{S_1}{\rho_1}, \ldots, \ctxvar{x_n}{S_n}{\rho_n}$ are 62 | interpreted as if they were 63 | $\sem{(\cdots(1 \otimes \oc_{\rho_1}S_1) \cdots \otimes \oc_{\rho_n} 64 | S_n)}$. 65 | With these definitions, we can prove the following fundamental lemma 66 | for our Kripke-indexed relational semantics. 67 | 68 | % $R \times S$ and $R \uplus S$ are defined pointwise on relations. 69 | % Particularly, there are two cases for $R \uplus S$: 70 | 71 | % \begin{itemize} 72 | % \item $R(r, r')$ implies $(R \uplus S)(\mathrm{inl}~r, \mathrm{inl}~r')$. 73 | % \item $S(s, s')$ implies $(R \uplus S)(\mathrm{inr}~s, \mathrm{inr}~s')$. 74 | % \end{itemize} 75 | 76 | % We assume a family of natural transformations $\oc$ satisfying the following laws. 77 | 78 | % \begin{mathpar} 79 | % \rho \leq \pi \implies (\oc_\pi R~w \implies \oc_\rho R~w) \and 80 | % \oc_0 R~w \implies J w \and 81 | % \oc_{\rho+\pi} R~w~(a, b) \implies \exists x,y. P(x,y)w \wedge \oc_\rho R~x~a \wedge \oc_\pi R~y~b \and 82 | % \oc_1 R~w \iff R~w \and 83 | % \oc_{\rho \cdot \pi} R~w \iff \oc_\rho(\oc_\pi R)~w 84 | % \end{mathpar} 85 | 86 | % The semantics of a context $\ctxvar{x_1}{S_1}{\rho_1}, \ldots, \ctxvar{x_n}{S_n}{\rho_n}$ is given by $\tensor{\llbracket \excl{\rho_1}{S_1}}{\tensor{\ldots}{\excl{\rho_n}{S_n}}} \rrbracket^R$. 87 | 88 | % This indexed relational semantics gives us a family of logical relations. 89 | % The fundamental lemma is as follows. 90 | 91 | \begin{theorem}[Fundamental Lemma] 92 | \begin{displaymath} 93 | \ctx{\Gamma}{\Delta} \vdash t : T \implies \llbracket \ctx{\Gamma}{\Delta} \rrbracket^R w~(\gamma, \gamma') \implies \llbracket T \rrbracket^R w~(\sem{t}\gamma, \sem{t}\gamma') 94 | \end{displaymath} 95 | \end{theorem} 96 | 97 | 98 | % Local Variables: 99 | % TeX-master: "quantitative" 100 | % End: 101 | -------------------------------------------------------------------------------- /tex/abstract/syntax.tex: -------------------------------------------------------------------------------- 1 | We define our type system with respect to a partially ordered semiring 2 | $R$ for tracking how variables are used. A \emph{partially ordered 3 | semiring} $(R, \leq, 0, +, 1 , \cdot)$ is a poset $(R, \leq)$, 4 | commutative monoid $(R, 0, +)$, and monoid $(R, 1, \cdot)$, such that 5 | $\cdot$ distributes over $0$ and $+$, and $+$ and $\cdot$ are 6 | monotonic with respect to $\leq$. We take $\rho,\pi \in R$. 7 | 8 | \paragraph{Examples} 9 | \begin{inparaenum} 10 | \item The zero-one-many semiring $\{0,1,\omega\}$ simulates linear 11 | typing in our system. 12 | \item Monotonicity typing uses the semiring with carrier 13 | $\{0,\uparrow,\downarrow,\equiv\}$, where the multiplicative unit is 14 | $\uparrow$ (covariance). The $\downarrow$ represents 15 | contra-variance, and $\equiv$ represents invariance. 16 | \item Sensitivity analysis uses the semiring with carrier 17 | $\mathbb{R} \cup \{\infty\}$ with $\min$ and $+$ as the addition and 18 | multiplication. 19 | \item Information flow analysis uses the semiring with carrier 20 | $\mathcal{P}(L)$, where $L$ is some set of labels. 21 | \end{inparaenum} 22 | 23 | \medskip 24 | 25 | The base language we consider is a 26 | bidirectional \citep{DBLP:journals/toplas/PierceT00} simply typed 27 | $\lambda$ calculus with the following types, where $\iota$ ranges over 28 | some set of base types: 29 | \begin{displaymath} 30 | S,T ::= \fun{S}{T} \mid \excl{\rho}{S} \mid \tensorOne \mid \tensor{S}{T} \mid 31 | \withTOne \mid \withT{S}{T} \mid \sumTZero \mid \sumT{S}{T} \mid \iota 32 | \end{displaymath} 33 | Bidirectional typing reduces the type annotations required. Since our 34 | language is bidirectionally typed, we have two syntactic categories of 35 | terms: $s$ ranges over checked terms, and $e$ ranges over synthesising 36 | terms. We use $t$ for both. 37 | \begin{displaymath} 38 | \begin{aligned} 39 | s & ::= \lam{x}{s} \mid \bang{s} \mid \unit \mid \ten{s_0}{s_1} \mid \eat \mid \wth{s_0}{s_1} \mid \inj{i}{s} \mid \emb{e} \\ 40 | e & ::= x \mid \app{e}{s} \mid \bm{T}{e}{\bind{x}{s}} \mid \del{T}{e}{s} \mid \prm{T}{e}{\bind{x,y}{s}} \\ 41 | & \quad \mid \proj{i}{e} \mid \exf{T}{e} \mid 42 | \cse{T}{e}{\bind{x}{s_0}}{\bind{y}{s_1}} \mid \ann{s}{S} 43 | \end{aligned} 44 | \end{displaymath} 45 | where curly braces and $\lambda$ denote variable binding and we take 46 | $i \in \{0,1\}$ wherever it appears. 47 | 48 | Contexts $\Gamma$ assign to each variable a type $S$ and a usage 49 | $\rho \in R$: 50 | $\Gamma = \ctxvar{x_1}{S_1}{\rho_1}, \ldots, 51 | \ctxvar{x_n}{S_n}{\rho_n}$. 52 | Contexts whose variables and types match form a left $R$-semimodule, 53 | by pointwise addition and scaling of the usage annotations. The 54 | partial order on $R$ is extended pointwise to contexts. 55 | % Contexts $\ctx{\Gamma}{\Delta}$ over variables $x_1, \ldots, x_n$ are to be understood as a \emph{typing context} $\Gamma$ of the form $x_1 : S_1, \ldots, x_n : S_n$, and a \emph{resourcing context} $\Delta$ of the form $x_0^{\rescomment{\rho_0}}, \ldots, x_n^{\rescomment{\rho_n}}$. 56 | % %The resourcing context in which for each $k$, $\rho_k = 0$, will be abbreviated to $\rescomment{0}$. 57 | % Resourcing contexts with matching variables form a left $R$-semimodule, taking $(x_0^{\rescomment{\rho_0}}, \ldots, x_n^{\rescomment{\rho_n}}) + (x_0^{\rescomment{\pi_0}}, \ldots, x_n^{\rescomment{\pi_n}}) = (x_0^{\rescomment{\rho_0 + \pi_0}}, \ldots, x_n^{\rescomment{\rho_n + \pi_n}})$ and $\rho \cdot (x_0^{\rescomment{\rho \cdot \pi_0}}, \ldots, x_n^{\rescomment{\rho \cdot \pi_n}})$. 58 | Typing judgements for checked and synthesising terms have the same 59 | contexts, but either record that a term is checked against a type 60 | ($\Gamma \vdash T \ni s$) or synthesise a type 61 | ($\Gamma \vdash e \in T$). 62 | % When we only care about typability, we write $t : T$ in place of 63 | % either $T \ni t$ or $t \in T$. 64 | 65 | The typing rules consist of a variable rule, two rules for change of 66 | direction, and introduction and elimination rules for each type 67 | former. The following rules for variables and function and $\oc_\rho$ 68 | introduction and elimination illustrate how usage information is 69 | tracked: {\footnotesize\begin{mathpar} 70 | % Variables 71 | \inferrule 72 | {\Gamma \leq 0\Gamma_1, \ctxvar{x}{S}{1}, 0\Gamma_2} 73 | {\Gamma \vdash x \in S} 74 | 75 | % Functions 76 | \inferrule 77 | {\Gamma, \ctxvar{x}{S}{1} \vdash T \ni s[x]} 78 | {\Gamma \vdash \fun{S}{T} \ni \lam{x}{s[x]}} 79 | 80 | \inferrule 81 | {\Gamma_1 \vdash e \in \fun{S}{T} 82 | \\ \Gamma_2 \vdash S \ni s 83 | \\ \Gamma \leq \Gamma_1 + \Gamma_2} 84 | {\Gamma \vdash \app{e}{s} \in T} 85 | 86 | % Bang 87 | \inferrule 88 | {\Gamma_1 \vdash S \ni s \\ \Gamma \leq \rho \cdot \Gamma_1} 89 | {\Gamma \vdash \excl{\rho}{S} \ni \bang{s}} 90 | 91 | \mprset{flushleft} 92 | \inferrule 93 | {\Gamma_1 \vdash e \in \excl{\rho}{S} \\\\ 94 | \Gamma_2, \ctxvar{x}{S}{\rho} \vdash T \ni s[x] \\\\ 95 | \Gamma \leq \Gamma_1 + \Gamma_2} 96 | {\Gamma \vdash \bm{T}{e}{\{x\}s[x]} \in T} 97 | \end{mathpar}} 98 | Sub-resourcing, weakening (adding $0$-use variables to the context) 99 | and substitution are all admissible. In our Agda formalisation, we 100 | have constructed our type system in two levels: a non-usage tracked 101 | simply-typed $\lambda$-calculus, with a usage-tracking system layered 102 | above. This emphasises the use of coeffect annotations as an 103 | \emph{analysis} of programs, they do not affect the underlying 104 | semantics, but comment on it. We introduce our semantic framework in 105 | the next section. 106 | 107 | %commented out for space 108 | 109 | % The order acts as a sub-resourcing relation. 110 | % Well resourced terms can also be weakened by the introduction of $0$-resource 111 | % variables to the context, and satisfy a substitution principle. 112 | 113 | % \begin{mathpar} 114 | % \inferrule*[Right=subres] 115 | % {\ctx{\Gamma}{\Delta} \vdash t : T 116 | % \\ \rescomment{\Delta' \leq \Delta}} 117 | % {\ctx{\Gamma}{\Delta'} \vdash t : T} 118 | 119 | % \inferrule*[Right=weak] 120 | % {\ctx{\Gamma}{\Delta}, \ctx{\Gamma'}{\Delta'} \vdash t : T} 121 | % {\ctx{\Gamma}{\Delta}, \ctxvar{x}{S}{0}, \ctx{\Gamma'}{\Delta'} 122 | % \vdash t : T} 123 | 124 | % \inferrule*[Right=subst] 125 | % {\ctx{\Gamma}{\Delta} \vdash t : T 126 | % \\ \forall x.~ 127 | % \ctx{\Gamma'}{\vec{\Delta'_x}} \vdash \vec{t'_x} : \Gamma_x 128 | % \\ \rescomment{\Delta'' \leq \sum_x \Delta_x \cdot \vec{\Delta'_x}}} 129 | % {\ctx{\Gamma'}{\Delta''} \vdash t[\vec{t'}] : T} 130 | % \end{mathpar} 131 | 132 | 133 | % Local Variables: 134 | % TeX-master: "quantitative" 135 | % End: 136 | -------------------------------------------------------------------------------- /tex/lin-tlla-abstract/default.nix: -------------------------------------------------------------------------------- 1 | { nixpkgs ? import {} }: 2 | nixpkgs.pkgs.callPackage ./pkg.nix { } 3 | -------------------------------------------------------------------------------- /tex/lin-tlla-abstract/pkg.nix: -------------------------------------------------------------------------------- 1 | { stdenv, texlive }: 2 | let 3 | tex-env = texlive.combine { 4 | inherit (texlive) scheme-small latexmk chktex stmaryrd mathpartir rsfs 5 | cmll xcolor paralist makecell tikz-cd ncctools thmtools 6 | xifthen ifmtarg polytable etoolbox environ xkeyval 7 | lazylist trimspaces; 8 | }; 9 | in stdenv.mkDerivation { 10 | name = "lin-ttla"; 11 | src = ./.; 12 | buildInputs = [ tex-env ]; 13 | buildPhase = '' 14 | latexmk quant-subst.tex 15 | ''; 16 | installPhase = ""; 17 | } 18 | -------------------------------------------------------------------------------- /tex/lin-tlla-abstract/shell.nix: -------------------------------------------------------------------------------- 1 | { pkgs ? import {} }: 2 | pkgs.mkShell { 3 | inputsFrom = [ (pkgs.callPackage ./pkg.nix { }) ]; 4 | } 5 | -------------------------------------------------------------------------------- /tex/lin-tlla-paper/CHANGES: -------------------------------------------------------------------------------- 1 | Thanks to the reviewers for their careful reading of our 2 | submission. We have made the following changes in the light of their 3 | comments: 4 | 5 | - The types and summary of the judgement forms are now separated into 6 | their own figures instead of inlined in the text. We hope this makes 7 | them easier to read, and clarifies the meanings of the judgements. 8 | 9 | - We have spelt out some more of the details of the translations 10 | between Barber's DILL and the Pfenning-Davies system and ours. 11 | 12 | - Added two paragraphs to the conclusion to compare our method to 13 | existing techniques for formalising Linear Logic, expanded the 14 | comparison to Licata et al's work, and to Abel and Bernardy's ICFP 15 | 2020 work. 16 | 17 | - Explained the use of colour to highlight the Agda definitions and 18 | terms, and linked them to the GitHub repository. 19 | 20 | - Minor grammatical and notational tweaks throughout, and fixed 21 | various Lemma and Figure references identified by the reviewers. 22 | -------------------------------------------------------------------------------- /tex/lin-tlla-paper/default.nix: -------------------------------------------------------------------------------- 1 | { nixpkgs ? import {} }: 2 | nixpkgs.pkgs.callPackage ./pkg.nix { } 3 | -------------------------------------------------------------------------------- /tex/lin-tlla-paper/pkg.nix: -------------------------------------------------------------------------------- 1 | { stdenv, texlive }: 2 | let 3 | tex-env = texlive.combine { 4 | inherit (texlive) scheme-small latexmk chktex stmaryrd mathpartir rsfs 5 | cmll xcolor paralist makecell tikz-cd ncctools thmtools 6 | xifthen ifmtarg polytable etoolbox environ xkeyval 7 | lazylist trimspaces; 8 | }; 9 | in stdenv.mkDerivation { 10 | name = "lin-ttla"; 11 | src = ./.; 12 | buildInputs = [ tex-env ]; 13 | buildPhase = '' 14 | latexmk quant-subst.tex 15 | ''; 16 | installPhase = ""; 17 | } 18 | -------------------------------------------------------------------------------- /tex/lin-tlla-paper/shell.nix: -------------------------------------------------------------------------------- 1 | { pkgs ? import {} }: 2 | pkgs.mkShell { 3 | inputsFrom = [ (pkgs.callPackage ./pkg.nix { }) ]; 4 | } 5 | -------------------------------------------------------------------------------- /tex/lin-tlla-presentation/default.nix: -------------------------------------------------------------------------------- 1 | { nixpkgs ? import {} }: 2 | nixpkgs.pkgs.callPackage ./pkg.nix { } 3 | -------------------------------------------------------------------------------- /tex/lin-tlla-presentation/pkg.nix: -------------------------------------------------------------------------------- 1 | { stdenv, texlive }: 2 | let 3 | tex-env = texlive.combine { 4 | inherit (texlive) scheme-small latexmk beamer stmaryrd mathpartir rsfs 5 | cmll xcolor paralist makecell tikz-cd ncctools biblatex 6 | biber thmtools 7 | xifthen ifmtarg polytable etoolbox environ xkeyval 8 | lazylist trimspaces; 9 | }; 10 | in stdenv.mkDerivation { 11 | name = "lin-ttla-presentation"; 12 | src = ./.; 13 | buildInputs = [ tex-env ]; 14 | buildPhase = '' 15 | latexmk quant-subst.tex 16 | ''; 17 | installPhase = ""; 18 | } 19 | -------------------------------------------------------------------------------- /tex/macros.tex: -------------------------------------------------------------------------------- 1 | \ifx\newelims\undefined 2 | \def\newelims{0} 3 | \fi 4 | \ifx\multnotation\undefined 5 | \def\multnotation{0} 6 | \fi 7 | 8 | \newcommand{\fixme}[1]{{\color{red}\underline{\em{#1}}}} 9 | \newcommand{\fixmeBA}[1]{{\color{red}{[Bob:{\em #1}]}}} 10 | 11 | 12 | \newcommand{\bind}[2]{% 13 | \if\debruijn0% 14 | {#2}\{{#1}\}% 15 | \else% 16 | {#2}% 17 | \fi% 18 | } 19 | \newcommand{\ctx}[2]{% 20 | \if\multnotation0% 21 | #1^{\resctx{#2}}% 22 | \else% 23 | \resctx{#2}#1% 24 | \fi 25 | } 26 | \newcommand{\ctxvar}[3]{% 27 | \if\debruijn0% 28 | \if\multnotation0% 29 | #1 \stackrel{\rescomment{#3}}: #2% 30 | \else% 31 | \rescomment{#3}#1 : #2% 32 | \fi% 33 | \else% 34 | \rescomment{#3}#2% 35 | \fi% 36 | } 37 | \definecolor{res}{HTML}{008000} 38 | \definecolor{resperm}{HTML}{008000} 39 | \newcommand{\rescomment}[1]{{\color{res}#1}} 40 | \newcommand{\rescommentperm}[1]{{\color{resperm}#1}} 41 | \newcommand{\resctx}[1]{\rescomment{\mathcal{#1}}} 42 | \newcommand{\resctxperm}[1]{\rescommentperm{\mathcal{#1}}} 43 | \newcommand{\resmat}[1]{\rescomment{#1}} 44 | 45 | 46 | \newcommand{\ann}[2]{#1 : #2} 47 | \newcommand{\emb}[1]{\underline{#1}} 48 | 49 | \newcommand{\base}[0]{\iota} 50 | 51 | \newcommand{\fun}[2]{#1 \multimap #2} 52 | \newcommand{\lam}[2]{% 53 | \if\debruijn0% 54 | \lambda #1.~#2% 55 | \else% 56 | \lambda #2% 57 | \fi% 58 | } 59 | \newcommand{\app}[2]{#1~#2} 60 | 61 | \newcommand{\excl}[2]{\oc \rescomment{#1} #2} 62 | \newcommand{\bang}[1]{\left[#1\right]} 63 | \newcommand{\bm}[4]{% 64 | \if\newelims0% 65 | \operatorname{bm}_{#1}(#2, \bind{#3}{#4})% 66 | \else% 67 | \if\debruijn0% 68 | \mathrm{let}~\bang{#3} = #2~\mathrm{in}~#4% 69 | \else% 70 | \mathrm{let}~\bang{-} = #2~\mathrm{in}~#4% 71 | \fi% 72 | \fi% 73 | } 74 | 75 | \newcommand{\tensorOne}[0]{1} 76 | \newcommand{\unit}[0]{(\mathbin{_\otimes})} 77 | \newcommand{\del}[3]{\if\newelims0% 78 | \operatorname{del}_{#1}(#2, #3)% 79 | \else% 80 | \mathrm{let}~\unit = #2~\mathrm{in}~#3% 81 | \fi} 82 | 83 | \newcommand{\tensor}[2]{#1 \otimes #2} 84 | \newcommand{\ten}[2]{(#1 \mathbin{_{\otimes}} #2)} 85 | \newcommand{\prm}[5]{% 86 | \if\newelims0% 87 | \operatorname{pm}_{#1}(#2, \bind{#3, #4}{#5})% 88 | \else% 89 | \if\debruijn0% 90 | \mathrm{let}~\ten{#3}{#4} = #2~\mathrm{in}~#5% 91 | \else% 92 | \mathrm{let}~\ten{-}{-} = #2~\mathrm{in}~#5% 93 | \fi% 94 | \fi% 95 | } 96 | 97 | \newcommand{\withTOne}[0]{\top} 98 | \newcommand{\eat}[0]{(\mathbin{_{\with}})} 99 | 100 | \newcommand{\withT}[2]{#1 \with #2} 101 | \newcommand{\wth}[2]{(#1 \mathbin{_\with} #2)} 102 | \newcommand{\proj}[2]{\operatorname{proj}_{#1} #2} 103 | 104 | \newcommand{\sumTZero}[0]{0} 105 | \newcommand{\exf}[2]{\operatorname{ex-falso} #2} 106 | 107 | \newcommand{\sumT}[2]{#1 \oplus #2} 108 | \newcommand{\inj}[2]{\operatorname{inj}_{#1} #2} 109 | \newcommand{\cse}[6]{% 110 | \if\newelims0% 111 | \operatorname{case}_{#1}(#2, \bind{#3}{#4}, \bind{#5}{#6})% 112 | \else% 113 | \if\debruijn0% 114 | \mathrm{case}~#2~\mathrm{of}~\inj{L}{#3} \mapsto #4 115 | ~;~ \inj{R}{#5} \mapsto #6% 116 | \else% 117 | \mathrm{case}~#2~\mathrm{of}~\inj{L}{-} \mapsto #4 ~;~ \inj{R}{-} \mapsto #6 118 | \fi% 119 | \fi% 120 | } 121 | 122 | \newcommand{\listT}[1]{\operatorname{List} #1} 123 | \newcommand{\nil}[0]{[]} 124 | \newcommand{\cons}[2]{#1 :: #2} 125 | \newcommand{\fold}[5]{\operatorname{fold} #1~#2~(#3,#4.~#5)} 126 | 127 | 128 | \newcommand{\typed}[1]{\mathit{#1t}} 129 | \newcommand{\resourced}[1]{\mathit{#1r}} 130 | 131 | 132 | \newcommand{\lvar}{\mathrel{\mathrlap{\sqsupset}{\mathord{-}}}} 133 | \newcommand{\sem}[1]{\left\llbracket #1 \right\rrbracket} 134 | 135 | %\def\tobar{\mathrel{\mkern3mu \vcenter{\hbox{$\scriptscriptstyle+$}}% 136 | % \mkern-12mu{\to}}} 137 | \newcommand\tobar{\mathrel{\ooalign{\hfil$\mapstochar\mkern5mu$\hfil\cr$\to$\cr}}} 138 | 139 | 140 | \newcommand{\subres}{\trianglelefteq} 141 | \newcommand{\subrctx}[2]{{#1} \subres {#2}} 142 | \makeatletter 143 | \newcommand{\subst}[2][]{\ext@arrow 0359\Rightarrowfill@{#1}{#2}} 144 | \makeatother 145 | 146 | 147 | \newenvironment{eqns}{\begin{array}{r@{\hspace{0.3em}}c@{\hspace{0.3em}}l}}{\end{array}} 148 | 149 | 150 | \newcommand{\mat}[1]{\mathbf{#1}} 151 | \newcommand{\vct}[1]{\mathbf{#1}} 152 | \DeclarePairedDelimiter\basis{\langle}{\rvert} 153 | 154 | 155 | \newcommand{\name}{\ensuremath{\lambda \resctxperm R}} 156 | 157 | 158 | \DeclareMathOperator\kit{Kit} 159 | \newcommand{\kitrel}{\mathbin{\blacklozenge}} 160 | 161 | 162 | \DeclareMathOperator\id{id} 163 | \DeclareMathOperator\inl{inl} 164 | \DeclareMathOperator\inr{inr} 165 | \DeclareMathOperator\Idx{Idx} 166 | 167 | 168 | \newcommand{\zero}{\ensuremath{\rescomment 0}} 169 | \newcommand{\linear}{\ensuremath{\rescomment 1}} 170 | \newcommand{\unrestricted}{\ensuremath{\rescomment \omega}} 171 | \newcommand{\instDILL}{\rescomment{01\omega}} 172 | 173 | \newcommand{\unused}{\ensuremath{\rescomment{0}}} 174 | \newcommand{\true}{\ensuremath{\rescomment{1}}} 175 | \newcommand{\valid}{\ensuremath{\rescomment{\square}}} 176 | \newcommand{\instPD}{\ensuremath{\rescomment{01\square}}} 177 | -------------------------------------------------------------------------------- /tex/paper-lncs/future-work.tex: -------------------------------------------------------------------------------- 1 | % Future Work -------------------------------------------------------------------------------- /tex/paper-lncs/instantiations.tex: -------------------------------------------------------------------------------- 1 | The ingredients of our fundamental lemma are perhaps well known 2 | (relational interpretations, Kripke-indexing), but the value of our 3 | framework lies in the generality of being able to choose $\mathcal{W}$ 4 | and its promonoidal structure, and the interpretion the $\oc_\rho$ 5 | modality as a relation transformer. Examples include: 6 | 7 | \subsection{Permutation Types} 8 | With the $\{0,1,\omega\}$ semiring, we 9 | take the category $\mathcal{W}$ to consist of lists of some type of 10 | keys, and permutations between them. The relation transformer is 11 | defined as: $\oc_0 R~l = \top$, where $\top$ is the total relation, 12 | $\oc_1 R~l = R~l$ and $R_\omega~R~l = (l = []) \land R~l$. With 13 | suitable types of keys and lists of keys, the fundamental lemma states 14 | that all programs are permutations. This result has already been 15 | formalised in a one-off type system at 16 | \url{https://github.com/bobatkey/sorting-types}. 17 | 18 | \subsection{Monotonicity Types} 19 | With $R$ the partially ordered semiring 20 | with carrier $\{0,\uparrow,\downarrow,\equiv\}$ ordered 21 | ${\equiv} \leq {\uparrow},{\downarrow}$ and ${\uparrow}, {\downarrow} \leq 0$, 22 | we take $\mathcal{W}$ to be the one-object, one-arrow category, and 23 | define the relation transformer $\oc$ to be: 24 | \begin{mathpar} 25 | \oc_0~R = \top 26 | 27 | \oc_\uparrow~R = R 28 | 29 | \oc_\downarrow~R = R^{op} 30 | 31 | \oc_\equiv~R = R \cap R^{op} 32 | \end{mathpar} 33 | If we let our base type be natural numbers with the relational 34 | interpretation $R_{\mathrm{nat}}(n,n') \Leftrightarrow n \leq n'$, 35 | then the fundamental lemma states that a program of type 36 | $\ctxvar{x}{\mathrm{nat}}{\uparrow} \vdash t : \mathrm{nat}$ is 37 | covariant (and similarly for contravariant and invariant). 38 | 39 | \subsection{Sensitivity Analysis} 40 | With the 41 | $R = \mathbb{R} \cup \{\infty\}$ semiring, we let $\mathcal{W}$ be $R$ 42 | as well. The relation transformer is given by scaling. With a base 43 | type of real numbers with relational intepretation 44 | $R_{\mathrm{real}}~k~(x,x') \Leftrightarrow |x-x'| \leq k$, then the 45 | fundamental lemma states that the usage annotations on the input 46 | variables tracks the extent to which the program is sensitive to 47 | changes in those variables. 48 | 49 | \subsection{Information Flow} 50 | With the $R = \mathcal{P}(L)$ semiring, 51 | we again take $\mathcal{W} = R$, and let the relation transformer to 52 | be 53 | $\oc_l R~l' = \{\top~\textrm{when }l \geq l'; R~\textrm{otherwise}\}$. 54 | Then the fundamental lemma yields the same non-interference properties 55 | as stated by Abadi et al. for the DCC \cite{abadi99core}. -------------------------------------------------------------------------------- /tex/paper-lncs/introduction.tex: -------------------------------------------------------------------------------- 1 | When verifying functional programs in a dependently typed programming language, 2 | we often end up writing a program twice --- once to implement the program we 3 | want to run, and once to prove a simple property of the program. 4 | A standard example of this is proving that a sorting function is a permutation. 5 | Below we have the $\textrm{sort}$ function, defined idiomatically by iteration, 6 | and a conventional proof $\textrm{sort-perm}$ showing that it is a permutation. 7 | We assume an $\textrm{insert}$ function, and proof $\textrm{insert-perm} : 8 | \forall x,\mathit{xs}.~\textrm{insert}~x~xs \simeq x :: xs$. 9 | 10 | \[ 11 | \begin{array}{ll} 12 | \begin{array}[t]{l} 13 | \textrm{sort} : \textrm{List}~A \to \textrm{List}~A \\ 14 | \textrm{sort} = \textrm{foldr}~[]~\textrm{insert} 15 | \end{array} 16 | \quad & \quad 17 | \begin{array}[t]{l} 18 | \textrm{sort-perm} : \forall\mathit{xs}.~\textrm{sort}~\mathit{xs} \simeq \mathit{xs} \\ 19 | \textrm{sort-perm}~[] = [] \\ 20 | \textrm{sort-perm}~(x :: \mathit{xs}) = \\ 21 | \quad \textrm{insert-perm}~x~(\textrm{sort}~\mathit{xs}) \\ 22 | \quad {}\bullet (x :: \textrm{sort-perm}~\mathit{xs}) 23 | \end{array} 24 | \end{array} 25 | \] 26 | 27 | This proof is slightly tricky, but it is proving something obvious. 28 | The $\textrm{insert}$ function never drops or duplicates its inputs into the 29 | output, so it is obviously a permutation, and $\textrm{sort}$ is only made up of 30 | functions that do no dropping or duplicating, so is also a permutation. 31 | The aim of this paper is to make observations like this formal. 32 | 33 | \fixme{Old stuff} In normal typed $\lambda$-calculi, variables may be used multiple 34 | times, in multiple contexts, for multiple reasons, as long as the 35 | types agree. The disciplines of linear types \cite{girard87linear} and 36 | coeffects \cite{PetricekOM14,BrunelGMZ14,GhicaS14} refine this by 37 | tracking variable usage. We might track how many times a variable is 38 | used, or if it is used co-, contra-, or invariantly. Such a discipline 39 | yields a general framework of ``context constrained computing'', where 40 | constraints on variables in the context tell us something interesting 41 | about the computation being performed. 42 | % Thus we put the 43 | % type information to work to tell us facts about programs that might 44 | % not otherwise be apparent. 45 | 46 | We will present work in progress on capturing the ``intensional'' 47 | properties of programs via a family of Kripke indexed relational 48 | semantics that refines a simple set-theoretic semantics of 49 | programs. The value of our approach lies in its generality. We can 50 | accommodate the following examples: 51 | \begin{enumerate} 52 | \item Linear types that capture properties like ``all list 53 | manipulating programs are permutations''. This example uses the 54 | Kripke-indexing to track the collection of datums currently being 55 | manipulated by the program. 56 | \item Monotonicity coeffects that track whether a program uses inputs 57 | co-, contra-, or in-variantly (or not at all). 58 | \item Sensitivity typing, tracking the sensitivity of programs in 59 | terms of input changes. This forms the core of systems for 60 | differential privacy \cite{reed10distance}. 61 | \item Information flow typing, in the style of the Dependency Core 62 | Calculus \cite{abadi99core}. 63 | \end{enumerate} 64 | 65 | The syntax and semantics we present here have been formalised in Agda: 66 | \url{https://github.com/laMudri/quantitative/}. 67 | 68 | Our main contributions are: 69 | 70 | \begin{itemize} 71 | \item A rigorous statement of substitution for a substructural type system 72 | \item A Kripke-indexed relational semantics providing strong free theorems 73 | \item A formalisation of this work in Agda 74 | \end{itemize} 75 | 76 | \subsection{Related work} 77 | 78 | We follow closely and extend the approaches of Petricek, Orchard, and Mycroft, 79 | and Ghica and Smith \cite{PetricekOM14,GhicaS14}. 80 | In particular, our framework is generic in a partially ordered semiring of 81 | \emph{usage constraints}, which are placed on each variable in the context. 82 | This is distinct from the approach taken by Brunel et al and the Granule project 83 | \cite{BrunelGMZ14,Granule18}, where \emph{unannotated} variables also exist, and 84 | are treated linearly. 85 | Both of the latter systems contain a \emph{dereliction} rule, stating that a 86 | variable annotated $\rescomment 1$ can be coerced to an unannotated variable. 87 | This rule can be justified by the graded comonad unit law 88 | $\fun{\excl{1}{A}}{A}$, so adds no new axioms to be met when we want to produce 89 | semantics. 90 | However, the distinction between annotated and unannotated variables does cause 91 | complexity in the syntax. 92 | For one thing, the dereliction rule is not syntax-directed. 93 | For another, the operation of merging contexts is now partial, and the scaling 94 | operation is not uniform or also partial. 95 | 96 | Beyond previous work that assumed all variables were annotated, we allow a full 97 | complement of propositional intuitionistic linear logic connectives. 98 | Particularly, we have tensor products, with products, sums, and the bang 99 | modality, as opposed to just functions. 100 | In contrast to this previous work, our functions are not annotated with a usage 101 | constraint; we prefer a combination of the function arrow and the bang modality 102 | so as to improve modularity of the connectives. 103 | 104 | Our main novel language feature compared to previous work on usage-constrained 105 | typing is an account of an inductive type (tensor lists) in both syntax and 106 | semantics. 107 | We believe that this account could be generalised to cover all strictly positive 108 | inductive types in a relatively straightforward manner. -------------------------------------------------------------------------------- /tex/paper-lncs/motivation.tex: -------------------------------------------------------------------------------- 1 | In this section, we give a brief informal overview of the properties we wish to 2 | obtain of programs for free by restricting usage. 3 | We use this to motivate the design of syntax and semantics detailed in 4 | \autoref{sec:syntax} and \autoref{sec:semantics} respectively. 5 | 6 | \subsection{Syntax} 7 | 8 | Our fundamental syntactic principle is that we will restrict use of the variable 9 | rule by encoding in contexts \emph{how} its variables can be used. 10 | To each variable in the context, we attach a \emph{usage annotation}. 11 | We may use a given variable only if it can be used in a plain manner, and all 12 | other variables in that context can be unused (discarded). 13 | 14 | Unlike types, we should expect the usage annotations of variables to change in a 15 | typing derivation. 16 | For example, suppose we want a linear $\lambda$-calculus, where each 17 | $\lambda$-bound variable has to be used exactly once. 18 | Then in this language, we want to write a curried function of two arguments that 19 | pairs those two arguments together. 20 | Na\"ively, we can write 21 | $\lambda x.\lambda y.~(x, y) : \fun{A}{\fun{B}{\tensor{A}{B}}}$. 22 | When we check this, we have to use $x$ and discard $y$ in the left of the pair, 23 | and use $y$ and discard $x$ in the right. 24 | Doing both of these must constitute using both $x$ and $y$, eventually 25 | discarding neither, implying some notion of accumulation of usages. 26 | 27 | To deal with this formally, we can set the usage annotations to be the natural 28 | numbers, with the intention of counting how many times a variable is used. 29 | We can discard a variable annotated $0$, and we can plainly use variables 30 | annotated $1$. 31 | So, to use $x$ in our example, we must be trying to conclude 32 | $\ctxvar{x}{A}{1}, \ctxvar{y}{B}{0} \vdash x : A$, and to use $y$, the 33 | annotations must be the other way round. 34 | Then, forming a pair lets us pointwise add together usage annotations, giving 35 | conclusion 36 | $\ctxvar{x}{A}{1+0}, \ctxvar{y}{B}{0+1} \vdash (x, y) : \tensor{A}{B}$. 37 | In general, we require the set of usage annotations to have an addition 38 | operator, as well as its unit $0$ and designated ``plain use'' annotation $1$. 39 | 40 | We also want a way to reify the idea of a variable usable in any particular way 41 | into a value in its own right. 42 | For example, we may want a value that can be used exactly twice, rather than 43 | exactly once. 44 | For this, we introduce the \emph{graded bang} type constructor $\excl{\rho}$, 45 | where $\rho$ is a usage annotation. 46 | This has value constructor $\bang{-}$, and using pattern matching notation 47 | allows us to write $\lambda\bang{x}.~(x, x) : \fun{\excl{2}{A}}{\tensor{A}{A}}$. 48 | Before pattern matching, we have a variable $\ctxvar{b}{\excl{2}{A}}{1}$, and 49 | after pattern matching, it is used up (has annotation $0$) and we get a new 50 | variable $\ctxvar{x}{A}{2}$. 51 | 52 | To produce an open term that can be used twice is to produce an open term that 53 | can be used once, but in a context where each variable can be used twice as many 54 | times as it was in producing the term once. 55 | Formally, if 56 | $\ctxvar{x_1}{A_1}{\pi_1}, \ldots, \ctxvar{x_n}{A_n}{\pi_n} \vdash t : B$, then 57 | $\ctxvar{x_1}{A_1}{\rho*\pi_1}, \ldots, \ctxvar{x_n}{A_n}{\rho*\pi_n} \vdash 58 | \bang t : \excl{\rho}{B}$. 59 | This means that we have multiplication on usage annotations, of which $1$ is the 60 | unit. 61 | 62 | \subsection{Semantics} 63 | 64 | The point of constraining the use of variables is to restrict ourselves to 65 | certain classes of well behaved terms. 66 | For example, we may be interested in the linear terms, or the monotonic terms, 67 | or the terms that are not too sensitive to change. 68 | We say that the terms that use variables in accordance with the rules given by 69 | the syntax and usage annotations are \emph{well provisioned}. 70 | We provide a unified denotational semantics as a tool to show that any well 71 | provisioned term really has the properties we wanted of it. 72 | 73 | We start by giving terms a standard $\mathrm{Set}$ semantics, written $\sem 74 | A$, $\sem \Gamma$, and $\sem t : \sem \Gamma \to \sem A$ for types, contexts and 75 | terms, respectively. 76 | When written in a dependently typed programming language, this is an interpreter 77 | that takes a metalanguage value for each variable in the context and produces a 78 | metalanguage value as a result. 79 | If we have a term in context with one variable, we can consider whether $\sem t 80 | : \sem A \to \sem B$ is monotonic. 81 | This would mean that if a bigger $\sem A$ value is given, the resulting $\sem B$ 82 | value will also be bigger. 83 | To capture this, we interpret each type $A$ as a binary endorelation over $\sem 84 | A$, written $\sem{A}^R \subseteq \sem A \times \sem A$. 85 | Then, given our one-variable term $\ctxvar{x}{A}{1} \vdash t : B$, its semantics 86 | says that it preserves the relation. 87 | Explicitly, $\sem{t}^R : 88 | \forall a, a'.~\sem{A}^R(a, a') \implies \sem{B}^R(\sem t~a, \sem t~a')$. 89 | % The property we want is the following square, where $\Gamma \vdash t : A$, 90 | % $\gamma, \gamma' \in \sem\Gamma$, $a, a' \in \sem A$. 91 | 92 | % \begin{tikzcd} 93 | % \gamma \arrow[r] \arrow[d, "\sem t" left] & \gamma' \arrow[d, "\sem t"] \\ 94 | % a \arrow[r] & a' 95 | % \end{tikzcd} 96 | 97 | For other properties, however, we also need the relation to be preserved 98 | relative to a \emph{world}. 99 | For example, in sensitivity analysis, we want say that a perturbation of 100 | $\varepsilon$ in the environment leads to at most a perturbation of 101 | $\varepsilon$ in the value produced. 102 | The world is this $\varepsilon$. 103 | \fixme{Too early to mention $\otimes$ vs $\&$?} 104 | Additionally, we want the perturbation of an environment to be the sum of 105 | perturbations of its variables, meaning that we need to be able to add worlds 106 | together. 107 | Finally, to work out the perturbation in a variable $\ctxvar{x}{A}{\rho}$, we 108 | must consider the usage annotation $\rho$. 109 | For sensitivity analysis, these annotations are going to be scaling factors, so 110 | the action of an annotation $\rho$ on a world $\varepsilon$ will produce the 111 | world $\rho\varepsilon$. 112 | In general, usage annotations will be interpreted as actions on worlds. 113 | -------------------------------------------------------------------------------- /tex/paper-lncs/quantitative.bib: -------------------------------------------------------------------------------- 1 | @inproceedings{BrunelGMZ14, 2 | author = {A. Brunel and 3 | M. Gaboardi and 4 | D. Mazza and 5 | S. Zdancewic}, 6 | title = {{A Core Quantitative Coeffect Calculus}}, 7 | booktitle = {{ESOP} 2014}, 8 | pages = {351--370}, 9 | year = 2014, 10 | timestamp = {Sun, 23 Mar 2014 10:48:25 +0100}, 11 | biburl = {http://dblp.uni-trier.de/rec/bib/conf/esop/BrunelGMZ14}, 12 | bibsource = {dblp computer science bibliography, http://dblp.org} 13 | } 14 | 15 | @inproceedings{GhicaS14, 16 | author = {Dan R. Ghica and 17 | Alex I. Smith}, 18 | title = {Bounded Linear Types in a Resource Semiring}, 19 | booktitle = {{ESOP} 2014}, 20 | pages = {331--350}, 21 | year = {2014}, 22 | timestamp = {Sun, 23 Mar 2014 10:48:25 +0100}, 23 | biburl = {http://dblp.uni-trier.de/rec/bib/conf/esop/GhicaS14}, 24 | bibsource = {dblp computer science bibliography, http://dblp.org} 25 | } 26 | 27 | @inproceedings{PetricekOM14, 28 | author = {Tomas Petricek and 29 | Dominic A. Orchard and 30 | Alan Mycroft}, 31 | title = {Coeffects: a calculus of context-dependent computation}, 32 | booktitle = {{ICFP} 2014}, 33 | pages = {123--135}, 34 | year = {2014}, 35 | timestamp = {Sun, 04 Jun 2017 10:05:10 +0200}, 36 | biburl = {http://dblp.uni-trier.de/rec/bib/conf/icfp/PetricekOM14}, 37 | bibsource = {dblp computer science bibliography, http://dblp.org} 38 | } 39 | 40 | @article{ girard87linear, 41 | author = "Jean-Yves Girard", 42 | title = "Linear Logic", 43 | journal = "Theor. Comp. Sci.", 44 | volume = 50, 45 | pages = "1--101", 46 | year = 1987 } 47 | 48 | @InProceedings{reed10distance, 49 | author = {J. Reed and B. C. Pierce}, 50 | title = {Distance Makes the Types Grow Stronger}, 51 | booktitle = {ICFP 2010}, 52 | year = 2010, 53 | editor = {P. Hudak and S. Weirich}, 54 | pages = {157-168}} 55 | 56 | @inproceedings{abadi99core, 57 | author = {M. Abadi and 58 | A. Banerjee and 59 | N. Heintze and 60 | J. G. Riecke}, 61 | title = {{A Core Calculus of Dependency}}, 62 | booktitle = {POPL '99}, 63 | year = 1999, 64 | pages = {147-160}, 65 | } 66 | 67 | @article{DBLP:journals/toplas/PierceT00, 68 | author = {Benjamin C. Pierce and David N. Turner}, 69 | title = {Local type inference}, 70 | journal = {{ACM} {TOPLAS}}, 71 | volume = {22}, 72 | number = {1}, 73 | pages = {1--44}, 74 | year = {2000}, 75 | timestamp = {Wed, 26 Nov 2003 14:26:46 +0100}, 76 | biburl = {https://dblp.org/rec/bib/journals/toplas/PierceT00}, 77 | bibsource = {dblp computer science bibliography, https://dblp.org} 78 | } 79 | 80 | @InProceedings{quantitative-type-theory, 81 | title = {The Syntax and Semantics of Quantitative Type Theory}, 82 | author = {Robert Atkey}, 83 | year = {2018}, 84 | doi = {10.1145/3209108.3209189}, 85 | booktitle = {LICS '18: 33rd Annual ACM/IEEE Symposium on Logic in Computer Science, July 9--12, 2018, Oxford, United Kingdom}}'}} 86 | 87 | @techreport{Barber1996, 88 | title = {Dual Intuitionistic Linear Logic}, 89 | author = {Andrew Barber}, 90 | year = {1996}, 91 | institution = {The University of Edinburgh} 92 | } 93 | 94 | @INPROCEEDINGS{judgmental, 95 | author = {Frank Pfenning and Rowan Davies}, 96 | title = {A Judgmental Reconstruction of Modal Logic}, 97 | booktitle = {Mathematical Structures in Computer Science}, 98 | year = {1999}, 99 | pages = {2001} 100 | } 101 | 102 | @InProceedings{context-constrained, 103 | title = {Context Constrained Computation}, 104 | author = {Robert Atkey and James Wood}, 105 | year = {2018}, 106 | booktitle = {3rd Workshop on Type-Driven Development (TyDe '18), Extended Abstract} 107 | } 108 | 109 | @article{Granule18, 110 | volume = {3}, 111 | month = {June}, 112 | title = {Quantitative program reasoning with graded modal types}, 113 | author = {Dominic A. Orchard and Vilem Liepelt and Harley Eades}, 114 | publisher = {ACM}, 115 | year = {2019}, 116 | journal = {Proceedings of the ACM on Programming Languages}, 117 | keywords = {graded modal types, linear types, coeffects, implementation}, 118 | url = {https://kar.kent.ac.uk/74450/}, 119 | abstract = {In programming, data is often considered to be infinitely copiable, arbitrarily discardable, and universally unconstrained. However this view is naive: some data encapsulates resources that are subject to protocols (e.g., file and device handles, channels); some data should not be arbitrarily copied or communicated (e.g., private data). Linear types provide a partial remedy by delineating data in two camps: "resources" to be used but never copied or discarded, and unconstrained values. However, this binary distinction is too coarse-grained. Instead, we propose the general notion of graded modal types, which in combination with linear and indexed types, provides an expressive type theory for enforcing fine-grained resource-like properties of data. We present a type system drawing together these aspects (linear, graded, and indexed) embodied in a fully-fledged functional language implementation, called Granule. We detail the type system, including its metatheoretic properties, and explore examples in the concrete language. This work advances the wider goal of expanding the reach of type systems to capture and verify a broader set of program properties.} 120 | } -------------------------------------------------------------------------------- /tex/paper-lncs/quantitative.tex: -------------------------------------------------------------------------------- 1 | % This is samplepaper.tex, a sample chapter demonstrating the 2 | % LLNCS macro package for Springer Computer Science proceedings; 3 | % Version 2.20 of 2017/10/04 4 | % 5 | \documentclass[runningheads]{llncs} 6 | % 7 | \usepackage{graphicx} 8 | % Used for displaying a sample figure. If possible, figure files should 9 | % be included in EPS format. 10 | 11 | \usepackage{stmaryrd} 12 | \usepackage{mathpartir} 13 | \usepackage{amssymb} 14 | \usepackage{cmll} 15 | \usepackage{xcolor} 16 | \usepackage{paralist} 17 | \usepackage{amsmath} 18 | \usepackage{mathrsfs} 19 | \usepackage{mathtools} 20 | \usepackage{multirow} 21 | \usepackage{relsize} 22 | \usepackage{tabularx} 23 | \usepackage{tikz-cd} 24 | \usepackage{breqn} 25 | 26 | \usetikzlibrary{arrows} 27 | 28 | \usepackage{hyperref} 29 | \renewcommand\UrlFont{\color{blue}\rmfamily} 30 | 31 | \def\newelims{1} 32 | \def\multnotation{1} 33 | \input{../macros.tex} 34 | 35 | \newcommand{\lemref}[1]{\hyperref[#1]{Lemma \ref*{#1}}} 36 | \newcommand{\exref}[1]{\hyperref[#1]{Example \ref*{#1}}} 37 | \newcommand{\judgeref}[1]{\hyperref[#1]{Judgement \ref*{#1}}} 38 | \newcommand{\propref}[1]{\hyperref[#1]{Proposition \ref*{#1}}} 39 | \newcommand{\diagref}[1]{\hyperref[#1]{Diagram \ref*{#1}}} 40 | \newcommand{\thmref}[1]{\hyperref[#1]{Theorem \ref*{#1}}} 41 | \newcommand{\defref}[1]{\hyperref[#1]{Definition \ref*{#1}}} 42 | 43 | \begin{document} 44 | % 45 | \title{A Relational Semantics Tracking Usage %Capturing Usage Restrictions 46 | \thanks{James Wood is supported by an EPSRC Studentship.}} 47 | % 48 | %\titlerunning{Abbreviated paper title} 49 | % If the paper title is too long for the running head, you can set 50 | % an abbreviated paper title here 51 | % 52 | \author{Robert Atkey\inst{1}\orcidID{0000-1111-2222-3333} \and 53 | James Wood\inst{1}\orcidID{1111-2222-3333-4444}} 54 | % 55 | \authorrunning{R. Atkey, J. Wood} 56 | % First names are abbreviated in the running head. 57 | % If there are more than two authors, 'et al.' is used. 58 | % 59 | \institute{University of Strathclyde, Glasgow, G1 1XQ, United Kingdom 60 | \email{\{robert.atkey,james.wood.100\}@strath.ac.uk}\\ 61 | \url{https://www.strath.ac.uk/}} 62 | % 63 | \maketitle % typeset the header of the contribution 64 | % 65 | \begin{abstract} 66 | The abstract should briefly summarize the contents of the paper in 67 | 150--250 words. 68 | 69 | \keywords{First keyword \and Second keyword \and Another keyword.} 70 | \end{abstract} 71 | 72 | \section{Introduction} 73 | \label{sec:introduction} 74 | \input{introduction} 75 | 76 | \section{Motivation} 77 | \label{sec:motivation} 78 | \input{motivation} 79 | 80 | \section{Syntax} 81 | \label{sec:syntax} 82 | \input{syntax} 83 | 84 | \section{Semantics} 85 | \label{sec:semantics} 86 | \input{semantics} 87 | 88 | \section{Instantiations} 89 | \label{sec:instantiations} 90 | \input{instantiations} 91 | 92 | \section{Conclusions} 93 | \label{sec:future-work} 94 | \input{future-work} 95 | % 96 | % ---- Bibliography ---- 97 | % 98 | % BibTeX users should specify bibliography style 'splncs04'. 99 | % References will then be sorted and formatted in the correct style. 100 | 101 | \bibliographystyle{splncs04} 102 | \bibliography{quantitative} 103 | 104 | \end{document} 105 | -------------------------------------------------------------------------------- /tex/paper/bib.bib: -------------------------------------------------------------------------------- 1 | @INPROCEEDINGS{Barendregt92lambdacalculi, 2 | author = {Henk Barendregt}, 3 | title = {Lambda Calculi with Types}, 4 | booktitle = {HANDBOOK OF LOGIC IN COMPUTER SCIENCE}, 5 | year = {1992}, 6 | pages = {117--309}, 7 | publisher = {Oxford University Press} 8 | } 9 | 10 | @Article{deBruijn:dummies, 11 | author = {Nicolas G. de~Bruijn}, 12 | title = "{Lambda Calculus notation with nameless dummies: a tool for automatic formula manipulation}", 13 | journal = {Indagationes Mathematic{\ae}}, 14 | year = {1972}, 15 | volume = {34}, 16 | pages = {381--392}, 17 | } 18 | 19 | @article{DBLP:journals/toplas/PierceT00, 20 | author = {Benjamin C. Pierce and David N. Turner}, 21 | title = {Local type inference}, 22 | journal = {{ACM} Trans. Program. Lang. Syst.}, 23 | volume = {22}, 24 | number = {1}, 25 | pages = {1--44}, 26 | year = {2000}, 27 | url = {http://doi.acm.org/10.1145/345099.345100}, 28 | doi = {10.1145/345099.345100}, 29 | timestamp = {Wed, 26 Nov 2003 14:26:46 +0100}, 30 | biburl = {https://dblp.org/rec/bib/journals/toplas/PierceT00}, 31 | bibsource = {dblp computer science bibliography, https://dblp.org} 32 | } -------------------------------------------------------------------------------- /tex/paper/future-work.tex: -------------------------------------------------------------------------------- 1 | % Future Work -------------------------------------------------------------------------------- /tex/paper/instantiations.tex: -------------------------------------------------------------------------------- 1 | The ingredients of our fundamental lemma are perhaps well known 2 | (relational interpretations, Kripke-indexing), but the value of our 3 | framework lies in the generality of being able to choose $\mathcal{W}$ 4 | and its promonoidal structure, and the interpretion the $\oc_\rho$ 5 | modality as a relation transformer. Examples include: 6 | 7 | \subsection{Permutation Types} 8 | With the $\{0,1,\omega\}$ semiring, we 9 | take the category $\mathcal{W}$ to consist of lists of some type of 10 | keys, and permutations between them. The relation transformer is 11 | defined as: $\oc_0 R~l = \top$, where $\top$ is the total relation, 12 | $\oc_1 R~l = R~l$ and $R_\omega~R~l = (l = []) \land R~l$. With 13 | suitable types of keys and lists of keys, the fundamental lemma states 14 | that all programs are permutations. This result has already been 15 | formalised in a one-off type system at 16 | \url{https://github.com/bobatkey/sorting-types}. 17 | 18 | \subsection{Monotonicity Types} 19 | With $R$ the partially ordered semiring 20 | with carrier $\{0,\uparrow,\downarrow,\equiv\}$ ordered 21 | ${\equiv} \leq {\uparrow},{\downarrow}$ and ${\uparrow}, {\downarrow} \leq 0$, 22 | we take $\mathcal{W}$ to be the one-object, one-arrow category, and 23 | define the relation transformer $\oc$ to be: 24 | \begin{mathpar} 25 | \oc_0~R = \top 26 | 27 | \oc_\uparrow~R = R 28 | 29 | \oc_\downarrow~R = R^{op} 30 | 31 | \oc_\equiv~R = R \cap R^{op} 32 | \end{mathpar} 33 | If we let our base type be natural numbers with the relational 34 | interpretation $R_{\mathrm{nat}}(n,n') \Leftrightarrow n \leq n'$, 35 | then the fundamental lemma states that a program of type 36 | $\ctxvar{x}{\mathrm{nat}}{\uparrow} \vdash t : \mathrm{nat}$ is 37 | covariant (and similarly for contravariant and invariant). 38 | 39 | \subsection{Sensitivity Analysis} 40 | With the 41 | $R = \mathbb{R} \cup \{\infty\}$ semiring, we let $\mathcal{W}$ be $R$ 42 | as well. The relation transformer is given by scaling. With a base 43 | type of real numbers with relational intepretation 44 | $R_{\mathrm{real}}~k~(x,x') \Leftrightarrow |x-x'| \leq k$, then the 45 | fundamental lemma states that the usage annotations on the input 46 | variables tracks the extent to which the program is sensitive to 47 | changes in those variables. 48 | 49 | \subsection{Information Flow} 50 | With the $R = \mathcal{P}(L)$ semiring, 51 | we again take $\mathcal{W} = R$, and let the relation transformer to 52 | be 53 | $\oc_l R~l' = \{\top~\textrm{when }l \geq l'; R~\textrm{otherwise}\}$. 54 | Then the fundamental lemma yields the same non-interference properties 55 | as stated by Abadi et al. for the DCC \cite{abadi99core}. -------------------------------------------------------------------------------- /tex/paper/introduction.tex: -------------------------------------------------------------------------------- 1 | In normal typed $\lambda$-calculi, variables may be used multiple 2 | times, in multiple contexts, for multiple reasons, as long as the 3 | types agree. The disciplines of linear types \cite{girard87linear} and 4 | coeffects \cite{PetricekOM14,BrunelGMZ14,GhicaS14} refine this by 5 | tracking variable usage. We might track how many times a variable is 6 | used, or if it is used co-, contra-, or invariantly. Such a discipline 7 | yields a general framework of ``context constrained computing'', where 8 | constraints on variables in the context tell us something interesting 9 | about the computation being performed. 10 | % Thus we put the 11 | % type information to work to tell us facts about programs that might 12 | % not otherwise be apparent. 13 | 14 | We will present work in progress on capturing the ``intensional'' 15 | properties of programs via a family of Kripke indexed relational 16 | semantics that refines a simple set-theoretic semantics of 17 | programs. The value of our approach lies in its generality. We can 18 | accommodate the following examples: 19 | \begin{enumerate} 20 | \item Linear types that capture properties like ``all list 21 | manipulating programs are permutations''. This example uses the 22 | Kripke-indexing to track the collection of datums currently being 23 | manipulated by the program. 24 | \item Monotonicity coeffects that track whether a program uses inputs 25 | co-, contra-, or in-variantly (or not at all). 26 | \item Sensitivity typing, tracking the sensitivity of programs in 27 | terms of input changes. This forms the core of systems for 28 | differential privacy \cite{reed10distance}. 29 | \item Information flow typing, in the style of the Dependency Core 30 | Calculus \cite{abadi99core}. 31 | \end{enumerate} 32 | 33 | The syntax and semantics we present here have been formalised in Agda: 34 | \url{https://github.com/laMudri/quantitative/}. 35 | 36 | Our main contributions are: 37 | 38 | \begin{itemize} 39 | \item A rigorous statement of substitution for a substructural type system 40 | \item A Kripke-indexed relational semantics providing strong free theorems 41 | \item A formalisation of this work in Agda 42 | \end{itemize} -------------------------------------------------------------------------------- /tex/paper/motivation.tex: -------------------------------------------------------------------------------- 1 | In this section, we give a brief informal overview of the properties we wish to 2 | obtain of programs for free by restricting usage. 3 | We use this to motivate the design of syntax and semantics detailed in 4 | \autoref{sec:syntax} and \autoref{sec:semantics} respectively. 5 | 6 | \subsection{Syntax} 7 | 8 | Our fundamental syntactic principle is that we will restrict use of the variable 9 | rule by encoding in contexts \emph{how} its variables can be used. 10 | To each variable in the context, we attach a \emph{usage annotation}. 11 | We may use a given variable only if it can be used in a plain manner, and all 12 | other variables in that context can be unused (discarded). 13 | 14 | Unlike types, we should expect the usage annotations of variables to change in a 15 | typing derivation. 16 | For example, suppose we want a linear $\lambda$-calculus, where each 17 | $\lambda$-bound variable has to be used exactly once. 18 | Then in this language, we want to write a curried function of two arguments that 19 | pairs those two arguments together. 20 | Na\"ively, we can write 21 | $\lambda x.\lambda y.~(x, y) : \fun{A}{\fun{B}{\tensor{A}{B}}}$. 22 | When we check this, we have to use $x$ and discard $y$ in the left of the pair, 23 | and use $y$ and discard $x$ in the right. 24 | Doing both of these must constitute using both $x$ and $y$, eventually 25 | discarding neither, implying some notion of accumulation of usages. 26 | 27 | To deal with this formally, we can set the usage annotations to be the natural 28 | numbers, with the intention of counting how many times a variable is used. 29 | We can discard a variable annotated $0$, and we can plainly use variables 30 | annotated $1$. 31 | So, to use $x$ in our example, we must be trying to conclude 32 | $\ctxvar{x}{A}{1}, \ctxvar{y}{B}{0} \vdash x : A$, and to use $y$, the 33 | annotations must be the other way round. 34 | Then, forming a pair lets us pointwise add together usage annotations, giving 35 | conclusion 36 | $\ctxvar{x}{A}{1+0}, \ctxvar{y}{B}{0+1} \vdash (x, y) : \tensor{A}{B}$. 37 | In general, we require the set of usage annotations to have an addition 38 | operator, as well as its unit $0$ and designated ``plain use'' annotation $1$. 39 | 40 | We also want a way to reify the idea of a variable usable in any particular way 41 | into a value in its own right. 42 | For example, we may want a value that can be used exactly twice, rather than 43 | exactly once. 44 | For this, we introduce the \emph{graded bang} type constructor $\excl{\rho}$, 45 | where $\rho$ is a usage annotation. 46 | This has value constructor $\bang{-}$, and using pattern matching notation 47 | allows us to write $\lambda\bang{x}.~(x, x) : \fun{\excl{2}{A}}{\tensor{A}{A}}$. 48 | Before pattern matching, we have a variable $\ctxvar{b}{\excl{2}{A}}{1}$, and 49 | after pattern matching, it is used up (has annotation $0$) and we get a new 50 | variable $\ctxvar{x}{A}{2}$. 51 | 52 | To produce an open term that can be used twice is to produce an open term that 53 | can be used once, but in a context where each variable can be used twice as many 54 | times as it was in producing the term once. 55 | Formally, if 56 | $\ctxvar{x_1}{A_1}{\pi_1}, \ldots, \ctxvar{x_n}{A_n}{\pi_n} \vdash t : B$, then 57 | $\ctxvar{x_1}{A_1}{\rho*\pi_1}, \ldots, \ctxvar{x_n}{A_n}{\rho*\pi_n} \vdash 58 | \bang t : \excl{\rho}{B}$. 59 | This means that we have multiplication on usage annotations, of which $1$ is the 60 | unit. 61 | 62 | \subsection{Semantics} 63 | 64 | The point of constraining the use of variables is to restrict ourselves to 65 | certain classes of well behaved terms. 66 | For example, we may be interested in the linear terms, or the monotonic terms, 67 | or the terms that are not too sensitive to change. 68 | We say that the terms that use variables in accordance with the rules given by 69 | the syntax and usage annotations are \emph{well provisioned}. 70 | We provide a unified denotational semantics as a tool to show that any well 71 | provisioned term really has the properties we wanted of it. 72 | 73 | We start by giving terms a standard $\mathrm{Set}$ semantics, written $\sem 74 | A$, $\sem \Gamma$, and $\sem t : \sem \Gamma \to \sem A$ for types, contexts and 75 | terms, respectively. 76 | When written in a dependently typed programming language, this is an interpreter 77 | that takes a metalanguage value for each variable in the context and produces a 78 | metalanguage value as a result. 79 | If we have a term in context with one variable, we can consider whether $\sem t 80 | : \sem A \to \sem B$ is monotonic. 81 | This would mean that if a bigger $\sem A$ value is given, the resulting $\sem B$ 82 | value will also be bigger. 83 | To capture this, we interpret each type $A$ as a binary endorelation over $\sem 84 | A$, written $\sem{A}^R \subseteq \sem A \times \sem A$. 85 | Then, given our one-variable term $\ctxvar{x}{A}{1} \vdash t : B$, its semantics 86 | says that it preserves the relation. 87 | Explicitly, $\sem{t}^R : 88 | \forall a, a'.~\sem{A}^R(a, a') \implies \sem{B}^R(\sem t~a, \sem t~a')$. 89 | % The property we want is the following square, where $\Gamma \vdash t : A$, 90 | % $\gamma, \gamma' \in \sem\Gamma$, $a, a' \in \sem A$. 91 | 92 | % \begin{tikzcd} 93 | % \gamma \arrow[r] \arrow[d, "\sem t" left] & \gamma' \arrow[d, "\sem t"] \\ 94 | % a \arrow[r] & a' 95 | % \end{tikzcd} 96 | 97 | For other properties, however, we also need the relation to be preserved 98 | relative to a \emph{world}. 99 | For example, in sensitivity analysis, we want say that a perturbation of 100 | $\varepsilon$ in the environment leads to at most a perturbation of 101 | $\varepsilon$ in the value produced. 102 | The world is this $\varepsilon$. 103 | \fixme{Too early to mention $\otimes$ vs $\&$?} 104 | Additionally, we want the perturbation of an environment to be the sum of 105 | perturbations of its variables, meaning that we need to be able to add worlds 106 | together. 107 | Finally, to work out the perturbation in a variable $\ctxvar{x}{A}{\rho}$, we 108 | must consider the usage annotation $\rho$. 109 | For sensitivity analysis, these annotations are going to be scaling factors, so 110 | the action of an annotation $\rho$ on a world $\varepsilon$ will produce the 111 | world $\rho\varepsilon$. 112 | In general, usage annotations will be interpreted as actions on worlds. 113 | -------------------------------------------------------------------------------- /tex/paper/quantitative.bib: -------------------------------------------------------------------------------- 1 | @inproceedings{BrunelGMZ14, 2 | author = {A. Brunel and 3 | M. Gaboardi and 4 | D. Mazza and 5 | S. Zdancewic}, 6 | title = {{A Core Quantitative Coeffect Calculus}}, 7 | booktitle = {{ESOP} 2014}, 8 | pages = {351--370}, 9 | year = 2014, 10 | timestamp = {Sun, 23 Mar 2014 10:48:25 +0100}, 11 | biburl = {http://dblp.uni-trier.de/rec/bib/conf/esop/BrunelGMZ14}, 12 | bibsource = {dblp computer science bibliography, http://dblp.org} 13 | } 14 | 15 | @inproceedings{GhicaS14, 16 | author = {Dan R. Ghica and 17 | Alex I. Smith}, 18 | title = {Bounded Linear Types in a Resource Semiring}, 19 | booktitle = {{ESOP} 2014}, 20 | pages = {331--350}, 21 | year = {2014}, 22 | timestamp = {Sun, 23 Mar 2014 10:48:25 +0100}, 23 | biburl = {http://dblp.uni-trier.de/rec/bib/conf/esop/GhicaS14}, 24 | bibsource = {dblp computer science bibliography, http://dblp.org} 25 | } 26 | 27 | @inproceedings{PetricekOM14, 28 | author = {Tomas Petricek and 29 | Dominic A. Orchard and 30 | Alan Mycroft}, 31 | title = {Coeffects: a calculus of context-dependent computation}, 32 | booktitle = {{ICFP} 2014}, 33 | pages = {123--135}, 34 | year = {2014}, 35 | timestamp = {Sun, 04 Jun 2017 10:05:10 +0200}, 36 | biburl = {http://dblp.uni-trier.de/rec/bib/conf/icfp/PetricekOM14}, 37 | bibsource = {dblp computer science bibliography, http://dblp.org} 38 | } 39 | 40 | @article{ girard87linear, 41 | author = "Jean-Yves Girard", 42 | title = "Linear Logic", 43 | journal = "Theor. Comp. Sci.", 44 | volume = 50, 45 | pages = "1--101", 46 | year = 1987 } 47 | 48 | @InProceedings{reed10distance, 49 | author = {J. Reed and B. C. Pierce}, 50 | title = {Distance Makes the Types Grow Stronger}, 51 | booktitle = {ICFP 2010}, 52 | year = 2010, 53 | editor = {P. Hudak and S. Weirich}, 54 | pages = {157-168}} 55 | 56 | @inproceedings{abadi99core, 57 | author = {M. Abadi and 58 | A. Banerjee and 59 | N. Heintze and 60 | J. G. Riecke}, 61 | title = {{A Core Calculus of Dependency}}, 62 | booktitle = {POPL '99}, 63 | year = 1999, 64 | pages = {147-160}, 65 | } 66 | 67 | @article{DBLP:journals/toplas/PierceT00, 68 | author = {Benjamin C. Pierce and David N. Turner}, 69 | title = {Local type inference}, 70 | journal = {{ACM} {TOPLAS}}, 71 | volume = {22}, 72 | number = {1}, 73 | pages = {1--44}, 74 | year = {2000}, 75 | timestamp = {Wed, 26 Nov 2003 14:26:46 +0100}, 76 | biburl = {https://dblp.org/rec/bib/journals/toplas/PierceT00}, 77 | bibsource = {dblp computer science bibliography, https://dblp.org} 78 | } 79 | 80 | @InProceedings{quantitative-type-theory, 81 | title = {The Syntax and Semantics of Quantitative Type Theory}, 82 | author = {Robert Atkey}, 83 | year = {2018}, 84 | doi = {10.1145/3209108.3209189}, 85 | booktitle = {LICS '18: 33rd Annual ACM/IEEE Symposium on Logic in Computer Science, July 9--12, 2018, Oxford, United Kingdom}}'}} 86 | 87 | @techreport{Barber1996, 88 | title = {Dual Intuitionistic Linear Logic}, 89 | author = {Andrew Barber}, 90 | year = {1996}, 91 | institution = {The University of Edinburgh} 92 | } 93 | 94 | @INPROCEEDINGS{judgmental, 95 | author = {Frank Pfenning and Rowan Davies}, 96 | title = {A Judgmental Reconstruction of Modal Logic}, 97 | booktitle = {Mathematical Structures in Computer Science}, 98 | year = {1999}, 99 | pages = {2001} 100 | } 101 | 102 | @InProceedings{context-constrained, 103 | title = {Context Constrained Computation}, 104 | author = {Robert Atkey and James Wood}, 105 | year = {2018}, 106 | booktitle = {3rd Workshop on Type-Driven Development (TyDe '18), Extended Abstract} 107 | } -------------------------------------------------------------------------------- /tex/presentation/default.nix: -------------------------------------------------------------------------------- 1 | { nixpkgs ? import {} }: 2 | nixpkgs.pkgs.callPackage ./pkg.nix { } 3 | -------------------------------------------------------------------------------- /tex/presentation/icfp-presentation-final.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/laMudri/quantitative/2f1d4e5fa21bdb054496e1d202bcd14a76f9fae1/tex/presentation/icfp-presentation-final.pdf -------------------------------------------------------------------------------- /tex/presentation/pkg.nix: -------------------------------------------------------------------------------- 1 | { stdenv, texlive }: 2 | let 3 | tex-env = texlive.combine { 4 | inherit (texlive) scheme-small latexmk beamer stmaryrd mathpartir 5 | cmll xcolor paralist makecell tikz-cd ncctools; 6 | }; 7 | in stdenv.mkDerivation { 8 | name = "quantitative-presentation"; 9 | src = ./.; 10 | buildInputs = [ tex-env ]; 11 | buildPhase = '' 12 | latexmk 13 | ''; 14 | installPhase = ""; 15 | } 16 | -------------------------------------------------------------------------------- /tex/presentation/spls-presentation-final.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/laMudri/quantitative/2f1d4e5fa21bdb054496e1d202bcd14a76f9fae1/tex/presentation/spls-presentation-final.pdf -------------------------------------------------------------------------------- /tex/types-presentation/bread.jpeg: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/laMudri/quantitative/2f1d4e5fa21bdb054496e1d202bcd14a76f9fae1/tex/types-presentation/bread.jpeg -------------------------------------------------------------------------------- /tex/types-presentation/default.nix: -------------------------------------------------------------------------------- 1 | { nixpkgs ? import {} }: 2 | nixpkgs.pkgs.callPackage ./pkg.nix { } 3 | -------------------------------------------------------------------------------- /tex/types-presentation/pkg.nix: -------------------------------------------------------------------------------- 1 | { stdenv, texlive }: 2 | let 3 | tex-env = texlive.combine { 4 | inherit (texlive) scheme-small latexmk beamer stmaryrd mathpartir rsfs 5 | cmll xcolor paralist makecell tikz-cd ncctools; 6 | }; 7 | in stdenv.mkDerivation { 8 | name = "quantitative-types-presentation"; 9 | src = ./.; 10 | buildInputs = [ tex-env ]; 11 | buildPhase = '' 12 | latexmk quantitative.tex 13 | ''; 14 | installPhase = ""; 15 | } 16 | -------------------------------------------------------------------------------- /tex/types-presentation/spls-presentation-final.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/laMudri/quantitative/2f1d4e5fa21bdb054496e1d202bcd14a76f9fae1/tex/types-presentation/spls-presentation-final.pdf -------------------------------------------------------------------------------- /tex/types-presentation/types-presentation-final.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/laMudri/quantitative/2f1d4e5fa21bdb054496e1d202bcd14a76f9fae1/tex/types-presentation/types-presentation-final.pdf --------------------------------------------------------------------------------