├── .gitignore ├── LICENSE ├── README.md └── src └── Web └── Semantic ├── DL ├── ABox.agda ├── ABox │ ├── Interp.agda │ ├── Interp │ │ ├── Meet.agda │ │ └── Morphism.agda │ ├── Model.agda │ └── Skolemization.agda ├── Category.agda ├── Category │ ├── Composition.agda │ ├── Morphism.agda │ ├── Object.agda │ ├── Properties.agda │ ├── Properties │ │ ├── Composition.agda │ │ ├── Composition │ │ │ ├── Assoc.agda │ │ │ ├── LeftUnit.agda │ │ │ ├── Lemmas.agda │ │ │ ├── RespectsEquiv.agda │ │ │ ├── RespectsWiring.agda │ │ │ └── RightUnit.agda │ │ ├── Equivalence.agda │ │ ├── Tensor.agda │ │ ├── Tensor │ │ │ ├── AssocNatural.agda │ │ │ ├── Coherence.agda │ │ │ ├── Functor.agda │ │ │ ├── Isomorphisms.agda │ │ │ ├── Lemmas.agda │ │ │ ├── RespectsEquiv.agda │ │ │ ├── RespectsWiring.agda │ │ │ ├── SymmNatural.agda │ │ │ └── UnitNatural.agda │ │ └── Wiring.agda │ ├── Tensor.agda │ ├── Unit.agda │ └── Wiring.agda ├── Concept.agda ├── Concept │ ├── Model.agda │ └── Skolemization.agda ├── FOL.agda ├── FOL │ └── Model.agda ├── Integrity.agda ├── Integrity │ ├── Closed.agda │ └── Closed │ │ ├── Alternate.agda │ │ └── Properties.agda ├── KB.agda ├── KB │ ├── Model.agda │ └── Skolemization.agda ├── Role.agda ├── Role │ ├── Model.agda │ └── Skolemization.agda ├── Sequent.agda ├── Sequent │ └── Model.agda ├── Signature.agda ├── TBox.agda └── TBox │ ├── Interp.agda │ ├── Interp │ └── Morphism.agda │ ├── Minimizable.agda │ ├── Model.agda │ └── Skolemization.agda ├── Everything.agda └── Util.agda /.gitignore: -------------------------------------------------------------------------------- 1 | # Keep this file in alphabetic order please! 2 | 3 | *~ 4 | .*.swp 5 | *.agdai 6 | *.agda.el 7 | .DS_Store 8 | *.lagda.el 9 | *.hi 10 | *.o 11 | *.tix 12 | *.vim 13 | dist 14 | Everything.agda 15 | EverythingSafe.agda 16 | html -------------------------------------------------------------------------------- /LICENSE: -------------------------------------------------------------------------------- 1 | Copyright (c) 2011 Alcatel-Lucent. 2 | 3 | Permission is hereby granted, free of charge, to any person obtaining a copy 4 | of this software and associated documentation files (the "Software"), to deal 5 | in the Software without restriction, including without limitation the rights 6 | to use, copy, modify, merge, publish, distribute, sublicense, and/or sell 7 | copies of the Software, and to permit persons to whom the Software is 8 | furnished to do so, subject to the following conditions: 9 | 10 | The above copyright notice and this permission notice shall be included in 11 | all copies or substantial portions of the Software. 12 | 13 | THE SOFTWARE IS PROVIDED "AS IS", WITHOUT WARRANTY OF ANY KIND, EXPRESS OR 14 | IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF MERCHANTABILITY, 15 | FITNESS FOR A PARTICULAR PURPOSE AND NONINFRINGEMENT. IN NO EVENT SHALL THE 16 | AUTHORS OR COPYRIGHT HOLDERS BE LIABLE FOR ANY CLAIM, DAMAGES OR OTHER 17 | LIABILITY, WHETHER IN AN ACTION OF CONTRACT, TORT OR OTHERWISE, ARISING FROM, 18 | OUT OF OR IN CONNECTION WITH THE SOFTWARE OR THE USE OR OTHER DEALINGS IN 19 | THE SOFTWARE. 20 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | Agda Semantic Web Libraries 2 | =========================== 3 | 4 | These libraries are intended to support processing of semantic web data in Agda, 5 | and to ensure that all such data satisfies its integrity constraints, expressed 6 | as an ontology. 7 | 8 | The data is expressed in terms of the W3C Resource Description Framework (RDF), 9 | and ontoloties are expressed in terms of the W3C Web Ontology Language (OWL). 10 | 11 | The semantics of RDF and OWL are given as Description Logics, 12 | formalized in the Web.Semantic.DL library. -------------------------------------------------------------------------------- /src/Web/Semantic/DL/ABox.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( ∃ ; _×_ ; _,_ ) 2 | open import Relation.Binary.PropositionalEquality using 3 | ( _≡_ ; refl ; trans ; cong₂ ) 4 | open import Relation.Unary using ( ∅ ; _∪_ ) 5 | open import Web.Semantic.DL.Signature using ( Signature ; CN ; RN ) 6 | open import Web.Semantic.Util using ( id ; _∘_ ; Subset ; ⁅_⁆ ) 7 | 8 | module Web.Semantic.DL.ABox where 9 | 10 | infixr 5 _∼_ _∈₁_ _∈₂_ 11 | infixr 4 _,_ 12 | 13 | data ABox (Σ : Signature) (X : Set) : Set where 14 | ε : ABox Σ X 15 | _,_ : (A B : ABox Σ X) → ABox Σ X 16 | _∼_ : (x y : X) → ABox Σ X 17 | _∈₁_ : (x : X) → (c : CN Σ) → ABox Σ X 18 | _∈₂_ : (xy : X × X) → (r : RN Σ) → ABox Σ X 19 | 20 | ⟨ABox⟩ : ∀ {Σ X Y} → (X → Y) → ABox Σ X → ABox Σ Y 21 | ⟨ABox⟩ f ε = ε 22 | ⟨ABox⟩ f (A , B) = (⟨ABox⟩ f A , ⟨ABox⟩ f B) 23 | ⟨ABox⟩ f (x ∼ y) = (f x ∼ f y) 24 | ⟨ABox⟩ f (x ∈₁ C) = (f x ∈₁ C) 25 | ⟨ABox⟩ f ((x , y) ∈₂ R) = ((f x , f y) ∈₂ R) 26 | 27 | ⟨ABox⟩-resp-id : ∀ {Σ X} (A : ABox Σ X) → (⟨ABox⟩ id A ≡ A) 28 | ⟨ABox⟩-resp-id ε = refl 29 | ⟨ABox⟩-resp-id (A , B) = cong₂ _,_ (⟨ABox⟩-resp-id A) (⟨ABox⟩-resp-id B) 30 | ⟨ABox⟩-resp-id (x ∼ y) = refl 31 | ⟨ABox⟩-resp-id (x ∈₁ c) = refl 32 | ⟨ABox⟩-resp-id ((x , y) ∈₂ r) = refl 33 | 34 | ⟨ABox⟩-resp-∘ : ∀ {Σ X Y Z} (f : Y → Z) (g : X → Y) (A : ABox Σ X) → (⟨ABox⟩ f (⟨ABox⟩ g A) ≡ ⟨ABox⟩ (f ∘ g) A) 35 | ⟨ABox⟩-resp-∘ f g ε = refl 36 | ⟨ABox⟩-resp-∘ f g (A , B) = cong₂ _,_ (⟨ABox⟩-resp-∘ f g A) (⟨ABox⟩-resp-∘ f g B) 37 | ⟨ABox⟩-resp-∘ f g (x ∼ y) = refl 38 | ⟨ABox⟩-resp-∘ f g (x ∈₁ c) = refl 39 | ⟨ABox⟩-resp-∘ f g ((x , y) ∈₂ r) = refl 40 | 41 | ⟨ABox⟩-resp-∘² : ∀ {Σ W X Y Z} (f : Y → Z) (g : X → Y) (h : W → X) → 42 | ∀ (A : ABox Σ W) → (⟨ABox⟩ f (⟨ABox⟩ g (⟨ABox⟩ h A)) ≡ ⟨ABox⟩ (f ∘ g ∘ h) A) 43 | ⟨ABox⟩-resp-∘² f g h A = 44 | trans (⟨ABox⟩-resp-∘ f g (⟨ABox⟩ h A)) (⟨ABox⟩-resp-∘ (f ∘ g) h A) 45 | 46 | Assertions : ∀ {Σ X} → ABox Σ X → Subset (ABox Σ X) 47 | Assertions ε = ∅ 48 | Assertions (A , B) = (Assertions A) ∪ (Assertions B) 49 | Assertions (x ∼ y) = ⁅ x ∼ y ⁆ 50 | Assertions (x ∈₁ c) = ⁅ x ∈₁ c ⁆ 51 | Assertions (xy ∈₂ r) = ⁅ xy ∈₂ r ⁆ 52 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/ABox/Interp.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( ∃ ; _×_ ; _,_ ; proj₁ ; proj₂ ) 2 | open import Relation.Unary using ( _∈_ ) 3 | open import Web.Semantic.DL.TBox.Interp using ( Δ ; _⊨_≈_ ) renaming 4 | ( Interp to Interp′ ; emp to emp′ ) 5 | open import Web.Semantic.DL.Signature using ( Signature ) 6 | open import Web.Semantic.Util using ( False ; id ) 7 | 8 | module Web.Semantic.DL.ABox.Interp where 9 | 10 | infixr 4 _,_ 11 | infixr 5 _*_ 12 | 13 | data Interp (Σ : Signature) (X : Set) : Set₁ where 14 | _,_ : ∀ I → (X → Δ {Σ} I) → (Interp Σ X) 15 | 16 | ⌊_⌋ : ∀ {Σ X} → Interp Σ X → Interp′ Σ 17 | ⌊ I , i ⌋ = I 18 | 19 | ind : ∀ {Σ X} → (I : Interp Σ X) → X → Δ ⌊ I ⌋ 20 | ind (I , i) = i 21 | 22 | ind² : ∀ {Σ X} → (I : Interp Σ X) → (X × X) → (Δ ⌊ I ⌋ × Δ ⌊ I ⌋) 23 | ind² I (x , y) = (ind I x , ind I y) 24 | 25 | _*_ : ∀ {Σ X Y} → (Y → X) → Interp Σ X → Interp Σ Y 26 | f * I = (⌊ I ⌋ , λ x → ind I (f x)) 27 | 28 | emp : ∀ {Σ} → Interp Σ False 29 | emp = (emp′ , id) 30 | 31 | data Surjective {Σ X} (I : Interp Σ X) : Set where 32 | surj : (∀ x → ∃ λ y → ⌊ I ⌋ ⊨ x ≈ ind I y) → (I ∈ Surjective) 33 | 34 | ind⁻¹ : ∀ {Σ X} {I : Interp Σ X} → (I ∈ Surjective) → (Δ ⌊ I ⌋ → X) 35 | ind⁻¹ (surj i) x = proj₁ (i x) 36 | 37 | surj✓ : ∀ {Σ X} {I : Interp Σ X} (I∈Surj : I ∈ Surjective) → ∀ x → (⌊ I ⌋ ⊨ x ≈ ind I (ind⁻¹ I∈Surj x)) 38 | surj✓ (surj i) x = proj₂ (i x) 39 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/ABox/Interp/Meet.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( _,_ ) 2 | open import Relation.Nullary using ( yes ; no ) 3 | open import Relation.Unary using ( _∈_ ) 4 | open import Web.Semantic.DL.ABox.Interp using 5 | ( Interp ; _,_ ; ⌊_⌋ ; ind ; ind² ; Surjective ; surj ; ind⁻¹ ; surj✓ ) 6 | open import Web.Semantic.DL.ABox.Interp.Morphism using ( _≲_ ; _,_ ; ≲⌊_⌋ ; ≲-resp-ind ; _≋_ ) 7 | open import Web.Semantic.DL.Signature using ( Signature ) 8 | open import Web.Semantic.DL.TBox.Interp using ( interp ; _⊨_≈_ ; ≈-refl ; ≈-sym ; ≈-trans ; con ; rol ; con-≈ ; rol-≈ ) 9 | open import Web.Semantic.DL.TBox.Interp.Morphism using ( morph ; ≲-image ; ≲-resp-≈ ; ≲-resp-con ; ≲-resp-rol ) 10 | open import Web.Semantic.Util using ( _∘_ ; ExclMiddle₁ ; ExclMiddle ; smaller-excl-middle ; id ; □ ; is ; is! ; is✓ ) 11 | 12 | -- Note that this construction uses excluded middle 13 | -- (for cardinality reasons, to embed a large set in a small set). 14 | 15 | module Web.Semantic.DL.ABox.Interp.Meet ( excl-middle₁ : ExclMiddle₁ ) { Σ : Signature } { X : Set } where 16 | 17 | -- This constructs the meet of a set of interpretations, 18 | -- which is the greatest lower bound on surjective interpretations 19 | 20 | meet : (Interp Σ X → Set) → Interp Σ X 21 | meet Is = 22 | ( interp X 23 | (λ x y → □ (is (excl-middle₁ (∀ I → (I ∈ Is) → 24 | (⌊ I ⌋ ⊨ ind I x ≈ ind I y))))) 25 | (λ {x} → is! (λ I I∈Is → 26 | ≈-refl ⌊ I ⌋)) 27 | (λ {x} {y} x≈y → is! (λ I I∈Is → 28 | ≈-sym ⌊ I ⌋ (is✓ x≈y I I∈Is))) 29 | (λ {x} {y} {z} x≈y y≈z → is! (λ I I∈Is → 30 | ≈-trans ⌊ I ⌋ (is✓ x≈y I I∈Is) (is✓ y≈z I I∈Is))) 31 | (λ c x → □ (is (excl-middle₁ (∀ I → (I ∈ Is) → 32 | (ind I x ∈ con ⌊ I ⌋ c))))) 33 | (λ r xy → □ (is (excl-middle₁ (∀ I → (I ∈ Is) → 34 | ind² I xy ∈ rol ⌊ I ⌋ r)))) 35 | (λ {x} {y} c x∈c x≈y → is! (λ I I∈Is → 36 | con-≈ ⌊ I ⌋ c (is✓ x∈c I I∈Is) (is✓ x≈y I I∈Is))) 37 | (λ {w} {x} {y} {z} r w≈x xy∈r y≈z → is! (λ I I∈Is → 38 | rol-≈ ⌊ I ⌋ r (is✓ w≈x I I∈Is) (is✓ xy∈r I I∈Is) (is✓ y≈z I I∈Is))) 39 | , id ) 40 | 41 | -- meets are surjective 42 | 43 | meet-surj : ∀ Is → (meet Is ∈ Surjective) 44 | meet-surj Is = surj (λ x → (x , ≈-refl ⌊ meet Is ⌋)) 45 | 46 | -- meets are lower bounds 47 | 48 | meet-lb : ∀ Is I → (I ∈ Is) → (meet Is ≲ I) 49 | meet-lb Is I I∈Is = 50 | ( morph (ind I) 51 | (λ x≈y → is✓ x≈y I I∈Is) 52 | (λ x∈⟦c⟧ → is✓ x∈⟦c⟧ I I∈Is) 53 | (λ xy∈⟦r⟧ → is✓ xy∈⟦r⟧ I I∈Is) 54 | , λ x → ≈-refl ⌊ I ⌋ ) 55 | 56 | -- meets are greatest lower bounds on surjective interpretations 57 | 58 | meet-glb : ∀ Is J → (J ∈ Surjective) → (∀ I → (I ∈ Is) → (J ≲ I)) → (J ≲ meet Is) 59 | meet-glb Is J J∈Surj J≲Is = 60 | ( morph (ind⁻¹ J∈Surj) 61 | (λ {x} {y} x≈y → is! (λ I I∈Is → 62 | ≈-trans ⌊ I ⌋ (≈-sym ⌊ I ⌋ (lemma I I∈Is x)) (≈-trans ⌊ I ⌋ (≲-resp-≈ ≲⌊ J≲Is I I∈Is ⌋ x≈y) (lemma I I∈Is y)))) 63 | (λ {c} {x} x∈⟦c⟧ → is! (λ I I∈Is → 64 | con-≈ ⌊ I ⌋ c (≲-resp-con ≲⌊ J≲Is I I∈Is ⌋ x∈⟦c⟧) (lemma I I∈Is x))) 65 | (λ {r} {x} {y} xy∈⟦r⟧ → is! (λ I I∈Is → 66 | rol-≈ ⌊ I ⌋ r (≈-sym ⌊ I ⌋ (lemma I I∈Is x)) (≲-resp-rol ≲⌊ J≲Is I I∈Is ⌋ xy∈⟦r⟧) (lemma I I∈Is y))) 67 | , λ x → is! (λ I I∈Is → 68 | ≈-trans ⌊ I ⌋ (≈-sym ⌊ I ⌋ (lemma I I∈Is (ind J x))) (≲-resp-ind (J≲Is I I∈Is) x)) ) where 69 | 70 | lemma : ∀ I I∈Is x → ⌊ I ⌋ ⊨ ≲-image ≲⌊ J≲Is I I∈Is ⌋ x ≈ ind I (ind⁻¹ J∈Surj x) 71 | lemma I I∈Is x = 72 | ≈-trans ⌊ I ⌋ (≲-resp-≈ ≲⌊ J≲Is I I∈Is ⌋ (surj✓ J∈Surj x)) 73 | (≲-resp-ind (J≲Is I I∈Is) (ind⁻¹ J∈Surj x)) 74 | 75 | -- Mediating morphisms for meets are unique 76 | 77 | meet-uniq : ∀ Is (I : Interp Σ X) → (I ∈ Is) → (J≲₁I J≲₂I : meet Is ≲ I) → (J≲₁I ≋ J≲₂I) 78 | meet-uniq Is I I∈Is J≲₁I J≲₂I x = ≈-trans ⌊ I ⌋ (≲-resp-ind J≲₁I x) (≈-sym ⌊ I ⌋ (≲-resp-ind J≲₂I x)) 79 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/ABox/Interp/Morphism.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( ∃ ; _×_ ; _,_ ; proj₁ ; proj₂ ) 2 | open import Relation.Unary using ( _∈_ ) 3 | open import Relation.Binary.PropositionalEquality using ( _≡_ ; refl ; cong ) 4 | open import Web.Semantic.DL.ABox.Interp using ( Interp ; _,_ ; ⌊_⌋ ; ind ; _*_ ) 5 | open import Web.Semantic.DL.Signature using ( Signature ) 6 | open import Web.Semantic.DL.TBox.Interp using 7 | ( Δ ; _⊨_≈_ ; ≈-refl ; ≈-sym ; ≈-trans ) renaming 8 | ( Interp to Interp′ ) 9 | open import Web.Semantic.DL.TBox.Interp.Morphism using 10 | ( ≲-image ; ≃-image ; ≃-image⁻¹ ; ≲-resp-≈ ; ≃-resp-≈ ; ≃-iso ) renaming 11 | ( _≲_ to _≲′_ ; _≃_ to _≃′_ 12 | ; ≲-refl to ≲′-refl ; ≲-trans to ≲′-trans 13 | ; ≃-refl to ≃′-refl ; ≃-sym to ≃′-sym ; ≃-trans to ≃′-trans 14 | ; ≃-impl-≲ to ≃-impl-≲′ ) 15 | open import Web.Semantic.Util using 16 | ( _∘_ ; _⊕_⊕_ ; inode ; bnode ; enode ; →-dist-⊕ ) 17 | 18 | module Web.Semantic.DL.ABox.Interp.Morphism {Σ : Signature} where 19 | 20 | infix 2 _≲_ _≃_ _≋_ 21 | infixr 4 _,_ 22 | infix 5 _**_ 23 | 24 | data _≲_ {X} (I J : Interp Σ X) : Set where 25 | _,_ : ∀ (I≲J : ⌊ I ⌋ ≲′ ⌊ J ⌋) → 26 | (i≲j : ∀ x → (⌊ J ⌋ ⊨ ≲-image I≲J (ind I x) ≈ ind J x)) → 27 | (I ≲ J) 28 | 29 | ≲⌊_⌋ : ∀ {X} {I J : Interp Σ X} → (I ≲ J) → (⌊ I ⌋ ≲′ ⌊ J ⌋) 30 | ≲⌊_⌋ (I≲J , i≲j) = I≲J 31 | 32 | ≲-resp-ind : ∀ {X} {I J : Interp Σ X} → (I≲J : I ≲ J) → 33 | ∀ x → (⌊ J ⌋ ⊨ ≲-image ≲⌊ I≲J ⌋ (ind I x) ≈ ind J x) 34 | ≲-resp-ind (I≲J , i≲j) = i≲j 35 | 36 | ≲-refl : ∀ {X} (I : Interp Σ X) → (I ≲ I) 37 | ≲-refl I = (≲′-refl ⌊ I ⌋ , λ x → ≈-refl ⌊ I ⌋) 38 | 39 | ≲-trans : ∀ {X} {I J K : Interp Σ X} → (I ≲ J) → (J ≲ K) → (I ≲ K) 40 | ≲-trans {X} {I} {J} {K} I≲J J≲K = 41 | (≲′-trans ≲⌊ I≲J ⌋ ≲⌊ J≲K ⌋ , λ x → ≈-trans ⌊ K ⌋ (≲-resp-≈ ≲⌊ J≲K ⌋ (≲-resp-ind I≲J x)) (≲-resp-ind J≲K x)) 42 | 43 | ≡-impl-≈ : ∀ {X : Set} (I : Interp′ Σ) {i j : X → Δ I} → 44 | (i ≡ j) → (∀ x → I ⊨ i x ≈ j x) 45 | ≡-impl-≈ I refl x = ≈-refl I 46 | 47 | ≡³-impl-≈ : ∀ {V X Y : Set} (I : Interp′ Σ) (i j : (X ⊕ V ⊕ Y) → Δ I) → 48 | (→-dist-⊕ i ≡ →-dist-⊕ j) → (∀ x → I ⊨ i x ≈ j x) 49 | ≡³-impl-≈ I i j i≡j (inode x) = ≡-impl-≈ I (cong proj₁ i≡j) x 50 | ≡³-impl-≈ I i j i≡j (bnode v) = ≡-impl-≈ I (cong proj₁ (cong proj₂ i≡j)) v 51 | ≡³-impl-≈ I i j i≡j (enode y) = ≡-impl-≈ I (cong proj₂ (cong proj₂ i≡j)) y 52 | 53 | ≡-impl-≲ : ∀ {X : Set} (I : Interp Σ X) j → 54 | (ind I ≡ j) → (I ≲ (⌊ I ⌋ , j)) 55 | ≡-impl-≲ (I , i) .i refl = ≲-refl (I , i) 56 | 57 | ≡³-impl-≲ : ∀ {V X Y : Set} (I : Interp Σ (X ⊕ V ⊕ Y)) j → 58 | (→-dist-⊕ (ind I) ≡ →-dist-⊕ j) → (I ≲ (⌊ I ⌋ , j)) 59 | ≡³-impl-≲ (I , i) j i≡j = (≲′-refl I , ≡³-impl-≈ I i j i≡j) 60 | 61 | _**_ : ∀ {X Y I J} (f : Y → X) → (I ≲ J) → (f * I ≲ f * J) 62 | f ** I≲J = (≲⌊ I≲J ⌋ , λ x → ≲-resp-ind I≲J (f x)) 63 | 64 | _≋_ : ∀ {X} {I J : Interp Σ X} → (I ≲ J) → (I ≲ J) → Set 65 | _≋_ {X} {I} {J} I≲₁J I≲₂J = ∀ x → (⌊ J ⌋ ⊨ ≲-image ≲⌊ I≲₁J ⌋ x ≈ ≲-image ≲⌊ I≲₂J ⌋ x) 66 | 67 | data _≃_ {X} (I J : Interp Σ X) : Set₁ where 68 | _,_ : (I≃J : ⌊ I ⌋ ≃′ ⌊ J ⌋) → 69 | (∀ x → (⌊ J ⌋ ⊨ ≃-image I≃J (ind I x) ≈ ind J x)) → 70 | (I ≃ J) 71 | 72 | ≃⌊_⌋ : ∀ {X} {I J : Interp Σ X} → (I ≃ J) → (⌊ I ⌋ ≃′ ⌊ J ⌋) 73 | ≃⌊_⌋ (I≃J , i≃j) = I≃J 74 | 75 | ≃-impl-≲ : ∀ {X} {I J : Interp Σ X} → (I ≃ J) → (I ≲ J) 76 | ≃-impl-≲ (I≃J , i≃j) = (≃-impl-≲′ I≃J , i≃j) 77 | 78 | ≃-resp-ind : ∀ {X} {I J : Interp Σ X} → (I≃J : I ≃ J) → 79 | ∀ x → (⌊ J ⌋ ⊨ ≃-image ≃⌊ I≃J ⌋ (ind I x) ≈ ind J x) 80 | ≃-resp-ind (I≃J , i≃j) = i≃j 81 | 82 | ≃-resp-ind⁻¹ : ∀ {X} {I J : Interp Σ X} → (I≃J : I ≃ J) → 83 | ∀ x → (⌊ I ⌋ ⊨ ≃-image⁻¹ ≃⌊ I≃J ⌋ (ind J x) ≈ ind I x) 84 | ≃-resp-ind⁻¹ {X} {I , i} {J , j} (I≃J , i≃j) x = 85 | ≈-trans I (≃-resp-≈ (≃′-sym I≃J) (≈-sym J (i≃j x))) (≃-iso I≃J (i x)) 86 | 87 | ≃-refl : ∀ {X} (I : Interp Σ X) → (I ≃ I) 88 | ≃-refl (I , i) = (≃′-refl I , λ x → ≈-refl I) 89 | 90 | ≃-symm : ∀ {X} {I J : Interp Σ X} → (I ≃ J) → (J ≃ I) 91 | ≃-symm {X} {I , i} {J , j} (I≃J , i≃j) = 92 | (≃′-sym I≃J , ≃-resp-ind⁻¹ {X} {I , i} {J , j} (I≃J , i≃j)) 93 | 94 | ≃-trans : ∀ {X} {I J K : Interp Σ X} → (I ≃ J) → (J ≃ K) → (I ≃ K) 95 | ≃-trans {X} {I , i} {J , j} {K , k} (I≃J , i≃j) (J≃K , j≃k) = 96 | (≃′-trans I≃J J≃K , λ x → ≈-trans K (≃-resp-≈ J≃K (i≃j x)) (j≃k x)) 97 | 98 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/ABox/Model.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( _×_ ; _,_ ) 2 | open import Data.Sum using ( inj₁ ; inj₂ ) 3 | open import Relation.Binary.PropositionalEquality using ( _≡_ ; refl ) 4 | open import Relation.Unary using ( _∈_ ; _⊆_ ) 5 | open import Web.Semantic.DL.ABox using 6 | ( ABox ; Assertions ; ⟨ABox⟩ ; ε ; _,_ ; _∼_ ; _∈₁_ ; _∈₂_ ) 7 | open import Web.Semantic.DL.ABox.Interp using ( Interp ; _,_ ; ⌊_⌋ ; ind ; _*_ ) 8 | open import Web.Semantic.DL.ABox.Interp.Morphism using 9 | ( _≲_ ; _,_ ; ≲⌊_⌋ ; ≲-resp-ind ; ≡³-impl-≲ ) 10 | open import Web.Semantic.DL.Signature using ( Signature ) 11 | open import Web.Semantic.DL.TBox.Interp using 12 | ( Δ ; _⊨_≈_ ; ≈-refl ; ≈-sym ; ≈-trans ; con ; rol ; con-≈ ; rol-≈ ) 13 | open import Web.Semantic.DL.TBox.Interp.Morphism using 14 | ( ≲-image ; ≲-resp-≈ ; ≲-resp-con ; ≲-resp-rol ; ≲-refl ) 15 | open import Web.Semantic.Util using 16 | ( True ; tt ; id ; _∘_ ; _⊕_⊕_ ; inode ; bnode ; enode ; →-dist-⊕ ) 17 | 18 | module Web.Semantic.DL.ABox.Model {Σ : Signature} where 19 | 20 | infix 2 _⊨a_ _⊨b_ 21 | infixr 5 _,_ 22 | 23 | _⟦_⟧₀ : ∀ {X} (I : Interp Σ X) → X → (Δ ⌊ I ⌋) 24 | I ⟦ x ⟧₀ = ind I x 25 | 26 | _⊨a_ : ∀ {X} → Interp Σ X → ABox Σ X → Set 27 | I ⊨a ε = True 28 | I ⊨a (A , B) = (I ⊨a A) × (I ⊨a B) 29 | I ⊨a x ∼ y = ⌊ I ⌋ ⊨ ind I x ≈ ind I y 30 | I ⊨a x ∈₁ c = ind I x ∈ con ⌊ I ⌋ c 31 | I ⊨a (x , y) ∈₂ r = (ind I x , ind I y) ∈ rol ⌊ I ⌋ r 32 | 33 | Assertions✓ : ∀ {X} (I : Interp Σ X) A {a} → 34 | (a ∈ Assertions A) → (I ⊨a A) → (I ⊨a a) 35 | Assertions✓ I ε () I⊨A 36 | Assertions✓ I (A , B) (inj₁ a∈A) (I⊨A , I⊨B) = Assertions✓ I A a∈A I⊨A 37 | Assertions✓ I (A , B) (inj₂ a∈B) (I⊨A , I⊨B) = Assertions✓ I B a∈B I⊨B 38 | Assertions✓ I (i ∼ j) refl I⊨A = I⊨A 39 | Assertions✓ I (i ∈₁ c) refl I⊨A = I⊨A 40 | Assertions✓ I (ij ∈₂ r) refl I⊨A = I⊨A 41 | 42 | ⊨a-resp-⊇ : ∀ {X} (I : Interp Σ X) A B → 43 | (Assertions A ⊆ Assertions B) → (I ⊨a B) → (I ⊨a A) 44 | ⊨a-resp-⊇ I ε B A⊆B I⊨B = tt 45 | ⊨a-resp-⊇ I (A₁ , A₂) B A⊆B I⊨B = 46 | ( ⊨a-resp-⊇ I A₁ B (A⊆B ∘ inj₁) I⊨B 47 | , ⊨a-resp-⊇ I A₂ B (A⊆B ∘ inj₂) I⊨B ) 48 | ⊨a-resp-⊇ I (x ∼ y) B A⊆B I⊨B = Assertions✓ I B (A⊆B refl) I⊨B 49 | ⊨a-resp-⊇ I (x ∈₁ c) B A⊆B I⊨B = Assertions✓ I B (A⊆B refl) I⊨B 50 | ⊨a-resp-⊇ I (xy ∈₂ r) B A⊆B I⊨B = Assertions✓ I B (A⊆B refl) I⊨B 51 | 52 | ⊨a-resp-≲ : ∀ {X} {I J : Interp Σ X} → (I ≲ J) → ∀ A → (I ⊨a A) → (J ⊨a A) 53 | ⊨a-resp-≲ {X} {I} {J} I≲J ε I⊨A = 54 | tt 55 | ⊨a-resp-≲ {X} {I} {J} I≲J (A , B) (I⊨A , I⊨B) = 56 | (⊨a-resp-≲ I≲J A I⊨A , ⊨a-resp-≲ I≲J B I⊨B) 57 | ⊨a-resp-≲ {X} {I} {J} I≲J (x ∼ y) I⊨x∼y = 58 | ≈-trans ⌊ J ⌋ (≈-sym ⌊ J ⌋ (≲-resp-ind I≲J x)) 59 | (≈-trans ⌊ J ⌋ (≲-resp-≈ ≲⌊ I≲J ⌋ I⊨x∼y) 60 | (≲-resp-ind I≲J y)) 61 | ⊨a-resp-≲ {X} {I} {J} I≲J (x ∈₁ c) I⊨x∈c = 62 | con-≈ ⌊ J ⌋ c (≲-resp-con ≲⌊ I≲J ⌋ I⊨x∈c) (≲-resp-ind I≲J x) 63 | ⊨a-resp-≲ {X} {I} {J} I≲J ((x , y) ∈₂ r) I⊨xy∈r = 64 | rol-≈ ⌊ J ⌋ r (≈-sym ⌊ J ⌋ (≲-resp-ind I≲J x)) 65 | (≲-resp-rol ≲⌊ I≲J ⌋ I⊨xy∈r) (≲-resp-ind I≲J y) 66 | 67 | ⊨a-resp-≡ : ∀ {X : Set} (I : Interp Σ X) j → 68 | (ind I ≡ j) → ∀ A → (I ⊨a A) → (⌊ I ⌋ , j ⊨a A) 69 | ⊨a-resp-≡ (I , i) .i refl A I⊨A = I⊨A 70 | 71 | ⊨a-resp-≡³ : ∀ {V X Y : Set} (I : Interp Σ (X ⊕ V ⊕ Y)) j → 72 | (→-dist-⊕ (ind I) ≡ →-dist-⊕ j) → ∀ A → (I ⊨a A) → (⌊ I ⌋ , j ⊨a A) 73 | ⊨a-resp-≡³ I j i≡j = ⊨a-resp-≲ (≡³-impl-≲ I j i≡j) 74 | 75 | ⟨ABox⟩-Assertions : ∀ {X Y a} (f : X → Y) (A : ABox Σ X) → 76 | (a ∈ Assertions A) → (⟨ABox⟩ f a ∈ Assertions (⟨ABox⟩ f A)) 77 | ⟨ABox⟩-Assertions f ε () 78 | ⟨ABox⟩-Assertions f (A , B) (inj₁ a∈A) = inj₁ (⟨ABox⟩-Assertions f A a∈A) 79 | ⟨ABox⟩-Assertions f (A , B) (inj₂ a∈B) = inj₂ (⟨ABox⟩-Assertions f B a∈B) 80 | ⟨ABox⟩-Assertions f (x ∼ y) refl = refl 81 | ⟨ABox⟩-Assertions f (x ∈₁ c) refl = refl 82 | ⟨ABox⟩-Assertions f ((x , y) ∈₂ r) refl = refl 83 | 84 | ⟨ABox⟩-resp-⊨ : ∀ {X Y} {I : Interp Σ X} {j : Y → Δ ⌊ I ⌋} 85 | (f : X → Y) → (∀ x → ⌊ I ⌋ ⊨ ind I x ≈ j (f x)) → 86 | ∀ A → (I ⊨a A) → (⌊ I ⌋ , j ⊨a ⟨ABox⟩ f A) 87 | ⟨ABox⟩-resp-⊨ {X} {Y} {I} f i≈j∘f ε I⊨ε = 88 | tt 89 | ⟨ABox⟩-resp-⊨ {X} {Y} {I} f i≈j∘f (A , B) (I⊨A , I⊨B) = 90 | (⟨ABox⟩-resp-⊨ f i≈j∘f A I⊨A , ⟨ABox⟩-resp-⊨ f i≈j∘f B I⊨B) 91 | ⟨ABox⟩-resp-⊨ {X} {Y} {I} f i≈j∘f (x ∼ y) x≈y = 92 | ≈-trans ⌊ I ⌋ (≈-sym ⌊ I ⌋ (i≈j∘f x)) (≈-trans ⌊ I ⌋ x≈y (i≈j∘f y)) 93 | ⟨ABox⟩-resp-⊨ {X} {Y} {I} f i≈j∘f (x ∈₁ c) x∈⟦c⟧ = 94 | con-≈ ⌊ I ⌋ c x∈⟦c⟧ (i≈j∘f x) 95 | ⟨ABox⟩-resp-⊨ {X} {Y} {I} f i≈j∘f ((x , y) ∈₂ r) xy∈⟦r⟧ = 96 | rol-≈ ⌊ I ⌋ r (≈-sym ⌊ I ⌋ (i≈j∘f x)) xy∈⟦r⟧ (i≈j∘f y) 97 | 98 | *-resp-⟨ABox⟩ : ∀ {X Y} (f : Y → X) I A → 99 | (I ⊨a ⟨ABox⟩ f A) → (f * I ⊨a A) 100 | *-resp-⟨ABox⟩ f (I , i) ε I⊨ε = 101 | tt 102 | *-resp-⟨ABox⟩ f (I , i) (A , B) (I⊨A , I⊨B) = 103 | (*-resp-⟨ABox⟩ f (I , i) A I⊨A , *-resp-⟨ABox⟩ f (I , i) B I⊨B ) 104 | *-resp-⟨ABox⟩ f (I , i) (x ∼ y) x≈y = 105 | x≈y 106 | *-resp-⟨ABox⟩ f (I , i) (x ∈₁ c) x∈⟦c⟧ = 107 | x∈⟦c⟧ 108 | *-resp-⟨ABox⟩ f (I , i) ((x , y) ∈₂ r) xy∈⟦c⟧ = 109 | xy∈⟦c⟧ 110 | 111 | -- bnodes I f is the same as I, except that f is used as the interpretation 112 | -- for bnodes. 113 | 114 | on-bnode : ∀ {V W X Y Z : Set} → (W → Z) → ((X ⊕ V ⊕ Y) → Z) → 115 | ((X ⊕ W ⊕ Y) → Z) 116 | on-bnode f g (inode x) = g (inode x) 117 | on-bnode f g (bnode w) = f w 118 | on-bnode f g (enode y) = g (enode y) 119 | 120 | bnodes : ∀ {V W X Y} → (I : Interp Σ (X ⊕ V ⊕ Y)) → (W → Δ ⌊ I ⌋) → 121 | Interp Σ (X ⊕ W ⊕ Y) 122 | bnodes I f = (⌊ I ⌋ , on-bnode f (ind I)) 123 | 124 | bnodes-resp-≲ : ∀ {V W X Y} (I J : Interp Σ (X ⊕ V ⊕ Y)) → 125 | (I≲J : I ≲ J) → (f : W → Δ ⌊ I ⌋) → 126 | (bnodes I f ≲ bnodes J (≲-image ≲⌊ I≲J ⌋ ∘ f)) 127 | bnodes-resp-≲ (I , i) (J , j) (I≲J , i≲j) f = (I≲J , lemma) where 128 | 129 | lemma : ∀ x → 130 | J ⊨ ≲-image I≲J (on-bnode f i x) ≈ on-bnode (≲-image I≲J ∘ f) j x 131 | lemma (inode x) = i≲j (inode x) 132 | lemma (bnode v) = ≈-refl J 133 | lemma (enode y) = i≲j (enode y) 134 | 135 | -- I ⊨b A whenever there exists an f such that bnodes I f ⊨a A 136 | 137 | data _⊨b_ {V W X Y} (I : Interp Σ (X ⊕ V ⊕ Y)) 138 | (A : ABox Σ (X ⊕ W ⊕ Y)) : Set where 139 | _,_ : ∀ f → (bnodes I f ⊨a A) → (I ⊨b A) 140 | 141 | inb : ∀ {V W X Y} {I : Interp Σ (X ⊕ V ⊕ Y)} {A : ABox Σ (X ⊕ W ⊕ Y)} → 142 | (I ⊨b A) → W → Δ ⌊ I ⌋ 143 | inb (f , I⊨A) = f 144 | 145 | ⊨b-impl-⊨a : ∀ {V W X Y} {I : Interp Σ (X ⊕ V ⊕ Y)} {A : ABox Σ (X ⊕ W ⊕ Y)} → 146 | (I⊨A : I ⊨b A) → (bnodes I (inb I⊨A) ⊨a A) 147 | ⊨b-impl-⊨a (f , I⊨A) = I⊨A 148 | 149 | ⊨a-impl-⊨b : ∀ {V X Y} (I : Interp Σ (X ⊕ V ⊕ Y)) A → (I ⊨a A) → (I ⊨b A) 150 | ⊨a-impl-⊨b I A I⊨A = 151 | (ind I ∘ bnode , ⊨a-resp-≲ (≲-refl ⌊ I ⌋ , lemma) A I⊨A) where 152 | 153 | lemma : ∀ x → ⌊ I ⌋ ⊨ ind I x ≈ on-bnode (ind I ∘ bnode) (ind I) x 154 | lemma (inode x) = ≈-refl ⌊ I ⌋ 155 | lemma (bnode v) = ≈-refl ⌊ I ⌋ 156 | lemma (enode y) = ≈-refl ⌊ I ⌋ 157 | 158 | ⊨b-resp-≲ : ∀ {V W X Y} {I J : Interp Σ (X ⊕ V ⊕ Y)} → (I ≲ J) 159 | → ∀ (A : ABox Σ (X ⊕ W ⊕ Y)) → (I ⊨b A) → (J ⊨b A) 160 | ⊨b-resp-≲ I≲J A (f , I⊨A) = 161 | ((≲-image ≲⌊ I≲J ⌋ ∘ f) , ⊨a-resp-≲ (bnodes-resp-≲ _ _ I≲J f) A I⊨A) 162 | 163 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/ABox/Skolemization.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( _×_ ; _,_ ) 2 | open import Relation.Unary using ( _∈_ ) 3 | open import Web.Semantic.DL.ABox using ( ABox ; ε ; _,_ ; _∼_ ; _∈₁_ ; _∈₂_ ) 4 | open import Web.Semantic.DL.ABox.Interp using ( ⌊_⌋ ; ind ) 5 | open import Web.Semantic.DL.ABox.Model using ( _⊨a_ ) 6 | open import Web.Semantic.DL.Concept.Model using ( ⟦⟧₁-resp-≈ ) 7 | open import Web.Semantic.DL.Concept.Skolemization using ( CSkolems ; cskolem ; cskolem-sound ) 8 | open import Web.Semantic.DL.FOL using ( Formula ; true ; _∧_ ; _∈₁_ ; _∈₂_ ; _∼_ ) 9 | open import Web.Semantic.DL.FOL.Model using ( _⊨f_ ) 10 | open import Web.Semantic.DL.Role.Skolemization using ( rskolem ; rskolem-sound ) 11 | open import Web.Semantic.DL.Role.Model using ( ⟦⟧₂-resp-≈ ) 12 | open import Web.Semantic.DL.Signature using ( Signature ) 13 | open import Web.Semantic.Util using ( True ; tt ) 14 | 15 | module Web.Semantic.DL.ABox.Skolemization {Σ : Signature} {X : Set} where 16 | 17 | askolem : ∀ {Δ} → (X → Δ) → ABox Σ X → Formula Σ Δ 18 | askolem i ε = true 19 | askolem i (A , B) = (askolem i A) ∧ (askolem i B) 20 | askolem i (x ∼ y) = i x ∼ i y 21 | askolem i (x ∈₁ c) = i x ∈₁ c 22 | askolem i ((x , y) ∈₂ r) = (i x , i y) ∈₂ r 23 | 24 | askolem-sound : ∀ I A → (⌊ I ⌋ ⊨f askolem (ind I) A) → (I ⊨a A) 25 | askolem-sound I ε _ = tt 26 | askolem-sound I (A , B) (I⊨A , I⊨B) = (askolem-sound I A I⊨A , askolem-sound I B I⊨B) 27 | askolem-sound I (x ∼ y) I⊨x∼y = I⊨x∼y 28 | askolem-sound I (x ∈₁ c) I⊨x∈c = I⊨x∈c 29 | askolem-sound I ((x , y) ∈₂ r) I⊨xy∈r = I⊨xy∈r 30 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Category.agda: -------------------------------------------------------------------------------- 1 | open import Web.Semantic.DL.Signature using ( Signature ) 2 | open import Web.Semantic.DL.TBox using ( TBox ) 3 | 4 | import Web.Semantic.DL.Category.Composition 5 | import Web.Semantic.DL.Category.Morphism 6 | import Web.Semantic.DL.Category.Object 7 | import Web.Semantic.DL.Category.Properties 8 | import Web.Semantic.DL.Category.Tensor 9 | import Web.Semantic.DL.Category.Unit 10 | import Web.Semantic.DL.Category.Wiring 11 | 12 | module Web.Semantic.DL.Category {Σ : Signature} where 13 | 14 | infix 2 _≣_ _⇒_ 15 | infixr 5 _∙_ _⊗_ _⟨⊗⟩_ 16 | 17 | -- Categorial structure 18 | 19 | Object : TBox Σ → TBox Σ → Set₁ 20 | Object = Web.Semantic.DL.Category.Object.Object 21 | 22 | _⇒_ : ∀ {S T} → (Object S T) → (Object S T) → Set₁ 23 | _⇒_ = Web.Semantic.DL.Category.Morphism._⇒_ 24 | 25 | identity : ∀ {S T} (A : Object S T) → (A ⇒ A) 26 | identity = Web.Semantic.DL.Category.Wiring.identity 27 | 28 | _∙_ : ∀ {S T} {A B C : Object S T} → (A ⇒ B) → (B ⇒ C) → (A ⇒ C) 29 | _∙_ = Web.Semantic.DL.Category.Composition._∙_ 30 | 31 | -- Symmetric monoidal structure 32 | 33 | I : ∀ {S T} → Object S T 34 | I = Web.Semantic.DL.Category.Unit.I 35 | 36 | _⊗_ : ∀ {S T} → Object S T → Object S T → Object S T 37 | _⊗_ = Web.Semantic.DL.Category.Tensor._⊗_ 38 | 39 | _⟨⊗⟩_ : ∀ {S T : TBox Σ} {A₁ A₂ B₁ B₂ : Object S T} → 40 | (A₁ ⇒ B₁) → (A₂ ⇒ B₂) → ((A₁ ⊗ A₂) ⇒ (B₁ ⊗ B₂)) 41 | _⟨⊗⟩_ = Web.Semantic.DL.Category.Tensor._⟨⊗⟩_ 42 | 43 | symm : ∀ {S T : TBox Σ} → (A B : Object S T) → ((A ⊗ B) ⇒ (B ⊗ A)) 44 | symm = Web.Semantic.DL.Category.Wiring.symm 45 | 46 | assoc : ∀ {S T : TBox Σ} → (A B C : Object S T) → 47 | (((A ⊗ B) ⊗ C) ⇒ (A ⊗ (B ⊗ C))) 48 | assoc = Web.Semantic.DL.Category.Wiring.assoc 49 | 50 | assoc⁻¹ : ∀ {S T : TBox Σ} → (A B C : Object S T) → 51 | ((A ⊗ (B ⊗ C)) ⇒ ((A ⊗ B) ⊗ C)) 52 | assoc⁻¹ = Web.Semantic.DL.Category.Wiring.assoc⁻¹ 53 | 54 | unit₁ : ∀ {S T : TBox Σ} (A : Object S T) → ((I ⊗ A) ⇒ A) 55 | unit₁ = Web.Semantic.DL.Category.Wiring.unit₁ 56 | 57 | unit₁⁻¹ : ∀ {S T : TBox Σ} (A : Object S T) → (A ⇒ (I ⊗ A)) 58 | unit₁⁻¹ = Web.Semantic.DL.Category.Wiring.unit₁⁻¹ 59 | 60 | unit₂ : ∀ {S T : TBox Σ} (A : Object S T) → ((A ⊗ I) ⇒ A) 61 | unit₂ = Web.Semantic.DL.Category.Wiring.unit₂ 62 | 63 | unit₂⁻¹ : ∀ {S T : TBox Σ} (A : Object S T) → (A ⇒ (A ⊗ I)) 64 | unit₂⁻¹ = Web.Semantic.DL.Category.Wiring.unit₂⁻¹ 65 | 66 | -- Equivalence of morphisms 67 | 68 | _≣_ : ∀ {S T} {A B : Object S T} → (A ⇒ B) → (A ⇒ B) → Set₁ 69 | _≣_ = Web.Semantic.DL.Category.Morphism._≣_ 70 | 71 | ≣-refl : ∀ {S T} {A B : Object S T} (F : A ⇒ B) → 72 | (F ≣ F) 73 | ≣-refl = Web.Semantic.DL.Category.Properties.≣-refl 74 | 75 | ≣-sym : ∀ {S T} {A B : Object S T} {F G : A ⇒ B} → 76 | (F ≣ G) → (G ≣ F) 77 | ≣-sym = Web.Semantic.DL.Category.Properties.≣-sym 78 | 79 | ≣-trans : ∀ {S T} {A B : Object S T} {F G H : A ⇒ B} → 80 | (F ≣ G) → (G ≣ H) → (F ≣ H) 81 | ≣-trans = Web.Semantic.DL.Category.Properties.≣-trans 82 | 83 | -- Equations of a category 84 | 85 | compose-resp-≣ : ∀ {S T} {A B C : Object S T} {F₁ F₂ : A ⇒ B} {G₁ G₂ : B ⇒ C} → 86 | (F₁ ≣ F₂) → (G₁ ≣ G₂) → (F₁ ∙ G₁ ≣ F₂ ∙ G₂) 87 | compose-resp-≣ = Web.Semantic.DL.Category.Properties.compose-resp-≣ 88 | 89 | compose-unit₁ : ∀ {S T} {A B C : Object S T} (F : A ⇒ B) → 90 | (identity A ∙ F ≣ F) 91 | compose-unit₁ = Web.Semantic.DL.Category.Properties.compose-unit₁ 92 | 93 | compose-unit₂ : ∀ {S T} {A B C : Object S T} (F : A ⇒ B) → 94 | (F ∙ identity B ≣ F) 95 | compose-unit₂ = Web.Semantic.DL.Category.Properties.compose-unit₂ 96 | 97 | compose-assoc : ∀ {S T} {A B C D : Object S T} 98 | (F : A ⇒ B) (G : B ⇒ C) (H : C ⇒ D) → 99 | ((F ∙ G) ∙ H ≣ F ∙ (G ∙ H)) 100 | compose-assoc = Web.Semantic.DL.Category.Properties.compose-assoc 101 | 102 | -- Tensor is a bifunctor 103 | 104 | tensor-resp-≣ : ∀ {S T} {A₁ A₂ B₁ B₂ : Object S T} 105 | {F₁ G₁ : A₁ ⇒ B₁} {F₂ G₂ : A₂ ⇒ B₂} → 106 | (F₁ ≣ G₁) → (F₂ ≣ G₂) → (F₁ ⟨⊗⟩ F₂ ≣ G₁ ⟨⊗⟩ G₂) 107 | tensor-resp-≣ = Web.Semantic.DL.Category.Properties.tensor-resp-≣ 108 | 109 | tensor-resp-compose : ∀ {S T} {A₁ A₂ B₁ B₂ C₁ C₂ : Object S T} 110 | (F₁ : A₁ ⇒ B₁) (F₂ : A₂ ⇒ B₂) (G₁ : B₁ ⇒ C₁) (G₂ : B₂ ⇒ C₂) → 111 | (((F₁ ∙ G₁) ⟨⊗⟩ (F₂ ∙ G₂)) ≣ ((F₁ ⟨⊗⟩ F₂) ∙ (G₁ ⟨⊗⟩ G₂))) 112 | tensor-resp-compose = Web.Semantic.DL.Category.Properties.tensor-resp-compose 113 | 114 | tensor-resp-id : ∀ {S T} (A₁ A₂ : Object S T) → 115 | ((identity A₁ ⟨⊗⟩ identity A₂) ≣ identity (A₁ ⊗ A₂)) 116 | tensor-resp-id = Web.Semantic.DL.Category.Properties.tensor-resp-id 117 | 118 | -- Isomorphisms of a symmetric monoidal category 119 | 120 | symm-iso : ∀ {S T} (A₁ A₂ : Object S T) → 121 | (symm A₁ A₂ ∙ symm A₂ A₁ ≣ identity (A₁ ⊗ A₂)) 122 | symm-iso = Web.Semantic.DL.Category.Properties.symm-iso 123 | 124 | assoc-iso : ∀ {S T} (A₁ A₂ A₃ : Object S T) → 125 | (assoc A₁ A₂ A₃ ∙ assoc⁻¹ A₁ A₂ A₃ ≣ identity ((A₁ ⊗ A₂) ⊗ A₃)) 126 | assoc-iso = Web.Semantic.DL.Category.Properties.assoc-iso 127 | 128 | assoc⁻¹-iso : ∀ {S T} (A₁ A₂ A₃ : Object S T) → 129 | (assoc⁻¹ A₁ A₂ A₃ ∙ assoc A₁ A₂ A₃ ≣ identity (A₁ ⊗ (A₂ ⊗ A₃))) 130 | assoc⁻¹-iso = Web.Semantic.DL.Category.Properties.assoc⁻¹-iso 131 | 132 | unit₁-iso : ∀ {S T} (A : Object S T) → 133 | (unit₁ A ∙ unit₁⁻¹ A ≣ identity (I ⊗ A)) 134 | unit₁-iso = Web.Semantic.DL.Category.Properties.unit₁-iso 135 | 136 | unit₁⁻¹-iso : ∀ {S T} (A : Object S T) → 137 | (unit₁⁻¹ A ∙ unit₁ A ≣ identity A) 138 | unit₁⁻¹-iso = Web.Semantic.DL.Category.Properties.unit₁⁻¹-iso 139 | 140 | unit₂-iso : ∀ {S T} (A : Object S T) → 141 | (unit₂ A ∙ unit₂⁻¹ A ≣ identity (A ⊗ I)) 142 | unit₂-iso = Web.Semantic.DL.Category.Properties.unit₂-iso 143 | 144 | unit₂⁻¹-iso : ∀ {S T} (A : Object S T) → 145 | (unit₂⁻¹ A ∙ unit₂ A ≣ identity A) 146 | unit₂⁻¹-iso = Web.Semantic.DL.Category.Properties.unit₂⁻¹-iso 147 | 148 | -- Coherence conditions of a symmetric monoidal category 149 | 150 | assoc-unit : ∀ {S T} (A₁ A₂ : Object S T) → 151 | (assoc A₁ I A₂ ∙ (identity A₁ ⟨⊗⟩ unit₁ A₂) ≣ unit₂ A₁ ⟨⊗⟩ identity A₂) 152 | assoc-unit = Web.Semantic.DL.Category.Properties.assoc-unit 153 | 154 | assoc-assoc : ∀ {S T} (A₁ A₂ A₃ A₄ : Object S T) → 155 | (assoc (A₁ ⊗ A₂) A₃ A₄ ∙ assoc A₁ A₂ (A₃ ⊗ A₄) ≣ 156 | (assoc A₁ A₂ A₃ ⟨⊗⟩ identity A₄) ∙ assoc A₁ (A₂ ⊗ A₃) A₄ ∙ (identity A₁ ⟨⊗⟩ assoc A₂ A₃ A₄)) 157 | assoc-assoc = Web.Semantic.DL.Category.Properties.assoc-assoc 158 | 159 | assoc-symm : ∀ {S T} (A₁ A₂ A₃ : Object S T) → 160 | (symm (A₁ ⊗ A₂) A₃ ∙ assoc⁻¹ A₃ A₁ A₂ ≣ 161 | assoc A₁ A₂ A₃ ∙ (identity A₁ ⟨⊗⟩ symm A₂ A₃) ∙ assoc⁻¹ A₁ A₃ A₂ ∙ (symm A₁ A₃ ⟨⊗⟩ identity A₂)) 162 | assoc-symm = Web.Semantic.DL.Category.Properties.assoc-symm 163 | 164 | -- Naturality conditions of a symmetric monoidal category 165 | 166 | unit₁-natural : ∀ {S T} {A B : Object S T} (F : A ⇒ B) → 167 | ((identity I ⟨⊗⟩ F) ∙ unit₁ B ≣ unit₁ A ∙ F) 168 | unit₁-natural = Web.Semantic.DL.Category.Properties.unit₁-natural 169 | 170 | unit₂-natural : ∀ {S T} {A B : Object S T} (F : A ⇒ B) → 171 | ((F ⟨⊗⟩ identity I) ∙ unit₂ B ≣ unit₂ A ∙ F) 172 | unit₂-natural = Web.Semantic.DL.Category.Properties.unit₂-natural 173 | 174 | symm-natural : ∀ {S T} {A₁ A₂ B₁ B₂ : Object S T} 175 | (F₁ : A₁ ⇒ B₁) (F₂ : A₂ ⇒ B₂) → 176 | ((F₁ ⟨⊗⟩ F₂) ∙ symm B₁ B₂ ≣ symm A₁ A₂ ∙ (F₂ ⟨⊗⟩ F₁)) 177 | symm-natural = Web.Semantic.DL.Category.Properties.symm-natural 178 | 179 | assoc-natural : ∀ {S T} {A₁ A₂ A₃ B₁ B₂ B₃ : Object S T} 180 | (F₁ : A₁ ⇒ B₁) (F₂ : A₂ ⇒ B₂) (F₃ : A₃ ⇒ B₃) → 181 | (((F₁ ⟨⊗⟩ F₂) ⟨⊗⟩ F₃) ∙ assoc B₁ B₂ B₃ ≣ 182 | assoc A₁ A₂ A₃ ∙ (F₁ ⟨⊗⟩ (F₂ ⟨⊗⟩ F₃))) 183 | assoc-natural = Web.Semantic.DL.Category.Properties.assoc-natural 184 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Category/Composition.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( ∃ ; _×_ ; _,_ ; proj₁ ; proj₂ ) 2 | open import Relation.Unary using ( _∈_ ) 3 | open import Web.Semantic.DL.ABox using ( ABox ; ⟨ABox⟩ ; _,_ ) 4 | open import Web.Semantic.DL.ABox.Interp using 5 | ( Interp ; _,_ ; ⌊_⌋ ; ind ; _*_ ) 6 | open import Web.Semantic.DL.ABox.Interp.Morphism using 7 | ( _≲_ ; _,_ ; _**_ ; ≲⌊_⌋ ; ≲-resp-ind ; _≋_ ) 8 | open import Web.Semantic.DL.ABox.Model using 9 | ( _⊨a_ ; ⊨a-resp-≲ ; ⟨ABox⟩-resp-⊨ ; *-resp-⟨ABox⟩ 10 | ; on-bnode ; bnodes ; _⊨b_ ; _,_ ) 11 | open import Web.Semantic.DL.Category.Morphism using 12 | ( _⇒_ ; _⇒_w/_ ; _,_ ; BN ; impl ; impl✓ ) 13 | open import Web.Semantic.DL.Category.Object using 14 | ( Object ; _,_ ; IN ; iface ) 15 | open import Web.Semantic.DL.Integrity using 16 | ( Unique ; Mediated ; Mediator ; Initial ; _⊕_⊨_ ; _>>_ ; _,_ 17 | ; med-≲ ; med-≋ ; med-uniq ; init-≲ ; init-⊨ ; init-med 18 | ; extension ; ext-⊨ ; ext-init ; ext✓ ) 19 | open import Web.Semantic.DL.KB using ( KB ; _,_ ) 20 | open import Web.Semantic.DL.KB.Model using ( _⊨_ ) 21 | open import Web.Semantic.DL.Signature using ( Signature ) 22 | open import Web.Semantic.DL.TBox using ( TBox ; _,_ ) 23 | open import Web.Semantic.DL.TBox.Interp using 24 | ( Δ ; _⊨_≈_ ; ≈-refl ; ≈-sym ; ≈-trans ) renaming 25 | ( Interp to Interp′ ) 26 | open import Web.Semantic.DL.TBox.Model using ( _⊨t_ ) 27 | open import Web.Semantic.DL.TBox.Interp.Morphism using 28 | ( ≲-image ; ≲-refl ; ≲-trans ; ≲-resp-≈ ) renaming 29 | ( _≲_ to _≲′_ ) 30 | open import Web.Semantic.Util using 31 | ( _⊕_⊕_ ; inode ; bnode ; enode ; left ; right ; hmerge ; _∘_ ) 32 | 33 | module Web.Semantic.DL.Category.Composition {Σ : Signature} where 34 | 35 | infixr 5 _⟫_ _∙_ 36 | 37 | _⟫_ : ∀ {V W X Y Z} → ABox Σ (X ⊕ V ⊕ Y) → ABox Σ (Y ⊕ W ⊕ Z) → 38 | ABox Σ (X ⊕ (V ⊕ Y ⊕ W) ⊕ Z) 39 | F ⟫ G = (⟨ABox⟩ left F , ⟨ABox⟩ right G) 40 | 41 | pipe : ∀ {V W X Y Z} → 42 | (J : Interp Σ (X ⊕ V ⊕ Y)) → (K : Interp Σ (Y ⊕ W ⊕ Z)) → 43 | (enode * J ≲ inode * K) → (Interp Σ (X ⊕ (V ⊕ Y ⊕ W) ⊕ Z)) 44 | pipe (J , j) (K , k) (J≲K , j≲k) = (K , hmerge (≲-image J≲K ∘ j) k) 45 | 46 | pipe-≳ : ∀ {V W X Y Z} → (I : Interp Σ X) → 47 | (J : Interp Σ (X ⊕ V ⊕ Y)) → (K : Interp Σ (Y ⊕ W ⊕ Z)) → 48 | (I≲J : I ≲ inode * J) → (J≲K : enode * J ≲ inode * K) → 49 | (I ≲ inode * (pipe J K J≲K)) 50 | pipe-≳ (I , i) (J , j) (K , k) (I≲J , i≲j) (J≲K , j≲k) = 51 | (≲-trans I≲J J≲K , λ x → ≲-resp-≈ J≲K (i≲j x)) 52 | 53 | pipe-left : ∀ {V W X Y Z} → (J : Interp Σ (X ⊕ V ⊕ Y)) → 54 | (K : Interp Σ (Y ⊕ W ⊕ Z)) → (J≲K : enode * J ≲ inode * K) → 55 | (J ≲ left * (pipe J K J≲K)) 56 | pipe-left (J , j) (K , k) (J≲K , j≲k) = (J≲K , lemma) where 57 | 58 | lemma : ∀ x → 59 | K ⊨ ≲-image J≲K (j x) ≈ hmerge (≲-image J≲K ∘ j) k (left x) 60 | lemma (inode x) = ≈-refl K 61 | lemma (bnode v) = ≈-refl K 62 | lemma (enode y) = j≲k y 63 | 64 | pipe-right : ∀ {V W X Y Z} → (J : Interp Σ (X ⊕ V ⊕ Y)) → 65 | (K : Interp Σ (Y ⊕ W ⊕ Z)) → (J≲K : enode * J ≲ inode * K) → 66 | (K ≲ right * (pipe J K J≲K)) 67 | pipe-right (J , j) (K , k) (J≲K , j≲k) = (≲-refl K , lemma) where 68 | 69 | lemma : ∀ x → 70 | K ⊨ k x ≈ hmerge (≲-image J≲K ∘ j) k (right x) 71 | lemma (inode y) = ≈-refl K 72 | lemma (bnode w) = ≈-refl K 73 | lemma (enode z) = ≈-refl K 74 | 75 | pipe-exp : ∀ {V W X Y Z} → (J : Interp Σ (X ⊕ V ⊕ Y)) → 76 | (K : Interp Σ (Y ⊕ W ⊕ Z)) → (J≲K : enode * J ≲ inode * K) → 77 | ∀ KB → (enode * K ⊨ KB) → (enode * pipe J K J≲K ⊨ KB) 78 | pipe-exp (J , j) (K , k) (J≲K , j≲k) KB K⊨KB = K⊨KB 79 | 80 | ⊨a-intro-⟫ : ∀ {V W X Y Z} → (I : Interp Σ (X ⊕ (V ⊕ Y ⊕ W) ⊕ Z)) → 81 | (F : ABox Σ (X ⊕ V ⊕ Y)) → (G : ABox Σ (Y ⊕ W ⊕ Z)) → 82 | (left * I ⊨a F) → (right * I ⊨a G) → (I ⊨a F ⟫ G) 83 | ⊨a-intro-⟫ (I , i) F G I⊨F I⊨G = 84 | ( ⟨ABox⟩-resp-⊨ left (λ x → ≈-refl I) F I⊨F 85 | , ⟨ABox⟩-resp-⊨ right (λ x → ≈-refl I) G I⊨G ) 86 | 87 | ⊨b-intro-⟫ : ∀ {V₁ W₁ V₂ W₂ X Y Z} → (I : Interp Σ (X ⊕ (V₁ ⊕ Y ⊕ W₁) ⊕ Z)) → 88 | (F : ABox Σ (X ⊕ V₂ ⊕ Y)) → (G : ABox Σ (Y ⊕ W₂ ⊕ Z)) → 89 | (left * I ⊨b F) → (right * I ⊨b G) → (I ⊨b F ⟫ G) 90 | ⊨b-intro-⟫ {V₂ = V₂} {W₂ = W₂} {Y = Y} (I , i) F G (f , I⊨F) (g , I⊨G) = 91 | (h , I⊨F⟫G) where 92 | 93 | h : (V₂ ⊕ Y ⊕ W₂) → Δ I 94 | h (inode v) = f v 95 | h (bnode y) = i (bnode (bnode y)) 96 | h (enode w) = g w 97 | 98 | lemmaˡ : ∀ x → 99 | I ⊨ on-bnode f (i ∘ left) x ≈ on-bnode h i (left x) 100 | lemmaˡ (inode x) = ≈-refl I 101 | lemmaˡ (bnode v) = ≈-refl I 102 | lemmaˡ (enode y) = ≈-refl I 103 | 104 | lemmaʳ : ∀ x → 105 | I ⊨ on-bnode g (i ∘ right) x ≈ on-bnode h i (right x) 106 | lemmaʳ (inode x) = ≈-refl I 107 | lemmaʳ (bnode v) = ≈-refl I 108 | lemmaʳ (enode y) = ≈-refl I 109 | 110 | I⊨F⟫G : bnodes (I , i) h ⊨a F ⟫ G 111 | I⊨F⟫G = ⊨a-intro-⟫ (bnodes (I , i) h) F G 112 | (⊨a-resp-≲ (≲-refl I , lemmaˡ) F I⊨F) 113 | (⊨a-resp-≲ (≲-refl I , lemmaʳ) G I⊨G) 114 | 115 | pipe-uniq : ∀ {V W X Y Z} I J K (M : Interp Σ (X ⊕ (V ⊕ Y ⊕ W) ⊕ Z)) 116 | (I≲J : I ≲ inode * J) (I≲M : I ≲ inode * M) → 117 | (J≲K : enode * J ≲ inode * K) (J≲M : J ≲ left * M) → 118 | (I≲M ≋ I≲J >> J≲M) → 119 | (Unique I J (left * M) I≲J I≲M) → 120 | (Unique (enode * J) K (right * M) J≲K (enode ** J≲M)) → 121 | (Unique I (pipe J K J≲K) M (pipe-≳ I J K I≲J J≲K) I≲M) 122 | pipe-uniq {V} {W} {X} {Y} {Z} (I , i) (J , j) (K , k) (M , m) 123 | (I≲J , i≲j) (I≲M , i≲m) (J≲K , j≲k) (J≲M , j≲m) I≲M≋I≲J≲M J≲M-uniq K≲M-uniq 124 | (L≲₁M , l≲₁m) (L≲₂M , l≲₂m) I≲M≋I≲L≲₁M I≲M≋I≲L≲₂M = 125 | K≲M-uniq 126 | (L≲₁M , lemmaʳ L≲₁M l≲₁m I≲M≋I≲L≲₁M) 127 | (L≲₂M , lemmaʳ L≲₂M l≲₂m I≲M≋I≲L≲₂M) 128 | (J≲M-uniq (J≲M , j≲m) 129 | (≲-trans J≲K L≲₁M , lemmaˡ L≲₁M l≲₁m I≲M≋I≲L≲₁M) 130 | I≲M≋I≲J≲M 131 | I≲M≋I≲L≲₁M ) 132 | (J≲M-uniq (J≲M , j≲m) 133 | ( ≲-trans J≲K L≲₂M , lemmaˡ L≲₂M l≲₂m I≲M≋I≲L≲₂M) 134 | I≲M≋I≲J≲M 135 | I≲M≋I≲L≲₂M ) where 136 | 137 | L : Interp′ Σ 138 | L = K 139 | 140 | l : (X ⊕ (V ⊕ Y ⊕ W) ⊕ Z) → Δ L 141 | l = hmerge (≲-image J≲K ∘ j) k 142 | 143 | I≲L : I ≲′ L 144 | I≲L = ≲-trans I≲J J≲K 145 | 146 | lemmaˡ : ∀ L≲M → (∀ x → M ⊨ ≲-image L≲M (l x) ≈ m x) → 147 | (∀ x → M ⊨ ≲-image I≲M x ≈ ≲-image L≲M (≲-image I≲L x)) → 148 | ∀ x → M ⊨ ≲-image L≲M (≲-image J≲K (j x)) ≈ m (left x) 149 | lemmaˡ L≲M l≲m I≲M≋I≲L≲M (inode x) = l≲m (inode x) 150 | lemmaˡ L≲M l≲m I≲M≋I≲L≲M (bnode v) = l≲m (bnode (inode v)) 151 | lemmaˡ L≲M l≲m I≲M≋I≲L≲M (enode y) = 152 | ≈-trans M (≲-resp-≈ L≲M (j≲k y)) (l≲m (bnode (bnode y))) 153 | 154 | lemmaʳ : ∀ L≲M → (∀ x → M ⊨ ≲-image L≲M (l x) ≈ m x) → 155 | (∀ x → M ⊨ ≲-image I≲M x ≈ ≲-image L≲M (≲-image I≲L x)) → 156 | ∀ x → M ⊨ ≲-image L≲M (k x) ≈ m (right x) 157 | lemmaʳ L≲M l≲m I≲M≋I≲L≲M (inode y) = l≲m (bnode (bnode y)) 158 | lemmaʳ L≲M l≲m I≲M≋I≲L≲M (bnode w) = l≲m (bnode (enode w)) 159 | lemmaʳ L≲M l≲m I≲M≋I≲L≲M (enode z) = l≲m (enode z) 160 | 161 | pipe-mediated : ∀ {V W X Y Z} I J K (M : Interp Σ (X ⊕ (V ⊕ Y ⊕ W) ⊕ Z)) 162 | (I≲J : I ≲ inode * J) (I≲M : I ≲ inode * M) → 163 | (J≲K : enode * J ≲ inode * K) → 164 | (m : Mediated I J (left * M) I≲J I≲M) → 165 | (Mediated (enode * J) K (right * M) J≲K (enode ** (med-≲ m))) → 166 | (Mediated I (pipe J K J≲K) M (pipe-≳ I J K I≲J J≲K) I≲M) 167 | pipe-mediated {V} {W} {X} {Y} {Z} (I , i) (J , j) (K , k) (M , m) 168 | (I≲J , i≲j) (I≲M , i≲m) (J≲K , j≲k) 169 | ((J≲M , j≲m) , I≲M≋I≲J≲M , J≲M-uniq) 170 | ((K≲M , k≲m) , J≲M≋J≲K≲M , K≲M-uniq) = 171 | ( (L≲M , l≲m) , I≲M≋I≲L≲M , L≲M-uniq) where 172 | 173 | L : Interp′ Σ 174 | L = K 175 | 176 | l : (X ⊕ (V ⊕ Y ⊕ W) ⊕ Z) → Δ L 177 | l = hmerge (≲-image J≲K ∘ j) k 178 | 179 | I≲L : I ≲′ L 180 | I≲L = ≲-trans I≲J J≲K 181 | 182 | i≲l : ∀ x → L ⊨ ≲-image I≲L (i x) ≈ l (inode x) 183 | i≲l x = ≲-resp-≈ J≲K (i≲j x) 184 | 185 | L≲M : L ≲′ M 186 | L≲M = K≲M 187 | 188 | l≲m : ∀ x → M ⊨ ≲-image K≲M (l x) ≈ m x 189 | l≲m (inode x) = 190 | ≈-trans M (≈-sym M (J≲M≋J≲K≲M (j (inode x)))) (j≲m (inode x)) 191 | l≲m (bnode (inode v)) = 192 | ≈-trans M (≈-sym M (J≲M≋J≲K≲M (j (bnode v)))) (j≲m (bnode v)) 193 | l≲m (bnode (bnode y)) = k≲m (inode y) 194 | l≲m (bnode (enode w)) = k≲m (bnode w) 195 | l≲m (enode z) = k≲m (enode z) 196 | 197 | I≲M≋I≲L≲M : ∀ x → M ⊨ ≲-image I≲M x ≈ ≲-image L≲M (≲-image I≲L x) 198 | I≲M≋I≲L≲M x = ≈-trans M (I≲M≋I≲J≲M x) (J≲M≋J≲K≲M (≲-image I≲J x)) 199 | 200 | L≲M-uniq : Unique (I , i) (L , l) (M , m) (I≲L , i≲l) (I≲M , i≲m) 201 | L≲M-uniq = pipe-uniq (I , i) (J , j) (K , k) (M , m) 202 | (I≲J , i≲j) (I≲M , i≲m) (J≲K , j≲k) (J≲M , j≲m) 203 | I≲M≋I≲J≲M J≲M-uniq K≲M-uniq 204 | 205 | pipe-mediator : ∀ S {V W X Y Z I} 206 | {J : Interp Σ (X ⊕ V ⊕ Y)} {K : Interp Σ (Y ⊕ W ⊕ Z)} {I≲J J≲K} F G → 207 | (Mediator I J I≲J (S , F)) → (Mediator (enode * J) K J≲K (S , G)) → 208 | (Mediator I (pipe J K J≲K) (pipe-≳ I J K I≲J J≲K) (S , F ⟫ G)) 209 | pipe-mediator S {V} {W} {X} {Y} {Z} 210 | {I , i} {J , j} {K , k} {I≲J , i≲j} {J≲K , j≲k} 211 | F G J-med K-med (M , m) (I≲M , i≲m) (M⊨S , M⊨F , M⊨G) = 212 | pipe-mediated (I , i) (J , j) (K , k) (M , m) 213 | (I≲J , i≲j) (I≲M , i≲m) (J≲K , j≲k) 214 | I≲J≲M-med J≲K≲M-med where 215 | 216 | mˡ : (X ⊕ V ⊕ Y) → Δ M 217 | mˡ x = m (left x) 218 | 219 | mʳ : (Y ⊕ W ⊕ Z) → Δ M 220 | mʳ x = m (right x) 221 | 222 | I≲J≲M-med : Mediated (I , i) (J , j) (M , mˡ) (I≲J , i≲j) (I≲M , i≲m) 223 | I≲J≲M-med = J-med (M , mˡ) (I≲M , i≲m) 224 | (M⊨S , *-resp-⟨ABox⟩ left (M , m) F M⊨F) 225 | 226 | J≲M : J ≲′ M 227 | J≲M = ≲⌊ med-≲ I≲J≲M-med ⌋ 228 | 229 | j≲m : ∀ x → M ⊨ ≲-image J≲M (j x) ≈ mˡ x 230 | j≲m = ≲-resp-ind (med-≲ I≲J≲M-med) 231 | 232 | j′ : Y → Δ J 233 | j′ y = j (enode y) 234 | 235 | j≲m′ : ∀ x → M ⊨ ≲-image J≲M (j′ x) ≈ m (bnode (bnode x)) 236 | j≲m′ x = j≲m (enode x) 237 | 238 | J≲K≲M-med : Mediated (J , j′) (K , k) (M , mʳ) (J≲K , j≲k) (J≲M , j≲m′) 239 | J≲K≲M-med = K-med (M , mʳ) (J≲M , j≲m′) 240 | (M⊨S , *-resp-⟨ABox⟩ right (M , m) G M⊨G) 241 | 242 | pipe-init : ∀ {S V W X Y Z I} 243 | {J : Interp Σ (X ⊕ V ⊕ Y)} {K : Interp Σ (Y ⊕ W ⊕ Z)} {F G} → 244 | (J-init : J ∈ Initial I (S , F)) → 245 | (K-init : K ∈ Initial (enode * J) (S , G)) → 246 | (pipe J K (init-≲ K-init) ∈ Initial I (S , F ⟫ G)) 247 | pipe-init {S} {V} {W} {X} {Y} {Z} {I , i} {J , j} {K , k} {F} {G} 248 | ((I≲J , i≲j) , (J⊨S , J⊨F) , J-med) ((J≲K , j≲k) , (K⊨S , K⊨G) , K-med) = 249 | ( (I≲L , i≲l) , (L⊨S , L⊨F⟫G) , L-med) where 250 | 251 | L : Interp′ Σ 252 | L = K 253 | 254 | l : (X ⊕ (V ⊕ Y ⊕ W) ⊕ Z) → Δ L 255 | l = hmerge (≲-image J≲K ∘ j) k 256 | 257 | I≲L : I ≲′ L 258 | I≲L = ≲-trans I≲J J≲K 259 | 260 | i≲l : ∀ x → L ⊨ ≲-image I≲L (i x) ≈ l (inode x) 261 | i≲l x = ≲-resp-≈ J≲K (i≲j x) 262 | 263 | L⊨S : L ⊨t S 264 | L⊨S = K⊨S 265 | 266 | L⊨F⟫G : (L , l) ⊨a F ⟫ G 267 | L⊨F⟫G = ⊨a-intro-⟫ (L , l) F G 268 | (⊨a-resp-≲ (pipe-left (J , j) (K , k) (J≲K , j≲k)) F J⊨F) 269 | (⊨a-resp-≲ (pipe-right (J , j) (K , k) (J≲K , j≲k)) G K⊨G) 270 | 271 | L-med : Mediator (I , i) (L , l) (I≲L , i≲l) (S , F ⟫ G) 272 | L-med = pipe-mediator S F G J-med K-med 273 | 274 | compose-⊨ : ∀ {S T V W X Y Z} A B C 275 | (F : ABox Σ (X ⊕ V ⊕ Y)) (G : ABox Σ (Y ⊕ W ⊕ Z)) → 276 | (∀ I → (I ⊨ (S , T) , A) → (I ⊕ S , F ⊨ T , B)) → 277 | (∀ J → (J ⊨ (S , T) , B) → (J ⊕ S , G ⊨ T , C)) → 278 | (∀ I → (I ⊨ (S , T) , A) → (I ⊕ S , F ⟫ G ⊨ T , C)) 279 | compose-⊨ {S} {T} {V} {W} {X} {Y} {Z} A B C F G F✓ G✓ I I⊨STA = 280 | ( pipe J K J≲K 281 | , pipe-init J-init K-init 282 | , pipe-exp J K J≲K (T , C) K⊨TC ) where 283 | 284 | I⊕SF⊨TB : I ⊕ S , F ⊨ T , B 285 | I⊕SF⊨TB = F✓ I I⊨STA 286 | 287 | J : Interp Σ (X ⊕ V ⊕ Y) 288 | J = extension I⊕SF⊨TB 289 | 290 | J-init : J ∈ Initial I (S , F) 291 | J-init = ext-init I⊕SF⊨TB 292 | 293 | J⊕SG⊨TC : enode * J ⊕ S , G ⊨ T , C 294 | J⊕SG⊨TC = G✓ (enode * J) (ext✓ I⊕SF⊨TB) 295 | 296 | K : Interp Σ (Y ⊕ W ⊕ Z) 297 | K = extension J⊕SG⊨TC 298 | 299 | K⊨TC : enode * K ⊨ (T , C) 300 | K⊨TC = ext-⊨ J⊕SG⊨TC 301 | 302 | K-init : K ∈ Initial (enode * J) (S , G) 303 | K-init = ext-init J⊕SG⊨TC 304 | 305 | J≲K : enode * J ≲ inode * K 306 | J≲K = init-≲ K-init 307 | 308 | _∙′_ : ∀ {S T} {A B C : Object S T} → (F : A ⇒ B) → (G : B ⇒ C) → 309 | (A ⇒ C w/ (BN F ⊕ IN B ⊕ BN G)) 310 | _∙′_ {S} {T} {X , X∈Fin , A} {Y , Y∈Fin , B} {Z , Z∈Fin , C} (V , F , F✓) (W , G , G✓) = 311 | (F ⟫ G , compose-⊨ A B C F G F✓ G✓) 312 | 313 | _∙_ : ∀ {S T} {A B C : Object S T} → (A ⇒ B) → (B ⇒ C) → (A ⇒ C) 314 | F ∙ G = ( _ , F ∙′ G ) 315 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Category/Morphism.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( ∃ ; _×_ ; _,_ ; proj₁ ; proj₂ ) 2 | open import Relation.Unary using ( _∈_ ) 3 | open import Web.Semantic.DL.ABox using ( ABox ) 4 | open import Web.Semantic.DL.ABox.Model using ( _⊨a_ ; _⊨b_ ; ⊨a-resp-≲ ; ⊨b-resp-≲ ) 5 | open import Web.Semantic.DL.ABox.Interp using ( Interp ; ⌊_⌋ ; _*_ ) 6 | open import Web.Semantic.DL.ABox.Interp.Morphism using ( _≲_ ; ≲-refl ) 7 | open import Web.Semantic.DL.KB using ( _,_ ) 8 | open import Web.Semantic.DL.KB.Model using ( _⊨_ ) 9 | open import Web.Semantic.DL.Integrity using ( Initial ; _⊕_⊨_ ; extension ; ext-init ; ext-⊨ ; ext✓ ; init-≲ ; init-⊨ ; init-med ; med-≲ ) 10 | open import Web.Semantic.DL.Signature using ( Signature ) 11 | open import Web.Semantic.DL.TBox using ( TBox ; _,_ ) 12 | open import Web.Semantic.DL.TBox.Model using ( _⊨t_ ) 13 | open import Web.Semantic.DL.Category.Object using ( Object ; _,_ ; IN ; iface ) 14 | open import Web.Semantic.Util using ( _⊕_⊕_ ; inode ; enode ) 15 | 16 | module Web.Semantic.DL.Category.Morphism {Σ : Signature} {S T : TBox Σ} where 17 | 18 | infixr 4 _,_ 19 | 20 | -- A morphism A ⇒ B is an abox F such that for every I ⊨ S , T , A 21 | -- there is a J which is the initial extension of I satisfying (S , F), 22 | -- and moreover J satisfies (T , B). 23 | 24 | data _⇒_w/_ (A B : Object S T) (V : Set) : Set₁ where 25 | _,_ : (F : ABox Σ (IN A ⊕ V ⊕ IN B)) → 26 | (∀ I → (I ⊨ (S , T) , iface A) → (I ⊕ (S , F) ⊨ (T , iface B))) → 27 | (A ⇒ B w/ V) 28 | 29 | data _⇒_ (A B : Object S T) : Set₁ where 30 | _,_ : ∀ V → (A ⇒ B w/ V) → (A ⇒ B) 31 | 32 | BN : ∀ {A B} → (F : A ⇒ B) → Set 33 | BN (V , F,F✓) = V 34 | 35 | impl : ∀ {A B} → (F : A ⇒ B) → ABox Σ (IN A ⊕ BN F ⊕ IN B) 36 | impl (V , F , F✓) = F 37 | 38 | impl✓ : ∀ {A B} → (F : A ⇒ B) → ∀ I → (I ⊨ (S , T) , iface A) → (I ⊕ (S , impl F) ⊨ (T , iface B)) 39 | impl✓ (V , F , F✓) = F✓ 40 | 41 | apply : ∀ {A B} (F : A ⇒ B) I → (I ⊨ (S , T) , iface A) → 42 | Interp Σ (IN A ⊕ BN F ⊕ IN B) 43 | apply F I I⊨STA = extension (impl✓ F I I⊨STA) 44 | 45 | apply-init : ∀ {A B} (F : A ⇒ B) I I⊨STA → 46 | (apply F I I⊨STA ∈ Initial I (S , impl F)) 47 | apply-init F I I⊨STA = ext-init (impl✓ F I I⊨STA) 48 | 49 | apply-⊨ : ∀ {A B} (F : A ⇒ B) I I⊨STA → 50 | (enode * (apply F I I⊨STA) ⊨ (T , iface B)) 51 | apply-⊨ F I I⊨STA = ext-⊨ (impl✓ F I I⊨STA) 52 | 53 | apply-≲ : ∀ {A B} (F : A ⇒ B) I I⊨STA → (I ⊨a impl F) → 54 | (apply F (inode * I) I⊨STA ≲ I) 55 | apply-≲ F I ((I⊨S , I⊨T) , I⊨A) I⊨F = 56 | med-≲ (init-med 57 | (apply-init F (inode * I) ((I⊨S , I⊨T) , I⊨A)) 58 | I 59 | (≲-refl (inode * I)) 60 | (I⊨S , I⊨F)) 61 | 62 | apply✓ : ∀ {A B} (F : A ⇒ B) I I⊨STA → 63 | (enode * apply F I I⊨STA ⊨ (S , T) , iface B) 64 | apply✓ F I I⊨STA = ext✓ (impl✓ F I I⊨STA) 65 | 66 | -- Morphisms F and G are equivalent whenever 67 | -- in any interpretation I ⊨ S,T 68 | -- we have I ⊨ F iff I ⊨ G. 69 | 70 | infix 2 _⊑_ _⊑′_ _≣_ 71 | 72 | _⊑_ : ∀ {A B : Object S T} → (A ⇒ B) → (A ⇒ B) → Set₁ 73 | _⊑_ {A} F G = 74 | ∀ I → (inode * I ⊨ (S , T) , iface A) → (I ⊨a impl F) → (I ⊨b impl G) 75 | 76 | data _≣_ {A B : Object S T} (F G : A ⇒ B) : Set₁ where 77 | _,_ : (F ⊑ G) → (G ⊑ F) → (F ≣ G) 78 | 79 | -- An alternative characterization, which may be easier 80 | -- to work with. 81 | 82 | _⊑′_ : ∀ {A B : Object S T} → (A ⇒ B) → (A ⇒ B) → Set₁ 83 | F ⊑′ G = ∀ I I⊨STA → (apply F I I⊨STA) ⊨b (impl G) 84 | 85 | ⊑′-impl-⊑ : ∀ {A B : Object S T} → (F G : A ⇒ B) → (F ⊑′ G) → (F ⊑ G) 86 | ⊑′-impl-⊑ F G F⊑′G I I⊨STA I⊨F = 87 | ⊨b-resp-≲ (apply-≲ F I I⊨STA I⊨F) (impl G) (F⊑′G (inode * I) I⊨STA) 88 | 89 | ⊑-impl-⊑′ : ∀ {A B : Object S T} → (F G : A ⇒ B) → (F ⊑ G) → (F ⊑′ G) 90 | ⊑-impl-⊑′ {A} {B} F G F⊑G I (I⊨ST , I⊨A) = J⊨G where 91 | 92 | J : Interp Σ (IN A ⊕ BN F ⊕ IN B) 93 | J = apply F I (I⊨ST , I⊨A) 94 | 95 | J⊨S : ⌊ J ⌋ ⊨t S 96 | J⊨S = proj₁ (init-⊨ (apply-init F I (I⊨ST , I⊨A))) 97 | 98 | J⊨T : ⌊ J ⌋ ⊨t T 99 | J⊨T = proj₁ (apply-⊨ F I (I⊨ST , I⊨A)) 100 | 101 | J⊨A : inode * J ⊨a iface A 102 | J⊨A = ⊨a-resp-≲ (init-≲ (apply-init F I (I⊨ST , I⊨A))) (iface A) I⊨A 103 | 104 | J⊨F : J ⊨a impl F 105 | J⊨F = proj₂ (init-⊨ (apply-init F I (I⊨ST , I⊨A))) 106 | 107 | J⊨G : J ⊨b impl G 108 | J⊨G = F⊑G J ((J⊨S , J⊨T) , J⊨A) J⊨F 109 | 110 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Category/Object.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( _×_ ; _,_ ) 2 | open import Relation.Unary using ( _∈_ ) 3 | open import Web.Semantic.DL.ABox using ( ABox ) 4 | open import Web.Semantic.DL.Signature using ( Signature ) 5 | open import Web.Semantic.DL.TBox using ( TBox ) 6 | open import Web.Semantic.Util using ( Finite ) 7 | 8 | module Web.Semantic.DL.Category.Object {Σ : Signature} where 9 | 10 | infixr 4 _,_ 11 | 12 | data Object (S T : TBox Σ) : Set₁ where 13 | _,_ : ∀ X → (X ∈ Finite × ABox Σ X) → Object S T 14 | 15 | IN : ∀ {S T} → Object S T → Set 16 | IN (X , X∈Fin , A) = X 17 | 18 | fin : ∀ {S T} → (A : Object S T) → (IN A ∈ Finite) 19 | fin (X , X∈Fin , A) = X∈Fin 20 | 21 | iface : ∀ {S T} → (A : Object S T) → (ABox Σ (IN A)) 22 | iface (X , X∈Fin , A) = A 23 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Category/Properties.agda: -------------------------------------------------------------------------------- 1 | module Web.Semantic.DL.Category.Properties where 2 | 3 | open import Web.Semantic.DL.Category.Properties.Equivalence public using 4 | ( ≣-refl ; ≣-sym ; ≣-trans ) 5 | 6 | open import Web.Semantic.DL.Category.Properties.Composition public using 7 | ( compose-resp-≣ ; compose-unit₁ ; compose-unit₂ ; compose-assoc ) 8 | 9 | open import Web.Semantic.DL.Category.Properties.Tensor public using 10 | ( tensor-resp-≣ ; tensor-resp-id ; tensor-resp-compose 11 | ; symm-iso ; assoc-iso ; assoc⁻¹-iso 12 | ; unit₁-iso ; unit₁⁻¹-iso ; unit₂-iso ; unit₂⁻¹-iso 13 | ; assoc-unit ; assoc-assoc ; assoc-symm 14 | ; unit₁-natural ; unit₂-natural ; symm-natural ; assoc-natural ) 15 | 16 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Category/Properties/Composition.agda: -------------------------------------------------------------------------------- 1 | module Web.Semantic.DL.Category.Properties.Composition where 2 | 3 | open import Web.Semantic.DL.Category.Properties.Composition.RespectsEquiv 4 | public using ( compose-resp-≣ ) 5 | 6 | open import Web.Semantic.DL.Category.Properties.Composition.LeftUnit 7 | public using ( compose-unit₁ ) 8 | 9 | open import Web.Semantic.DL.Category.Properties.Composition.RightUnit 10 | public using ( compose-unit₂ ) 11 | 12 | open import Web.Semantic.DL.Category.Properties.Composition.Assoc 13 | public using ( compose-assoc ) 14 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Category/Properties/Composition/Assoc.agda: -------------------------------------------------------------------------------- 1 | open import Relation.Binary.PropositionalEquality using ( refl ) 2 | open import Web.Semantic.DL.ABox.Interp using ( _*_ ; ⌊_⌋ ; ind ) 3 | open import Web.Semantic.DL.ABox.Interp.Morphism using ( _,_ ) 4 | open import Web.Semantic.DL.ABox.Model using 5 | ( _⊨a_ ; on-bnode ; bnodes ; ⊨a-resp-≲ ; ⊨a-resp-≡³ ; _,_ ) 6 | open import Web.Semantic.DL.Category.Object using 7 | ( Object ; IN ; iface ) 8 | open import Web.Semantic.DL.Category.Morphism using 9 | ( _⇒_ ; BN ; impl ; _≣_ ; _⊑_ ; _,_ ) 10 | open import Web.Semantic.DL.Category.Composition using ( _∙_ ) 11 | open import Web.Semantic.DL.Category.Properties.Composition.Lemmas using 12 | ( compose-left ; compose-right ; compose-resp-⊨a ) 13 | open import Web.Semantic.DL.Signature using ( Signature ) 14 | open import Web.Semantic.DL.TBox using ( TBox ) 15 | open import Web.Semantic.DL.TBox.Interp using ( Δ ; _⊨_≈_ ; ≈-refl ; ≈-sym ) 16 | open import Web.Semantic.DL.TBox.Interp.Morphism using ( ≲-refl ) 17 | open import Web.Semantic.Util using 18 | ( _∘_ ; _⊕_⊕_ ; False ; inode ; bnode ; enode ; left ; right ) 19 | 20 | module Web.Semantic.DL.Category.Properties.Composition.Assoc 21 | {Σ : Signature} {S T : TBox Σ} where 22 | 23 | compose-assoc : ∀ {A B C D : Object S T} (F : A ⇒ B) (G : B ⇒ C) (H : C ⇒ D) → 24 | ((F ∙ G) ∙ H ≣ F ∙ (G ∙ H)) 25 | compose-assoc {A} {B} {C} {D} F G H = (LHS⊑RHS , RHS⊑LHS) where 26 | 27 | LHS⊑RHS : (F ∙ G) ∙ H ⊑ F ∙ (G ∙ H) 28 | LHS⊑RHS I I⊨STA I⊨LHS = (f , I⊨RHS) where 29 | 30 | f : (BN F ⊕ IN B ⊕ (BN G ⊕ IN C ⊕ BN H)) → Δ ⌊ I ⌋ 31 | f (inode u) = ind I (bnode (inode (inode u))) 32 | f (bnode y) = ind I (bnode (inode (bnode y))) 33 | f (enode (inode v)) = ind I (bnode (inode (enode v))) 34 | f (enode (bnode z)) = ind I (bnode (bnode z)) 35 | f (enode (enode w)) = ind I (bnode (enode w)) 36 | 37 | I⊨F : left * bnodes I f ⊨a impl F 38 | I⊨F = ⊨a-resp-≡³ (left * left * I) (on-bnode f (ind I) ∘ left) refl 39 | (impl F) (compose-left F G (left * I) (compose-left (F ∙ G) H I I⊨LHS)) 40 | 41 | I⊨G : left * right * bnodes I f ⊨a impl G 42 | I⊨G = ⊨a-resp-≡³ (right * left * I) (on-bnode f (ind I) ∘ right ∘ left) refl 43 | (impl G) (compose-right F G (left * I) (compose-left (F ∙ G) H I I⊨LHS)) 44 | 45 | I⊨H : right * right * bnodes I f ⊨a impl H 46 | I⊨H = ⊨a-resp-≡³ (right * I) (on-bnode f (ind I) ∘ right ∘ right) refl 47 | (impl H) (compose-right (F ∙ G) H I I⊨LHS) 48 | 49 | I⊨RHS : bnodes I f ⊨a impl (F ∙ (G ∙ H)) 50 | I⊨RHS = compose-resp-⊨a F (G ∙ H) (bnodes I f) I⊨F 51 | (compose-resp-⊨a G H (right * bnodes I f) I⊨G I⊨H) 52 | 53 | RHS⊑LHS : F ∙ (G ∙ H) ⊑ (F ∙ G) ∙ H 54 | RHS⊑LHS I I⊨STA I⊨RHS = (f , I⊨LHS) where 55 | 56 | f : ((BN F ⊕ IN B ⊕ BN G) ⊕ IN C ⊕ BN H) → Δ ⌊ I ⌋ 57 | f (inode (inode u)) = ind I (bnode (inode u)) 58 | f (inode (bnode y)) = ind I (bnode (bnode y)) 59 | f (inode (enode v)) = ind I (bnode (enode (inode v))) 60 | f (bnode z) = ind I (bnode (enode (bnode z))) 61 | f (enode w) = ind I (bnode (enode (enode w))) 62 | 63 | I⊨F : left * left * bnodes I f ⊨a impl F 64 | I⊨F = ⊨a-resp-≡³ (left * I) (on-bnode f (ind I) ∘ left ∘ left) refl 65 | (impl F) (compose-left F (G ∙ H) I I⊨RHS) 66 | 67 | I⊨G : right * left * bnodes I f ⊨a impl G 68 | I⊨G = ⊨a-resp-≡³ (left * right * I) (on-bnode f (ind I) ∘ left ∘ right) refl 69 | (impl G) (compose-left G H (right * I) (compose-right F (G ∙ H) I I⊨RHS)) 70 | 71 | I⊨H : right * bnodes I f ⊨a impl H 72 | I⊨H = ⊨a-resp-≡³ (right * right * I) (on-bnode f (ind I) ∘ right) refl 73 | (impl H) (compose-right G H (right * I) (compose-right F (G ∙ H) I I⊨RHS)) 74 | 75 | I⊨LHS : bnodes I f ⊨a impl ((F ∙ G) ∙ H) 76 | I⊨LHS = compose-resp-⊨a (F ∙ G) H (bnodes I f) 77 | (compose-resp-⊨a F G (left * bnodes I f) I⊨F I⊨G) I⊨H 78 | 79 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Category/Properties/Composition/LeftUnit.agda: -------------------------------------------------------------------------------- 1 | open import Relation.Binary.PropositionalEquality using ( refl ) 2 | open import Web.Semantic.DL.ABox.Interp using ( _*_ ; ⌊_⌋ ; ind ) 3 | open import Web.Semantic.DL.ABox.Interp.Morphism using ( _,_ ) 4 | open import Web.Semantic.DL.ABox.Model using 5 | ( _⊨a_ ; on-bnode ; bnodes ; ⊨a-resp-≲ ; ⊨a-resp-≡³ ; _,_ ) 6 | open import Web.Semantic.DL.Category.Object using 7 | ( Object ; IN ; iface ) 8 | open import Web.Semantic.DL.Category.Morphism using 9 | ( _⇒_ ; BN ; impl ; _≣_ ; _⊑_ ; _,_ ) 10 | open import Web.Semantic.DL.Category.Composition using ( _∙_ ) 11 | open import Web.Semantic.DL.Category.Wiring using ( identity ) 12 | open import Web.Semantic.DL.Category.Properties.Composition.Lemmas using 13 | ( compose-left ; compose-right ; compose-resp-⊨a 14 | ; identity-intro ; identity-elim ) 15 | open import Web.Semantic.DL.Signature using ( Signature ) 16 | open import Web.Semantic.DL.TBox using ( TBox ) 17 | open import Web.Semantic.DL.TBox.Interp using ( Δ ; _⊨_≈_ ; ≈-refl ; ≈-sym ) 18 | open import Web.Semantic.DL.TBox.Interp.Morphism using ( ≲-refl ) 19 | open import Web.Semantic.Util using 20 | ( _∘_ ; _⊕_⊕_ ; False ; inode ; bnode ; enode ; left ; right ) 21 | 22 | module Web.Semantic.DL.Category.Properties.Composition.LeftUnit 23 | {Σ : Signature} {S T : TBox Σ} where 24 | 25 | compose-unit₁ : ∀ {A B : Object S T} (F : A ⇒ B) → 26 | (identity A ∙ F ≣ F) 27 | compose-unit₁ {A} {B} F = ( idF⊑F , F⊑idF ) where 28 | 29 | idF⊑F : identity A ∙ F ⊑ F 30 | idF⊑F I I⊨STA I⊨idF = (f , I⊨F) where 31 | 32 | Iˡ⊨id : left * I ⊨a impl (identity A) 33 | Iˡ⊨id = compose-left (identity A) F I I⊨idF 34 | 35 | Iʳ⊨F : right * I ⊨a impl F 36 | Iʳ⊨F = compose-right (identity A) F I I⊨idF 37 | 38 | f : BN F → Δ ⌊ I ⌋ 39 | f w = ind I (bnode (enode w)) 40 | 41 | f✓ : ∀ x → ⌊ I ⌋ ⊨ ind I (right x) ≈ on-bnode f (ind I) x 42 | f✓ (inode x) = ≈-sym ⌊ I ⌋ (identity-elim A (left * I) Iˡ⊨id x) 43 | f✓ (bnode v) = ≈-refl ⌊ I ⌋ 44 | f✓ (enode y) = ≈-refl ⌊ I ⌋ 45 | 46 | I⊨F : bnodes I f ⊨a impl F 47 | I⊨F = ⊨a-resp-≲ (≲-refl ⌊ I ⌋ , f✓) (impl F) Iʳ⊨F 48 | 49 | F⊑idF : F ⊑ identity A ∙ F 50 | F⊑idF I I⊨STA I⊨F = (f , I⊨idF) where 51 | 52 | f : (False ⊕ IN A ⊕ BN F) → Δ ⌊ I ⌋ 53 | f (inode ()) 54 | f (bnode x) = ind I (inode x) 55 | f (enode v) = ind I (bnode v) 56 | 57 | Iˡ⊨id : left * bnodes I f ⊨a impl (identity A) 58 | Iˡ⊨id = identity-intro A (left * bnodes I f) (λ x → ≈-refl ⌊ I ⌋) 59 | 60 | Iʳ⊨F : right * bnodes I f ⊨a impl F 61 | Iʳ⊨F = ⊨a-resp-≡³ I (on-bnode f (ind I) ∘ right) refl (impl F) I⊨F 62 | 63 | I⊨idF : bnodes I f ⊨a impl (identity A ∙ F) 64 | I⊨idF = compose-resp-⊨a (identity A) F (bnodes I f) Iˡ⊨id Iʳ⊨F 65 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Category/Properties/Composition/Lemmas.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( _,_ ; proj₁ ; proj₂ ) 2 | open import Web.Semantic.DL.ABox.Interp using 3 | ( Interp ; _*_ ; ⌊_⌋ ; ind ) 4 | open import Web.Semantic.DL.ABox.Interp.Morphism using 5 | ( _≲_ ; _**_ ; ≲-refl ) 6 | open import Web.Semantic.DL.ABox.Model using 7 | ( _⊨a_ ; _⊨b_ ; *-resp-⟨ABox⟩ ; ⊨a-resp-≲ ) 8 | open import Web.Semantic.DL.Category.Object using 9 | ( Object ; _,_ ; IN ; fin ; iface ) 10 | open import Web.Semantic.DL.Category.Morphism using 11 | ( _⇒_ ; _,_ ; BN ; impl 12 | ; apply ; apply✓ ; apply-init ) 13 | open import Web.Semantic.DL.Category.Wiring using 14 | ( identity ; wires-≈ ; wires-≈⁻¹ ) 15 | open import Web.Semantic.DL.Category.Composition using 16 | ( _∙_ ; ⊨a-intro-⟫ ; ⊨b-intro-⟫ ) 17 | open import Web.Semantic.DL.Integrity using ( med-≲ ; init-med ) 18 | open import Web.Semantic.DL.KB using ( _,_ ) 19 | open import Web.Semantic.DL.KB.Model using ( _⊨_ ) 20 | open import Web.Semantic.DL.Signature using ( Signature ) 21 | open import Web.Semantic.DL.TBox using ( TBox ; _,_ ) 22 | open import Web.Semantic.DL.TBox.Interp using ( _⊨_≈_ ) 23 | open import Web.Semantic.Util using 24 | ( id ; _⊕_⊕_ ; inode ; bnode ; enode ; left ; right ) 25 | 26 | module Web.Semantic.DL.Category.Properties.Composition.Lemmas 27 | {Σ : Signature} {S T : TBox Σ} where 28 | 29 | identity-elim : ∀ (A : Object S T) I → (I ⊨a impl (identity A)) → 30 | ∀ x → (⌊ I ⌋ ⊨ ind I (inode x) ≈ ind I (enode x)) 31 | identity-elim (X , X∈Fin , A) I I⊨id x = wires-≈ id (proj₂ X∈Fin x) I⊨id 32 | 33 | identity-intro : ∀ (A : Object S T) I → 34 | (∀ x → (⌊ I ⌋ ⊨ ind I (inode x) ≈ ind I (enode x))) → 35 | (I ⊨a impl (identity A)) 36 | identity-intro (X , X∈Fin , A) I ∀x∙x₁≈x₂ = wires-≈⁻¹ id ∀x∙x₁≈x₂ (proj₁ X∈Fin) 37 | 38 | compose-left : ∀ {A B C : Object S T} (F : A ⇒ B) (G : B ⇒ C) I → 39 | (I ⊨a impl (F ∙ G)) → (left * I ⊨a impl F) 40 | compose-left {X , X∈Fin , A} {Y , Y∈Fin , B} {Z , Z∈Fin , C} 41 | (V , F , F✓) (W , G , G✓) I (I⊨F , I⊨G) = 42 | *-resp-⟨ABox⟩ left I F I⊨F 43 | 44 | compose-right : ∀ {A B C : Object S T} (F : A ⇒ B) (G : B ⇒ C) I → 45 | (I ⊨a impl (F ∙ G)) → (right * I ⊨a impl G) 46 | compose-right {X , X∈Fin , A} {Y , Y∈Fin , B} {Z , Z∈Fin , C} 47 | (V , F , F✓) (W , G , G✓) I (I⊨F , I⊨G) = 48 | *-resp-⟨ABox⟩ right I G I⊨G 49 | 50 | compose-mid : ∀ {A B C : Object S T} (F : A ⇒ B) (G : B ⇒ C) 51 | (I : Interp Σ (IN A ⊕ (BN F ⊕ IN B ⊕ BN G) ⊕ IN C)) → 52 | (inode * I ⊨ (S , T) , iface A) → (I ⊨a impl (F ∙ G)) → 53 | (bnode * bnode * I ⊨ (S , T) , iface B) 54 | compose-mid {A} {B} {C} F G I ((I⊨S , I⊨T) , I⊨A) I⊨F⟫G = 55 | ((I⊨S , I⊨T) , ⊨a-resp-≲ (enode ** J≲I) (iface B) J⊨B) where 56 | 57 | J : Interp Σ (IN A ⊕ BN F ⊕ IN B) 58 | J = apply F (inode * I) ((I⊨S , I⊨T) , I⊨A) 59 | 60 | J⊨B : enode * J ⊨a iface B 61 | J⊨B = proj₂ (apply✓ F (inode * I) ((I⊨S , I⊨T) , I⊨A)) 62 | 63 | J≲I : J ≲ left * I 64 | J≲I = med-≲ (init-med (apply-init F (inode * I) ((I⊨S , I⊨T) , I⊨A)) 65 | (left * I) (≲-refl (inode * I)) (I⊨S , compose-left F G I I⊨F⟫G)) 66 | 67 | compose-resp-⊨a : ∀ {A B C : Object S T} (F : A ⇒ B) (G : B ⇒ C) I → 68 | (left * I ⊨a impl F) → (right * I ⊨a impl G) → (I ⊨a impl (F ∙ G)) 69 | compose-resp-⊨a {X , X∈Fin , A} {Y , Y∈Fin , B} {Z , Z∈Fin , C} 70 | (V , F , F✓) (W , G , G✓) I I⊨F I⊨G = 71 | ⊨a-intro-⟫ I F G I⊨F I⊨G 72 | 73 | compose-resp-⊨b : ∀ {A B C : Object S T} {V W : Set} (F : A ⇒ B) (G : B ⇒ C) 74 | (I : Interp Σ (IN A ⊕ (V ⊕ IN B ⊕ W) ⊕ IN C)) → 75 | (left * I ⊨b impl F) → (right * I ⊨b impl G) → (I ⊨b impl (F ∙ G)) 76 | compose-resp-⊨b {X , X∈Fin , A} {Y , Y∈Fin , B} {Z , Z∈Fin , C} 77 | (V , F , F✓) (W , G , G✓) I I⊨F I⊨G = 78 | ⊨b-intro-⟫ I F G I⊨F I⊨G 79 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Category/Properties/Composition/RespectsEquiv.agda: -------------------------------------------------------------------------------- 1 | open import Web.Semantic.DL.ABox.Interp using ( _*_ ) 2 | open import Web.Semantic.DL.Category.Object using ( Object ) 3 | open import Web.Semantic.DL.Category.Morphism using ( _⇒_ ; _≣_ ; _⊑_ ; _,_ ) 4 | open import Web.Semantic.DL.Category.Composition using ( _∙_ ) 5 | open import Web.Semantic.DL.Category.Properties.Composition.Lemmas using 6 | ( compose-left ; compose-mid ; compose-right ; compose-resp-⊨b ) 7 | open import Web.Semantic.DL.Signature using ( Signature ) 8 | open import Web.Semantic.DL.TBox using ( TBox ) 9 | open import Web.Semantic.Util using ( left ; right ) 10 | 11 | module Web.Semantic.DL.Category.Properties.Composition.RespectsEquiv 12 | {Σ : Signature} {S T : TBox Σ} where 13 | 14 | compose-resp-⊑ : ∀ {A B C : Object S T} (F₁ F₂ : A ⇒ B) (G₁ G₂ : B ⇒ C) → 15 | (F₁ ⊑ F₂) → (G₁ ⊑ G₂) → (F₁ ∙ G₁ ⊑ F₂ ∙ G₂) 16 | compose-resp-⊑ F₁ F₂ G₁ G₂ F₁⊑F₂ G₁⊑G₂ I I⊨STA I⊨F₁⟫G₁ = 17 | compose-resp-⊨b F₂ G₂ I 18 | (F₁⊑F₂ (left * I) I⊨STA 19 | (compose-left F₁ G₁ I I⊨F₁⟫G₁)) 20 | (G₁⊑G₂ (right * I) (compose-mid F₁ G₁ I I⊨STA I⊨F₁⟫G₁) 21 | (compose-right F₁ G₁ I I⊨F₁⟫G₁)) 22 | 23 | compose-resp-≣ : ∀ {A B C : Object S T} {F₁ F₂ : A ⇒ B} {G₁ G₂ : B ⇒ C} → 24 | (F₁ ≣ F₂) → (G₁ ≣ G₂) → (F₁ ∙ G₁ ≣ F₂ ∙ G₂) 25 | compose-resp-≣ {A} {B} {C} {F₁} {F₂} {G₁} {G₂} 26 | (F₁⊑F₂ , F₂⊑F₁) (G₁⊑G₂ , G₂⊑G₁) = 27 | ( compose-resp-⊑ F₁ F₂ G₁ G₂ F₁⊑F₂ G₁⊑G₂ 28 | , compose-resp-⊑ F₂ F₁ G₂ G₁ F₂⊑F₁ G₂⊑G₁ ) 29 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Category/Properties/Composition/RespectsWiring.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( proj₁ ; proj₂ ) 2 | open import Relation.Binary.PropositionalEquality using ( _≡_ ; sym ; cong ) 3 | open import Relation.Unary using ( _⊆_ ) 4 | open import Web.Semantic.DL.ABox using ( ABox ; ⟨ABox⟩ ; Assertions ) 5 | open import Web.Semantic.DL.ABox.Interp using ( ⌊_⌋ ; ind ) 6 | open import Web.Semantic.DL.ABox.Model using ( _⊨a_ ; bnodes ; _,_ ) 7 | open import Web.Semantic.DL.Category.Object using ( Object ; IN ; fin ; iface ) 8 | open import Web.Semantic.DL.Category.Morphism using ( impl ; _≣_ ; _⊑_ ; _,_ ) 9 | open import Web.Semantic.DL.Category.Composition using ( _∙_ ) 10 | open import Web.Semantic.DL.Category.Properties.Composition.Lemmas using 11 | ( compose-left ; compose-right ; compose-resp-⊨a ) 12 | open import Web.Semantic.DL.Category.Wiring using 13 | ( wiring ; wires-≈ ; wires-≈⁻¹ ) 14 | open import Web.Semantic.DL.Signature using ( Signature ) 15 | open import Web.Semantic.DL.TBox using ( TBox ) 16 | open import Web.Semantic.DL.TBox.Interp using 17 | ( Δ ; _⊨_≈_ ; ≈-refl ; ≈-refl′ ; ≈-trans ) 18 | open import Web.Semantic.Util using 19 | ( _∘_ ; False ; elim ; _⊕_⊕_ ; inode ; bnode ; enode ) 20 | 21 | module Web.Semantic.DL.Category.Properties.Composition.RespectsWiring 22 | {Σ : Signature} {S T : TBox Σ} where 23 | 24 | compose-resp-wiring : ∀ (A B C : Object S T) → 25 | (f : IN B → IN A) → 26 | (f✓ : Assertions (⟨ABox⟩ f (iface B)) ⊆ Assertions (iface A)) → 27 | (g : IN C → IN B) → 28 | (g✓ : Assertions (⟨ABox⟩ g (iface C)) ⊆ Assertions (iface B)) → 29 | (h : IN C → IN A) → 30 | (h✓ : Assertions (⟨ABox⟩ h (iface C)) ⊆ Assertions (iface A)) → 31 | (∀ x → f (g x) ≡ h x) → 32 | (wiring A B f f✓ ∙ wiring B C g g✓ ≣ wiring A C h h✓) 33 | compose-resp-wiring A B C f f✓ g g✓ h h✓ fg≡h = 34 | (LHS⊑RHS , RHS⊑LHS) where 35 | 36 | LHS⊑RHS : wiring A B f f✓ ∙ wiring B C g g✓ ⊑ wiring A C h h✓ 37 | LHS⊑RHS I I⊨STA I⊨F = (elim , I⊨RHS) where 38 | 39 | lemma : ∀ x → ⌊ I ⌋ ⊨ ind I (inode (h x)) ≈ ind I (enode x) 40 | lemma x = ≈-trans ⌊ I ⌋ 41 | (≈-refl′ ⌊ I ⌋ (cong (ind I ∘ inode) (sym (fg≡h x)))) 42 | (≈-trans ⌊ I ⌋ 43 | (wires-≈ f (proj₂ (fin B) (g x)) 44 | (compose-left (wiring A B f f✓) (wiring B C g g✓) I I⊨F)) 45 | (wires-≈ g (proj₂ (fin C) x) 46 | (compose-right (wiring A B f f✓) (wiring B C g g✓) I I⊨F))) 47 | 48 | I⊨RHS : bnodes I elim ⊨a impl (wiring A C h h✓) 49 | I⊨RHS = wires-≈⁻¹ h lemma (proj₁ (fin C)) 50 | 51 | RHS⊑LHS : wiring A C h h✓ ⊑ wiring A B f f✓ ∙ wiring B C g g✓ 52 | RHS⊑LHS I I⊨STA I⊨F = (i , I⊨LHS) where 53 | 54 | i : (False ⊕ IN B ⊕ False) → Δ ⌊ I ⌋ 55 | i (inode ()) 56 | i (bnode y) = ind I (inode (f y)) 57 | i (enode ()) 58 | 59 | lemma : ∀ x → ⌊ I ⌋ ⊨ ind I (inode (f (g x))) ≈ ind I (enode x) 60 | lemma x = ≈-trans ⌊ I ⌋ 61 | (≈-refl′ ⌊ I ⌋ (cong (ind I ∘ inode) (fg≡h x))) 62 | (wires-≈ h (proj₂ (fin C) x) I⊨F) 63 | 64 | I⊨LHS : bnodes I i ⊨a impl (wiring A B f f✓ ∙ wiring B C g g✓) 65 | I⊨LHS = compose-resp-⊨a (wiring A B f f✓) (wiring B C g g✓) (bnodes I i) 66 | (wires-≈⁻¹ f (λ x → ≈-refl ⌊ I ⌋) (proj₁ (fin B))) 67 | (wires-≈⁻¹ g lemma (proj₁ (fin C))) 68 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Category/Properties/Composition/RightUnit.agda: -------------------------------------------------------------------------------- 1 | open import Relation.Binary.PropositionalEquality using ( refl ) 2 | open import Web.Semantic.DL.ABox.Interp using ( _*_ ; ⌊_⌋ ; ind ) 3 | open import Web.Semantic.DL.ABox.Interp.Morphism using ( _,_ ) 4 | open import Web.Semantic.DL.ABox.Model using 5 | ( _⊨a_ ; on-bnode ; bnodes ; ⊨a-resp-≲ ; ⊨a-resp-≡³ ; _,_ ) 6 | open import Web.Semantic.DL.Category.Object using 7 | ( Object ; IN ; iface ) 8 | open import Web.Semantic.DL.Category.Morphism using 9 | ( _⇒_ ; BN ; impl ; _≣_ ; _⊑_ ; _,_ ) 10 | open import Web.Semantic.DL.Category.Composition using ( _∙_ ) 11 | open import Web.Semantic.DL.Category.Wiring using ( identity ) 12 | open import Web.Semantic.DL.Category.Properties.Composition.Lemmas using 13 | ( compose-left ; compose-right ; compose-resp-⊨a 14 | ; identity-intro ; identity-elim ) 15 | open import Web.Semantic.DL.Signature using ( Signature ) 16 | open import Web.Semantic.DL.TBox using ( TBox ) 17 | open import Web.Semantic.DL.TBox.Interp using ( Δ ; _⊨_≈_ ; ≈-refl ; ≈-sym ) 18 | open import Web.Semantic.DL.TBox.Interp.Morphism using ( ≲-refl ) 19 | open import Web.Semantic.Util using 20 | ( _∘_ ; _⊕_⊕_ ; False ; inode ; bnode ; enode ; left ; right ) 21 | 22 | module Web.Semantic.DL.Category.Properties.Composition.RightUnit 23 | {Σ : Signature} {S T : TBox Σ} where 24 | 25 | compose-unit₂ : ∀ {A B : Object S T} (F : A ⇒ B) → 26 | (F ∙ identity B ≣ F) 27 | compose-unit₂ {A} {B} F = ( Fid⊑F , F⊑Fid ) where 28 | 29 | Fid⊑F : F ∙ identity B ⊑ F 30 | Fid⊑F I I⊨STA I⊨Fid = (f , I⊨F) where 31 | 32 | Iˡ⊨F : left * I ⊨a impl F 33 | Iˡ⊨F = compose-left F (identity B) I I⊨Fid 34 | 35 | Iʳ⊨id : right * I ⊨a impl (identity B) 36 | Iʳ⊨id = compose-right F (identity B) I I⊨Fid 37 | 38 | f : BN F → Δ ⌊ I ⌋ 39 | f w = ind I (bnode (inode w)) 40 | 41 | f✓ : ∀ x → ⌊ I ⌋ ⊨ ind I (left x) ≈ on-bnode f (ind I) x 42 | f✓ (inode x) = ≈-refl ⌊ I ⌋ 43 | f✓ (bnode v) = ≈-refl ⌊ I ⌋ 44 | f✓ (enode y) = identity-elim B (right * I) Iʳ⊨id y 45 | 46 | I⊨F : bnodes I f ⊨a impl F 47 | I⊨F = ⊨a-resp-≲ (≲-refl ⌊ I ⌋ , f✓) (impl F) Iˡ⊨F 48 | 49 | F⊑Fid : F ⊑ F ∙ identity B 50 | F⊑Fid I I⊨STA I⊨F = (f , I⊨Fid) where 51 | 52 | f : (BN F ⊕ IN B ⊕ False) → Δ ⌊ I ⌋ 53 | f (inode v) = ind I (bnode v) 54 | f (bnode y) = ind I (enode y) 55 | f (enode ()) 56 | 57 | Iˡ⊨F : left * bnodes I f ⊨a impl F 58 | Iˡ⊨F = ⊨a-resp-≡³ I (on-bnode f (ind I) ∘ left) refl (impl F) I⊨F 59 | 60 | Iʳ⊨id : right * bnodes I f ⊨a impl (identity B) 61 | Iʳ⊨id = identity-intro B (right * bnodes I f) (λ x → ≈-refl ⌊ I ⌋) 62 | 63 | I⊨Fid : bnodes I f ⊨a impl (F ∙ identity B) 64 | I⊨Fid = compose-resp-⊨a F (identity B) (bnodes I f) Iˡ⊨F Iʳ⊨id 65 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Category/Properties/Equivalence.agda: -------------------------------------------------------------------------------- 1 | open import Relation.Binary.PropositionalEquality using ( refl ) 2 | open import Web.Semantic.DL.ABox.Interp using ( Interp ; ⌊_⌋ ; ind ; _*_ ) 3 | open import Web.Semantic.DL.ABox.Model using 4 | ( _⊨a_ ; ⊨a-resp-≡³ ; ⊨a-impl-⊨b ; ⊨b-impl-⊨a 5 | ; _,_ ; inb ; on-bnode ; bnodes ) 6 | open import Web.Semantic.DL.Category.Object using ( Object ; IN ; iface ) 7 | open import Web.Semantic.DL.Category.Morphism using 8 | ( _⇒_ ; _,_ ; BN ; impl ; _⊑_ ; _≣_ ) 9 | open import Web.Semantic.DL.KB using ( _,_ ) 10 | open import Web.Semantic.DL.KB.Model using ( _⊨_ ) 11 | open import Web.Semantic.DL.Signature using ( Signature ) 12 | open import Web.Semantic.DL.TBox using ( TBox ; _,_ ) 13 | open import Web.Semantic.DL.TBox.Interp using ( Δ ) 14 | open import Web.Semantic.Util using ( _⊕_⊕_ ; inode ) 15 | 16 | module Web.Semantic.DL.Category.Properties.Equivalence 17 | {Σ : Signature} {S T : TBox Σ} where 18 | 19 | ⊑-refl : ∀ {A B : Object S T} (F : A ⇒ B) → (F ⊑ F) 20 | ⊑-refl F I I⊨STA I⊨F = ⊨a-impl-⊨b I (impl F) I⊨F 21 | 22 | ⊑-trans : ∀ {A B : Object S T} (F G H : A ⇒ B) → (F ⊑ G) → (G ⊑ H) → (F ⊑ H) 23 | ⊑-trans {A} {B} F G H F⊑G G⊑H I I⊨STA I⊨F = (g , I⊨H) where 24 | 25 | f : BN G → Δ ⌊ I ⌋ 26 | f = inb (F⊑G I I⊨STA I⊨F) 27 | 28 | J : Interp Σ (IN A ⊕ BN G ⊕ IN B) 29 | J = bnodes I f 30 | 31 | J⊨STA : inode * J ⊨ (S , T) , iface A 32 | J⊨STA = I⊨STA 33 | 34 | J⊨G : J ⊨a impl G 35 | J⊨G = ⊨b-impl-⊨a (F⊑G I I⊨STA I⊨F) 36 | 37 | g : BN H → Δ ⌊ I ⌋ 38 | g = inb (G⊑H J J⊨STA J⊨G) 39 | 40 | K : Interp Σ (IN A ⊕ BN H ⊕ IN B) 41 | K = bnodes J g 42 | 43 | K⊨H : K ⊨a impl H 44 | K⊨H = ⊨b-impl-⊨a (G⊑H J J⊨STA J⊨G) 45 | 46 | I⊨H : bnodes I g ⊨a impl H 47 | I⊨H = ⊨a-resp-≡³ K (on-bnode g (ind I)) refl (impl H) K⊨H 48 | 49 | ≣-refl : ∀ {A B : Object S T} (F : A ⇒ B) → (F ≣ F) 50 | ≣-refl F = (⊑-refl F , ⊑-refl F) 51 | 52 | ≣-sym : ∀ {A B : Object S T} {F G : A ⇒ B} → (F ≣ G) → (G ≣ F) 53 | ≣-sym (F⊑G , G⊑F) = (G⊑F , F⊑G) 54 | 55 | ≣-trans : ∀ {A B : Object S T} {F G H : A ⇒ B} → (F ≣ G) → (G ≣ H) → (F ≣ H) 56 | ≣-trans {A} {B} {F} {G} {H} (F⊑G , G⊑F) (G⊑H , H⊑G) = 57 | (⊑-trans F G H F⊑G G⊑H , ⊑-trans H G F H⊑G G⊑F) 58 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Category/Properties/Tensor.agda: -------------------------------------------------------------------------------- 1 | module Web.Semantic.DL.Category.Properties.Tensor where 2 | 3 | open import Web.Semantic.DL.Category.Properties.Tensor.RespectsEquiv public using 4 | ( tensor-resp-≣ ) 5 | 6 | open import Web.Semantic.DL.Category.Properties.Tensor.Functor public using 7 | ( tensor-resp-id ; tensor-resp-compose ) 8 | 9 | open import Web.Semantic.DL.Category.Properties.Tensor.Isomorphisms public using 10 | ( symm-iso ; assoc-iso ; assoc⁻¹-iso 11 | ; unit₁-iso ; unit₁⁻¹-iso ; unit₂-iso ; unit₂⁻¹-iso ) 12 | 13 | open import Web.Semantic.DL.Category.Properties.Tensor.Coherence public using 14 | ( assoc-unit ; assoc-assoc ; assoc-symm ) 15 | 16 | open import Web.Semantic.DL.Category.Properties.Tensor.UnitNatural public using 17 | ( unit₁-natural ; unit₂-natural ) 18 | 19 | open import Web.Semantic.DL.Category.Properties.Tensor.SymmNatural public using 20 | ( symm-natural ) 21 | 22 | open import Web.Semantic.DL.Category.Properties.Tensor.AssocNatural public using 23 | ( assoc-natural ) 24 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Category/Properties/Tensor/AssocNatural.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( proj₁ ; proj₂ ) 2 | open import Data.Sum using ( _⊎_ ; inj₁ ; inj₂ ) 3 | open import Relation.Binary.PropositionalEquality using ( refl ) 4 | open import Web.Semantic.DL.ABox.Interp using ( ⌊_⌋ ; ind ; _*_ ) 5 | open import Web.Semantic.DL.ABox.Interp.Morphism using ( _,_ ) 6 | open import Web.Semantic.DL.ABox.Model using 7 | ( _⊨a_ ; on-bnode ; bnodes ; _,_ ; ⊨a-resp-≲ ) 8 | open import Web.Semantic.DL.Category.Composition using ( _∙_ ) 9 | open import Web.Semantic.DL.Category.Properties.Composition.Lemmas using 10 | ( compose-left ; compose-right ; compose-resp-⊨a ) 11 | open import Web.Semantic.DL.Category.Properties.Tensor.Lemmas using 12 | ( tensor-up ; tensor-down ; tensor-resp-⊨a ) 13 | open import Web.Semantic.DL.Category.Object using ( Object ; IN ; fin ) 14 | open import Web.Semantic.DL.Category.Morphism using 15 | ( _⇒_ ; BN ; impl ; _⊑_ ; _≣_ ; _,_ ) 16 | open import Web.Semantic.DL.Category.Tensor using ( _⊗_ ; _⟨⊗⟩_ ) 17 | open import Web.Semantic.DL.Category.Unit using ( I ) 18 | open import Web.Semantic.DL.Category.Wiring using 19 | ( wires-≈ ; wires-≈⁻¹ ; assoc ) 20 | open import Web.Semantic.DL.Signature using ( Signature ) 21 | open import Web.Semantic.DL.TBox using ( TBox ) 22 | open import Web.Semantic.DL.TBox.Interp using 23 | ( Δ ; _⊨_≈_ ; ≈-refl ; ≈-refl′ ; ≈-sym ) 24 | open import Web.Semantic.DL.TBox.Interp.Morphism using ( ≲-refl ) 25 | open import Web.Semantic.Util using 26 | ( _∘_ ; False ; ⊎-assoc ; ⊎-assoc⁻¹ 27 | ; _⊕_⊕_ ; inode ; bnode ; enode ; left ; right ; up ; down ) 28 | 29 | module Web.Semantic.DL.Category.Properties.Tensor.AssocNatural 30 | {Σ : Signature} {S T : TBox Σ} where 31 | 32 | assoc-natural : ∀ {A₁ A₂ A₃ B₁ B₂ B₃ : Object S T} 33 | (F₁ : A₁ ⇒ B₁) (F₂ : A₂ ⇒ B₂) (F₃ : A₃ ⇒ B₃) → 34 | (((F₁ ⟨⊗⟩ F₂) ⟨⊗⟩ F₃) ∙ assoc B₁ B₂ B₃ ≣ 35 | assoc A₁ A₂ A₃ ∙ (F₁ ⟨⊗⟩ (F₂ ⟨⊗⟩ F₃))) 36 | assoc-natural {A₁} {A₂} {A₃} {B₁} {B₂} {B₃} F₁ F₂ F₃ = (LHS⊑RHS , RHS⊑LHS) where 37 | 38 | LHS⊑RHS : ((F₁ ⟨⊗⟩ F₂) ⟨⊗⟩ F₃) ∙ assoc B₁ B₂ B₃ ⊑ 39 | assoc A₁ A₂ A₃ ∙ (F₁ ⟨⊗⟩ (F₂ ⟨⊗⟩ F₃)) 40 | LHS⊑RHS J J⊨STA J⊨LHS = (f , J⊨RHS) where 41 | 42 | f : (False ⊕ (IN A₁ ⊎ (IN A₂ ⊎ IN A₃)) ⊕ (BN F₁ ⊎ (BN F₂ ⊎ BN F₃))) → Δ ⌊ J ⌋ 43 | f (inode ()) 44 | f (bnode x) = ind J (inode (⊎-assoc⁻¹ x)) 45 | f (enode v) = ind J (bnode (inode (⊎-assoc⁻¹ v))) 46 | 47 | lemma₁ : ∀ x → 48 | ⌊ J ⌋ ⊨ ind J (left (up (up x))) ≈ 49 | on-bnode f (ind J) (right (up x)) 50 | lemma₁ (inode x) = ≈-refl ⌊ J ⌋ 51 | lemma₁ (bnode v) = ≈-refl ⌊ J ⌋ 52 | lemma₁ (enode y) = wires-≈ ⊎-assoc⁻¹ 53 | (proj₂ (fin (B₁ ⊗ (B₂ ⊗ B₃))) (inj₁ y)) 54 | (compose-right ((F₁ ⟨⊗⟩ F₂) ⟨⊗⟩ F₃) (assoc B₁ B₂ B₃) J J⊨LHS) 55 | 56 | lemma₂ : ∀ x → 57 | ⌊ J ⌋ ⊨ ind J (left (up (down x))) ≈ 58 | on-bnode f (ind J) (right (down (up x))) 59 | lemma₂ (inode x) = ≈-refl ⌊ J ⌋ 60 | lemma₂ (bnode v) = ≈-refl ⌊ J ⌋ 61 | lemma₂ (enode y) = wires-≈ ⊎-assoc⁻¹ 62 | (proj₂ (fin (B₁ ⊗ (B₂ ⊗ B₃))) (inj₂ (inj₁ y))) 63 | (compose-right ((F₁ ⟨⊗⟩ F₂) ⟨⊗⟩ F₃) (assoc B₁ B₂ B₃) J J⊨LHS) 64 | 65 | lemma₃ : ∀ x → 66 | ⌊ J ⌋ ⊨ ind J (left (down x)) ≈ 67 | on-bnode f (ind J) (right (down (down x))) 68 | lemma₃ (inode x) = ≈-refl ⌊ J ⌋ 69 | lemma₃ (bnode v) = ≈-refl ⌊ J ⌋ 70 | lemma₃ (enode y) = wires-≈ ⊎-assoc⁻¹ 71 | (proj₂ (fin (B₁ ⊗ (B₂ ⊗ B₃))) (inj₂ (inj₂ y))) 72 | (compose-right ((F₁ ⟨⊗⟩ F₂) ⟨⊗⟩ F₃) (assoc B₁ B₂ B₃) J J⊨LHS) 73 | 74 | J⊨RHS : bnodes J f ⊨a impl (assoc A₁ A₂ A₃ ∙ (F₁ ⟨⊗⟩ (F₂ ⟨⊗⟩ F₃))) 75 | J⊨RHS = compose-resp-⊨a (assoc A₁ A₂ A₃) (F₁ ⟨⊗⟩ (F₂ ⟨⊗⟩ F₃)) (bnodes J f) 76 | (wires-≈⁻¹ ⊎-assoc⁻¹ (λ x → ≈-refl ⌊ J ⌋) (proj₁ (fin (A₁ ⊗ (A₂ ⊗ A₃))))) 77 | (tensor-resp-⊨a F₁ (F₂ ⟨⊗⟩ F₃) (right * bnodes J f) 78 | (⊨a-resp-≲ (≲-refl ⌊ J ⌋ , lemma₁) (impl F₁) 79 | (tensor-up F₁ F₂ (up * left * J) 80 | (tensor-up (F₁ ⟨⊗⟩ F₂) F₃ (left * J) 81 | (compose-left ((F₁ ⟨⊗⟩ F₂) ⟨⊗⟩ F₃) (assoc B₁ B₂ B₃) J J⊨LHS)))) 82 | (tensor-resp-⊨a F₂ F₃ (down * right * bnodes J f) 83 | (⊨a-resp-≲ (≲-refl ⌊ J ⌋ , lemma₂) (impl F₂) 84 | (tensor-down F₁ F₂ (up * left * J) 85 | (tensor-up (F₁ ⟨⊗⟩ F₂) F₃ (left * J) 86 | (compose-left ((F₁ ⟨⊗⟩ F₂) ⟨⊗⟩ F₃) (assoc B₁ B₂ B₃) J J⊨LHS)))) 87 | (⊨a-resp-≲ (≲-refl ⌊ J ⌋ , lemma₃) (impl F₃) 88 | (tensor-down (F₁ ⟨⊗⟩ F₂) F₃ (left * J) 89 | (compose-left ((F₁ ⟨⊗⟩ F₂) ⟨⊗⟩ F₃) (assoc B₁ B₂ B₃) J J⊨LHS))))) 90 | 91 | RHS⊑LHS : assoc A₁ A₂ A₃ ∙ (F₁ ⟨⊗⟩ (F₂ ⟨⊗⟩ F₃)) ⊑ 92 | ((F₁ ⟨⊗⟩ F₂) ⟨⊗⟩ F₃) ∙ assoc B₁ B₂ B₃ 93 | RHS⊑LHS J J⊨STA J⊨RHS = (f , J⊨LHS) where 94 | 95 | f : (((BN F₁ ⊎ BN F₂) ⊎ BN F₃) ⊕ ((IN B₁ ⊎ IN B₂) ⊎ IN B₃) ⊕ False) → Δ ⌊ J ⌋ 96 | f (inode v) = ind J (bnode (enode (⊎-assoc v))) 97 | f (bnode y) = ind J (enode (⊎-assoc y)) 98 | f (enode ()) 99 | 100 | lemma₀ : ∀ x → 101 | ⌊ J ⌋ ⊨ ind J (enode (⊎-assoc (⊎-assoc⁻¹ x))) ≈ ind J (enode x) 102 | lemma₀ (inj₁ x) = ≈-refl ⌊ J ⌋ 103 | lemma₀ (inj₂ (inj₁ x)) = ≈-refl ⌊ J ⌋ 104 | lemma₀ (inj₂ (inj₂ y)) = ≈-refl ⌊ J ⌋ 105 | 106 | lemma₁ : ∀ x → ⌊ J ⌋ ⊨ ind J (right (up x)) ≈ 107 | on-bnode f (ind J) (left (up (up x))) 108 | lemma₁ (inode x) = ≈-sym ⌊ J ⌋ (wires-≈ ⊎-assoc⁻¹ 109 | (proj₂ (fin (A₁ ⊗ (A₂ ⊗ A₃))) (inj₁ x)) 110 | (compose-left (assoc A₁ A₂ A₃) (F₁ ⟨⊗⟩ (F₂ ⟨⊗⟩ F₃)) J J⊨RHS)) 111 | lemma₁ (bnode v) = ≈-refl ⌊ J ⌋ 112 | lemma₁ (enode y) = ≈-refl ⌊ J ⌋ 113 | 114 | lemma₂ : ∀ x → ⌊ J ⌋ ⊨ ind J (right (down (up x))) ≈ 115 | on-bnode f (ind J) (left (up (down x))) 116 | lemma₂ (inode x) = ≈-sym ⌊ J ⌋ (wires-≈ ⊎-assoc⁻¹ 117 | (proj₂ (fin (A₁ ⊗ (A₂ ⊗ A₃))) (inj₂ (inj₁ x))) 118 | (compose-left (assoc A₁ A₂ A₃) (F₁ ⟨⊗⟩ (F₂ ⟨⊗⟩ F₃)) J J⊨RHS)) 119 | lemma₂ (bnode v) = ≈-refl ⌊ J ⌋ 120 | lemma₂ (enode y) = ≈-refl ⌊ J ⌋ 121 | 122 | lemma₃ : ∀ x → ⌊ J ⌋ ⊨ ind J (right (down (down x))) ≈ 123 | on-bnode f (ind J) (left (down x)) 124 | lemma₃ (inode x) = ≈-sym ⌊ J ⌋ (wires-≈ ⊎-assoc⁻¹ 125 | (proj₂ (fin (A₁ ⊗ (A₂ ⊗ A₃))) (inj₂ (inj₂ x))) 126 | (compose-left (assoc A₁ A₂ A₃) (F₁ ⟨⊗⟩ (F₂ ⟨⊗⟩ F₃)) J J⊨RHS)) 127 | lemma₃ (bnode v) = ≈-refl ⌊ J ⌋ 128 | lemma₃ (enode y) = ≈-refl ⌊ J ⌋ 129 | 130 | J⊨LHS : bnodes J f ⊨a impl (((F₁ ⟨⊗⟩ F₂) ⟨⊗⟩ F₃) ∙ assoc B₁ B₂ B₃) 131 | J⊨LHS = compose-resp-⊨a ((F₁ ⟨⊗⟩ F₂) ⟨⊗⟩ F₃) (assoc B₁ B₂ B₃) (bnodes J f) 132 | (tensor-resp-⊨a (F₁ ⟨⊗⟩ F₂) F₃ (left * bnodes J f) 133 | (tensor-resp-⊨a F₁ F₂ (up * left * bnodes J f) 134 | (⊨a-resp-≲ (≲-refl ⌊ J ⌋ , lemma₁) (impl F₁) 135 | (tensor-up F₁ (F₂ ⟨⊗⟩ F₃) (right * J) 136 | (compose-right (assoc A₁ A₂ A₃) (F₁ ⟨⊗⟩ (F₂ ⟨⊗⟩ F₃)) J J⊨RHS))) 137 | (⊨a-resp-≲ (≲-refl ⌊ J ⌋ , lemma₂) (impl F₂) 138 | (tensor-up F₂ F₃ (down * right * J) 139 | (tensor-down F₁ (F₂ ⟨⊗⟩ F₃) (right * J) 140 | (compose-right (assoc A₁ A₂ A₃) (F₁ ⟨⊗⟩ (F₂ ⟨⊗⟩ F₃)) J J⊨RHS))))) 141 | (⊨a-resp-≲ (≲-refl ⌊ J ⌋ , lemma₃) (impl F₃) 142 | (tensor-down F₂ F₃ (down * right * J) 143 | (tensor-down F₁ (F₂ ⟨⊗⟩ F₃) (right * J) 144 | (compose-right (assoc A₁ A₂ A₃) (F₁ ⟨⊗⟩ (F₂ ⟨⊗⟩ F₃)) J J⊨RHS))))) 145 | (wires-≈⁻¹ ⊎-assoc⁻¹ lemma₀ (proj₁ (fin (B₁ ⊗ (B₂ ⊗ B₃))))) 146 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Category/Properties/Tensor/Coherence.agda: -------------------------------------------------------------------------------- 1 | open import Relation.Binary.PropositionalEquality using ( refl ) 2 | open import Web.Semantic.DL.Category.Object using ( Object ) 3 | open import Web.Semantic.DL.Category.Morphism using ( _≣_ ) 4 | open import Web.Semantic.DL.Category.Composition using ( _∙_ ) 5 | open import Web.Semantic.DL.Category.Tensor using ( _⊗_ ; _⟨⊗⟩_ ) 6 | open import Web.Semantic.DL.Category.Unit using ( I ) 7 | open import Web.Semantic.DL.Category.Wiring using 8 | ( identity ; symm ; assoc ; assoc⁻¹ ; unit₁ ; unit₁⁻¹ ; unit₂ ; unit₂⁻¹ ) 9 | open import Web.Semantic.DL.Category.Properties.Wiring using 10 | ( rewriting ; rewrite-compose ; rewrite-tensor 11 | ; rewrite-identity ; rewrite-symm ; rewrite-assoc ; rewrite-assoc⁻¹ 12 | ; rewrite-unit₁ ; rewrite-unit₁⁻¹ ; rewrite-unit₂ ; rewrite-unit₂⁻¹ ) 13 | open import Web.Semantic.DL.Signature using ( Signature ) 14 | open import Web.Semantic.DL.TBox using ( TBox ) 15 | open import Web.Semantic.Util using ( _≡⊎≡_ ) 16 | 17 | module Web.Semantic.DL.Category.Properties.Tensor.Coherence 18 | {Σ : Signature} {S T : TBox Σ} where 19 | 20 | assoc-unit : ∀ (A₁ A₂ : Object S T) → 21 | (assoc A₁ I A₂ ∙ (identity A₁ ⟨⊗⟩ unit₁ A₂) ≣ unit₂ A₁ ⟨⊗⟩ identity A₂) 22 | assoc-unit A₁ A₂ = rewriting 23 | (rewrite-compose (rewrite-assoc A₁ I A₂) 24 | (rewrite-tensor (rewrite-identity A₁) (rewrite-unit₁ A₂))) 25 | ((λ x → refl) ≡⊎≡ (λ x → refl)) 26 | (rewrite-tensor (rewrite-unit₂ A₁) (rewrite-identity A₂)) 27 | 28 | assoc-assoc : ∀ (A₁ A₂ A₃ A₄ : Object S T) → 29 | (assoc (A₁ ⊗ A₂) A₃ A₄ ∙ assoc A₁ A₂ (A₃ ⊗ A₄) ≣ 30 | (assoc A₁ A₂ A₃ ⟨⊗⟩ identity A₄) ∙ assoc A₁ (A₂ ⊗ A₃) A₄ ∙ (identity A₁ ⟨⊗⟩ assoc A₂ A₃ A₄)) 31 | assoc-assoc A₁ A₂ A₃ A₄ = rewriting 32 | (rewrite-compose (rewrite-assoc (A₁ ⊗ A₂) A₃ A₄) (rewrite-assoc A₁ A₂ (A₃ ⊗ A₄))) 33 | ((λ x → refl) ≡⊎≡ ((λ x → refl) ≡⊎≡ ((λ x → refl) ≡⊎≡ (λ x → refl)))) 34 | (rewrite-compose (rewrite-tensor (rewrite-assoc A₁ A₂ A₃) (rewrite-identity A₄)) 35 | (rewrite-compose (rewrite-assoc A₁ (A₂ ⊗ A₃) A₄) 36 | (rewrite-tensor (rewrite-identity A₁) (rewrite-assoc A₂ A₃ A₄)))) 37 | 38 | assoc-symm : ∀ (A₁ A₂ A₃ : Object S T) → 39 | (symm (A₁ ⊗ A₂) A₃ ∙ assoc⁻¹ A₃ A₁ A₂ ≣ 40 | assoc A₁ A₂ A₃ ∙ (identity A₁ ⟨⊗⟩ symm A₂ A₃) ∙ assoc⁻¹ A₁ A₃ A₂ ∙ (symm A₁ A₃ ⟨⊗⟩ identity A₂)) 41 | assoc-symm A₁ A₂ A₃ = rewriting 42 | (rewrite-compose (rewrite-symm (A₁ ⊗ A₂) A₃) (rewrite-assoc⁻¹ A₃ A₁ A₂)) 43 | (((λ x → refl) ≡⊎≡ (λ x → refl)) ≡⊎≡ (λ x → refl)) 44 | (rewrite-compose (rewrite-assoc A₁ A₂ A₃) 45 | (rewrite-compose (rewrite-tensor (rewrite-identity A₁) (rewrite-symm A₂ A₃)) 46 | (rewrite-compose (rewrite-assoc⁻¹ A₁ A₃ A₂) 47 | (rewrite-tensor (rewrite-symm A₁ A₃) (rewrite-identity A₂))))) 48 | 49 | symm-unit : ∀ (A : Object S T) → 50 | (symm A I ∙ unit₁ A ≣ unit₂ A) 51 | symm-unit A = rewriting 52 | (rewrite-compose (rewrite-symm A I) (rewrite-unit₁ A)) 53 | (λ x → refl) 54 | (rewrite-unit₂ A) 55 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Category/Properties/Tensor/Functor.agda: -------------------------------------------------------------------------------- 1 | open import Data.Sum using ( _⊎_ ; inj₁ ; inj₂ ) 2 | open import Relation.Binary.PropositionalEquality using ( refl ) 3 | open import Web.Semantic.DL.ABox.Interp using ( ⌊_⌋ ; ind ; _*_ ) 4 | open import Web.Semantic.DL.ABox.Model using 5 | ( _⊨a_ ; ⊨a-resp-≡³ ; on-bnode ; bnodes ; _,_ ) 6 | open import Web.Semantic.DL.Category.Object using ( Object ; IN ) 7 | open import Web.Semantic.DL.Category.Morphism using 8 | ( _⇒_ ; BN ; impl ; _⊑_ ; _≣_ ; _,_ ) 9 | open import Web.Semantic.DL.Category.Composition using ( _∙_ ) 10 | open import Web.Semantic.DL.Category.Tensor using ( _⊗_ ; _⟨⊗⟩_ ) 11 | open import Web.Semantic.DL.Category.Wiring using ( identity ; id✓ ) 12 | open import Web.Semantic.DL.Category.Properties.Composition.Lemmas using 13 | ( compose-left ; compose-right ; compose-resp-⊨a ) 14 | open import Web.Semantic.DL.Category.Properties.Tensor.Lemmas using 15 | ( tensor-up ; tensor-down ; tensor-resp-⊨a ) 16 | open import Web.Semantic.DL.Category.Properties.Tensor.RespectsWiring using 17 | ( tensor-resp-wiring ) 18 | open import Web.Semantic.DL.Signature using ( Signature ) 19 | open import Web.Semantic.DL.TBox using ( TBox ) 20 | open import Web.Semantic.DL.TBox.Interp using ( Δ ) 21 | open import Web.Semantic.Util using 22 | ( id ; _∘_ ; _⊕_⊕_ ; inode ; bnode ; enode ; left ; right ; up ; down ) 23 | 24 | module Web.Semantic.DL.Category.Properties.Tensor.Functor 25 | {Σ : Signature} {S T : TBox Σ} where 26 | 27 | tensor-resp-id : ∀ (A₁ A₂ : Object S T) → 28 | ((identity A₁ ⟨⊗⟩ identity A₂) ≣ identity (A₁ ⊗ A₂)) 29 | tensor-resp-id A₁ A₂ = 30 | tensor-resp-wiring A₁ A₂ A₁ A₂ 31 | id (id✓ A₁) id (id✓ A₂) id (id✓ (A₁ ⊗ A₂)) 32 | (λ x → refl) (λ x → refl) 33 | 34 | tensor-resp-compose : ∀ {A₁ A₂ B₁ B₂ C₁ C₂ : Object S T} 35 | (F₁ : A₁ ⇒ B₁) (F₂ : A₂ ⇒ B₂) (G₁ : B₁ ⇒ C₁) (G₂ : B₂ ⇒ C₂) → 36 | (((F₁ ∙ G₁) ⟨⊗⟩ (F₂ ∙ G₂)) ≣ ((F₁ ⟨⊗⟩ F₂) ∙ (G₁ ⟨⊗⟩ G₂))) 37 | tensor-resp-compose {A₁} {A₂} {B₁} {B₂} {C₁} {C₂} F₁ F₂ G₁ G₂ = 38 | (LHS⊑RHS , RHS⊑LHS) where 39 | 40 | LHS⊑RHS : (F₁ ∙ G₁) ⟨⊗⟩ (F₂ ∙ G₂) ⊑ (F₁ ⟨⊗⟩ F₂) ∙ (G₁ ⟨⊗⟩ G₂) 41 | LHS⊑RHS I I⊨STA₁A₂ I⊨LHS = (f , I⊨RHS) where 42 | 43 | f : ((BN F₁ ⊎ BN F₂) ⊕ (IN B₁ ⊎ IN B₂) ⊕ (BN G₁ ⊎ BN G₂)) → Δ ⌊ I ⌋ 44 | f (inode (inj₁ v)) = ind I (bnode (inj₁ (inode v))) 45 | f (inode (inj₂ v)) = ind I (bnode (inj₂ (inode v))) 46 | f (bnode (inj₁ y)) = ind I (bnode (inj₁ (bnode y))) 47 | f (bnode (inj₂ y)) = ind I (bnode (inj₂ (bnode y))) 48 | f (enode (inj₁ w)) = ind I (bnode (inj₁ (enode w))) 49 | f (enode (inj₂ w)) = ind I (bnode (inj₂ (enode w))) 50 | 51 | Iˡ₁⊨F₁ : up * left * bnodes I f ⊨a impl F₁ 52 | Iˡ₁⊨F₁ = ⊨a-resp-≡³ 53 | (left * up * I) (on-bnode f (ind I) ∘ left ∘ up) refl (impl F₁) 54 | (compose-left F₁ G₁ (up * I) (tensor-up (F₁ ∙ G₁) (F₂ ∙ G₂) I I⊨LHS)) 55 | 56 | Iˡ₂⊨F₂ : down * (left * bnodes I f) ⊨a impl F₂ 57 | Iˡ₂⊨F₂ = ⊨a-resp-≡³ 58 | (left * down * I) (on-bnode f (ind I) ∘ left ∘ down) refl (impl F₂) 59 | (compose-left F₂ G₂ (down * I) (tensor-down (F₁ ∙ G₁) (F₂ ∙ G₂) I I⊨LHS)) 60 | 61 | Iʳ₁⊨G₁ : up * (right * bnodes I f) ⊨a impl G₁ 62 | Iʳ₁⊨G₁ = ⊨a-resp-≡³ 63 | (right * up * I) (on-bnode f (ind I) ∘ right ∘ up) refl (impl G₁) 64 | (compose-right F₁ G₁ (up * I) (tensor-up (F₁ ∙ G₁) (F₂ ∙ G₂) I I⊨LHS)) 65 | 66 | Iʳ₂⊨G₂ : down * (right * bnodes I f) ⊨a impl G₂ 67 | Iʳ₂⊨G₂ = ⊨a-resp-≡³ 68 | (right * down * I) (on-bnode f (ind I) ∘ right ∘ down) refl (impl G₂) 69 | (compose-right F₂ G₂ (down * I) (tensor-down (F₁ ∙ G₁) (F₂ ∙ G₂) I I⊨LHS)) 70 | 71 | I⊨RHS : bnodes I f ⊨a impl ((F₁ ⟨⊗⟩ F₂) ∙ (G₁ ⟨⊗⟩ G₂)) 72 | I⊨RHS = compose-resp-⊨a (F₁ ⟨⊗⟩ F₂) (G₁ ⟨⊗⟩ G₂) (bnodes I f) 73 | (tensor-resp-⊨a F₁ F₂ (left * bnodes I f) Iˡ₁⊨F₁ Iˡ₂⊨F₂) 74 | (tensor-resp-⊨a G₁ G₂ (right * bnodes I f) Iʳ₁⊨G₁ Iʳ₂⊨G₂) 75 | 76 | RHS⊑LHS : (F₁ ⟨⊗⟩ F₂) ∙ (G₁ ⟨⊗⟩ G₂) ⊑ (F₁ ∙ G₁) ⟨⊗⟩ (F₂ ∙ G₂) 77 | RHS⊑LHS I I⊨STA₁A₂ I⊨RHS = (f , I⊨LHS) where 78 | 79 | f : ((BN F₁ ⊕ IN B₁ ⊕ BN G₁) ⊎ (BN F₂ ⊕ IN B₂ ⊕ BN G₂)) → Δ ⌊ I ⌋ 80 | f (inj₁ (inode v)) = ind I (bnode (inode (inj₁ v))) 81 | f (inj₁ (bnode y)) = ind I (bnode (bnode (inj₁ y))) 82 | f (inj₁ (enode w)) = ind I (bnode (enode (inj₁ w))) 83 | f (inj₂ (inode v)) = ind I (bnode (inode (inj₂ v))) 84 | f (inj₂ (bnode y)) = ind I (bnode (bnode (inj₂ y))) 85 | f (inj₂ (enode w)) = ind I (bnode (enode (inj₂ w))) 86 | 87 | I₁ˡ⊨F₁ : left * up * bnodes I f ⊨a impl F₁ 88 | I₁ˡ⊨F₁ = ⊨a-resp-≡³ 89 | (up * left * I) (on-bnode f (ind I) ∘ up ∘ left) refl (impl F₁) 90 | (tensor-up F₁ F₂ (left * I) (compose-left (F₁ ⟨⊗⟩ F₂) (G₁ ⟨⊗⟩ G₂) I I⊨RHS)) 91 | 92 | I₁ʳ⊨G₁ : right * up * bnodes I f ⊨a impl G₁ 93 | I₁ʳ⊨G₁ = ⊨a-resp-≡³ 94 | (up * right * I) (on-bnode f (ind I) ∘ up ∘ right) refl (impl G₁) 95 | (tensor-up G₁ G₂ (right * I) (compose-right (F₁ ⟨⊗⟩ F₂) (G₁ ⟨⊗⟩ G₂) I I⊨RHS)) 96 | 97 | I₂ˡ⊨F₂ : left * down * bnodes I f ⊨a impl F₂ 98 | I₂ˡ⊨F₂ = ⊨a-resp-≡³ 99 | (down * left * I) (on-bnode f (ind I) ∘ down ∘ left) refl (impl F₂) 100 | (tensor-down F₁ F₂ (left * I) (compose-left (F₁ ⟨⊗⟩ F₂) (G₁ ⟨⊗⟩ G₂) I I⊨RHS)) 101 | 102 | I₂ʳ⊨G₂ : right * down * bnodes I f ⊨a impl G₂ 103 | I₂ʳ⊨G₂ = ⊨a-resp-≡³ 104 | (down * right * I) (on-bnode f (ind I) ∘ down ∘ right) refl (impl G₂) 105 | (tensor-down G₁ G₂ (right * I) (compose-right (F₁ ⟨⊗⟩ F₂) (G₁ ⟨⊗⟩ G₂) I I⊨RHS)) 106 | 107 | I⊨LHS : bnodes I f ⊨a impl ((F₁ ∙ G₁) ⟨⊗⟩ (F₂ ∙ G₂)) 108 | I⊨LHS = tensor-resp-⊨a (F₁ ∙ G₁) (F₂ ∙ G₂) (bnodes I f) 109 | (compose-resp-⊨a F₁ G₁ (up * bnodes I f) I₁ˡ⊨F₁ I₁ʳ⊨G₁) 110 | (compose-resp-⊨a F₂ G₂ (down * bnodes I f) I₂ˡ⊨F₂ I₂ʳ⊨G₂) 111 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Category/Properties/Tensor/Isomorphisms.agda: -------------------------------------------------------------------------------- 1 | open import Relation.Binary.PropositionalEquality using ( refl ) 2 | open import Web.Semantic.DL.Category.Object using ( Object ) 3 | open import Web.Semantic.DL.Category.Morphism using ( _≣_ ) 4 | open import Web.Semantic.DL.Category.Composition using ( _∙_ ) 5 | open import Web.Semantic.DL.Category.Tensor using ( _⊗_ ; _⟨⊗⟩_ ) 6 | open import Web.Semantic.DL.Category.Unit using ( I ) 7 | open import Web.Semantic.DL.Category.Wiring using 8 | ( identity ; symm ; assoc ; assoc⁻¹ ; unit₁ ; unit₁⁻¹ ; unit₂ ; unit₂⁻¹ ) 9 | open import Web.Semantic.DL.Category.Properties.Wiring using 10 | ( rewriting ; rewrite-compose ; rewrite-tensor 11 | ; rewrite-identity ; rewrite-symm ; rewrite-assoc ; rewrite-assoc⁻¹ 12 | ; rewrite-unit₁ ; rewrite-unit₁⁻¹ ; rewrite-unit₂ ; rewrite-unit₂⁻¹ ) 13 | open import Web.Semantic.DL.Signature using ( Signature ) 14 | open import Web.Semantic.DL.TBox using ( TBox ) 15 | open import Web.Semantic.Util using ( _≡⊎≡_ ) 16 | 17 | module Web.Semantic.DL.Category.Properties.Tensor.Isomorphisms 18 | {Σ : Signature} {S T : TBox Σ} where 19 | 20 | symm-iso : ∀ (A₁ A₂ : Object S T) → 21 | (symm A₁ A₂ ∙ symm A₂ A₁ ≣ identity (A₁ ⊗ A₂)) 22 | symm-iso A₁ A₂ = 23 | rewriting 24 | (rewrite-compose (rewrite-symm A₁ A₂) (rewrite-symm A₂ A₁)) 25 | ((λ x → refl) ≡⊎≡ (λ x → refl)) 26 | (rewrite-identity (A₁ ⊗ A₂)) 27 | 28 | assoc-iso : ∀ (A₁ A₂ A₃ : Object S T) → 29 | (assoc A₁ A₂ A₃ ∙ assoc⁻¹ A₁ A₂ A₃ ≣ identity ((A₁ ⊗ A₂) ⊗ A₃)) 30 | assoc-iso A₁ A₂ A₃ = rewriting 31 | (rewrite-compose (rewrite-assoc A₁ A₂ A₃) (rewrite-assoc⁻¹ A₁ A₂ A₃)) 32 | (((λ x → refl) ≡⊎≡ (λ x → refl)) ≡⊎≡ (λ x → refl)) 33 | (rewrite-identity ((A₁ ⊗ A₂) ⊗ A₃)) 34 | 35 | assoc⁻¹-iso : ∀ (A₁ A₂ A₃ : Object S T) → 36 | (assoc⁻¹ A₁ A₂ A₃ ∙ assoc A₁ A₂ A₃ ≣ identity (A₁ ⊗ (A₂ ⊗ A₃))) 37 | assoc⁻¹-iso A₁ A₂ A₃ = rewriting 38 | (rewrite-compose (rewrite-assoc⁻¹ A₁ A₂ A₃) (rewrite-assoc A₁ A₂ A₃)) 39 | ((λ x → refl) ≡⊎≡ ((λ x → refl) ≡⊎≡ (λ x → refl))) 40 | (rewrite-identity (A₁ ⊗ (A₂ ⊗ A₃))) 41 | 42 | unit₁-iso : ∀ (A : Object S T) → 43 | (unit₁ A ∙ unit₁⁻¹ A ≣ identity (I ⊗ A)) 44 | unit₁-iso A = rewriting 45 | (rewrite-compose (rewrite-unit₁ A) (rewrite-unit₁⁻¹ A)) 46 | ((λ ()) ≡⊎≡ (λ x → refl)) 47 | (rewrite-identity (I ⊗ A)) 48 | 49 | unit₁⁻¹-iso : ∀ (A : Object S T) → 50 | (unit₁⁻¹ A ∙ unit₁ A ≣ identity A) 51 | unit₁⁻¹-iso A = rewriting 52 | (rewrite-compose (rewrite-unit₁⁻¹ A) (rewrite-unit₁ A)) 53 | (λ x → refl) 54 | (rewrite-identity A) 55 | 56 | unit₂-iso : ∀ (A : Object S T) → 57 | (unit₂ A ∙ unit₂⁻¹ A ≣ identity (A ⊗ I)) 58 | unit₂-iso A = rewriting 59 | (rewrite-compose (rewrite-unit₂ A) (rewrite-unit₂⁻¹ A)) 60 | ((λ x → refl) ≡⊎≡ (λ ())) 61 | (rewrite-identity (A ⊗ I)) 62 | 63 | unit₂⁻¹-iso : ∀ (A : Object S T) → 64 | (unit₂⁻¹ A ∙ unit₂ A ≣ identity A) 65 | unit₂⁻¹-iso A = rewriting 66 | (rewrite-compose (rewrite-unit₂⁻¹ A) (rewrite-unit₂ A)) 67 | (λ x → refl) 68 | (rewrite-identity A) 69 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Category/Properties/Tensor/Lemmas.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( _,_ ) 2 | open import Data.Sum using ( _⊎_ ) 3 | open import Web.Semantic.DL.ABox.Interp using ( Interp ; _*_ ) 4 | open import Web.Semantic.DL.ABox.Model using ( _⊨a_ ; _⊨b_ ; *-resp-⟨ABox⟩ ) 5 | open import Web.Semantic.DL.Category.Object using ( Object ; _,_ ; IN ) 6 | open import Web.Semantic.DL.Category.Morphism using ( _⇒_ ; _,_ ; BN ; impl ) 7 | open import Web.Semantic.DL.Category.Tensor using 8 | ( _⟨⊗⟩_ ; ⊨a-intro-⟨&⟩ ; ⊨b-intro-⟨&⟩ ) 9 | open import Web.Semantic.DL.Signature using ( Signature ) 10 | open import Web.Semantic.DL.TBox using ( TBox ) 11 | open import Web.Semantic.Util using ( _⊕_⊕_ ; up ; down ) 12 | 13 | module Web.Semantic.DL.Category.Properties.Tensor.Lemmas 14 | {Σ : Signature} {S T : TBox Σ} where 15 | 16 | tensor-up : ∀ {A₁ A₂ B₁ B₂ : Object S T} (F₁ : A₁ ⇒ B₁) (F₂ : A₂ ⇒ B₂) I → 17 | (I ⊨a impl (F₁ ⟨⊗⟩ F₂)) → (up * I ⊨a impl F₁) 18 | tensor-up {X₁ , X₁∈Fin , A₁} {X₂ , X₂∈Fin , A₂} {Y₁ , Y₁∈Fin , B₁} {Y₂ , Y₂∈Fin , B₂} 19 | (V₁ , F₁ , F₁✓) (V₂ , F₂ , F₂✓) I (I⊨F₁ , I⊨F₂) = *-resp-⟨ABox⟩ up I F₁ I⊨F₁ 20 | 21 | tensor-down : ∀ {A₁ A₂ B₁ B₂ : Object S T} (F₁ : A₁ ⇒ B₁) (F₂ : A₂ ⇒ B₂) I → 22 | (I ⊨a impl (F₁ ⟨⊗⟩ F₂)) → (down * I ⊨a impl F₂) 23 | tensor-down {X₁ , X₁∈Fin , A₁} {X₂ , X₂∈Fin , A₂} {Y₁ , Y₁∈Fin , B₁} {Y₂ , Y₂∈Fin , B₂} 24 | (V₁ , F₁ , F₁✓) (V₂ , F₂ , F₂✓) I (I⊨F₁ , I⊨F₂) = *-resp-⟨ABox⟩ down I F₂ I⊨F₂ 25 | 26 | tensor-resp-⊨a : ∀ {A₁ A₂ B₁ B₂ : Object S T} 27 | (F₁ : A₁ ⇒ B₁) (F₂ : A₂ ⇒ B₂) → 28 | (I : Interp Σ ((IN A₁ ⊎ IN A₂) ⊕ (BN F₁ ⊎ BN F₂) ⊕ (IN B₁ ⊎ IN B₂))) → 29 | (up * I ⊨a impl F₁) → (down * I ⊨a impl F₂) → (I ⊨a impl (F₁ ⟨⊗⟩ F₂)) 30 | tensor-resp-⊨a {X₁ , X₁∈Fin , A₁} {X₂ , X₂∈Fin , A₂} {Y₁ , Y₁∈Fin , B₁} {Y₂ , Y₂∈Fin , B₂} 31 | (V₁ , F₁ , F₁✓) (V₂ , F₂ , F₂✓) I I₁⊨F₁ I₂⊨F₂ = 32 | ⊨a-intro-⟨&⟩ I F₁ F₂ I₁⊨F₁ I₂⊨F₂ 33 | 34 | tensor-resp-⊨b : ∀ {A₁ A₂ B₁ B₂ : Object S T} {V₁ V₂ : Set} 35 | (F₁ : A₁ ⇒ B₁) (F₂ : A₂ ⇒ B₂) → 36 | (I : Interp Σ ((IN A₁ ⊎ IN A₂) ⊕ (V₁ ⊎ V₂) ⊕ (IN B₁ ⊎ IN B₂))) → 37 | (up * I ⊨b impl F₁) → (down * I ⊨b impl F₂) → (I ⊨b impl (F₁ ⟨⊗⟩ F₂)) 38 | tensor-resp-⊨b {X₁ , X₁∈Fin , A₁} {X₂ , X₂∈Fin , A₂} {Y₁ , Y₁∈Fin , B₁} {Y₂ , Y₂∈Fin , B₂} 39 | (V₁ , F₁ , F₁✓) (V₂ , F₂ , F₂✓) I I⊨F₁ I⊨F₂ = 40 | ⊨b-intro-⟨&⟩ I F₁ F₂ I⊨F₁ I⊨F₂ 41 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Category/Properties/Tensor/RespectsEquiv.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( _,_ ) 2 | open import Data.Sum using ( inj₁ ; inj₂ ) 3 | open import Web.Semantic.DL.ABox.Interp using ( _*_ ) 4 | open import Web.Semantic.DL.ABox.Model using ( *-resp-⟨ABox⟩ ) 5 | open import Web.Semantic.DL.Category.Object using ( Object ; iface ) 6 | open import Web.Semantic.DL.Category.Morphism using ( _⇒_ ; _≣_ ; _⊑_ ; _,_ ) 7 | open import Web.Semantic.DL.Category.Tensor using ( _⟨⊗⟩_ ) 8 | open import Web.Semantic.DL.Category.Properties.Tensor.Lemmas using 9 | ( tensor-up ; tensor-down ; tensor-resp-⊨b ) 10 | open import Web.Semantic.DL.Signature using ( Signature ) 11 | open import Web.Semantic.DL.TBox using ( TBox ) 12 | open import Web.Semantic.Util using ( inode ; up ; down ) 13 | 14 | module Web.Semantic.DL.Category.Properties.Tensor.RespectsEquiv 15 | {Σ : Signature} {S T : TBox Σ} where 16 | 17 | tensor-resp-⊑ : ∀ {A₁ A₂ B₁ B₂ : Object S T} 18 | (F₁ G₁ : A₁ ⇒ B₁) (F₂ G₂ : A₂ ⇒ B₂) → 19 | (F₁ ⊑ G₁) → (F₂ ⊑ G₂) → (F₁ ⟨⊗⟩ F₂ ⊑ G₁ ⟨⊗⟩ G₂) 20 | tensor-resp-⊑ {A₁} {A₂} F₁ G₁ F₂ G₂ F₁⊑G₁ F₂⊑G₂ 21 | I (I⊨ST , (I⊨A₁ , I⊨A₂)) I⊨F₁F₂ = 22 | tensor-resp-⊨b G₁ G₂ I 23 | (F₁⊑G₁ (up * I) 24 | (I⊨ST , *-resp-⟨ABox⟩ inj₁ (inode * I) (iface A₁) I⊨A₁) 25 | (tensor-up F₁ F₂ I I⊨F₁F₂)) 26 | (F₂⊑G₂ (down * I) 27 | (I⊨ST , *-resp-⟨ABox⟩ inj₂ (inode * I) (iface A₂) I⊨A₂) 28 | (tensor-down F₁ F₂ I I⊨F₁F₂)) 29 | 30 | tensor-resp-≣ : ∀ {A₁ A₂ B₁ B₂ : Object S T} 31 | {F₁ G₁ : A₁ ⇒ B₁} {F₂ G₂ : A₂ ⇒ B₂} → 32 | (F₁ ≣ G₁) → (F₂ ≣ G₂) → (F₁ ⟨⊗⟩ F₂ ≣ G₁ ⟨⊗⟩ G₂) 33 | tensor-resp-≣ {A₁} {A₂} {B₁} {B₂} {F₁} {G₁} {F₂} {G₂} 34 | (F₁⊑G₁ , G₁⊑F₁) (F₂⊑G₂ , G₂⊑F₂) = 35 | ( tensor-resp-⊑ F₁ G₁ F₂ G₂ F₁⊑G₁ F₂⊑G₂ 36 | , tensor-resp-⊑ G₁ F₁ G₂ F₂ G₁⊑F₁ G₂⊑F₂ ) 37 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Category/Properties/Tensor/RespectsWiring.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( proj₁ ; proj₂ ) 2 | open import Data.Sum using ( inj₁ ; inj₂ ) 3 | open import Relation.Binary.PropositionalEquality using 4 | ( _≡_ ; refl ; sym ; cong ) 5 | open import Relation.Unary using ( _⊆_ ) 6 | open import Web.Semantic.DL.ABox using ( ABox ; ⟨ABox⟩ ; Assertions ) 7 | open import Web.Semantic.DL.ABox.Interp using ( ⌊_⌋ ; ind ) 8 | open import Web.Semantic.DL.ABox.Model using ( _⊨a_ ; bnodes ; _,_ ) 9 | open import Web.Semantic.DL.Category.Object using ( Object ; IN ; fin ; iface ) 10 | open import Web.Semantic.DL.Category.Morphism using ( impl ; _≣_ ; _⊑_ ; _,_ ) 11 | open import Web.Semantic.DL.Category.Tensor using ( _⊗_ ; _⟨⊗⟩_ ) 12 | open import Web.Semantic.DL.Category.Properties.Tensor.Lemmas using 13 | ( tensor-up ; tensor-down ; tensor-resp-⊨a ) 14 | open import Web.Semantic.DL.Category.Wiring using 15 | ( wiring ; wires-≈ ; wires-≈⁻¹ ; identity ; id✓ ) 16 | open import Web.Semantic.DL.Signature using ( Signature ) 17 | open import Web.Semantic.DL.TBox using ( TBox ) 18 | open import Web.Semantic.DL.TBox.Interp using 19 | ( Δ ; _⊨_≈_ ; ≈-refl ; ≈-refl′ ; ≈-trans ) 20 | open import Web.Semantic.Util using 21 | ( id ; _∘_ ; False ; elim ; inj⁻¹ ; _⊕_⊕_ ; inode ; bnode ; enode ) 22 | 23 | module Web.Semantic.DL.Category.Properties.Tensor.RespectsWiring 24 | {Σ : Signature} {S T : TBox Σ} where 25 | 26 | tensor-resp-wiring : ∀ (A₁ A₂ B₁ B₂ : Object S T) → 27 | (f₁ : IN B₁ → IN A₁) → 28 | (f₁✓ : Assertions (⟨ABox⟩ f₁ (iface B₁)) ⊆ Assertions (iface A₁)) → 29 | (f₂ : IN B₂ → IN A₂) → 30 | (f₂✓ : Assertions (⟨ABox⟩ f₂ (iface B₂)) ⊆ Assertions (iface A₂)) → 31 | (g : IN (B₁ ⊗ B₂) → IN (A₁ ⊗ A₂)) → 32 | (g✓ : Assertions (⟨ABox⟩ g (iface (B₁ ⊗ B₂))) ⊆ Assertions (iface (A₁ ⊗ A₂))) → 33 | (∀ x → inj₁ (f₁ x) ≡ g (inj₁ x)) → 34 | (∀ x → inj₂ (f₂ x) ≡ g (inj₂ x)) → 35 | ((wiring A₁ B₁ f₁ f₁✓ ⟨⊗⟩ wiring A₂ B₂ f₂ f₂✓) ≣ 36 | (wiring (A₁ ⊗ A₂) (B₁ ⊗ B₂) g g✓)) 37 | tensor-resp-wiring A₁ A₂ B₁ B₂ f₁ f₁✓ f₂ f₂✓ g g✓ f₁≡g₁ f₂≡g₂ = 38 | (LHS⊑RHS , RHS⊑LHS) where 39 | 40 | LHS⊑RHS : 41 | wiring A₁ B₁ f₁ f₁✓ ⟨⊗⟩ wiring A₂ B₂ f₂ f₂✓ ⊑ 42 | wiring (A₁ ⊗ A₂) (B₁ ⊗ B₂) g g✓ 43 | LHS⊑RHS I I⊨STA I⊨F = (elim , I⊨RHS) where 44 | 45 | lemma : ∀ x → ⌊ I ⌋ ⊨ ind I (inode (g x)) ≈ ind I (enode x) 46 | lemma (inj₁ x) = ≈-trans ⌊ I ⌋ 47 | (≈-refl′ ⌊ I ⌋ (cong (ind I ∘ inode) (sym (f₁≡g₁ x)))) 48 | (wires-≈ f₁ (proj₂ (fin B₁) x) 49 | (tensor-up (wiring A₁ B₁ f₁ f₁✓) (wiring A₂ B₂ f₂ f₂✓) I I⊨F)) 50 | lemma (inj₂ x) = ≈-trans ⌊ I ⌋ 51 | (≈-refl′ ⌊ I ⌋ (cong (ind I ∘ inode) (sym (f₂≡g₂ x)))) 52 | (wires-≈ f₂ (proj₂ (fin B₂) x) 53 | (tensor-down (wiring A₁ B₁ f₁ f₁✓) (wiring A₂ B₂ f₂ f₂✓) I I⊨F)) 54 | 55 | I⊨RHS : bnodes I elim ⊨a impl (wiring (A₁ ⊗ A₂) (B₁ ⊗ B₂) g g✓) 56 | I⊨RHS = wires-≈⁻¹ g lemma (proj₁ (fin (B₁ ⊗ B₂))) 57 | 58 | RHS⊑LHS : 59 | wiring (A₁ ⊗ A₂) (B₁ ⊗ B₂) g g✓ ⊑ 60 | wiring A₁ B₁ f₁ f₁✓ ⟨⊗⟩ wiring A₂ B₂ f₂ f₂✓ 61 | RHS⊑LHS I I⊨STA I⊨F = (elim ∘ inj⁻¹ , I⊨LHS) where 62 | 63 | lemma₁ : ∀ x → ⌊ I ⌋ ⊨ ind I (inode (inj₁ (f₁ x))) ≈ ind I (enode (inj₁ x)) 64 | lemma₁ x = ≈-trans ⌊ I ⌋ 65 | (≈-refl′ ⌊ I ⌋ (cong (ind I ∘ inode) (f₁≡g₁ x))) 66 | (wires-≈ g (proj₂ (fin (B₁ ⊗ B₂)) (inj₁ x)) I⊨F) 67 | 68 | lemma₂ : ∀ x → ⌊ I ⌋ ⊨ ind I (inode (inj₂ (f₂ x))) ≈ ind I (enode (inj₂ x)) 69 | lemma₂ x = ≈-trans ⌊ I ⌋ 70 | (≈-refl′ ⌊ I ⌋ (cong (ind I ∘ inode) (f₂≡g₂ x))) 71 | (wires-≈ g (proj₂ (fin (B₁ ⊗ B₂)) (inj₂ x)) I⊨F) 72 | 73 | I⊨LHS : bnodes I (elim ∘ inj⁻¹) ⊨a impl (wiring A₁ B₁ f₁ f₁✓ ⟨⊗⟩ wiring A₂ B₂ f₂ f₂✓) 74 | I⊨LHS = tensor-resp-⊨a (wiring A₁ B₁ f₁ f₁✓) (wiring A₂ B₂ f₂ f₂✓) (bnodes I (elim ∘ inj⁻¹)) 75 | (wires-≈⁻¹ f₁ lemma₁ (proj₁ (fin B₁))) (wires-≈⁻¹ f₂ lemma₂ (proj₁ (fin B₂))) 76 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Category/Properties/Tensor/SymmNatural.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( proj₁ ; proj₂ ) 2 | open import Data.Sum using ( _⊎_ ; inj₁ ; inj₂ ) 3 | open import Relation.Binary.PropositionalEquality using ( refl ) 4 | open import Web.Semantic.DL.ABox.Interp using ( ⌊_⌋ ; ind ; _*_ ) 5 | open import Web.Semantic.DL.ABox.Interp.Morphism using ( _,_ ) 6 | open import Web.Semantic.DL.ABox.Model using 7 | ( _⊨a_ ; on-bnode ; bnodes ; _,_ ; ⊨a-resp-≲ ) 8 | open import Web.Semantic.DL.Category.Composition using ( _∙_ ) 9 | open import Web.Semantic.DL.Category.Properties.Composition.Lemmas using 10 | ( compose-left ; compose-right ; compose-resp-⊨a ) 11 | open import Web.Semantic.DL.Category.Properties.Tensor.Lemmas using 12 | ( tensor-up ; tensor-down ; tensor-resp-⊨a ) 13 | open import Web.Semantic.DL.Category.Object using ( Object ; IN ; fin ) 14 | open import Web.Semantic.DL.Category.Morphism using 15 | ( _⇒_ ; BN ; impl ; _⊑_ ; _≣_ ; _,_ ) 16 | open import Web.Semantic.DL.Category.Tensor using ( _⊗_ ; _⟨⊗⟩_ ) 17 | open import Web.Semantic.DL.Category.Unit using ( I ) 18 | open import Web.Semantic.DL.Category.Wiring using 19 | ( wires-≈ ; wires-≈⁻¹ ; symm ) 20 | open import Web.Semantic.DL.Signature using ( Signature ) 21 | open import Web.Semantic.DL.TBox using ( TBox ) 22 | open import Web.Semantic.DL.TBox.Interp using ( Δ ; _⊨_≈_ ; ≈-refl ; ≈-sym ) 23 | open import Web.Semantic.DL.TBox.Interp.Morphism using ( ≲-refl ) 24 | open import Web.Semantic.Util using 25 | ( _∘_ ; False ; ⊎-swap 26 | ; _⊕_⊕_ ; inode ; bnode ; enode ; left ; right ; up ; down ) 27 | 28 | module Web.Semantic.DL.Category.Properties.Tensor.SymmNatural 29 | {Σ : Signature} {S T : TBox Σ} where 30 | 31 | symm-natural : ∀ {A₁ A₂ B₁ B₂ : Object S T} (F₁ : A₁ ⇒ B₁) (F₂ : A₂ ⇒ B₂) → 32 | ((F₁ ⟨⊗⟩ F₂) ∙ symm B₁ B₂ ≣ symm A₁ A₂ ∙ (F₂ ⟨⊗⟩ F₁)) 33 | symm-natural {A₁} {A₂} {B₁} {B₂} F₁ F₂ = (LHS⊑RHS , RHS⊑LHS) where 34 | 35 | LHS⊑RHS : (F₁ ⟨⊗⟩ F₂) ∙ symm B₁ B₂ ⊑ symm A₁ A₂ ∙ (F₂ ⟨⊗⟩ F₁) 36 | LHS⊑RHS J J⊨STA J⊨LHS = (f , J⊨RHS) where 37 | 38 | f : False ⊕ (IN A₂ ⊎ IN A₁) ⊕ (BN F₂ ⊎ BN F₁) → Δ ⌊ J ⌋ 39 | f (inode ()) 40 | f (bnode (inj₁ x)) = ind J (inode (inj₂ x)) 41 | f (bnode (inj₂ x)) = ind J (inode (inj₁ x)) 42 | f (enode (inj₁ v)) = ind J (bnode (inode (inj₂ v))) 43 | f (enode (inj₂ v)) = ind J (bnode (inode (inj₁ v))) 44 | 45 | lemma₀ : ∀ x → 46 | ⌊ J ⌋ ⊨ ind J (inode (⊎-swap x)) ≈ f (bnode x) 47 | lemma₀ (inj₁ x) = ≈-refl ⌊ J ⌋ 48 | lemma₀ (inj₂ y) = ≈-refl ⌊ J ⌋ 49 | 50 | lemma₁ : ∀ x → 51 | ⌊ J ⌋ ⊨ ind J (left (up x)) ≈ on-bnode f (ind J) (right (down x)) 52 | lemma₁ (inode x) = ≈-refl ⌊ J ⌋ 53 | lemma₁ (bnode v) = ≈-refl ⌊ J ⌋ 54 | lemma₁ (enode y) = wires-≈ ⊎-swap (proj₂ (fin (B₂ ⊗ B₁)) (inj₂ y)) 55 | (compose-right (F₁ ⟨⊗⟩ F₂) (symm B₁ B₂) J J⊨LHS) 56 | 57 | lemma₂ : ∀ x → 58 | ⌊ J ⌋ ⊨ ind J (left (down x)) ≈ on-bnode f (ind J) (right (up x)) 59 | lemma₂ (inode x) = ≈-refl ⌊ J ⌋ 60 | lemma₂ (bnode v) = ≈-refl ⌊ J ⌋ 61 | lemma₂ (enode y) = wires-≈ ⊎-swap (proj₂ (fin (B₂ ⊗ B₁)) (inj₁ y)) 62 | (compose-right (F₁ ⟨⊗⟩ F₂) (symm B₁ B₂) J J⊨LHS) 63 | 64 | J⊨RHS : bnodes J f ⊨a impl (symm A₁ A₂ ∙ (F₂ ⟨⊗⟩ F₁)) 65 | J⊨RHS = compose-resp-⊨a (symm A₁ A₂) (F₂ ⟨⊗⟩ F₁) (bnodes J f) 66 | (wires-≈⁻¹ ⊎-swap lemma₀ (proj₁ (fin (A₂ ⊗ A₁)))) 67 | (tensor-resp-⊨a F₂ F₁ (right * bnodes J f) 68 | (⊨a-resp-≲ (≲-refl ⌊ J ⌋ , lemma₂) (impl F₂) 69 | (tensor-down F₁ F₂ (left * J) 70 | (compose-left (F₁ ⟨⊗⟩ F₂) (symm B₁ B₂) J J⊨LHS))) 71 | (⊨a-resp-≲ (≲-refl ⌊ J ⌋ , lemma₁) (impl F₁) 72 | (tensor-up F₁ F₂ (left * J) 73 | (compose-left (F₁ ⟨⊗⟩ F₂) (symm B₁ B₂) J J⊨LHS)))) 74 | 75 | RHS⊑LHS : symm A₁ A₂ ∙ (F₂ ⟨⊗⟩ F₁) ⊑ (F₁ ⟨⊗⟩ F₂) ∙ symm B₁ B₂ 76 | RHS⊑LHS J J⊨STA J⊨RHS = (f , J⊨LHS) where 77 | 78 | f : ((BN F₁ ⊎ BN F₂) ⊕ (IN B₁ ⊎ IN B₂) ⊕ False) → Δ ⌊ J ⌋ 79 | f (inode (inj₁ v)) = ind J (bnode (enode (inj₂ v))) 80 | f (inode (inj₂ v)) = ind J (bnode (enode (inj₁ v))) 81 | f (bnode (inj₁ y)) = ind J (enode (inj₂ y)) 82 | f (bnode (inj₂ y)) = ind J (enode (inj₁ y)) 83 | f (enode ()) 84 | 85 | lemma₀ : ∀ x → ⌊ J ⌋ ⊨ f (bnode (⊎-swap x)) ≈ ind J (enode x) 86 | lemma₀ (inj₁ y) = ≈-refl ⌊ J ⌋ 87 | lemma₀ (inj₂ y) = ≈-refl ⌊ J ⌋ 88 | 89 | lemma₁ : ∀ x → 90 | ⌊ J ⌋ ⊨ ind J (right (down x)) ≈ on-bnode f (ind J) (left (up x)) 91 | lemma₁ (inode x) = ≈-sym ⌊ J ⌋ 92 | (wires-≈ ⊎-swap (proj₂ (fin (A₂ ⊗ A₁)) (inj₂ x)) 93 | (compose-left (symm A₁ A₂) (F₂ ⟨⊗⟩ F₁) J J⊨RHS)) 94 | lemma₁ (bnode v) = ≈-refl ⌊ J ⌋ 95 | lemma₁ (enode y) = ≈-refl ⌊ J ⌋ 96 | 97 | lemma₂ : ∀ x → 98 | ⌊ J ⌋ ⊨ ind J (right (up x)) ≈ on-bnode f (ind J) (left (down x)) 99 | lemma₂ (inode x) = ≈-sym ⌊ J ⌋ 100 | (wires-≈ ⊎-swap (proj₂ (fin (A₂ ⊗ A₁)) (inj₁ x)) 101 | (compose-left (symm A₁ A₂) (F₂ ⟨⊗⟩ F₁) J J⊨RHS)) 102 | lemma₂ (bnode v) = ≈-refl ⌊ J ⌋ 103 | lemma₂ (enode y) = ≈-refl ⌊ J ⌋ 104 | 105 | J⊨LHS : bnodes J f ⊨a impl ((F₁ ⟨⊗⟩ F₂) ∙ symm B₁ B₂) 106 | J⊨LHS = compose-resp-⊨a (F₁ ⟨⊗⟩ F₂) (symm B₁ B₂) (bnodes J f) 107 | (tensor-resp-⊨a F₁ F₂ (left * bnodes J f) 108 | (⊨a-resp-≲ (≲-refl ⌊ J ⌋ , lemma₁) (impl F₁) 109 | (tensor-down F₂ F₁ (right * J) 110 | (compose-right (symm A₁ A₂) (F₂ ⟨⊗⟩ F₁) J J⊨RHS))) 111 | (⊨a-resp-≲ (≲-refl ⌊ J ⌋ , lemma₂) (impl F₂) 112 | (tensor-up F₂ F₁ (right * J) 113 | (compose-right (symm A₁ A₂) (F₂ ⟨⊗⟩ F₁) J J⊨RHS)))) 114 | (wires-≈⁻¹ ⊎-swap lemma₀ (proj₁ (fin (B₂ ⊗ B₁)))) -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Category/Properties/Tensor/UnitNatural.agda: -------------------------------------------------------------------------------- 1 | open import Data.List using ( [] ) 2 | open import Data.Product using ( proj₁ ; proj₂ ) 3 | open import Data.Sum using ( _⊎_ ; inj₁ ; inj₂ ) 4 | open import Relation.Binary.PropositionalEquality using ( refl ) 5 | open import Web.Semantic.DL.ABox.Interp using ( ⌊_⌋ ; ind ; _*_ ) 6 | open import Web.Semantic.DL.ABox.Interp.Morphism using ( _,_ ) 7 | open import Web.Semantic.DL.ABox.Model using 8 | ( _⊨a_ ; on-bnode ; bnodes ; _,_ ; ⊨a-resp-≲ ) 9 | open import Web.Semantic.DL.Category.Composition using ( _∙_ ) 10 | open import Web.Semantic.DL.Category.Properties.Composition.Assoc using 11 | ( compose-assoc ) 12 | open import Web.Semantic.DL.Category.Properties.Composition.Lemmas using 13 | ( compose-left ; compose-right ; compose-resp-⊨a ) 14 | open import Web.Semantic.DL.Category.Properties.Composition.RespectsEquiv using 15 | ( compose-resp-≣ ) 16 | open import Web.Semantic.DL.Category.Properties.Equivalence using 17 | ( ≣-refl ; ≣-sym ; ≣-trans ) 18 | open import Web.Semantic.DL.Category.Properties.Tensor.Coherence using 19 | ( symm-unit ) 20 | open import Web.Semantic.DL.Category.Properties.Tensor.Lemmas using 21 | ( tensor-up ; tensor-down ; tensor-resp-⊨a ) 22 | open import Web.Semantic.DL.Category.Properties.Tensor.SymmNatural using 23 | ( symm-natural ) 24 | open import Web.Semantic.DL.Category.Object using ( Object ; IN ; fin ) 25 | open import Web.Semantic.DL.Category.Morphism using 26 | ( _⇒_ ; BN ; impl ; _⊑_ ; _≣_ ; _,_ ) 27 | open import Web.Semantic.DL.Category.Tensor using ( _⟨⊗⟩_ ) 28 | open import Web.Semantic.DL.Category.Unit using ( I ) 29 | open import Web.Semantic.DL.Category.Wiring using 30 | ( wires-≈ ; wires-≈⁻¹ ; identity ; unit₁ ; unit₂ ; symm ) 31 | open import Web.Semantic.DL.Signature using ( Signature ) 32 | open import Web.Semantic.DL.TBox using ( TBox ) 33 | open import Web.Semantic.DL.TBox.Interp using ( Δ ; _⊨_≈_ ; ≈-refl ; ≈-sym ) 34 | open import Web.Semantic.DL.TBox.Interp.Morphism using ( ≲-refl ) 35 | open import Web.Semantic.Util using 36 | ( _∘_ ; False ; _⊕_⊕_ ; inode ; bnode ; enode ; left ; right ; up ; down ) 37 | 38 | module Web.Semantic.DL.Category.Properties.Tensor.UnitNatural 39 | {Σ : Signature} {S T : TBox Σ} where 40 | 41 | unit₁-natural : ∀ {A B : Object S T} (F : A ⇒ B) → 42 | ((identity I ⟨⊗⟩ F) ∙ unit₁ B ≣ unit₁ A ∙ F) 43 | unit₁-natural {A} {B} F = (LHS⊑RHS , RHS⊑LHS) where 44 | 45 | LHS⊑RHS : ((identity I ⟨⊗⟩ F) ∙ unit₁ B ⊑ unit₁ A ∙ F) 46 | LHS⊑RHS J J⊨STA J⊨LHS = (f , J⊨RHS) where 47 | 48 | f : False ⊕ IN A ⊕ BN F → Δ ⌊ J ⌋ 49 | f (inode ()) 50 | f (bnode x) = ind J (inode (inj₂ x)) 51 | f (enode v) = ind J (bnode (inode (inj₂ v))) 52 | 53 | f✓ : ∀ x → ⌊ J ⌋ ⊨ ind J (left (down x)) ≈ on-bnode f (ind J) (right x) 54 | f✓ (inode x) = ≈-refl ⌊ J ⌋ 55 | f✓ (bnode v) = ≈-refl ⌊ J ⌋ 56 | f✓ (enode y) = wires-≈ inj₂ (proj₂ (fin B) y) 57 | (compose-right (identity I ⟨⊗⟩ F) (unit₁ B) J J⊨LHS) 58 | 59 | J⊨RHS : bnodes J f ⊨a impl (unit₁ A ∙ F) 60 | J⊨RHS = compose-resp-⊨a (unit₁ A) F (bnodes J f) 61 | (wires-≈⁻¹ inj₂ (λ x → ≈-refl ⌊ J ⌋) (proj₁ (fin A))) 62 | (⊨a-resp-≲ (≲-refl ⌊ J ⌋ , f✓) (impl F) 63 | (tensor-down (identity I) F (left * J) 64 | (compose-left (identity I ⟨⊗⟩ F) (unit₁ B) J J⊨LHS))) 65 | 66 | RHS⊑LHS : (unit₁ A ∙ F ⊑ (identity I ⟨⊗⟩ F) ∙ unit₁ B) 67 | RHS⊑LHS J J⊨STA J⊨RHS = (f , J⊨LHS) where 68 | 69 | f : ((False ⊎ BN F) ⊕ (False ⊎ IN B) ⊕ False) → Δ ⌊ J ⌋ 70 | f (inode (inj₁ ())) 71 | f (inode (inj₂ v)) = ind J (bnode (enode v)) 72 | f (bnode (inj₁ ())) 73 | f (bnode (inj₂ y)) = ind J (enode y) 74 | f (enode ()) 75 | 76 | f✓ : ∀ x → ⌊ J ⌋ ⊨ ind J (right x) ≈ on-bnode f (ind J) (left (down x)) 77 | f✓ (inode x) = ≈-sym ⌊ J ⌋ (wires-≈ inj₂ (proj₂ (fin A) x) (compose-left (unit₁ A) F J J⊨RHS)) 78 | f✓ (bnode v) = ≈-refl ⌊ J ⌋ 79 | f✓ (enode y) = ≈-refl ⌊ J ⌋ 80 | 81 | J⊨LHS : bnodes J f ⊨a impl ((identity I ⟨⊗⟩ F) ∙ unit₁ B) 82 | J⊨LHS = compose-resp-⊨a (identity I ⟨⊗⟩ F) (unit₁ B) (bnodes J f) 83 | (tensor-resp-⊨a (identity I) F (left * bnodes J f) 84 | (wires-≈⁻¹ {I = up * left * bnodes J f} (λ ()) (λ ()) []) 85 | (⊨a-resp-≲ (≲-refl ⌊ J ⌋ , f✓) (impl F) 86 | (compose-right (unit₁ A) F J J⊨RHS))) 87 | (wires-≈⁻¹ inj₂ (λ x → ≈-refl ⌊ J ⌋) (proj₁ (fin B))) 88 | 89 | unit₂-natural : ∀ {A B : Object S T} (F : A ⇒ B) → 90 | ((F ⟨⊗⟩ identity I) ∙ unit₂ B ≣ unit₂ A ∙ F) 91 | unit₂-natural {A} {B} F = 92 | ≣-trans 93 | (compose-resp-≣ (≣-refl (F ⟨⊗⟩ identity I)) (≣-sym (symm-unit B))) 94 | (≣-trans 95 | (≣-sym (compose-assoc (F ⟨⊗⟩ identity I) (symm B I) (unit₁ B))) 96 | (≣-trans 97 | (compose-resp-≣ (symm-natural F (identity I)) (≣-refl (unit₁ B))) 98 | (≣-trans 99 | (compose-assoc (symm A I) (identity I ⟨⊗⟩ F) (unit₁ B)) 100 | (≣-trans 101 | (compose-resp-≣ (≣-refl (symm A I)) (unit₁-natural F)) 102 | (≣-trans 103 | (≣-sym (compose-assoc (symm A I) (unit₁ A) F)) 104 | (compose-resp-≣ (symm-unit A) (≣-refl F))))))) 105 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Category/Properties/Wiring.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( _,_ ; proj₁ ; proj₂ ) 2 | open import Data.Sum using ( inj₁ ; inj₂ ) 3 | open import Relation.Unary using ( _⊆_ ) 4 | open import Relation.Binary.PropositionalEquality using ( _≡_ ; refl ; sym ; cong ) 5 | open import Web.Semantic.DL.ABox using ( ε ; _,_ ; _∼_ ; _∈₁_ ; _∈₂_ ; ⟨ABox⟩ ; Assertions ) 6 | open import Web.Semantic.DL.ABox.Interp using ( ⌊_⌋ ; ind ) 7 | open import Web.Semantic.DL.ABox.Model using ( ⊨a-impl-⊨b ; ⟨ABox⟩-Assertions ) 8 | open import Web.Semantic.DL.Category.Composition using ( _∙_ ) 9 | open import Web.Semantic.DL.Category.Object using ( Object ; _,_ ; IN ; fin ; iface ) 10 | open import Web.Semantic.DL.Category.Morphism using ( _⇒_ ; _⊑_ ; _≣_ ; _,_ ) 11 | open import Web.Semantic.DL.Category.Properties.Composition.RespectsEquiv using ( compose-resp-≣ ) 12 | open import Web.Semantic.DL.Category.Properties.Composition.RespectsWiring using ( compose-resp-wiring ) 13 | open import Web.Semantic.DL.Category.Properties.Equivalence using ( ≣-refl ; ≣-sym ; ≣-trans ) 14 | open import Web.Semantic.DL.Category.Properties.Tensor.RespectsEquiv using ( tensor-resp-≣ ) 15 | open import Web.Semantic.DL.Category.Properties.Tensor.RespectsWiring using ( tensor-resp-wiring ) 16 | open import Web.Semantic.DL.Category.Tensor using ( _⊗_ ; _⟨⊗⟩_ ) 17 | open import Web.Semantic.DL.Category.Unit using ( I ) 18 | open import Web.Semantic.DL.Category.Wiring using 19 | ( wiring ; wires ; wires-≈ ; wires-≈⁻¹ 20 | ; identity ; symm ; assoc ; assoc⁻¹ ; unit₁ ; unit₁⁻¹ ; unit₂ ; unit₂⁻¹ 21 | ; id✓ ; ⊎-swap✓ ; ⊎-assoc✓ ; ⊎-assoc⁻¹✓ ; inj₁✓ ; inj₂✓ ; ⊎-unit₁✓ ; ⊎-unit₂✓ ) 22 | open import Web.Semantic.DL.Signature using ( Signature ) 23 | open import Web.Semantic.DL.TBox using ( TBox ) 24 | open import Web.Semantic.DL.TBox.Interp using ( ≈-refl′ ; ≈-trans ) 25 | open import Web.Semantic.Util using 26 | ( id ; _∘_ ; inode ; _⟨⊎⟩_ ; ⊎-swap ; ⊎-assoc ; ⊎-assoc⁻¹ ; ⊎-unit₁ ; ⊎-unit₂ ) 27 | 28 | module Web.Semantic.DL.Category.Properties.Wiring 29 | {Σ : Signature} {S T : TBox Σ} where 30 | 31 | wiring-⊑-refl : ∀ (A B : Object S T) (f g : IN B → IN A) 32 | (f✓ : Assertions (⟨ABox⟩ f (iface B)) ⊆ Assertions (iface A)) 33 | (g✓ : Assertions (⟨ABox⟩ g (iface B)) ⊆ Assertions (iface A)) → 34 | (∀ x → f x ≡ g x) → (wiring A B f f✓ ⊑ wiring A B g g✓) 35 | wiring-⊑-refl A B f g f✓ g✓ f≡g I I⊨STA I⊨F = 36 | ⊨a-impl-⊨b I (wires g (proj₁ (fin B))) 37 | (wires-≈⁻¹ g 38 | (λ x → ≈-trans ⌊ I ⌋ 39 | (≈-refl′ ⌊ I ⌋ (cong (ind I ∘ inode) (sym (f≡g x)))) 40 | (wires-≈ f (proj₂ (fin B) x) I⊨F)) 41 | (proj₁ (fin B))) 42 | 43 | wiring-≣-refl : ∀ (A B : Object S T) (f g : IN B → IN A) 44 | (f✓ : Assertions (⟨ABox⟩ f (iface B)) ⊆ Assertions (iface A)) 45 | (g✓ : Assertions (⟨ABox⟩ g (iface B)) ⊆ Assertions (iface A)) → 46 | (∀ x → f x ≡ g x) → (wiring A B f f✓ ≣ wiring A B g g✓) 47 | wiring-≣-refl A B f g f✓ g✓ f≡g = 48 | ( wiring-⊑-refl A B f g f✓ g✓ f≡g 49 | , wiring-⊑-refl A B g f g✓ f✓ (λ x → sym (f≡g x)) ) 50 | 51 | ∘✓ : (A B C : Object S T) → 52 | (f : IN B → IN A) → (g : IN C → IN B) → 53 | (Assertions (⟨ABox⟩ f (iface B)) ⊆ Assertions (iface A)) → 54 | (Assertions (⟨ABox⟩ g (iface C)) ⊆ Assertions (iface B)) → 55 | (Assertions (⟨ABox⟩ (f ∘ g) (iface C)) ⊆ Assertions (iface A)) 56 | ∘✓ A B (X , X∈Fin , ε) f g f✓ g✓ () 57 | ∘✓ A B (X , X∈Fin , C₁ , C₂) f g f✓ g✓ (inj₁ a∈C₁) = ∘✓ A B (X , X∈Fin , C₁) f g f✓ (g✓ ∘ inj₁) a∈C₁ 58 | ∘✓ A B (X , X∈Fin , C₁ , C₂) f g f✓ g✓ (inj₂ a∈C₂) = ∘✓ A B (X , X∈Fin , C₂) f g f✓ (g✓ ∘ inj₂) a∈C₂ 59 | ∘✓ A B (X , X∈Fin , x ∼ y) f g f✓ g✓ refl = f✓ (⟨ABox⟩-Assertions f (iface B) (g✓ refl)) 60 | ∘✓ A B (X , X∈Fin , x ∈₁ c) f g f✓ g✓ refl = f✓ (⟨ABox⟩-Assertions f (iface B) (g✓ refl)) 61 | ∘✓ A B (X , X∈Fin , (x , y) ∈₂ r) f g f✓ g✓ refl = f✓ (⟨ABox⟩-Assertions f (iface B) (g✓ refl)) 62 | 63 | ⟨⊎⟩✓ : (A₁ A₂ B₁ B₂ : Object S T) → 64 | (f₁ : IN B₁ → IN A₁) → (f₂ : IN B₂ → IN A₂) → 65 | (Assertions (⟨ABox⟩ f₁ (iface B₁)) ⊆ Assertions (iface A₁)) → 66 | (Assertions (⟨ABox⟩ f₂ (iface B₂)) ⊆ Assertions (iface A₂)) → 67 | (Assertions (⟨ABox⟩ (f₁ ⟨⊎⟩ f₂) (iface (B₁ ⊗ B₂))) ⊆ Assertions (iface (A₁ ⊗ A₂))) 68 | ⟨⊎⟩✓ A₁ A₂ (X₁ , X₁∈Fin , ε) B₂ f₁ f₂ f₁✓ f₂✓ (inj₁ ()) 69 | ⟨⊎⟩✓ A₁ A₂ (X₁ , X₁∈Fin , B₁₁ , B₁₂) B₂ f₁ f₂ f₁✓ f₂✓ (inj₁ (inj₁ b∈B₁₁)) = 70 | ⟨⊎⟩✓ A₁ A₂ (X₁ , X₁∈Fin , B₁₁) B₂ f₁ f₂ (f₁✓ ∘ inj₁) f₂✓ (inj₁ b∈B₁₁) 71 | ⟨⊎⟩✓ A₁ A₂ (X₁ , X₁∈Fin , B₁₁ , B₁₂) B₂ f₁ f₂ f₁✓ f₂✓ (inj₁ (inj₂ b∈B₁₂)) = 72 | ⟨⊎⟩✓ A₁ A₂ (X₁ , X₁∈Fin , B₁₂) B₂ f₁ f₂ (f₁✓ ∘ inj₂) f₂✓ (inj₁ b∈B₁₂) 73 | ⟨⊎⟩✓ A₁ A₂ (X₁ , X₁∈Fin , x ∼ y) B₂ f₁ f₂ f₁✓ f₂✓ (inj₁ refl) = 74 | inj₁ (⟨ABox⟩-Assertions inj₁ (iface A₁) (f₁✓ refl)) 75 | ⟨⊎⟩✓ A₁ A₂ (X₁ , X₁∈Fin , x ∈₁ c) B₂ f₁ f₂ f₁✓ f₂✓ (inj₁ refl) = 76 | inj₁ (⟨ABox⟩-Assertions inj₁ (iface A₁) (f₁✓ refl)) 77 | ⟨⊎⟩✓ A₁ A₂ (X₁ , X₁∈Fin , (x , y) ∈₂ r) B₂ f₁ f₂ f₁✓ f₂✓ (inj₁ refl) = 78 | inj₁ (⟨ABox⟩-Assertions inj₁ (iface A₁) (f₁✓ refl)) 79 | ⟨⊎⟩✓ A₁ A₂ B₁ (X₂ , X₂∈Fin , ε) f₁ f₂ f₁✓ f₂✓ (inj₂ ()) 80 | ⟨⊎⟩✓ A₁ A₂ B₁ (X₂ , X₂∈Fin , B₂₁ , B₂₂) f₁ f₂ f₁✓ f₂✓ (inj₂ (inj₁ b∈B₂₁)) = 81 | ⟨⊎⟩✓ A₁ A₂ B₁ (X₂ , X₂∈Fin , B₂₁) f₁ f₂ f₁✓ (f₂✓ ∘ inj₁) (inj₂ b∈B₂₁) 82 | ⟨⊎⟩✓ A₁ A₂ B₁ (X₂ , X₂∈Fin , B₂₁ , B₂₂) f₁ f₂ f₁✓ f₂✓ (inj₂ (inj₂ b∈B₂₂)) = 83 | ⟨⊎⟩✓ A₁ A₂ B₁ (X₂ , X₂∈Fin , B₂₂) f₁ f₂ f₁✓ (f₂✓ ∘ inj₂) (inj₂ b∈B₂₂) 84 | ⟨⊎⟩✓ A₁ A₂ B₁ (X₂ , X₂∈Fin , x ∼ y) f₁ f₂ f₁✓ f₂✓ (inj₂ refl) = 85 | inj₂ (⟨ABox⟩-Assertions inj₂ (iface A₂) (f₂✓ refl)) 86 | ⟨⊎⟩✓ A₁ A₂ B₁ (X₂ , X₂∈Fin , x ∈₁ c) f₁ f₂ f₁✓ f₂✓ (inj₂ refl) = 87 | inj₂ (⟨ABox⟩-Assertions inj₂ (iface A₂) (f₂✓ refl)) 88 | ⟨⊎⟩✓ A₁ A₂ B₁ (X₂ , X₂∈Fin , (x , y) ∈₂ r) f₁ f₂ f₁✓ f₂✓ (inj₂ refl) = 89 | inj₂ (⟨ABox⟩-Assertions inj₂ (iface A₂) (f₂✓ refl)) 90 | 91 | data Rewrite : (A B : Object S T) (F : A ⇒ B) (f : IN B → IN A) → 92 | (Assertions (⟨ABox⟩ f (iface B)) ⊆ Assertions (iface A)) → Set₁ where 93 | rewrite-wiring : ∀ A B f (f✓ : Assertions (⟨ABox⟩ f (iface B)) ⊆ Assertions (iface A)) → 94 | Rewrite A B (wiring A B f f✓) f f✓ 95 | rewrite-compose : ∀ {A B C F G f g} 96 | {f✓ : Assertions (⟨ABox⟩ f (iface B)) ⊆ Assertions (iface A)} 97 | {g✓ : Assertions (⟨ABox⟩ g (iface C)) ⊆ Assertions (iface B)} → 98 | Rewrite A B F f f✓ → Rewrite B C G g g✓ → Rewrite A C (F ∙ G) (f ∘ g) (∘✓ A B C f g f✓ g✓) 99 | rewrite-tensor : ∀ {A₁ A₂ B₁ B₂ F₁ F₂ f₁ f₂} 100 | {f₁✓ : Assertions (⟨ABox⟩ f₁ (iface B₁)) ⊆ Assertions (iface A₁)} 101 | {f₂✓ : Assertions (⟨ABox⟩ f₂ (iface B₂)) ⊆ Assertions (iface A₂)} → 102 | Rewrite A₁ B₁ F₁ f₁ f₁✓ → Rewrite A₂ B₂ F₂ f₂ f₂✓ → 103 | Rewrite (A₁ ⊗ A₂) (B₁ ⊗ B₂) (F₁ ⟨⊗⟩ F₂) (f₁ ⟨⊎⟩ f₂) (⟨⊎⟩✓ A₁ A₂ B₁ B₂ f₁ f₂ f₁✓ f₂✓) 104 | 105 | rewrite-identity : ∀ A → Rewrite A A (identity A) id (id✓ A) 106 | rewrite-identity A = rewrite-wiring A A id (id✓ A) 107 | 108 | rewrite-symm : ∀ A B → Rewrite (A ⊗ B) (B ⊗ A) (symm A B) ⊎-swap (⊎-swap✓ A B) 109 | rewrite-symm A B = rewrite-wiring (A ⊗ B) (B ⊗ A) ⊎-swap (⊎-swap✓ A B) 110 | 111 | rewrite-assoc : ∀ A B C → Rewrite ((A ⊗ B) ⊗ C) (A ⊗ (B ⊗ C)) (assoc A B C) ⊎-assoc⁻¹ (⊎-assoc⁻¹✓ A B C) 112 | rewrite-assoc A B C = rewrite-wiring ((A ⊗ B) ⊗ C) (A ⊗ (B ⊗ C)) ⊎-assoc⁻¹ (⊎-assoc⁻¹✓ A B C) 113 | 114 | rewrite-assoc⁻¹ : ∀ A B C → Rewrite (A ⊗ (B ⊗ C)) ((A ⊗ B) ⊗ C) (assoc⁻¹ A B C) ⊎-assoc (⊎-assoc✓ A B C) 115 | rewrite-assoc⁻¹ A B C = rewrite-wiring (A ⊗ (B ⊗ C)) ((A ⊗ B) ⊗ C) ⊎-assoc (⊎-assoc✓ A B C) 116 | 117 | rewrite-unit₁ : ∀ A → Rewrite (I ⊗ A) A (unit₁ A) inj₂ (inj₂✓ A) 118 | rewrite-unit₁ A = rewrite-wiring (I ⊗ A) A inj₂ (inj₂✓ A) 119 | 120 | rewrite-unit₁⁻¹ : ∀ A → Rewrite A (I ⊗ A) (unit₁⁻¹ A) ⊎-unit₁ (⊎-unit₁✓ A) 121 | rewrite-unit₁⁻¹ A = rewrite-wiring A (I ⊗ A) ⊎-unit₁ (⊎-unit₁✓ A) 122 | 123 | rewrite-unit₂ : ∀ A → Rewrite (A ⊗ I) A (unit₂ A) inj₁ (inj₁✓ A) 124 | rewrite-unit₂ A = rewrite-wiring (A ⊗ I) A inj₁ (inj₁✓ A) 125 | 126 | rewrite-unit₂⁻¹ : ∀ A → Rewrite A (A ⊗ I) (unit₂⁻¹ A) ⊎-unit₂ (⊎-unit₂✓ A) 127 | rewrite-unit₂⁻¹ A = rewrite-wiring A (A ⊗ I) ⊎-unit₂ (⊎-unit₂✓ A) 128 | 129 | rewriting : ∀ {A B F G f g} 130 | {f✓ : Assertions (⟨ABox⟩ f (iface B)) ⊆ Assertions (iface A)} 131 | {g✓ : Assertions (⟨ABox⟩ g (iface B)) ⊆ Assertions (iface A)} → 132 | (Rewrite A B F f f✓) → (∀ x → f x ≡ g x) → (Rewrite A B G g g✓) → 133 | (F ≣ G) 134 | rewriting {A} {B} {F} {G} {f} {g} {f✓} {g✓} F⇓f f≡g G⇓g = 135 | ≣-trans (lemma F⇓f) 136 | (≣-trans (wiring-≣-refl A B f g f✓ g✓ f≡g) 137 | (≣-sym (lemma G⇓g))) where 138 | 139 | lemma : ∀ {A B F f} 140 | {f✓ : Assertions (⟨ABox⟩ f (iface B)) ⊆ Assertions (iface A)} → 141 | (Rewrite A B F f f✓) → 142 | (F ≣ wiring A B f f✓) 143 | lemma (rewrite-wiring A B f f✓) = ≣-refl (wiring A B f f✓) 144 | lemma (rewrite-compose {A} {B} {C} {F} {G} {f} {g} {f✓} {g✓} F⇓f G⇓g) = 145 | ≣-trans (compose-resp-≣ (lemma F⇓f) (lemma G⇓g)) 146 | (compose-resp-wiring A B C f f✓ g g✓ (f ∘ g) 147 | (∘✓ A B C f g f✓ g✓) (λ x → refl)) 148 | lemma (rewrite-tensor {A₁} {A₂} {B₁} {B₂} {F₁} {F₂} {f₁} {f₂} {f₁✓} {f₂✓} F₁⇓f₁ F₂⇓f₂) = 149 | ≣-trans (tensor-resp-≣ (lemma F₁⇓f₁) (lemma F₂⇓f₂)) 150 | (tensor-resp-wiring A₁ A₂ B₁ B₂ f₁ f₁✓ f₂ f₂✓ (f₁ ⟨⊎⟩ f₂) 151 | (⟨⊎⟩✓ A₁ A₂ B₁ B₂ f₁ f₂ f₁✓ f₂✓) (λ x → refl) (λ x → refl)) 152 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Category/Tensor.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( _,_ ; proj₁ ; proj₂ ) 2 | open import Data.Sum using ( _⊎_ ; inj₁ ; inj₂ ) 3 | open import Relation.Binary.PropositionalEquality using ( refl ) 4 | open import Relation.Unary using ( _∈_ ) 5 | open import Web.Semantic.DL.ABox using ( ABox ; _,_ ; ⟨ABox⟩ ) 6 | open import Web.Semantic.DL.ABox.Interp using ( Interp ; _,_ ; ⌊_⌋ ; ind ; _*_ ) 7 | open import Web.Semantic.DL.ABox.Interp.Morphism using 8 | ( _≲_ ; _,_ ; ≲⌊_⌋ ; ≲-resp-ind ; _**_ ; ≡³-impl-≈ ; ≡³-impl-≲ ) 9 | open import Web.Semantic.DL.ABox.Model using 10 | ( _⊨a_ ; _⊨b_ ; bnodes ; on-bnode ; _,_ 11 | ; ⟨ABox⟩-resp-⊨ ; ⊨a-resp-≲ ; *-resp-⟨ABox⟩ ) 12 | open import Web.Semantic.DL.Category.Morphism using 13 | ( _⇒_ ; _⇒_w/_ ; _,_ ; BN ; impl ; impl✓ ) 14 | open import Web.Semantic.DL.Category.Object using 15 | ( Object ; _,_ ; IN ; fin ; iface ) 16 | open import Web.Semantic.DL.Integrity using 17 | ( _⊕_⊨_ ; Unique ; Mediated ; Mediator ; Initial 18 | ; _,_ ; extension ; ext-init ; ext-⊨ ; ext✓ ; init-≲ ) 19 | open import Web.Semantic.DL.KB using ( _,_ ) 20 | open import Web.Semantic.DL.KB.Model using ( _⊨_ ) 21 | open import Web.Semantic.DL.Signature using ( Signature ) 22 | open import Web.Semantic.DL.TBox using ( TBox ; _,_ ) 23 | open import Web.Semantic.DL.TBox.Interp using 24 | ( Δ ; _⊨_≈_ ; ≈-refl ; ≈-sym ; ≈-trans ) 25 | open import Web.Semantic.DL.TBox.Model using ( _⊨t_ ) 26 | open import Web.Semantic.DL.TBox.Interp.Morphism using 27 | ( ≲-image ; ≲-resp-≈ ; ≲-trans ) renaming ( _≲_ to _≲′_ ) 28 | open import Web.Semantic.Util using 29 | ( _∘_ ; Finite ; ⊎-resp-Fin ; _⊕_⊕_ ; inode ; bnode ; enode ; up ; down ; vmerge ) 30 | 31 | module Web.Semantic.DL.Category.Tensor {Σ : Signature} where 32 | 33 | _&_ : ∀ {X₁ X₂} → ABox Σ X₁ → ABox Σ X₂ → ABox Σ (X₁ ⊎ X₂) 34 | A₁ & A₂ = (⟨ABox⟩ inj₁ A₁ , ⟨ABox⟩ inj₂ A₂) 35 | 36 | _⊗_ : ∀ {S T : TBox Σ} → Object S T → Object S T → Object S T 37 | A₁ ⊗ A₂ = 38 | ( (IN A₁ ⊎ IN A₂) 39 | , ⊎-resp-Fin (fin A₁) (fin A₂) 40 | , iface A₁ & iface A₂ ) 41 | 42 | _⟨&⟩_ : ∀ {V₁ V₂ X₁ X₂ Y₁ Y₂} → 43 | ABox Σ (X₁ ⊕ V₁ ⊕ Y₁) → ABox Σ (X₂ ⊕ V₂ ⊕ Y₂) → 44 | ABox Σ ((X₁ ⊎ X₂) ⊕ (V₁ ⊎ V₂) ⊕ (Y₁ ⊎ Y₂)) 45 | F₁ ⟨&⟩ F₂ = (⟨ABox⟩ up F₁ , ⟨ABox⟩ down F₂) 46 | 47 | ⊨a-intro-⟨&⟩ : ∀ {V₁ V₂ X₁ X₂ Y₁ Y₂} → 48 | (I : Interp Σ ((X₁ ⊎ X₂) ⊕ (V₁ ⊎ V₂) ⊕ (Y₁ ⊎ Y₂))) → 49 | (F₁ : ABox Σ (X₁ ⊕ V₁ ⊕ Y₁)) → (F₂ : ABox Σ (X₂ ⊕ V₂ ⊕ Y₂)) → 50 | (up * I ⊨a F₁) → (down * I ⊨a F₂) → (I ⊨a F₁ ⟨&⟩ F₂) 51 | ⊨a-intro-⟨&⟩ (I , i) F₁ F₂ I₁⊨F₁ I₂⊨F₂ = 52 | ( ⟨ABox⟩-resp-⊨ up (λ x → ≈-refl I) F₁ I₁⊨F₁ 53 | , ⟨ABox⟩-resp-⊨ down (λ x → ≈-refl I) F₂ I₂⊨F₂ ) 54 | 55 | ⊨b-intro-⟨&⟩ : ∀ {V₁ V₂ W₁ W₂ X₁ X₂ Y₁ Y₂} → 56 | (I : Interp Σ ((X₁ ⊎ X₂) ⊕ (W₁ ⊎ W₂) ⊕ (Y₁ ⊎ Y₂))) → 57 | (F₁ : ABox Σ (X₁ ⊕ V₁ ⊕ Y₁)) → (F₂ : ABox Σ (X₂ ⊕ V₂ ⊕ Y₂)) → 58 | (up * I ⊨b F₁) → (down * I ⊨b F₂) → (I ⊨b F₁ ⟨&⟩ F₂) 59 | ⊨b-intro-⟨&⟩ {V₁} {V₂} I F₁ F₂ (f₁ , I₁⊨F₁) (f₂ , I₂⊨F₂) = 60 | (f , I⊨F₁F₂) where 61 | 62 | f : (V₁ ⊎ V₂) → Δ ⌊ I ⌋ 63 | f (inj₁ v) = f₁ v 64 | f (inj₂ v) = f₂ v 65 | 66 | I⊨F₁F₂ : bnodes I f ⊨a F₁ ⟨&⟩ F₂ 67 | I⊨F₁F₂ = 68 | ( ⟨ABox⟩-resp-⊨ up 69 | (≡³-impl-≈ ⌊ I ⌋ (on-bnode f₁ (ind I ∘ up)) 70 | (on-bnode f (ind I) ∘ up) refl) 71 | F₁ I₁⊨F₁ 72 | , ⟨ABox⟩-resp-⊨ down 73 | (≡³-impl-≈ ⌊ I ⌋ (on-bnode f₂ (ind I ∘ down)) 74 | (on-bnode f (ind I) ∘ down) refl) 75 | F₂ I₂⊨F₂ ) 76 | 77 | go₂ : ∀ {V₁ X₁ X₂ Y₁} → (I : Interp Σ (X₁ ⊎ X₂)) → 78 | (J₁ : Interp Σ (X₁ ⊕ V₁ ⊕ Y₁)) → (⌊ I ⌋ ≲′ ⌊ J₁ ⌋) → 79 | Interp Σ X₂ 80 | go₂ I J₁ I≲J = (⌊ J₁ ⌋ , ≲-image I≲J ∘ ind I ∘ inj₂) 81 | 82 | go₂-≳ : ∀ {V₁ X₁ X₂ Y₁} → (I : Interp Σ (X₁ ⊎ X₂)) → 83 | (J₁ : Interp Σ (X₁ ⊕ V₁ ⊕ Y₁)) → (I≲J : ⌊ I ⌋ ≲′ ⌊ J₁ ⌋) → 84 | (inj₂ * I ≲ go₂ I J₁ I≲J) 85 | go₂-≳ I (J , j₁) I≲J = (I≲J , λ x → ≈-refl J) 86 | 87 | go₂-≲ : ∀ {V₁ V₂ X₁ X₂ Y₁ Y₂} → (I : Interp Σ (X₁ ⊎ X₂)) → 88 | (J₁ : Interp Σ (X₁ ⊕ V₁ ⊕ Y₁)) → 89 | (L : Interp Σ ((X₁ ⊎ X₂) ⊕ (V₁ ⊎ V₂) ⊕ (Y₁ ⊎ Y₂))) → 90 | (I₁≲J₁ : inj₁ * I ≲ inode * J₁) → (I≲L : I ≲ inode * L) → 91 | (Mediated (inj₁ * I) J₁ (up * L) I₁≲J₁ (inj₁ ** I≲L)) → 92 | (go₂ I J₁ ≲⌊ I₁≲J₁ ⌋ ≲ inode * (down * L)) 93 | go₂-≲ (I , i) (J , j₁) (L , l) (I≲J , i₁≲j₁) (I≲L , i≲l) 94 | ((J≲L , j₁≲l₁) , I≲L≋I≲J≲L , J₁≲L₁-uniq) = 95 | (J≲L , λ x → ≈-trans L (≈-sym L (I≲L≋I≲J≲L (i (inj₂ x)))) (i≲l (inj₂ x))) 96 | 97 | par : ∀ {V₁ V₂ X₁ X₂ Y₁ Y₂} → (J₁ : Interp Σ (X₁ ⊕ V₁ ⊕ Y₁)) → 98 | (K₂ : Interp Σ (X₂ ⊕ V₂ ⊕ Y₂)) → (⌊ J₁ ⌋ ≲′ ⌊ K₂ ⌋) → 99 | Interp Σ ((X₁ ⊎ X₂) ⊕ (V₁ ⊎ V₂) ⊕ (Y₁ ⊎ Y₂)) 100 | par J₁ K₂ J≲K = 101 | (⌊ K₂ ⌋ , vmerge (≲-image J≲K ∘ ind J₁) (ind K₂)) 102 | 103 | par-inj₁ : ∀ {V₁ V₂ X₁ X₂ Y₁ Y₂} → (J₁ : Interp Σ (X₁ ⊕ V₁ ⊕ Y₁)) → 104 | (K₂ : Interp Σ (X₂ ⊕ V₂ ⊕ Y₂)) → (J≲K : ⌊ J₁ ⌋ ≲′ ⌊ K₂ ⌋) → 105 | (J₁ ≲ up * par J₁ K₂ J≲K) 106 | par-inj₁ (J , j₁) (K , k₂) J≲K = 107 | ( J≲K 108 | , ≡³-impl-≈ K 109 | (≲-image J≲K ∘ j₁) 110 | (vmerge (≲-image J≲K ∘ j₁) k₂ ∘ up) 111 | refl ) 112 | 113 | par-inj₂ : ∀ {V₁ V₂ X₁ X₂ Y₁ Y₂} → (J₁ : Interp Σ (X₁ ⊕ V₁ ⊕ Y₁)) → 114 | (K₂ : Interp Σ (X₂ ⊕ V₂ ⊕ Y₂)) → (J≲K : ⌊ J₁ ⌋ ≲′ ⌊ K₂ ⌋) → 115 | (K₂ ≲ down * par J₁ K₂ J≲K) 116 | par-inj₂ (J , j₁) (K , k₂) J≲K = 117 | ≡³-impl-≲ (K , k₂) (vmerge (≲-image J≲K ∘ j₁) k₂ ∘ down) refl 118 | 119 | par-exp : ∀ {V₁ V₂ X₁ X₂ Y₁ Y₂} → (J₁ : Interp Σ (X₁ ⊕ V₁ ⊕ Y₁)) → 120 | (K₂ : Interp Σ (X₂ ⊕ V₂ ⊕ Y₂)) → (J≲K : ⌊ J₁ ⌋ ≲′ ⌊ K₂ ⌋) → 121 | (T : TBox Σ) → (B₁ : ABox Σ Y₁) → (B₂ : ABox Σ Y₂) → 122 | (enode * J₁ ⊨ T , B₁) → (enode * K₂ ⊨ T , B₂) → 123 | (enode * (par J₁ K₂ J≲K) ⊨ T , (B₁ & B₂)) 124 | par-exp J₁ K₂ J≲K T B₁ B₂ (J⊨T , J₁⊨B₁) (K⊨T , K₂⊨B₂) = 125 | ( K⊨T 126 | , ⟨ABox⟩-resp-⊨ inj₁ (λ y → ≈-refl ⌊ K₂ ⌋) B₁ 127 | (⊨a-resp-≲ (enode ** par-inj₁ J₁ K₂ J≲K) B₁ J₁⊨B₁) 128 | , ⟨ABox⟩-resp-⊨ inj₂ (λ y → ≈-refl ⌊ K₂ ⌋) B₂ K₂⊨B₂) 129 | 130 | par-≳ : ∀ {V₁ V₂ X₁ X₂ Y₁ Y₂} → (I : Interp Σ (X₁ ⊎ X₂)) → 131 | (J₁ : Interp Σ (X₁ ⊕ V₁ ⊕ Y₁)) → (K₂ : Interp Σ (X₂ ⊕ V₂ ⊕ Y₂)) → 132 | (I₁≲J₁ : inj₁ * I ≲ inode * J₁) → 133 | (J₂≲K₂ : go₂ I J₁ ≲⌊ I₁≲J₁ ⌋ ≲ inode * K₂) → 134 | (I ≲ inode * par J₁ K₂ ≲⌊ J₂≲K₂ ⌋) 135 | par-≳ I J₁ K₂ (I≲J , i₁≲j₁) (J≲K , j₂≲k₂) = 136 | ( ≲-trans I≲J J≲K , lemma ) where 137 | 138 | lemma : ∀ x → ⌊ K₂ ⌋ ⊨ 139 | ≲-image J≲K (≲-image I≲J (ind I x)) ≈ 140 | vmerge (≲-image J≲K ∘ ind J₁) (ind K₂) (inode x) 141 | lemma (inj₁ x) = ≲-resp-≈ J≲K (i₁≲j₁ x) 142 | lemma (inj₂ x) = j₂≲k₂ x 143 | 144 | par-impl : ∀ {V₁ V₂ X₁ X₂ Y₁ Y₂} → (J₁ : Interp Σ (X₁ ⊕ V₁ ⊕ Y₁)) → 145 | (K₂ : Interp Σ (X₂ ⊕ V₂ ⊕ Y₂)) → (J≲K : ⌊ J₁ ⌋ ≲′ ⌊ K₂ ⌋) → 146 | (F₁ : ABox Σ (X₁ ⊕ V₁ ⊕ Y₁)) → (F₂ : ABox Σ (X₂ ⊕ V₂ ⊕ Y₂)) → 147 | (J₁ ⊨a F₁) → (K₂ ⊨a F₂) → 148 | (par J₁ K₂ J≲K ⊨a F₁ ⟨&⟩ F₂) 149 | par-impl J₁ K₂ J≲K F₁ F₂ J₁⊨F₁ K₂⊨F₂ = 150 | ( ⟨ABox⟩-resp-⊨ up (λ x → ≈-refl ⌊ K₂ ⌋) F₁ 151 | (⊨a-resp-≲ (par-inj₁ J₁ K₂ J≲K) F₁ J₁⊨F₁) 152 | , ⟨ABox⟩-resp-⊨ down (λ x → ≈-refl ⌊ K₂ ⌋) F₂ 153 | (⊨a-resp-≲ (par-inj₂ J₁ K₂ J≲K) F₂ K₂⊨F₂) ) 154 | 155 | par-mediated : ∀ {V₁ V₂ X₁ X₂ Y₁ Y₂} → (I : Interp Σ (X₁ ⊎ X₂)) → 156 | (J₁ : Interp Σ (X₁ ⊕ V₁ ⊕ Y₁)) → (K₂ : Interp Σ (X₂ ⊕ V₂ ⊕ Y₂)) → 157 | (L : Interp Σ ((X₁ ⊎ X₂) ⊕ (V₁ ⊎ V₂) ⊕ (Y₁ ⊎ Y₂))) → 158 | (I₁≲J₁ : inj₁ * I ≲ inode * J₁) → (I≲L : I ≲ inode * L) → 159 | (J₂≲K₂ : go₂ I J₁ ≲⌊ I₁≲J₁ ⌋ ≲ inode * K₂) → 160 | (m : Mediated (inj₁ * I) J₁ (up * L) I₁≲J₁ (inj₁ ** I≲L)) → 161 | (Mediated (go₂ I J₁ ≲⌊ I₁≲J₁ ⌋) K₂ (down * L) J₂≲K₂ (go₂-≲ I J₁ L I₁≲J₁ I≲L m)) → 162 | (Mediated I (par J₁ K₂ ≲⌊ J₂≲K₂ ⌋) L (par-≳ I J₁ K₂ I₁≲J₁ J₂≲K₂) I≲L) 163 | par-mediated (I , i) (J , j₁) (K , k₂) (L , l) (I≲J , i₁≲j₁) (I≲L , i≲l) (J≲K , j₂≲k₂) 164 | ((J≲L , j₁≲l₁) , I≲L≋I≲J≲L , J₁≲L₁-uniq) ((K≲L , k₂≲l₂) , J≲L≋J≲K≲L , K₂≲L₂-uniq) = 165 | ((K≲L , k≲l) , I≲L≋I≲K≲L , K≲L-uniq) where 166 | 167 | k = vmerge (≲-image J≲K ∘ j₁) k₂ 168 | 169 | I≲K : I ≲′ K 170 | I≲K = ≲-trans I≲J J≲K 171 | 172 | i≲k : ∀ x → K ⊨ ≲-image I≲K (i x) ≈ k (inode x) 173 | i≲k = ≲-resp-ind (par-≳ (I , i) (J , j₁) (K , k₂) (I≲J , i₁≲j₁) (J≲K , j₂≲k₂)) 174 | 175 | k≲l : ∀ x → L ⊨ ≲-image K≲L (k x) ≈ l x 176 | k≲l (inode (inj₁ x)) = ≈-trans L (≈-sym L (J≲L≋J≲K≲L (j₁ (inode x)))) (j₁≲l₁ (inode x)) 177 | k≲l (inode (inj₂ x)) = k₂≲l₂ (inode x) 178 | k≲l (bnode (inj₁ v)) = ≈-trans L (≈-sym L (J≲L≋J≲K≲L (j₁ (bnode v)))) (j₁≲l₁ (bnode v)) 179 | k≲l (bnode (inj₂ v)) = k₂≲l₂ (bnode v) 180 | k≲l (enode (inj₁ y)) = ≈-trans L (≈-sym L (J≲L≋J≲K≲L (j₁ (enode y)))) (j₁≲l₁ (enode y)) 181 | k≲l (enode (inj₂ y)) = k₂≲l₂ (enode y) 182 | 183 | I≲L≋I≲K≲L : ∀ x → L ⊨ ≲-image I≲L x ≈ ≲-image K≲L (≲-image I≲K x) 184 | I≲L≋I≲K≲L x = ≈-trans L (I≲L≋I≲J≲L x) (J≲L≋J≲K≲L (≲-image I≲J x)) 185 | 186 | lemma₁ : ∀ (K≲L : (K , k) ≲ (L , l)) x → (L ⊨ ≲-image ≲⌊ K≲L ⌋ (≲-image J≲K (j₁ x)) ≈ l (up x)) 187 | lemma₁ (K≲L' , k≲l) (inode x) = k≲l (inode (inj₁ x)) 188 | lemma₁ (K≲L' , k≲l) (bnode v) = k≲l (bnode (inj₁ v)) 189 | lemma₁ (K≲L' , k≲l) (enode y) = k≲l (enode (inj₁ y)) 190 | 191 | lemma₂ : ∀ (K≲L : (K , k) ≲ (L , l)) x → (L ⊨ ≲-image ≲⌊ K≲L ⌋ (k₂ x) ≈ l (down x)) 192 | lemma₂ (K≲L , k≲l) (inode x) = k≲l (inode (inj₂ x)) 193 | lemma₂ (K≲L , k≲l) (bnode v) = k≲l (bnode (inj₂ v)) 194 | lemma₂ (K≲L , k≲l) (enode y) = k≲l (enode (inj₂ y)) 195 | 196 | K≲L-uniq : Unique (I , i) (K , k) (L , l) (I≲K , i≲k) (I≲L , i≲l) 197 | K≲L-uniq (K≲₁L , k≲₁l) (K≲₂L , k≲₂l) I≲L≋I≲K≲₁L I≲L≋I≲K≲₂L = 198 | K₂≲L₂-uniq (K≲₁L , lemma₂ (K≲₁L , k≲₁l)) (K≲₂L , lemma₂ (K≲₂L , k≲₂l)) 199 | (J₁≲L₁-uniq (J≲L , j₁≲l₁) (≲-trans J≲K K≲₁L , lemma₁ (K≲₁L , k≲₁l)) I≲L≋I≲J≲L I≲L≋I≲K≲₁L) 200 | (J₁≲L₁-uniq (J≲L , j₁≲l₁) (≲-trans J≲K K≲₂L , lemma₁ (K≲₂L , k≲₂l)) I≲L≋I≲J≲L I≲L≋I≲K≲₂L) 201 | 202 | par-mediator : ∀ {V₁ V₂ X₁ X₂ Y₁ Y₂} → (I : Interp Σ (X₁ ⊎ X₂)) → 203 | (J₁ : Interp Σ (X₁ ⊕ V₁ ⊕ Y₁)) → (K₂ : Interp Σ (X₂ ⊕ V₂ ⊕ Y₂)) → 204 | (I₁≲J₁ : inj₁ * I ≲ inode * J₁) → 205 | (J₂≲K₂ : go₂ I J₁ ≲⌊ I₁≲J₁ ⌋ ≲ inode * K₂) → (S : TBox Σ) → 206 | (F₁ : ABox Σ (X₁ ⊕ V₁ ⊕ Y₁)) → (F₂ : ABox Σ (X₂ ⊕ V₂ ⊕ Y₂)) → 207 | (Mediator (inj₁ * I) J₁ I₁≲J₁ (S , F₁)) → 208 | (Mediator (go₂ I J₁ ≲⌊ I₁≲J₁ ⌋) K₂ J₂≲K₂ (S , F₂)) → 209 | (Mediator I (par J₁ K₂ ≲⌊ J₂≲K₂ ⌋) 210 | (par-≳ I J₁ K₂ I₁≲J₁ J₂≲K₂) (S , (F₁ ⟨&⟩ F₂))) 211 | par-mediator I J₁ K₂ I₁≲J₁ J₂≲K₂ S F₁ F₂ J₁-med K₂-med L I≲L (L⊨S , L⊨F₁ , L⊨F₂) = 212 | par-mediated I J₁ K₂ L I₁≲J₁ I≲L J₂≲K₂ I₁≲J₁≲L₁-med J₂≲K₂≲L₂-med where 213 | 214 | I₁≲J₁≲L₁-med : Mediated (inj₁ * I) J₁ (up * L) I₁≲J₁ (inj₁ ** I≲L) 215 | I₁≲J₁≲L₁-med = J₁-med (up * L) (inj₁ ** I≲L) (L⊨S , *-resp-⟨ABox⟩ up L F₁ L⊨F₁) 216 | 217 | J₂≲K₂≲L₂-med : Mediated (go₂ I J₁ ≲⌊ I₁≲J₁ ⌋) K₂ (down * L) J₂≲K₂ (go₂-≲ I J₁ L I₁≲J₁ I≲L I₁≲J₁≲L₁-med) 218 | J₂≲K₂≲L₂-med = K₂-med (down * L) (go₂-≲ I J₁ L I₁≲J₁ I≲L I₁≲J₁≲L₁-med) (L⊨S , *-resp-⟨ABox⟩ down L F₂ L⊨F₂) 219 | 220 | par-init : ∀ {V₁ V₂ X₁ X₂ Y₁ Y₂} → (I : Interp Σ (X₁ ⊎ X₂)) → 221 | (J₁ : Interp Σ (X₁ ⊕ V₁ ⊕ Y₁)) → (K₂ : Interp Σ (X₂ ⊕ V₂ ⊕ Y₂)) → 222 | (S : TBox Σ) → (F₁ : ABox Σ (X₁ ⊕ V₁ ⊕ Y₁)) → (F₂ : ABox Σ (X₂ ⊕ V₂ ⊕ Y₂)) → 223 | (J₁-init : J₁ ∈ Initial (inj₁ * I) (S , F₁)) → 224 | (K₂-init : K₂ ∈ Initial (go₂ I J₁ ≲⌊ init-≲ J₁-init ⌋) (S , F₂)) → 225 | (par J₁ K₂ ≲⌊ init-≲ K₂-init ⌋ ∈ 226 | Initial I (S , (F₁ ⟨&⟩ F₂))) 227 | par-init I J₁ K₂ S F₁ F₂ 228 | (I₁≲J₁ , (J⊨S , J₁⊨F₁) , J₁-med) (J₂≲K₂ , (K⊨S , K₂⊨F₂) , K₂-med) = 229 | ( par-≳ I J₁ K₂ I₁≲J₁ J₂≲K₂ 230 | , (K⊨S , par-impl J₁ K₂ ≲⌊ J₂≲K₂ ⌋ F₁ F₂ J₁⊨F₁ K₂⊨F₂) 231 | , par-mediator I J₁ K₂ I₁≲J₁ J₂≲K₂ S F₁ F₂ J₁-med K₂-med ) 232 | 233 | tensor-⊨ : ∀ {V₁ V₂ X₁ X₂ Y₁ Y₂} → (S T : TBox Σ) → 234 | (A₁ : ABox Σ X₁) → (A₂ : ABox Σ X₂) → (B₁ : ABox Σ Y₁) → (B₂ : ABox Σ Y₂) → 235 | (F₁ : ABox Σ (X₁ ⊕ V₁ ⊕ Y₁)) (F₂ : ABox Σ (X₂ ⊕ V₂ ⊕ Y₂)) → 236 | (∀ I → (I ⊨ (S , T) , A₁) → (I ⊕ S , F₁ ⊨ T , B₁)) → 237 | (∀ I → (I ⊨ (S , T) , A₂) → (I ⊕ S , F₂ ⊨ T , B₂)) → 238 | (∀ I → (I ⊨ (S , T) , (A₁ & A₂)) → 239 | (I ⊕ S , F₁ ⟨&⟩ F₂ ⊨ T , (B₁ & B₂))) 240 | tensor-⊨ {V₁} {V₂} {X₁} {X₂} {Y₁} {Y₂} 241 | S T A₁ A₂ B₁ B₂ F₁ F₂ F₁✓ F₂✓ I (I⊨ST , (I⊨A₁ , I⊨A₂)) = 242 | ( par J₁ K₂ ≲⌊ J₂≲K₂ ⌋ 243 | , par-init I J₁ K₂ S F₁ F₂ J₁-init K₂-init 244 | , par-exp J₁ K₂ ≲⌊ J₂≲K₂ ⌋ T B₁ B₂ 245 | (ext-⊨ I₁⊕SF₁⊨TB₁) (ext-⊨ J₂⊕SF₂⊨TB₂) ) where 246 | 247 | I₁ : Interp Σ X₁ 248 | I₁ = inj₁ * I 249 | 250 | I₂ : Interp Σ X₂ 251 | I₂ = inj₂ * I 252 | 253 | I₁⊨A₁ : I₁ ⊨a A₁ 254 | I₁⊨A₁ = *-resp-⟨ABox⟩ inj₁ I A₁ I⊨A₁ 255 | 256 | I₂⊨A₂ : I₂ ⊨a A₂ 257 | I₂⊨A₂ = *-resp-⟨ABox⟩ inj₂ I A₂ I⊨A₂ 258 | 259 | I₁⊕SF₁⊨TB₁ : I₁ ⊕ S , F₁ ⊨ T , B₁ 260 | I₁⊕SF₁⊨TB₁ = F₁✓ I₁ (I⊨ST , I₁⊨A₁) 261 | 262 | J₁ : Interp Σ (X₁ ⊕ V₁ ⊕ Y₁) 263 | J₁ = extension I₁⊕SF₁⊨TB₁ 264 | 265 | J₁-init : J₁ ∈ Initial I₁ (S , F₁) 266 | J₁-init = ext-init I₁⊕SF₁⊨TB₁ 267 | 268 | I₁≲J₁ : I₁ ≲ inode * J₁ 269 | I₁≲J₁ = init-≲ J₁-init 270 | 271 | J₂ : Interp Σ X₂ 272 | J₂ = go₂ I J₁ ≲⌊ I₁≲J₁ ⌋ 273 | 274 | I₂≲J₂ : I₂ ≲ J₂ 275 | I₂≲J₂ = go₂-≳ I J₁ ≲⌊ I₁≲J₁ ⌋ 276 | 277 | J⊨ST : ⌊ J₂ ⌋ ⊨t (S , T) 278 | J⊨ST = proj₁ (ext✓ I₁⊕SF₁⊨TB₁) 279 | 280 | J₂⊨A₂ : J₂ ⊨a A₂ 281 | J₂⊨A₂ = ⊨a-resp-≲ I₂≲J₂ A₂ I₂⊨A₂ 282 | 283 | J₂⊕SF₂⊨TB₂ : J₂ ⊕ S , F₂ ⊨ T , B₂ 284 | J₂⊕SF₂⊨TB₂ = F₂✓ J₂ (J⊨ST , J₂⊨A₂) 285 | 286 | K₂ : Interp Σ (X₂ ⊕ V₂ ⊕ Y₂) 287 | K₂ = extension J₂⊕SF₂⊨TB₂ 288 | 289 | K₂-init : K₂ ∈ Initial J₂ (S , F₂) 290 | K₂-init = ext-init J₂⊕SF₂⊨TB₂ 291 | 292 | J₂≲K₂ : J₂ ≲ inode * K₂ 293 | J₂≲K₂ = init-≲ K₂-init 294 | 295 | _⟨⊗⟩′_ : ∀ {S T : TBox Σ} {A₁ A₂ B₁ B₂ : Object S T} → 296 | (F₁ : A₁ ⇒ B₁) → (F₂ : A₂ ⇒ B₂) → 297 | ((A₁ ⊗ A₂) ⇒ (B₁ ⊗ B₂) w/ (BN F₁ ⊎ BN F₂)) 298 | _⟨⊗⟩′_ {S} {T} {X₁ , X₁∈Fin , A₁} {X₂ , X₂∈Fin , A₂} {Y₁ , Y₁∈Fin , B₁} {Y₂ , Y₂∈Fin , B₂} 299 | (V₁ , F₁ , F₁✓) (V₂ , F₂ , F₂✓) = 300 | (F₁ ⟨&⟩ F₂ , tensor-⊨ S T A₁ A₂ B₁ B₂ F₁ F₂ F₁✓ F₂✓) 301 | 302 | _⟨⊗⟩_ : ∀ {S T : TBox Σ} {A₁ A₂ B₁ B₂ : Object S T} → 303 | (A₁ ⇒ B₁) → (A₂ ⇒ B₂) → ((A₁ ⊗ A₂) ⇒ (B₁ ⊗ B₂)) 304 | F₁ ⟨⊗⟩ F₂ = ( (BN F₁ ⊎ BN F₂) , F₁ ⟨⊗⟩′ F₂ ) 305 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Category/Unit.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( _,_ ) 2 | open import Web.Semantic.DL.ABox using ( ε ) 3 | open import Web.Semantic.DL.Category.Object using ( Object ; _,_ ) 4 | open import Web.Semantic.DL.Signature using ( Signature ) 5 | open import Web.Semantic.DL.TBox using ( TBox ) 6 | open import Web.Semantic.Util using ( False ; False∈Fin ) 7 | 8 | module Web.Semantic.DL.Category.Unit {Σ : Signature} where 9 | 10 | I : ∀ {S T : TBox Σ} → Object S T 11 | I = (False , False∈Fin , ε) 12 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Category/Wiring.agda: -------------------------------------------------------------------------------- 1 | open import Data.List using ( List ; [] ; _∷_ ) 2 | open import Data.List.Any using ( here ; there ) 3 | open import Data.Product using ( _,_ ; proj₁ ; proj₂ ) 4 | open import Data.Sum using ( inj₁ ; inj₂ ) 5 | open import Relation.Binary.PropositionalEquality using 6 | ( refl ; sym ; trans ; subst₂ ) 7 | open import Relation.Unary using ( _∈_ ; _⊆_ ) 8 | open import Web.Semantic.DL.ABox using 9 | ( ABox ; ε ; _,_ ; _∼_ ; ⟨ABox⟩ ; Assertions 10 | ; ⟨ABox⟩-resp-id ; ⟨ABox⟩-resp-∘ ; ⟨ABox⟩-resp-∘² ) 11 | open import Web.Semantic.DL.ABox.Interp using ( Interp ; _,_ ; ⌊_⌋ ; ind ; _*_ ) 12 | open import Web.Semantic.DL.ABox.Interp.Morphism using ( _,_ ) 13 | open import Web.Semantic.DL.ABox.Model using 14 | ( _⊨a_ ; *-resp-⟨ABox⟩ ; ⊨a-resp-⊇ ) 15 | open import Web.Semantic.DL.Category.Object using 16 | ( Object ; _,_ ; IN ; fin ; iface ) 17 | open import Web.Semantic.DL.Category.Morphism using 18 | ( _⇒_ ; _,_ ; BN ; impl ) 19 | open import Web.Semantic.DL.Category.Tensor using ( _⊗_ ) 20 | open import Web.Semantic.DL.Category.Unit using ( I ) 21 | open import Web.Semantic.DL.Integrity using 22 | ( Unique ; Mediator ; Initial ; _⊕_⊨_ ; _,_ ) 23 | open import Web.Semantic.DL.KB using ( _,_ ) 24 | open import Web.Semantic.DL.KB.Model using ( _⊨_ ) 25 | open import Web.Semantic.DL.Signature using ( Signature ) 26 | open import Web.Semantic.DL.TBox using ( TBox ; _,_ ) 27 | open import Web.Semantic.DL.TBox.Interp using 28 | ( _⊨_≈_ ; ≈-refl ; ≈-sym ; ≈-trans ) 29 | open import Web.Semantic.DL.TBox.Interp.Morphism using ( ≲-refl ; ≲-image ) 30 | open import Web.Semantic.DL.TBox.Model using ( _⊨t_ ) 31 | open import Web.Semantic.Util using 32 | ( id ; _∘_ ; False ; tt ; ⊎-swap ; ⊎-assoc ; ⊎-assoc⁻¹ ; ⊎-unit₁ ; ⊎-unit₂ 33 | ; _∈ˡ_ ; _⊕_⊕_ ; inode ; bnode ; enode ) 34 | 35 | module Web.Semantic.DL.Category.Wiring {Σ : Signature} where 36 | 37 | wires : ∀ {X Y} → (Y → X) → List Y → ABox Σ (X ⊕ False ⊕ Y) 38 | wires f [] = ε 39 | wires f (y ∷ ys) = (inode (f y) ∼ enode y , wires f ys) 40 | 41 | wires-≈ : ∀ {X Y y ys I} (f : Y → X) → (y ∈ˡ ys) → (I ⊨a wires f ys) → 42 | (⌊ I ⌋ ⊨ ind I (inode (f y)) ≈ ind I (enode y)) 43 | wires-≈ f (here refl) (x≈y , xs≈ys) = x≈y 44 | wires-≈ f (there y∈ys) (x≈y , xs≈ys) = wires-≈ f y∈ys xs≈ys 45 | 46 | wires-≈⁻¹ : ∀ {X Y I} (f : Y → X) → 47 | (∀ y → ⌊ I ⌋ ⊨ ind I (inode (f y)) ≈ ind I (enode y)) → 48 | ∀ ys → (I ⊨a wires f ys) 49 | wires-≈⁻¹ f ∀y∙x≈y [] = tt 50 | wires-≈⁻¹ f ∀y∙x≈y (y ∷ ys) = (∀y∙x≈y y , wires-≈⁻¹ f ∀y∙x≈y ys) 51 | 52 | wired : ∀ {X Y} → (Y → X) → (X ⊕ False ⊕ Y) → X 53 | wired f (inode x) = x 54 | wired f (bnode ()) 55 | wired f (enode y) = f y 56 | 57 | wired-init : ∀ {X Y} (f : Y → X) ys → (∀ y → y ∈ˡ ys) → 58 | (I : Interp Σ X) → (S : TBox Σ) → (⌊ I ⌋ ⊨t S) → 59 | (wired f * I ∈ Initial I (S , wires f ys)) 60 | wired-init f ys ∀y∙y∈ys (I , i) S I⊨S = 61 | ( (≲-refl I , λ x → ≈-refl I) 62 | , (I⊨S , wires-≈⁻¹ f (λ x → ≈-refl I) ys) 63 | , mediator ) where 64 | 65 | mediator : Mediator (I , i) (I , i ∘ wired f) 66 | (≲-refl I , λ x → ≈-refl I) (S , wires f ys) 67 | mediator (K , k) (I≲K , i≲k) (K⊨S , K⊨F) = 68 | ((I≲K , j≲k) , (λ x → ≈-refl K) , uniq ) where 69 | 70 | j≲k : ∀ x → K ⊨ ≲-image I≲K (i (wired f x)) ≈ k x 71 | j≲k (inode x) = i≲k x 72 | j≲k (bnode ()) 73 | j≲k (enode y) = ≈-trans K (i≲k (f y)) (wires-≈ f (∀y∙y∈ys y) K⊨F) 74 | 75 | uniq : Unique (I , i) (I , i ∘ wired f) (K , k) 76 | (≲-refl I , λ x → ≈-refl I) (I≲K , i≲k) 77 | uniq J≲₁K J≲₂K I≲K≋I≲J≲₁K I≲K≋I≲J≲₂K x = 78 | ≈-trans K (≈-sym K (I≲K≋I≲J≲₁K x)) (I≲K≋I≲J≲₂K x) 79 | 80 | wires✓ : ∀ {X Y ys} (f : Y → X) → (∀ y → y ∈ˡ ys) → 81 | (S T : TBox Σ) (A : ABox Σ X) (B : ABox Σ Y) → 82 | (Assertions (⟨ABox⟩ f B) ⊆ Assertions A) → 83 | (I : Interp Σ X) → (I ⊨ (S , T) , A) → (I ⊕ S , wires f ys ⊨ T , B) 84 | wires✓ {X} {Y} {ys} f ∀y∙y∈ys S T A B B⊆A I I⊨STA = 85 | (J , J-init , (J⊨T , J⊨B)) where 86 | 87 | J : Interp Σ (X ⊕ False ⊕ Y) 88 | J = wired f * I 89 | 90 | J⊨T : ⌊ J ⌋ ⊨t T 91 | J⊨T = proj₂ (proj₁ I⊨STA) 92 | 93 | J⊨B : enode * J ⊨a B 94 | J⊨B = *-resp-⟨ABox⟩ f I B (⊨a-resp-⊇ I (⟨ABox⟩ f B) A B⊆A (proj₂ I⊨STA)) 95 | 96 | J-init : J ∈ Initial I (S , wires f ys) 97 | J-init = wired-init f ys ∀y∙y∈ys I S (proj₁ (proj₁ I⊨STA)) 98 | 99 | wiring-impl : ∀ {S T : TBox Σ} → (A B : Object S T) (f : IN B → IN A) → 100 | ABox Σ (IN A ⊕ False ⊕ IN B) 101 | wiring-impl A B f = 102 | wires f (proj₁ (fin B)) 103 | 104 | wiring-impl✓ : ∀ {S T : TBox Σ} → (A B : Object S T) (f : IN B → IN A) → 105 | (Assertions (⟨ABox⟩ f (iface B)) ⊆ Assertions (iface A)) → 106 | ∀ I → (I ⊨ (S , T) , iface A) → (I ⊕ (S , wiring-impl A B f) ⊨ (T , iface B)) 107 | wiring-impl✓ {S} {T} (X , X∈Fin , A) (Y , Y∈Fin , B) f B⊆A = 108 | wires✓ f (proj₂ Y∈Fin) S T A B B⊆A 109 | 110 | wiring : ∀ {S T : TBox Σ} → (A B : Object S T) (f : IN B → IN A) → 111 | (Assertions (⟨ABox⟩ f (iface B)) ⊆ Assertions (iface A)) → 112 | (A ⇒ B) 113 | wiring A B f B⊆A = 114 | ( False , wiring-impl A B f , wiring-impl✓ A B f B⊆A ) 115 | 116 | id✓ : ∀ {S T : TBox Σ} → (A : Object S T) → 117 | (Assertions (⟨ABox⟩ id (iface A)) ⊆ Assertions (iface A)) 118 | id✓ A = subst₂ Assertions (⟨ABox⟩-resp-id (iface A)) refl 119 | 120 | identity : ∀ {S T : TBox Σ} → (A : Object S T) → (A ⇒ A) 121 | identity A = wiring A A id (id✓ A) 122 | 123 | ⊎-swap✓ : ∀ {S T : TBox Σ} → (A B : Object S T) → 124 | Assertions (⟨ABox⟩ ⊎-swap (iface (B ⊗ A))) ⊆ Assertions (iface (A ⊗ B)) 125 | ⊎-swap✓ A B (inj₁ b∈B) = 126 | inj₂ (subst₂ Assertions (⟨ABox⟩-resp-∘ ⊎-swap inj₁ (iface B)) refl b∈B) 127 | ⊎-swap✓ A B (inj₂ a∈A) = 128 | inj₁ (subst₂ Assertions (⟨ABox⟩-resp-∘ ⊎-swap inj₂ (iface A)) refl a∈A) 129 | 130 | symm : ∀ {S T : TBox Σ} → (A B : Object S T) → ((A ⊗ B) ⇒ (B ⊗ A)) 131 | symm A B = wiring (A ⊗ B) (B ⊗ A) ⊎-swap (⊎-swap✓ A B) 132 | 133 | ⊎-assoc⁻¹✓ : ∀ {S T : TBox Σ} → (A B C : Object S T) → 134 | (Assertions (⟨ABox⟩ ⊎-assoc⁻¹ (iface (A ⊗ (B ⊗ C))))) ⊆ 135 | (Assertions (iface ((A ⊗ B) ⊗ C))) 136 | ⊎-assoc⁻¹✓ A B C (inj₁ a∈A) = inj₁ (inj₁ (subst₂ Assertions 137 | (trans (⟨ABox⟩-resp-∘ ⊎-assoc⁻¹ inj₁ (iface A)) 138 | (sym (⟨ABox⟩-resp-∘ inj₁ inj₁ (iface A)))) 139 | refl a∈A)) 140 | ⊎-assoc⁻¹✓ A B C (inj₂ (inj₁ b∈B)) = inj₁ (inj₂ (subst₂ Assertions 141 | (trans (⟨ABox⟩-resp-∘² ⊎-assoc⁻¹ inj₂ inj₁ (iface B)) 142 | (sym (⟨ABox⟩-resp-∘ inj₁ inj₂ (iface B)))) 143 | refl b∈B)) 144 | ⊎-assoc⁻¹✓ A B C (inj₂ (inj₂ c∈C)) = inj₂ (subst₂ Assertions 145 | (⟨ABox⟩-resp-∘² ⊎-assoc⁻¹ inj₂ inj₂ (iface C)) 146 | refl c∈C) 147 | 148 | assoc : ∀ {S T : TBox Σ} → (A B C : Object S T) → 149 | (((A ⊗ B) ⊗ C) ⇒ (A ⊗ (B ⊗ C))) 150 | assoc A B C = wiring ((A ⊗ B) ⊗ C) (A ⊗ (B ⊗ C)) ⊎-assoc⁻¹ (⊎-assoc⁻¹✓ A B C) 151 | 152 | ⊎-assoc✓ : ∀ {S T : TBox Σ} → (A B C : Object S T) → 153 | (Assertions (⟨ABox⟩ ⊎-assoc (iface ((A ⊗ B) ⊗ C)))) ⊆ 154 | (Assertions (iface (A ⊗ (B ⊗ C)))) 155 | ⊎-assoc✓ A B C (inj₁ (inj₁ a∈A)) = inj₁ (subst₂ Assertions 156 | (⟨ABox⟩-resp-∘² ⊎-assoc inj₁ inj₁ (iface A)) 157 | refl a∈A) 158 | ⊎-assoc✓ A B C (inj₁ (inj₂ b∈B)) = inj₂ (inj₁ (subst₂ Assertions 159 | (trans (⟨ABox⟩-resp-∘² ⊎-assoc inj₁ inj₂ (iface B)) 160 | (sym (⟨ABox⟩-resp-∘ inj₂ inj₁ (iface B)))) 161 | refl b∈B)) 162 | ⊎-assoc✓ A B C (inj₂ c∈C) = inj₂ (inj₂ (subst₂ Assertions 163 | (trans (⟨ABox⟩-resp-∘ ⊎-assoc inj₂ (iface C)) 164 | (sym (⟨ABox⟩-resp-∘ inj₂ inj₂ (iface C)))) 165 | refl c∈C)) 166 | 167 | assoc⁻¹ : ∀ {S T : TBox Σ} → (A B C : Object S T) → 168 | ((A ⊗ (B ⊗ C)) ⇒ ((A ⊗ B) ⊗ C)) 169 | assoc⁻¹ A B C = wiring (A ⊗ (B ⊗ C)) ((A ⊗ B) ⊗ C) ⊎-assoc (⊎-assoc✓ A B C) 170 | 171 | inj₂✓ : ∀ {S T : TBox Σ} → (A : Object S T) → 172 | (Assertions (⟨ABox⟩ inj₂ (iface A))) ⊆ 173 | (Assertions (iface (I ⊗ A))) 174 | inj₂✓ A = inj₂ 175 | 176 | unit₁ : ∀ {S T : TBox Σ} (A : Object S T) → ((I ⊗ A) ⇒ A) 177 | unit₁ A = wiring (I ⊗ A) A inj₂ (inj₂✓ A) 178 | 179 | ⊎-unit₁✓ : ∀ {S T : TBox Σ} → (A : Object S T) → 180 | Assertions (⟨ABox⟩ ⊎-unit₁ (iface (I ⊗ A))) ⊆ Assertions (iface A) 181 | ⊎-unit₁✓ A (inj₁ ()) 182 | ⊎-unit₁✓ A (inj₂ a∈A) = subst₂ Assertions 183 | (trans (⟨ABox⟩-resp-∘ ⊎-unit₁ inj₂ (iface A)) (⟨ABox⟩-resp-id (iface A))) 184 | refl a∈A 185 | 186 | unit₁⁻¹ : ∀ {S T : TBox Σ} (A : Object S T) → (A ⇒ (I ⊗ A)) 187 | unit₁⁻¹ A = wiring A (I ⊗ A) ⊎-unit₁ (⊎-unit₁✓ A) 188 | 189 | inj₁✓ : ∀ {S T : TBox Σ} → (A : Object S T) → 190 | (Assertions (⟨ABox⟩ inj₁ (iface A))) ⊆ 191 | (Assertions (iface (A ⊗ I))) 192 | inj₁✓ A = inj₁ 193 | 194 | unit₂ : ∀ {S T : TBox Σ} (A : Object S T) → ((A ⊗ I) ⇒ A) 195 | unit₂ A = wiring (A ⊗ I) A inj₁ (inj₁✓ A) 196 | 197 | ⊎-unit₂✓ : ∀ {S T : TBox Σ} (A : Object S T) → 198 | Assertions (⟨ABox⟩ ⊎-unit₂ (iface (A ⊗ I))) ⊆ Assertions (iface A) 199 | ⊎-unit₂✓ A (inj₁ a∈A) = subst₂ Assertions 200 | (trans (⟨ABox⟩-resp-∘ ⊎-unit₂ inj₁ (iface A)) (⟨ABox⟩-resp-id (iface A))) 201 | refl a∈A 202 | ⊎-unit₂✓ A (inj₂ ()) 203 | 204 | unit₂⁻¹ : ∀ {S T : TBox Σ} (A : Object S T) → (A ⇒ (A ⊗ I)) 205 | unit₂⁻¹ A = wiring A (A ⊗ I) ⊎-unit₂ (⊎-unit₂✓ A) 206 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Concept.agda: -------------------------------------------------------------------------------- 1 | open import Web.Semantic.DL.Role using ( Role ) 2 | open import Web.Semantic.DL.Signature using ( Signature ; CN ; RN ) 3 | 4 | module Web.Semantic.DL.Concept where 5 | 6 | infixr 9 ∀[_]_ 7 | infixr 8 ∃⟨_⟩_ 8 | infixl 7 _⊓_ 9 | infixl 6 _⊔_ 10 | 11 | data Concept (Σ : Signature) : Set where 12 | ⟨_⟩ : (c : CN Σ) → Concept Σ 13 | ¬⟨_⟩ : (c : CN Σ) → Concept Σ 14 | ⊤ : Concept Σ 15 | ⊥ : Concept Σ 16 | _⊓_ : (C D : Concept Σ) → Concept Σ 17 | _⊔_ : (C D : Concept Σ) → Concept Σ 18 | ∀[_]_ : (R : Role Σ) (C : Concept Σ) → Concept Σ 19 | ∃⟨_⟩_ : (R : Role Σ) (C : Concept Σ) → Concept Σ 20 | ≤1 : (R : Role Σ) → Concept Σ 21 | >1 : (R : Role Σ) → Concept Σ 22 | 23 | neg : ∀ {Σ} (C : Concept Σ) → Concept Σ 24 | neg ⟨ c ⟩ = ¬⟨ c ⟩ 25 | neg ¬⟨ c ⟩ = ⟨ c ⟩ 26 | neg ⊤ = ⊥ 27 | neg ⊥ = ⊤ 28 | neg (C ⊓ D) = neg C ⊔ neg D 29 | neg (C ⊔ D) = neg C ⊓ neg D 30 | neg (∀[ R ] C) = ∃⟨ R ⟩ neg C 31 | neg (∃⟨ R ⟩ C) = ∀[ R ] neg C 32 | neg (≤1 R) = >1 R 33 | neg (>1 R) = ≤1 R 34 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Concept/Model.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( ∃ ; _×_ ; _,_ ) 2 | open import Data.Sum using ( inj₁ ; inj₂ ) 3 | open import Relation.Nullary using ( Dec ; yes ; no ) 4 | open import Relation.Unary using ( _∈_ ; _∉_ ; ∅ ; U ; _∪_ ; _∩_ ; ∁ ) 5 | open import Web.Semantic.DL.Concept using ( Concept ; ⟨_⟩ ; ¬⟨_⟩ ; ⊤ ; ⊥ ; _⊓_ ; _⊔_ ; ∀[_]_ ; ∃⟨_⟩_ ; ≤1 ; >1 ; neg ) 6 | open import Web.Semantic.DL.Role.Model using ( _⟦_⟧₂ ; ⟦⟧₂-resp-≈ ; ⟦⟧₂-resp-≲ ; ⟦⟧₂-galois ) 7 | open import Web.Semantic.DL.Signature using ( Signature ) 8 | open import Web.Semantic.DL.TBox.Interp using ( Interp ; Δ ; _⊨_≈_ ; _⊨_≉_ ; con ; con-≈ ; ≈-refl ; ≈-sym ; ≈-trans ) 9 | open import Web.Semantic.DL.TBox.Interp.Morphism using ( _≲_ ; _≃_ ; ≲-resp-con ; ≲-resp-≈ ; ≃-impl-≲ ; ≃-impl-≳ ; ≃-iso ; ≃-iso⁻¹ ; ≃-image ; ≃-image⁻¹ ; ≃-refl-≈ ; ≃-sym ) 10 | open import Web.Semantic.Util using ( Subset ; tt ; _∘_ ; ExclMiddle ; elim ) 11 | 12 | module Web.Semantic.DL.Concept.Model {Σ : Signature} where 13 | 14 | _⟦_⟧₁ : ∀ (I : Interp Σ) → Concept Σ → Subset (Δ I) 15 | I ⟦ ⟨ c ⟩ ⟧₁ = con I c 16 | I ⟦ ¬⟨ c ⟩ ⟧₁ = ∁ (con I c) 17 | I ⟦ ⊤ ⟧₁ = U 18 | I ⟦ ⊥ ⟧₁ = ∅ 19 | I ⟦ C ⊓ D ⟧₁ = I ⟦ C ⟧₁ ∩ I ⟦ D ⟧₁ 20 | I ⟦ C ⊔ D ⟧₁ = I ⟦ C ⟧₁ ∪ I ⟦ D ⟧₁ 21 | I ⟦ ∀[ R ] C ⟧₁ = λ x → ∀ y → ((x , y) ∈ I ⟦ R ⟧₂) → (y ∈ I ⟦ C ⟧₁) 22 | I ⟦ ∃⟨ R ⟩ C ⟧₁ = λ x → ∃ λ y → ((x , y) ∈ I ⟦ R ⟧₂) × (y ∈ I ⟦ C ⟧₁) 23 | I ⟦ ≤1 R ⟧₁ = λ x → ∀ y z → ((x , y) ∈ I ⟦ R ⟧₂) → ((x , z) ∈ I ⟦ R ⟧₂) → (I ⊨ y ≈ z) 24 | I ⟦ >1 R ⟧₁ = λ x → ∃ λ y → ∃ λ z → ((x , y) ∈ I ⟦ R ⟧₂) × ((x , z) ∈ I ⟦ R ⟧₂) × (I ⊨ y ≉ z) 25 | 26 | ⟦⟧₁-resp-≈ : ∀ I C {x y} → (x ∈ I ⟦ C ⟧₁) → (I ⊨ x ≈ y) → (y ∈ I ⟦ C ⟧₁) 27 | ⟦⟧₁-resp-≈ I ⟨ c ⟩ x∈⟦c⟧ x≈y = 28 | con-≈ I c x∈⟦c⟧ x≈y 29 | ⟦⟧₁-resp-≈ I ¬⟨ c ⟩ x∉⟦c⟧ x≈y = 30 | λ y∈⟦c⟧ → x∉⟦c⟧ (con-≈ I c y∈⟦c⟧ (≈-sym I x≈y)) 31 | ⟦⟧₁-resp-≈ I ⊤ x∈⊤ x≈y = 32 | tt 33 | ⟦⟧₁-resp-≈ I (C ⊓ D) (x∈⟦C⟧ , x∈⟦D⟧) x≈y = 34 | (⟦⟧₁-resp-≈ I C x∈⟦C⟧ x≈y , ⟦⟧₁-resp-≈ I D x∈⟦D⟧ x≈y) 35 | ⟦⟧₁-resp-≈ I (C ⊔ D) (inj₁ x∈⟦C⟧) x≈y = 36 | inj₁ (⟦⟧₁-resp-≈ I C x∈⟦C⟧ x≈y) 37 | ⟦⟧₁-resp-≈ I (C ⊔ D) (inj₂ x∈⟦D⟧) x≈y = 38 | inj₂ (⟦⟧₁-resp-≈ I D x∈⟦D⟧ x≈y) 39 | ⟦⟧₁-resp-≈ I (∀[ R ] C) x∈⟦∀RC⟧ x≈y = 40 | λ z yz∈⟦R⟧ → x∈⟦∀RC⟧ z (⟦⟧₂-resp-≈ I R x≈y yz∈⟦R⟧ (≈-refl I)) 41 | ⟦⟧₁-resp-≈ I (∃⟨ R ⟩ C) (z , xz∈⟦R⟧ , z∈⟦C⟧) x≈y = 42 | (z , ⟦⟧₂-resp-≈ I R (≈-sym I x≈y) xz∈⟦R⟧ (≈-refl I) , z∈⟦C⟧) 43 | ⟦⟧₁-resp-≈ I (≤1 R) x∈⟦≤1R⟧ x≈y = 44 | λ w z yw∈⟦R⟧ yz∈⟦R⟧ → x∈⟦≤1R⟧ w z 45 | (⟦⟧₂-resp-≈ I R x≈y yw∈⟦R⟧ (≈-refl I)) 46 | (⟦⟧₂-resp-≈ I R x≈y yz∈⟦R⟧ (≈-refl I)) 47 | ⟦⟧₁-resp-≈ I (>1 R) (w , z , xw∈⟦R⟧ , xz∈⟦R⟧ , w≉z) x≈y = 48 | ( w , z 49 | , ⟦⟧₂-resp-≈ I R (≈-sym I x≈y) xw∈⟦R⟧ (≈-refl I) 50 | , ⟦⟧₂-resp-≈ I R (≈-sym I x≈y) xz∈⟦R⟧ (≈-refl I) 51 | , w≉z) 52 | ⟦⟧₁-resp-≈ I ⊥ () x≈y 53 | 54 | mutual 55 | 56 | ⟦⟧₁-refl-≃ : ∀ {I J : Interp Σ} → (I≃J : I ≃ J) → ∀ {x} C → (≃-image⁻¹ I≃J x ∈ I ⟦ C ⟧₁) → (x ∈ J ⟦ C ⟧₁) 57 | ⟦⟧₁-refl-≃ {I} {J} I≃J {x} C x∈⟦C⟧ = ⟦⟧₁-resp-≈ J C (⟦⟧₁-resp-≃ I≃J C x∈⟦C⟧) (≃-iso⁻¹ I≃J x) 58 | 59 | ⟦⟧₁-resp-≃ : ∀ {I J : Interp Σ} → (I≃J : I ≃ J) → ∀ {x} C → (x ∈ I ⟦ C ⟧₁) → (≃-image I≃J x ∈ J ⟦ C ⟧₁) 60 | ⟦⟧₁-resp-≃ {I} {J} I≃J ⟨ c ⟩ x∈⟦c⟧ = 61 | ≲-resp-con (≃-impl-≲ I≃J) x∈⟦c⟧ 62 | ⟦⟧₁-resp-≃ {I} {J} I≃J {x} ¬⟨ c ⟩ x∉⟦c⟧ = 63 | λ x∈⟦c⟧ → x∉⟦c⟧ (con-≈ I c (≲-resp-con (≃-impl-≳ I≃J) x∈⟦c⟧) (≃-iso I≃J x)) 64 | ⟦⟧₁-resp-≃ {I} {J} I≃J ⊤ _ = 65 | tt 66 | ⟦⟧₁-resp-≃ {I} {J} I≃J ⊥ () 67 | ⟦⟧₁-resp-≃ {I} {J} I≃J (C ⊓ D) (x∈⟦C⟧ , x∈⟦D⟧) = 68 | (⟦⟧₁-resp-≃ I≃J C x∈⟦C⟧ , ⟦⟧₁-resp-≃ I≃J D x∈⟦D⟧) 69 | ⟦⟧₁-resp-≃ {I} {J} I≃J (C ⊔ D) (inj₁ x∈⟦C⟧) = 70 | inj₁ (⟦⟧₁-resp-≃ I≃J C x∈⟦C⟧) 71 | ⟦⟧₁-resp-≃ {I} {J} I≃J (C ⊔ D) (inj₂ x∈⟦D⟧) = 72 | inj₂ (⟦⟧₁-resp-≃ I≃J D x∈⟦D⟧) 73 | ⟦⟧₁-resp-≃ {I} {J} I≃J (∀[ R ] C) x∈⟦∀RC⟧ = 74 | λ y xy∈⟦R⟧ → ⟦⟧₁-refl-≃ I≃J C 75 | (x∈⟦∀RC⟧ (≃-image⁻¹ I≃J y) (⟦⟧₂-galois (≃-sym I≃J) R xy∈⟦R⟧)) 76 | ⟦⟧₁-resp-≃ {I} {J} I≃J (∃⟨ R ⟩ C) (y , xy∈⟦R⟧ , y∈⟦C⟧) = 77 | ( ≃-image I≃J y 78 | , ⟦⟧₂-resp-≲ (≃-impl-≲ I≃J) R xy∈⟦R⟧ 79 | , ⟦⟧₁-resp-≃ I≃J C y∈⟦C⟧ ) 80 | ⟦⟧₁-resp-≃ {I} {J} I≃J (≤1 R) x∈⟦≤1R⟧ = 81 | λ y z xy∈⟦R⟧ xz∈⟦R⟧ → ≃-refl-≈ I≃J 82 | (x∈⟦≤1R⟧ (≃-image⁻¹ I≃J y) (≃-image⁻¹ I≃J z) 83 | (⟦⟧₂-galois (≃-sym I≃J) R xy∈⟦R⟧) (⟦⟧₂-galois (≃-sym I≃J) R xz∈⟦R⟧)) 84 | ⟦⟧₁-resp-≃ {I} {J} I≃J (>1 R) (y , z , xy∈⟦R⟧ , xz∈⟦R⟧ , y≉z) = 85 | ( ≃-image I≃J y , ≃-image I≃J z 86 | , ⟦⟧₂-resp-≲ (≃-impl-≲ I≃J) R xy∈⟦R⟧ 87 | , ⟦⟧₂-resp-≲ (≃-impl-≲ I≃J) R xz∈⟦R⟧ 88 | , y≉z ∘ ≃-refl-≈ (≃-sym I≃J) ) 89 | 90 | neg-sound : ∀ I {x} C → (x ∈ I ⟦ neg C ⟧₁) → (x ∉ I ⟦ C ⟧₁) 91 | neg-sound I ⟨ c ⟩ x∉⟦c⟧ x∈⟦c⟧ = x∉⟦c⟧ x∈⟦c⟧ 92 | neg-sound I ¬⟨ c ⟩ x∈⟦c⟧ x∉⟦c⟧ = x∉⟦c⟧ x∈⟦c⟧ 93 | neg-sound I ⊤ () _ 94 | neg-sound I ⊥ _ () 95 | neg-sound I (C ⊓ D) (inj₁ x∈⟦¬C⟧) (x∈⟦C⟧ , x∈⟦D⟧) = neg-sound I C x∈⟦¬C⟧ x∈⟦C⟧ 96 | neg-sound I (C ⊓ D) (inj₂ x∈⟦¬D⟧) (x∈⟦C⟧ , x∈⟦D⟧) = neg-sound I D x∈⟦¬D⟧ x∈⟦D⟧ 97 | neg-sound I (C ⊔ D) (x∈⟦¬C⟧ , x∈⟦¬D⟧) (inj₁ x∈⟦C⟧) = neg-sound I C x∈⟦¬C⟧ x∈⟦C⟧ 98 | neg-sound I (C ⊔ D) (x∈⟦¬C⟧ , x∈⟦¬D⟧) (inj₂ x∈⟦D⟧) = neg-sound I D x∈⟦¬D⟧ x∈⟦D⟧ 99 | neg-sound I (∀[ R ] C) (y , xy∈⟦R⟧ , y∈⟦¬C⟧) x∈⟦∀RC⟧ = neg-sound I C y∈⟦¬C⟧ (x∈⟦∀RC⟧ y xy∈⟦R⟧) 100 | neg-sound I (∃⟨ R ⟩ C) x∈⟦∀R¬C⟧ (y , xy∈⟦R⟧ , y∈⟦C⟧) = neg-sound I C (x∈⟦∀R¬C⟧ y xy∈⟦R⟧) y∈⟦C⟧ 101 | neg-sound I (≤1 R) (y , z , xy∈⟦R⟧ , xz∈⟦R⟧ , y≉z) x∈⟦≤1R⟧ = y≉z (x∈⟦≤1R⟧ y z xy∈⟦R⟧ xz∈⟦R⟧) 102 | neg-sound I (>1 R) x∈⟦≤1R⟧ (y , z , xy∈⟦R⟧ , xz∈⟦R⟧ , y≉z) = y≉z (x∈⟦≤1R⟧ y z xy∈⟦R⟧ xz∈⟦R⟧) 103 | 104 | neg-complete : ExclMiddle → ∀ I {x} C → (x ∉ I ⟦ C ⟧₁) → (x ∈ I ⟦ neg C ⟧₁) 105 | neg-complete excl-middle I ⟨ c ⟩ x∉⟦c⟧ = x∉⟦c⟧ 106 | neg-complete excl-middle I {x} ¬⟨ c ⟩ ¬x∉⟦c⟧ with excl-middle (x ∈ con I c) 107 | neg-complete excl-middle I ¬⟨ c ⟩ ¬x∉⟦c⟧ | yes x∈⟦c⟧ = x∈⟦c⟧ 108 | neg-complete excl-middle I ¬⟨ c ⟩ ¬x∉⟦c⟧ | no x∉⟦c⟧ = elim (¬x∉⟦c⟧ x∉⟦c⟧) 109 | neg-complete excl-middle I ⊤ x∉⟦⊤⟧ = x∉⟦⊤⟧ tt 110 | neg-complete excl-middle I ⊥ x∉⟦⊥⟧ = tt 111 | neg-complete excl-middle I {x} (C ⊓ D) x∉⟦C⊓D⟧ with excl-middle (x ∈ I ⟦ C ⟧₁) 112 | neg-complete excl-middle I (C ⊓ D) x∉⟦C⊓D⟧ | yes x∈⟦C⟧ = 113 | inj₂ (neg-complete excl-middle I D (λ x∈⟦D⟧ → x∉⟦C⊓D⟧ (x∈⟦C⟧ , x∈⟦D⟧))) 114 | neg-complete excl-middle I (C ⊓ D) x∉⟦C⊓D⟧ | no x∉⟦C⟧ = 115 | inj₁ (neg-complete excl-middle I C x∉⟦C⟧) 116 | neg-complete excl-middle I (C ⊔ D) x∉⟦C⊔D⟧ = 117 | ( neg-complete excl-middle I C (x∉⟦C⊔D⟧ ∘ inj₁) 118 | , neg-complete excl-middle I D (x∉⟦C⊔D⟧ ∘ inj₂)) 119 | neg-complete excl-middle I {x} (∀[ R ] C) x∉⟦∀RC⟧ 120 | with excl-middle (∃ λ y → ((x , y) ∈ I ⟦ R ⟧₂) × (y ∉ I ⟦ C ⟧₁)) 121 | neg-complete excl-middle I {x} (∀[ R ] C) x∉⟦∀RC⟧ | yes (y , xy∈⟦R⟧ , y∉⟦C⟧) = 122 | (y , xy∈⟦R⟧ , neg-complete excl-middle I C y∉⟦C⟧) 123 | neg-complete excl-middle I {x} (∀[ R ] C) x∉⟦∀RC⟧ | no ∄xy∈⟦R⟧∙y∉⟦C⟧ = 124 | elim (x∉⟦∀RC⟧ lemma) where 125 | lemma : x ∈ I ⟦ ∀[ R ] C ⟧₁ 126 | lemma y xy∈⟦R⟧ with excl-middle (y ∈ I ⟦ C ⟧₁) 127 | lemma y xy∈⟦R⟧ | yes y∈⟦C⟧ = y∈⟦C⟧ 128 | lemma y xy∈⟦R⟧ | no y∉⟦C⟧ = elim (∄xy∈⟦R⟧∙y∉⟦C⟧ (y , xy∈⟦R⟧ , y∉⟦C⟧)) 129 | neg-complete excl-middle I (∃⟨ R ⟩ C) x∉⟦∃RC⟧ = 130 | λ y xy∈⟦R⟧ → neg-complete excl-middle I C 131 | (λ y∈⟦C⟧ → x∉⟦∃RC⟧ (y , xy∈⟦R⟧ , y∈⟦C⟧)) 132 | neg-complete excl-middle I {x} (≤1 R) x∉⟦≤1R⟧ = lemma where 133 | lemma : x ∈ I ⟦ >1 R ⟧₁ 134 | lemma with excl-middle (x ∈ I ⟦ >1 R ⟧₁) 135 | lemma | yes x∈⟦>1R⟧ = x∈⟦>1R⟧ 136 | lemma | no x∉⟦>1R⟧ = elim (x∉⟦≤1R⟧ lemma′) where 137 | lemma′ : x ∈ I ⟦ ≤1 R ⟧₁ 138 | lemma′ y z xy∈⟦R⟧ xz∈⟦R⟧ with excl-middle (I ⊨ y ≈ z) 139 | lemma′ y z xy∈⟦R⟧ xz∈⟦R⟧ | yes y≈z = y≈z 140 | lemma′ y z xy∈⟦R⟧ xz∈⟦R⟧ | no y≉z = 141 | elim (x∉⟦>1R⟧ (y , z , xy∈⟦R⟧ , xz∈⟦R⟧ , y≉z)) 142 | neg-complete excl-middle I {x} (>1 R) x∉⟦>1R⟧ = lemma where 143 | lemma : x ∈ I ⟦ ≤1 R ⟧₁ 144 | lemma y z xy∈⟦R⟧ xz∈⟦R⟧ with excl-middle (I ⊨ y ≈ z) 145 | lemma y z xy∈⟦R⟧ xz∈⟦R⟧ | yes y≈z = y≈z 146 | lemma y z xy∈⟦R⟧ xz∈⟦R⟧ | no y≉z = 147 | elim (x∉⟦>1R⟧ (y , z , xy∈⟦R⟧ , xz∈⟦R⟧ , y≉z)) 148 | 149 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Concept/Skolemization.agda: -------------------------------------------------------------------------------- 1 | open import Data.Bool using ( Bool ; true ; false ; if_then_else_ ) 2 | open import Data.Product using ( _×_ ; _,_ ) 3 | open import Data.Sum using ( inj₁ ; inj₂ ) 4 | open import Relation.Unary using ( _∈_ ) 5 | open import Web.Semantic.DL.Concept using ( Concept ; ⟨_⟩ ; ¬⟨_⟩ ; ⊤ ; ⊥ ; _⊓_ ; _⊔_ ; ∀[_]_ ; ∃⟨_⟩_ ; ≤1 ; >1 ; neg ) 6 | open import Web.Semantic.DL.Concept.Model using ( _⟦_⟧₁ ; ⟦⟧₁-resp-≈ ) 7 | open import Web.Semantic.DL.FOL using ( Formula ; true ; false ; _∧_ ; _∈₁_ ; _∈₂_ ; _∈₁_⇒_ ; _∼_ ; _∼_⇒_ ; ∀₁ ) 8 | open import Web.Semantic.DL.FOL.Model using ( _⊨f_ ) 9 | open import Web.Semantic.DL.Role.Skolemization using ( rskolem ; rskolem⇒ ; rskolem-sound ; rskolem⇒-sound ) 10 | open import Web.Semantic.DL.Signature using ( Signature ) 11 | open import Web.Semantic.Util using ( True ; tt ) 12 | 13 | module Web.Semantic.DL.Concept.Skolemization {Σ : Signature} where 14 | 15 | CSkolems : Set → Concept Σ → Set 16 | CSkolems Δ ⟨ c ⟩ = True 17 | CSkolems Δ ¬⟨ c ⟩ = True 18 | CSkolems Δ ⊤ = True 19 | CSkolems Δ ⊥ = True 20 | CSkolems Δ (C ⊓ D) = (CSkolems Δ C) × (CSkolems Δ D) 21 | CSkolems Δ (C ⊔ D) = (Δ → Bool) × (CSkolems Δ C) × (CSkolems Δ D) 22 | CSkolems Δ (∀[ R ] C) = CSkolems Δ C 23 | CSkolems Δ (∃⟨ R ⟩ C) = (Δ → Δ) × (CSkolems Δ C) 24 | CSkolems Δ (≤1 R) = True 25 | CSkolems Δ (>1 R) = (Δ → Δ) × (Δ → Δ) 26 | 27 | cskolem : ∀ {Δ} C → (CSkolems Δ C) → Δ → Formula Σ Δ 28 | cskolem ⟨ c ⟩ Φ x = x ∈₁ c 29 | cskolem ¬⟨ c ⟩ Φ x = x ∈₁ c ⇒ false 30 | cskolem ⊤ Φ x = true 31 | cskolem ⊥ Φ x = false 32 | cskolem (C ⊓ D) (Φ , Ψ) x = (cskolem C Φ x) ∧ (cskolem D Ψ x) 33 | cskolem (C ⊔ D) (φ , Φ , Ψ) x = if (φ x) then (cskolem C Φ x) else (cskolem D Ψ x) 34 | cskolem (∀[ R ] C) Φ x = ∀₁ λ y → rskolem⇒ R x y (cskolem C Φ y) 35 | cskolem (∃⟨ R ⟩ C) (φ , Ψ) x = (rskolem R x (φ x)) ∧ (cskolem C Ψ (φ x)) 36 | cskolem (≤1 R) Φ x = ∀₁ λ y → ∀₁ λ z → rskolem⇒ R x y (rskolem⇒ R x z (y ∼ z)) 37 | cskolem (>1 R) (φ , ψ) x = (rskolem R x (φ x)) ∧ (rskolem R x (ψ x)) ∧ (φ x ∼ ψ x ⇒ false) 38 | 39 | cskolem-sound : ∀ I C Φ x → (I ⊨f cskolem C Φ x) → (x ∈ I ⟦ C ⟧₁) 40 | cskolem-sound I ⟨ c ⟩ Φ x x∈⟦c⟧ = x∈⟦c⟧ 41 | cskolem-sound I ¬⟨ c ⟩ Φ x x∉⟦c⟧ = x∉⟦c⟧ 42 | cskolem-sound I ⊤ Φ x _ = tt 43 | cskolem-sound I ⊥ Φ x () 44 | cskolem-sound I (C ⊓ D) (Φ , Ψ) x (Fx , Gx) = 45 | (cskolem-sound I C Φ x Fx , cskolem-sound I D Ψ x Gx) 46 | cskolem-sound I (C ⊔ D) (φ , Φ , Ψ) x F∨Gx with φ x 47 | cskolem-sound I (C ⊔ D) (φ , Φ , Ψ) x Fx | true = inj₁ (cskolem-sound I C Φ x Fx) 48 | cskolem-sound I (C ⊔ D) (φ , Φ , Ψ) x Gx | false = inj₂ (cskolem-sound I D Ψ x Gx) 49 | cskolem-sound I (∀[ R ] C) Φ x ∀RFx = 50 | λ y xy∈R → cskolem-sound I C Φ y (rskolem⇒-sound I R x y (cskolem C Φ y) (∀RFx y) xy∈R) 51 | cskolem-sound I (∃⟨ R ⟩ C) (φ , Ψ) x (xy∈R , Fy) = 52 | ( φ x , rskolem-sound I R x (φ x) xy∈R , cskolem-sound I C Ψ (φ x) Fy ) 53 | cskolem-sound I (≤1 R) Φ x ≤1Rx = 54 | λ y z xy∈R xz∈R → rskolem⇒-sound I R x z (y ∼ z) 55 | (rskolem⇒-sound I R x y (rskolem⇒ R x z (y ∼ z)) (≤1Rx y z) xy∈R) xz∈R 56 | cskolem-sound I (>1 R) (φ , ψ) x (xy∈R , xz∈R , y≁z) = 57 | (φ x , ψ x , rskolem-sound I R x (φ x) xy∈R , rskolem-sound I R x (ψ x) xz∈R , y≁z) -------------------------------------------------------------------------------- /src/Web/Semantic/DL/FOL.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( _×_ ) 2 | open import Web.Semantic.DL.Signature using ( Signature ; CN ; RN ) 3 | 4 | module Web.Semantic.DL.FOL where 5 | 6 | -- A fragment of first order logic with no existentials or disjunctions. 7 | 8 | infixr 4 _∧_ 9 | 10 | data Formula (Σ : Signature) (Δ : Set) : Set where 11 | true : Formula Σ Δ 12 | false : Formula Σ Δ 13 | _∧_ : Formula Σ Δ → Formula Σ Δ → Formula Σ Δ 14 | _∈₁_ : Δ → CN Σ → Formula Σ Δ 15 | _∈₁_⇒_ : Δ → CN Σ → Formula Σ Δ → Formula Σ Δ 16 | _∈₂_ : (Δ × Δ) → RN Σ → Formula Σ Δ 17 | _∈₂_⇒_ : (Δ × Δ) → RN Σ → Formula Σ Δ → Formula Σ Δ 18 | _∼_ : Δ → Δ → Formula Σ Δ 19 | _∼_⇒_ : Δ → Δ → Formula Σ Δ → Formula Σ Δ 20 | ∀₁ : (Δ → Formula Σ Δ) → Formula Σ Δ 21 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/FOL/Model.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( _×_ ; _,_ ) 2 | open import Data.Sum using ( _⊎_ ) 3 | open import Relation.Unary using ( _∈_ ; _∉_ ) 4 | open import Web.Semantic.DL.FOL using ( Formula ; true ; false ; _∧_ ; _∈₁_ ; _∈₁_⇒_ ; _∈₂_ ; _∈₂_⇒_ ; _∼_ ; _∼_⇒_ ; ∀₁ ) 5 | open import Web.Semantic.DL.Signature using ( Signature ) 6 | open import Web.Semantic.DL.TBox.Interp using ( Interp ; Δ ; _⊨_≈_ ; con ; rol ) 7 | open import Web.Semantic.Util using ( True ; False ) 8 | 9 | module Web.Semantic.DL.FOL.Model {Σ : Signature} where 10 | 11 | _⊨f_ : (I : Interp Σ) → Formula Σ (Δ I) → Set 12 | I ⊨f true = True 13 | I ⊨f false = False 14 | I ⊨f (F ∧ G) = (I ⊨f F) × (I ⊨f G) 15 | I ⊨f (x ∈₁ c) = x ∈ (con I c) 16 | I ⊨f (x ∈₁ c ⇒ F) = (x ∈ (con I c)) → (I ⊨f F) 17 | I ⊨f ((x , y) ∈₂ r) = (x , y) ∈ (rol I r) 18 | I ⊨f ((x , y) ∈₂ r ⇒ F) = ((x , y) ∈ (rol I r)) → (I ⊨f F) 19 | I ⊨f (x ∼ y) = I ⊨ x ≈ y 20 | I ⊨f (x ∼ y ⇒ F) = (I ⊨ x ≈ y) → (I ⊨f F) 21 | I ⊨f ∀₁ F = ∀ x → (I ⊨f F x) 22 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Integrity.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( ∃ ; _×_ ; _,_ ; proj₁ ; proj₂ ) 2 | open import Relation.Unary using ( _∈_ ) 3 | open import Web.Semantic.DL.ABox using ( ABox ) 4 | open import Web.Semantic.DL.ABox.Interp using ( Interp ; _*_ ) 5 | open import Web.Semantic.DL.ABox.Interp.Morphism using 6 | ( _≲_ ; ≲-trans ; _**_ ; _≋_ ) 7 | open import Web.Semantic.DL.KB using ( KB ; _,_ ; tbox ) 8 | open import Web.Semantic.DL.KB.Model using ( _⊨_ ) 9 | open import Web.Semantic.DL.Signature using ( Signature ) 10 | open import Web.Semantic.DL.TBox using ( _,_ ) 11 | open import Web.Semantic.Util using ( _⊕_⊕_ ; inode ; enode ) 12 | 13 | module Web.Semantic.DL.Integrity {Σ : Signature} {X V Y : Set} where 14 | 15 | -- A variant of the notion of OWL integrity in: 16 | 17 | -- Boris Motik, Ian Horrocks, and Ulrike Sattler. Bridging the gap between OWL and 18 | -- relational databases. In Proc. Conf. World Wide Web (WWW2007), pp 807–816, 2007. 19 | 20 | -- They propose defining that an integrity constraint S (given as a T-Box) is valid 21 | -- for a knowlege base K whenever S is valid in every minimal Herbrand model. 22 | -- We strengthen this requirement to say that there is a minimal model of K, 23 | -- which validtes S. 24 | 25 | infixr 4 _>>_ _,_ 26 | infix 2 _⊕_⊨_ 27 | 28 | -- J ∈ Initial I K whenever J is the initial extension of I 29 | -- satisfying K, that is there is I≲J s.t. I ⊨ K 30 | -- and for any other such I≲K, there is a unique 31 | -- J≲K which commutes with I≲J and I≲K. 32 | 33 | _>>_ : ∀ {I : Interp Σ X} {J K : Interp Σ (X ⊕ V ⊕ Y)} → 34 | (I ≲ inode * J) → (J ≲ K) → (I ≲ inode * K) 35 | I≲J >> J≲K = ≲-trans I≲J (inode ** J≲K) 36 | 37 | Unique : ∀ (I : Interp Σ X) → (J K : Interp Σ (X ⊕ V ⊕ Y)) → 38 | (I ≲ inode * J) → (I ≲ inode * K) → Set 39 | Unique I J K I≲J I≲K = ∀ (J≲₁K J≲₂K : J ≲ K) → 40 | (I≲K ≋ I≲J >> J≲₁K) → (I≲K ≋ I≲J >> J≲₂K) → (J≲₁K ≋ J≲₂K) 41 | 42 | data Mediated (I : Interp Σ X) (J K : Interp Σ (X ⊕ V ⊕ Y)) 43 | (I≲J : I ≲ inode * J) (I≲K : I ≲ inode * K) : Set where 44 | _,_ : (J≲K : J ≲ K) → (I≲K ≋ I≲J >> J≲K) × (Unique I J K I≲J I≲K) 45 | → Mediated I J K I≲J I≲K 46 | 47 | med-≲ : ∀ {I} {J K : Interp Σ (X ⊕ V ⊕ Y)} {I≲J I≲K} → 48 | (Mediated I J K I≲J I≲K) → (J ≲ K) 49 | med-≲ (J≲K , I≲K≋I≲J≲K , J≲K-uniq) = J≲K 50 | 51 | med-≋ : ∀ {I} {J K : Interp Σ (X ⊕ V ⊕ Y)} {I≲J I≲K} → 52 | (m : Mediated I J K I≲J I≲K) → (I≲K ≋ I≲J >> med-≲ m) 53 | med-≋ (J≲K , I≲K≋I≲J≲K , J≲K-uniq) = I≲K≋I≲J≲K 54 | 55 | med-uniq : ∀ {I} {J K : Interp Σ (X ⊕ V ⊕ Y)} {I≲J I≲K} → 56 | (Mediated I J K I≲J I≲K) → Unique I J K I≲J I≲K 57 | med-uniq (J≲K , I≲K≋I≲J≲K , J≲K-uniq) = J≲K-uniq 58 | 59 | Mediator : ∀ (I : Interp Σ X) → (J : Interp Σ (X ⊕ V ⊕ Y)) → 60 | (I ≲ inode * J) → (KB Σ (X ⊕ V ⊕ Y)) → Set₁ 61 | Mediator I J I≲J KB = 62 | ∀ K (I≲K : I ≲ inode * K) → (K ⊨ KB) → Mediated I J K I≲J I≲K 63 | 64 | data Initial I KB (J : Interp Σ (X ⊕ V ⊕ Y)) : Set₁ where 65 | _,_ : ∀ (I≲J : I ≲ inode * J) → (J ⊨ KB) × (Mediator I J I≲J KB) → 66 | (J ∈ Initial I KB) 67 | 68 | init-≲ : ∀ {I KB} {J : Interp Σ (X ⊕ V ⊕ Y)} → 69 | (J ∈ Initial I KB) → (I ≲ inode * J) 70 | init-≲ (I≲J , J⊨KB , J-med) = I≲J 71 | 72 | init-⊨ : ∀ {I KB} {J : Interp Σ (X ⊕ V ⊕ Y)} → 73 | (J ∈ Initial I KB) → (J ⊨ KB) 74 | init-⊨ (I≲J , J⊨KB , J-med) = J⊨KB 75 | 76 | init-med : ∀ {I KB} {J : Interp Σ (X ⊕ V ⊕ Y)} → 77 | (J-init : J ∈ Initial I KB) → Mediator I J (init-≲ J-init) KB 78 | init-med (I≲J , J⊨KB , J-med) = J-med 79 | 80 | -- I ⊕ KB₁ ⊨ KB₂ whenever the initial extension of I satisfying 81 | -- KB₁ exists, and has exports satisfying KB₂. 82 | 83 | data _⊕_⊨_ I (KB₁ : KB Σ (X ⊕ V ⊕ Y)) KB₂ : Set₁ where 84 | _,_ : ∀ J → ((J ∈ Initial I KB₁) × (enode * J ⊨ KB₂)) → (I ⊕ KB₁ ⊨ KB₂) 85 | 86 | extension : ∀ {I KB₁ KB₂} → (I ⊕ KB₁ ⊨ KB₂) → (Interp Σ (X ⊕ V ⊕ Y)) 87 | extension (J , J-init , J⊨KB₂) = J 88 | 89 | ext-init : ∀ {I} {KB₁ : KB Σ (X ⊕ V ⊕ Y)} {KB₂} → 90 | (I⊕KB₁⊨KB₂ : I ⊕ KB₁ ⊨ KB₂) → (extension I⊕KB₁⊨KB₂ ∈ Initial I KB₁) 91 | ext-init (J , J-init , J⊨KB₂) = J-init 92 | 93 | ext-⊨ : ∀ {I} {KB₁ : KB Σ (X ⊕ V ⊕ Y)} {KB₂} → 94 | (I⊕KB₁⊨KB₂ : I ⊕ KB₁ ⊨ KB₂) → (enode * (extension I⊕KB₁⊨KB₂) ⊨ KB₂) 95 | ext-⊨ (J , J-init , J⊨KB₂) = J⊨KB₂ 96 | 97 | ext✓ : ∀ {I S T} {F : ABox Σ (X ⊕ V ⊕ Y)} {B} → 98 | (I⊕SF⊨TB : I ⊕ S , F ⊨ T , B) → 99 | (enode * (extension I⊕SF⊨TB) ⊨ (S , T) , B) 100 | ext✓ I⊕SF⊨TB = 101 | ( ( proj₁ (init-⊨ (ext-init I⊕SF⊨TB)) 102 | , proj₁ (ext-⊨ I⊕SF⊨TB) ) 103 | , proj₂ (ext-⊨ I⊕SF⊨TB) ) 104 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Integrity/Closed.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( _×_ ; _,_ ) 2 | open import Relation.Binary.PropositionalEquality using ( refl ) 3 | open import Relation.Unary using ( _∈_ ) 4 | open import Web.Semantic.DL.ABox using ( ABox ; ⟨ABox⟩ ) 5 | open import Web.Semantic.DL.ABox.Interp using ( Interp ; Surjective ; _,_ ; ⌊_⌋ ; ind ; _*_ ; emp ) 6 | open import Web.Semantic.DL.ABox.Interp.Morphism using ( _≲_ ; _,_ ; ≲⌊_⌋ ; ≲-resp-ind ; _≋_ ; _**_ ) 7 | open import Web.Semantic.DL.ABox.Model using ( _⊨a_ ; ⟨ABox⟩-resp-⊨ ; *-resp-⟨ABox⟩ ; ⊨a-resp-≡ ) 8 | open import Web.Semantic.DL.Integrity using ( _⊕_⊨_ ; Initial ; Mediator ; Mediated ; _,_ ) 9 | open import Web.Semantic.DL.KB using ( KB ; _,_ ) 10 | open import Web.Semantic.DL.KB.Model using ( _⊨_ ) 11 | open import Web.Semantic.DL.Signature using ( Signature ) 12 | open import Web.Semantic.DL.TBox.Interp using ( _⊨_≈_ ; ≈-refl ) 13 | open import Web.Semantic.DL.TBox.Interp.Morphism using ( morph ; ≲-image ) 14 | open import Web.Semantic.Util using ( False ; _∘_ ; _⊕_⊕_ ; bnode ; inode ; enode ) 15 | 16 | module Web.Semantic.DL.Integrity.Closed {Σ : Signature} {X : Set} where 17 | 18 | infix 2 _⊨₀_ sur_⊨₀_ 19 | infixr 4 _,_ 20 | 21 | -- A closed-world variant on integrity constraints. 22 | 23 | data Mediated₀ (I J : Interp Σ X) : Set where 24 | _,_ : (I≲J : I ≲ J) → (∀ (I≲₁J I≲₂J : I ≲ J) → (I≲₁J ≋ I≲₂J)) → Mediated₀ I J 25 | 26 | data Initial₀ KB (I : Interp Σ X) : Set₁ where 27 | _,_ : (I ⊨ KB) → (∀ J → (J ⊨ KB) → Mediated₀ I J) → (I ∈ Initial₀ KB) 28 | 29 | data _⊨₀_ (KB₁ KB₂ : KB Σ X) : Set₁ where 30 | _,_ : ∀ I → ((I ∈ Initial₀ KB₁) × (I ⊨ KB₂)) → (KB₁ ⊨₀ KB₂) 31 | 32 | -- Surjective variant 33 | 34 | data sur_⊨₀_ (KB₁ KB₂ : KB Σ X) : Set₁ where 35 | _,_ : ∀ I → ((I ∈ Surjective) × (I ∈ Initial₀ KB₁) × (I ⊨ KB₂)) → (sur KB₁ ⊨₀ KB₂) 36 | 37 | -- Closed-world ICs are given by specializing the open-world case 38 | -- to the empty imported interpretion. 39 | 40 | exp : KB Σ X → KB Σ (False ⊕ False ⊕ X) 41 | exp (T , A) = (T , ⟨ABox⟩ enode A) 42 | 43 | enode⁻¹ : (False ⊕ False ⊕ X) → X 44 | enode⁻¹ (inode ()) 45 | enode⁻¹ (bnode ()) 46 | enode⁻¹ (enode x) = x 47 | 48 | emp-≲ : ∀ (I : Interp Σ False) → (emp ≲ I) 49 | emp-≲ I = (morph (λ ()) (λ {}) (λ ()) (λ {r} → λ {}) , λ ()) 50 | 51 | ⊨₀-impl-⊨ : ∀ KB₁ KB₂ → (KB₁ ⊨₀ KB₂) → (emp ⊕ exp KB₁ ⊨ KB₂) 52 | ⊨₀-impl-⊨ (T , A) (U , B) (I , ((I⊨T , I⊨A) , I-med) , I⊨U , I⊨B) = 53 | (I′ , I′-init , I⊨U , ⊨a-resp-≡ I (ind I) refl B I⊨B) where 54 | 55 | A′ : ABox Σ (False ⊕ False ⊕ X) 56 | A′ = ⟨ABox⟩ enode A 57 | 58 | I′ : Interp Σ (False ⊕ False ⊕ X) 59 | I′ = enode⁻¹ * I 60 | 61 | emp≲I′ : emp ≲ inode * I′ 62 | emp≲I′ = emp-≲ (inode * I′) 63 | 64 | I′-med : Mediator emp I′ emp≲I′ (T , A′) 65 | I′-med J′ emp≲J′ (J′⊨T , J′⊨A′) = lemma (I-med J (J′⊨T , J⊨A)) where 66 | 67 | J : Interp Σ X 68 | J = enode * J′ 69 | 70 | J⊨A : J ⊨a A 71 | J⊨A = *-resp-⟨ABox⟩ enode J′ A J′⊨A′ 72 | 73 | lemma : Mediated₀ I J → Mediated emp I′ J′ emp≲I′ emp≲J′ 74 | lemma ((I≲J , i≲j) , I≲J-uniq) = ((I≲J , i′≲j′) , (λ ()) , I′≲J′-uniq) where 75 | 76 | i′≲j′ : ∀ x → ⌊ J′ ⌋ ⊨ ≲-image I≲J (ind I (enode⁻¹ x)) ≈ ind J′ x 77 | i′≲j′ (inode ()) 78 | i′≲j′ (bnode ()) 79 | i′≲j′ (enode x) = i≲j x 80 | 81 | I′≲J′-uniq : ∀ (I≲₁J I≲₂J : I′ ≲ J′)_ _ → I≲₁J ≋ I≲₂J 82 | I′≲J′-uniq I≲₁J I≲₂J _ _ = I≲J-uniq 83 | (≲⌊ I≲₁J ⌋ , λ x → ≲-resp-ind I≲₁J (enode x)) 84 | (≲⌊ I≲₂J ⌋ , λ x → ≲-resp-ind I≲₂J (enode x)) 85 | 86 | I′⊨A′ : I′ ⊨a A′ 87 | I′⊨A′ = ⟨ABox⟩-resp-⊨ enode (λ x → ≈-refl ⌊ I ⌋) A I⊨A 88 | 89 | I′-init : I′ ∈ Initial emp (T , A′) 90 | I′-init = ( emp-≲ (inode * I′) , (I⊨T , I′⊨A′) , I′-med ) 91 | 92 | ⊨-impl-⊨₀ : ∀ KB₁ KB₂ → (emp ⊕ exp KB₁ ⊨ KB₂) → (KB₁ ⊨₀ KB₂) 93 | ⊨-impl-⊨₀ (T , A) (U , B) 94 | (I′ , (emp≲I′ , (I′⊨T , I′⊨A) , I′-med) , I′⊨U , I′⊨B) = 95 | (I , ((I′⊨T , I⊨A) , I-med) , I′⊨U , I′⊨B) where 96 | 97 | I : Interp Σ X 98 | I = enode * I′ 99 | 100 | I⊨A : I ⊨a A 101 | I⊨A = *-resp-⟨ABox⟩ enode I′ A I′⊨A 102 | 103 | I-med : ∀ J → (J ⊨ T , A) → Mediated₀ I J 104 | I-med J (J⊨T , J⊨A) = lemma (I′-med J′ emp≲J′ (J⊨T , J′⊨A)) where 105 | 106 | J′ : Interp Σ (False ⊕ False ⊕ X) 107 | J′ = enode⁻¹ * J 108 | 109 | emp≲J′ : emp ≲ inode * J′ 110 | emp≲J′ = emp-≲ (inode * J′) 111 | 112 | J′⊨A : J′ ⊨a ⟨ABox⟩ enode A 113 | J′⊨A = ⟨ABox⟩-resp-⊨ enode (λ x → ≈-refl ⌊ J ⌋) A J⊨A 114 | 115 | I≲J-impl-I′≲J′ : (I ≲ J) → (I′ ≲ J′) 116 | I≲J-impl-I′≲J′ I≲J = (≲⌊ I≲J ⌋ , i≲j) where 117 | 118 | i≲j : ∀ x → ⌊ J ⌋ ⊨ ≲-image ≲⌊ I≲J ⌋ (ind I′ x) ≈ ind J (enode⁻¹ x) 119 | i≲j (inode ()) 120 | i≲j (bnode ()) 121 | i≲j (enode x) = ≲-resp-ind I≲J x 122 | 123 | lemma : Mediated emp I′ J′ emp≲I′ emp≲J′ → Mediated₀ I J 124 | lemma (I≲J , _ , I≲J-uniq) = 125 | ( (≲⌊ I≲J ⌋ , λ x → ≲-resp-ind I≲J (enode x)) 126 | , λ I≲₁J I≲₂J x → I≲J-uniq 127 | (I≲J-impl-I′≲J′ I≲₁J) (I≲J-impl-I′≲J′ I≲₂J) (λ ()) (λ ()) x) 128 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Integrity/Closed/Alternate.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( _×_ ; _,_ ; swap ) 2 | open import Relation.Nullary using ( ¬_ ) 3 | open import Web.Semantic.DL.ABox using ( ABox ; ε ; _,_ ; _∼_ ; _∈₁_ ; _∈₂_ ) 4 | open import Web.Semantic.DL.ABox.Model using ( _⊨a_ ) 5 | open import Web.Semantic.DL.Concept using 6 | ( Concept ; ⟨_⟩ ; ¬⟨_⟩ ; ⊤ ; ⊥ ; _⊔_ ; _⊓_ ; ∀[_]_ ; ∃⟨_⟩_ ; ≤1 ; >1 ; neg ) 7 | open import Web.Semantic.DL.KB using ( KB ; tbox ; abox ) 8 | open import Web.Semantic.DL.KB.Model using ( _⊨_ ) 9 | open import Web.Semantic.DL.Role using ( Role ; ⟨_⟩ ; ⟨_⟩⁻¹ ) 10 | open import Web.Semantic.DL.Signature using ( Signature ) 11 | open import Web.Semantic.DL.TBox using 12 | ( TBox ; ε ; _,_ ; _⊑₁_ ; _⊑₂_ ; Dis ; Ref ; Irr ; Tra ) 13 | 14 | module Web.Semantic.DL.Integrity.Closed.Alternate 15 | {Σ : Signature} {X : Set} where 16 | 17 | -- An alternate definition of closed-world integrity. 18 | 19 | infixr 4 _,_ 20 | infix 2 _⊫_∈₂_ _⊫_∈₁_ _⊫a_ _⊫t_ 21 | 22 | data _⊫_∼_ (K : KB Σ X) (x y : X) : Set₁ where 23 | eq : (∀ I → (I ⊨ K) → (I ⊨a x ∼ y)) → (K ⊫ x ∼ y) 24 | 25 | data _⊫_∈₂_ (K : KB Σ X) (xy : X × X) : Role Σ → Set₁ where 26 | rel : ∀ {r} → (∀ I → (I ⊨ K) → (I ⊨a xy ∈₂ r)) → (K ⊫ xy ∈₂ ⟨ r ⟩) 27 | rev : ∀ {r} → (∀ I → (I ⊨ K) → (I ⊨a swap xy ∈₂ r)) → (K ⊫ xy ∈₂ ⟨ r ⟩⁻¹) 28 | 29 | data _⊫_∈₁_ (K : KB Σ X) (x : X) : Concept Σ → Set₁ where 30 | +atom : ∀ {c} → (∀ I → (I ⊨ K) → (I ⊨a x ∈₁ c)) → (K ⊫ x ∈₁ ⟨ c ⟩) 31 | -atom : ∀ {c} → ¬(∀ I → (I ⊨ K) → (I ⊨a x ∈₁ c)) → (K ⊫ x ∈₁ ¬⟨ c ⟩) 32 | top : (K ⊫ x ∈₁ ⊤) 33 | inj₁ : ∀ {C D} → (K ⊫ x ∈₁ C) → (K ⊫ x ∈₁ C ⊔ D) 34 | inj₂ : ∀ {C D} → (K ⊫ x ∈₁ D) → (K ⊫ x ∈₁ C ⊔ D) 35 | _,_ : ∀ {C D} → (K ⊫ x ∈₁ C) → (K ⊫ x ∈₁ D) → (K ⊫ x ∈₁ C ⊓ D) 36 | all : ∀ {R C} → (∀ y → (K ⊫ (x , y) ∈₂ R) → (K ⊫ y ∈₁ C)) → (K ⊫ x ∈₁ ∀[ R ] C) 37 | ex : ∀ {R C} y → (K ⊫ (x , y) ∈₂ R) → (K ⊫ y ∈₁ C) → (K ⊫ x ∈₁ ∃⟨ R ⟩ C) 38 | uniq : ∀ {R} → (∀ y z → (K ⊫ (x , y) ∈₂ R) → (K ⊫ (x , z) ∈₂ R) → (K ⊫ y ∼ z)) → (K ⊫ x ∈₁ ≤1 R) 39 | ¬uniq : ∀ {R} y z → (K ⊫ (x , y) ∈₂ R) → (K ⊫ (x , z) ∈₂ R) → ¬(K ⊫ y ∼ z) → (K ⊫ x ∈₁ >1 R) 40 | 41 | data _⊫a_ (K : KB Σ X) : ABox Σ X → Set₁ where 42 | ε : (K ⊫a ε) 43 | _,_ : ∀ {A B} → (K ⊫a A) → (K ⊫a B) → (K ⊫a A , B) 44 | eq : ∀ x y → (K ⊫ x ∼ y) → (K ⊫a x ∼ y) 45 | rl : ∀ xy r → (K ⊫ xy ∈₂ ⟨ r ⟩) → (K ⊫a xy ∈₂ r) 46 | cn : ∀ x c → (K ⊫ x ∈₁ ⟨ c ⟩) → (K ⊫a x ∈₁ c) 47 | 48 | data _⊫t_ (K : KB Σ X) : TBox Σ → Set₁ where 49 | ε : (K ⊫t ε) 50 | _,_ : ∀ {T U} → (K ⊫t T) → (K ⊫t U) → (K ⊫t T , U) 51 | rl : ∀ Q R → (∀ xy → (K ⊫ xy ∈₂ Q) → (K ⊫ xy ∈₂ R)) → (K ⊫t Q ⊑₂ R) 52 | cn : ∀ C D → (∀ x → (K ⊫ x ∈₁ neg C ⊔ D)) → (K ⊫t C ⊑₁ D) 53 | dis : ∀ Q R → (∀ xy → (K ⊫ xy ∈₂ Q) → ¬(K ⊫ xy ∈₂ R)) → (K ⊫t Dis Q R) 54 | ref : ∀ R → (∀ x → (K ⊫ (x , x) ∈₂ R)) → (K ⊫t Ref R) 55 | irr : ∀ R → (∀ x → ¬(K ⊫ (x , x) ∈₂ R)) → (K ⊫t Irr R) 56 | tra : ∀ R → (∀ x y z → (K ⊫ (x , y) ∈₂ R) → (K ⊫ (y , z) ∈₂ R) → (K ⊫ (x , z) ∈₂ R)) → (K ⊫t Tra R) 57 | 58 | data _⊫k_ (K L : KB Σ X) : Set₁ where 59 | _,_ : (K ⊫t tbox L) → (K ⊫a abox L) → (K ⊫k L) 60 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Integrity/Closed/Properties.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( _×_ ; _,_ ; swap ) 2 | open import Data.Sum using ( inj₁ ; inj₂ ) 3 | open import Relation.Nullary using ( ¬_ ; yes ; no ) 4 | open import Relation.Unary using ( _∈_ ) 5 | open import Web.Semantic.DL.ABox using ( ABox ; ε ; _,_ ; _∼_ ; _∈₁_ ; _∈₂_ ) 6 | open import Web.Semantic.DL.ABox.Interp using 7 | ( Interp ; ⌊_⌋ ; ind ; ind⁻¹ ; Surjective ; surj✓ ) 8 | open import Web.Semantic.DL.ABox.Interp.Morphism using 9 | ( _≲_ ; _≃_ ; _,_ ; ≲⌊_⌋ ; ≲-resp-ind ) 10 | open import Web.Semantic.DL.ABox.Interp.Meet using 11 | ( meet ; meet-lb ; meet-glb ; meet-uniq ; meet-surj ) 12 | open import Web.Semantic.DL.ABox.Model using ( _⊨a_ ) 13 | open import Web.Semantic.DL.Concept using 14 | ( Concept ; ⟨_⟩ ; ¬⟨_⟩ ; ⊤ ; ⊥ ; _⊔_ ; _⊓_ ; ∀[_]_ ; ∃⟨_⟩_ ; ≤1 ; >1 ; neg ) 15 | open import Web.Semantic.DL.Concept.Model using 16 | ( _⟦_⟧₁ ; neg-sound ; neg-complete ) 17 | open import Web.Semantic.DL.Integrity.Closed using 18 | ( Mediated₀ ; Initial₀ ; sur_⊨₀_ ; _,_ ) 19 | open import Web.Semantic.DL.Integrity.Closed.Alternate using 20 | ( _⊫_∼_ ; _⊫_∈₁_ ; _⊫_∈₂_ ; _⊫t_ ; _⊫a_ ; _⊫k_ 21 | ; eq ; rel ; rev ; +atom ; -atom ; top ; inj₁ ; inj₂ 22 | ; all ; ex ; uniq ; ¬uniq ; cn ; rl ; dis ; ref ; irr ; tra ; ε ; _,_ ) 23 | open import Web.Semantic.DL.KB using ( KB ; tbox ; abox ) 24 | open import Web.Semantic.DL.KB.Model using ( _⊨_ ; Interps ; ⊨-resp-≃ ) 25 | open import Web.Semantic.DL.Role using ( Role ; ⟨_⟩ ; ⟨_⟩⁻¹ ) 26 | open import Web.Semantic.DL.Role.Model using ( _⟦_⟧₂ ) 27 | open import Web.Semantic.DL.Signature using ( Signature ) 28 | open import Web.Semantic.DL.TBox using 29 | ( TBox ; ε ; _,_ ; _⊑₁_ ; _⊑₂_ ; Dis ; Ref ; Irr ; Tra ) 30 | open import Web.Semantic.DL.TBox.Interp using ( _⊨_≈_ ; ≈-sym ; ≈-trans ) 31 | open import Web.Semantic.DL.TBox.Interp.Morphism using ( ≲-resp-≈ ; iso ) 32 | open import Web.Semantic.DL.TBox.Model using ( _⊨t_ ) 33 | open import Web.Semantic.Util using 34 | ( ExclMiddle ; ExclMiddle₁ ; smaller-excl-middle ; is! ; is✓ ; tt ; elim ) 35 | 36 | module Web.Semantic.DL.Integrity.Closed.Properties 37 | (excl-middle₁ : ExclMiddle₁) {Σ : Signature} {X : Set} where 38 | 39 | -- The two definitions of closed-world integrity coincide for surjective interpretations. 40 | -- Note that this requires excluded middle, as the alternate definition assumes a classical logic, 41 | -- for example interpreting C ⊑ D as (¬ C ⊔ D) being a tautology. 42 | 43 | min : KB Σ X → Interp Σ X 44 | min K = meet excl-middle₁ (Interps K) 45 | 46 | excl-middle : ExclMiddle 47 | excl-middle = smaller-excl-middle excl-middle₁ 48 | 49 | sound∼ : ∀ K x y → (K ⊫ x ∼ y) → (⌊ min K ⌋ ⊨ x ≈ y) 50 | sound∼ K x y (eq K⊫x∼y) = is! K⊫x∼y 51 | 52 | complete∼ : ∀ K x y → (⌊ min K ⌋ ⊨ x ≈ y) → (K ⊫ x ∼ y) 53 | complete∼ K x y x≈y = eq (is✓ x≈y) 54 | 55 | sound₂ : ∀ K R xy → (K ⊫ xy ∈₂ R) → (xy ∈ ⌊ min K ⌋ ⟦ R ⟧₂) 56 | sound₂ K ⟨ r ⟩ (x , y) (rel K⊫xy∈r) = is! K⊫xy∈r 57 | sound₂ K ⟨ r ⟩⁻¹ (x , y) (rev K⊫yx∈r) = is! K⊫yx∈r 58 | 59 | complete₂ : ∀ K R xy → (xy ∈ ⌊ min K ⌋ ⟦ R ⟧₂) → (K ⊫ xy ∈₂ R) 60 | complete₂ K ⟨ r ⟩ (x , y) xy∈⟦r⟧ = rel (is✓ xy∈⟦r⟧) 61 | complete₂ K ⟨ r ⟩⁻¹ (x , y) yx∈⟦r⟧ = rev (is✓ yx∈⟦r⟧) 62 | 63 | sound₁ : ∀ K C x → (K ⊫ x ∈₁ C) → (x ∈ ⌊ min K ⌋ ⟦ C ⟧₁) 64 | sound₁ K ⟨ c ⟩ x (+atom K⊨x∈c) = 65 | is! K⊨x∈c 66 | sound₁ K ¬⟨ c ⟩ x (-atom K⊭x∈c) = 67 | λ x∈⟦c⟧ → K⊭x∈c (is✓ x∈⟦c⟧) 68 | sound₁ K ⊤ x top = tt 69 | sound₁ K ⊥ x () 70 | sound₁ K (C ⊓ D) x (K⊫x∈C , K⊫x∈D) = 71 | (sound₁ K C x K⊫x∈C , sound₁ K D x K⊫x∈D) 72 | sound₁ K (C ⊔ D) x (inj₁ K⊫x∈C) = 73 | inj₁ (sound₁ K C x K⊫x∈C) 74 | sound₁ K (C ⊔ D) x (inj₂ K⊫x∈D) = 75 | inj₂ (sound₁ K D x K⊫x∈D) 76 | sound₁ K (∀[ R ] C) x (all K⊫x∈∀RC) = 77 | λ y xy∈⟦R⟧ → sound₁ K C y (K⊫x∈∀RC y (complete₂ K R _ xy∈⟦R⟧)) 78 | sound₁ K (∃⟨ R ⟩ C) x (ex y K⊫xy∈R K⊫y∈C) = 79 | (y , sound₂ K R _ K⊫xy∈R , sound₁ K C y K⊫y∈C) 80 | sound₁ K (≤1 R) x (uniq K⊫x∈≤1R) = 81 | λ y z xy∈⟦R⟧ xz∈⟦R⟧ → sound∼ K y z 82 | (K⊫x∈≤1R y z (complete₂ K R _ xy∈⟦R⟧) (complete₂ K R _ xz∈⟦R⟧)) 83 | sound₁ K (>1 R) x (¬uniq y z K⊫xy∈R K⊫xz∈R K⊯y∼z) = 84 | ( y , z , sound₂ K R _ K⊫xy∈R , sound₂ K R _ K⊫xz∈R 85 | , λ y≈z → K⊯y∼z (complete∼ K y z y≈z) ) 86 | 87 | complete₁ : ∀ K C x → (x ∈ ⌊ min K ⌋ ⟦ C ⟧₁) → (K ⊫ x ∈₁ C) 88 | complete₁ K ⟨ c ⟩ x x∈⟦c⟧ = 89 | +atom (is✓ x∈⟦c⟧) 90 | complete₁ K ¬⟨ c ⟩ x x∉⟦c⟧ = 91 | -atom (λ K⊫x∈c → x∉⟦c⟧ (is! K⊫x∈c)) 92 | complete₁ K ⊤ x x∈⟦C⟧ = 93 | top 94 | complete₁ K ⊥ x () 95 | complete₁ K (C ⊓ D) x (x∈⟦C⟧ , x∈⟦D⟧) = 96 | (complete₁ K C x x∈⟦C⟧ , complete₁ K D x x∈⟦D⟧) 97 | complete₁ K (C ⊔ D) x (inj₁ x∈⟦C⟧) = 98 | inj₁ (complete₁ K C x x∈⟦C⟧) 99 | complete₁ K (C ⊔ D) x (inj₂ x∈⟦D⟧) = 100 | inj₂ (complete₁ K D x x∈⟦D⟧) 101 | complete₁ K (∀[ R ] C) x x∈⟦∀RC⟧ = 102 | all (λ y K⊫xy∈R → complete₁ K C y (x∈⟦∀RC⟧ y (sound₂ K R _ K⊫xy∈R))) 103 | complete₁ K (∃⟨ R ⟩ C) x (y , xy∈⟦R⟧ , y∈⟦C⟧) = 104 | ex y (complete₂ K R (x , y) xy∈⟦R⟧) (complete₁ K C y y∈⟦C⟧) 105 | complete₁ K (≤1 R) x x∈⟦≤1R⟧ = 106 | uniq (λ y z K⊫xy∈R K⊫xz∈R → complete∼ K y z 107 | (x∈⟦≤1R⟧ y z (sound₂ K R _ K⊫xy∈R) (sound₂ K R _ K⊫xz∈R))) 108 | complete₁ K (>1 R) x (y , z , xy∈⟦R⟧ , xz∈⟦R⟧ , y≉z) = 109 | ¬uniq y z (complete₂ K R _ xy∈⟦R⟧) (complete₂ K R _ xz∈⟦R⟧) 110 | (λ K⊫y∼z → y≉z (sound∼ K y z K⊫y∼z)) 111 | 112 | ⊫-impl-min⊨ : ∀ K L → (K ⊫k L) → (min K ⊨ L) 113 | ⊫-impl-min⊨ K L (K⊫T , K⊫A) = ( J⊨T K⊫T , J⊨A K⊫A ) where 114 | 115 | J : Interp Σ X 116 | J = min K 117 | 118 | J⊨T : ∀ {T} → (K ⊫t T) → (⌊ J ⌋ ⊨t T) 119 | J⊨T ε = tt 120 | J⊨T (K⊫T , K⊫U) = (J⊨T K⊫T , J⊨T K⊫U) 121 | J⊨T (rl Q R K⊫Q⊑R) = 122 | λ {xy} xy∈⟦Q⟧ → sound₂ K R xy (K⊫Q⊑R xy (complete₂ K Q xy xy∈⟦Q⟧)) 123 | J⊨T (cn C D K⊫C⊑D) = λ {x} → lemma x (K⊫C⊑D x) where 124 | 125 | lemma : ∀ x → (K ⊫ x ∈₁ (neg C ⊔ D)) → 126 | (x ∈ ⌊ J ⌋ ⟦ C ⟧₁) → (x ∈ ⌊ J ⌋ ⟦ D ⟧₁) 127 | lemma x (inj₁ K⊫x∈¬C) x∈⟦C⟧ = 128 | elim (neg-sound ⌊ J ⌋ {x} C (sound₁ K (neg C) x K⊫x∈¬C) x∈⟦C⟧) 129 | lemma x (inj₂ K⊫x∈D) x∈⟦C⟧ = 130 | sound₁ K D x K⊫x∈D 131 | 132 | J⊨T (dis Q R K⊫DisQR) = λ {xy} xy∈⟦Q⟧ xy∈⟦R⟧ → 133 | K⊫DisQR xy (complete₂ K Q xy xy∈⟦Q⟧) (complete₂ K R xy xy∈⟦R⟧) 134 | J⊨T (ref R K⊫RefR) = λ x → sound₂ K R (x , x) (K⊫RefR x) 135 | J⊨T (irr R K⊫IrrR) = λ x xx∈⟦R⟧ → K⊫IrrR x (complete₂ K R (x , x) xx∈⟦R⟧) 136 | J⊨T (tra R K⊫TraR) = λ {x} {y} {z} xy∈⟦R⟧ yz∈⟦R⟧ → 137 | sound₂ K R (x , z) (K⊫TraR x y z 138 | (complete₂ K R (x , y) xy∈⟦R⟧) (complete₂ K R (y , z) yz∈⟦R⟧)) 139 | 140 | J⊨A : ∀ {A} → (K ⊫a A) → (J ⊨a A) 141 | J⊨A ε = tt 142 | J⊨A (K⊫A , K⊫B) = (J⊨A K⊫A , J⊨A K⊫B) 143 | J⊨A (eq x y K⊫x∼y) = sound∼ K x y K⊫x∼y 144 | J⊨A (rl (x , y) r K⊫xy∈r) = sound₂ K ⟨ r ⟩ (x , y) K⊫xy∈r 145 | J⊨A (cn x c K⊫x∈c) = sound₁ K ⟨ c ⟩ x K⊫x∈c 146 | 147 | min⊨-impl-⊫ : ∀ K L → (min K ⊨ L) → (K ⊫k L) 148 | min⊨-impl-⊫ K L (J⊨T , J⊨A) = 149 | ( K⊫T (tbox L) J⊨T , K⊫A (abox L) J⊨A ) where 150 | 151 | J : Interp Σ X 152 | J = min K 153 | 154 | K⊫T : ∀ T → (⌊ J ⌋ ⊨t T) → (K ⊫t T) 155 | K⊫T ε J⊨ε = ε 156 | K⊫T (T , U) (J⊨T , J⊨U) = (K⊫T T J⊨T , K⊫T U J⊨U) 157 | K⊫T (Q ⊑₂ R) J⊨Q⊑R = 158 | rl Q R (λ xy K⊫xy∈Q → complete₂ K R xy (J⊨Q⊑R (sound₂ K Q xy K⊫xy∈Q))) 159 | K⊫T (C ⊑₁ D) J⊨C⊑D = cn C D lemma where 160 | 161 | lemma : ∀ x → (K ⊫ x ∈₁ neg C ⊔ D) 162 | lemma x with excl-middle (x ∈ ⌊ J ⌋ ⟦ C ⟧₁) 163 | lemma x | yes x∈⟦C⟧ = 164 | inj₂ (complete₁ K D x (J⊨C⊑D x∈⟦C⟧)) 165 | lemma x | no x∉⟦C⟧ = 166 | inj₁ (complete₁ K (neg C) x (neg-complete excl-middle ⌊ J ⌋ C x∉⟦C⟧)) 167 | 168 | K⊫T (Dis Q R) J⊨DisQR = 169 | dis Q R (λ xy K⊫xy∈Q K⊫xy∈R → 170 | J⊨DisQR (sound₂ K Q xy K⊫xy∈Q) (sound₂ K R xy K⊫xy∈R)) 171 | K⊫T (Ref R) J⊨RefR = 172 | ref R (λ x → complete₂ K R (x , x) (J⊨RefR x)) 173 | K⊫T (Irr R) J⊨IrrR = 174 | irr R (λ x K⊫xx∈R → J⊨IrrR x (sound₂ K R (x , x) K⊫xx∈R)) 175 | K⊫T (Tra R) J⊨TrR = 176 | tra R (λ x y z K⊫xy∈R K⊫yz∈R → 177 | complete₂ K R (x , z) (J⊨TrR 178 | (sound₂ K R (x , y) K⊫xy∈R) (sound₂ K R (y , z) K⊫yz∈R))) 179 | 180 | K⊫A : ∀ A → (J ⊨a A) → (K ⊫a A) 181 | K⊫A ε J⊨ε = ε 182 | K⊫A (A , B) (J⊨A , J⊨B) = (K⊫A A J⊨A , K⊫A B J⊨B) 183 | K⊫A (x ∼ y) x≈y = eq x y (complete∼ K x y x≈y) 184 | K⊫A (x ∈₁ c) x∈⟦c⟧ = cn x c (complete₁ K ⟨ c ⟩ x x∈⟦c⟧) 185 | K⊫A ((x , y) ∈₂ r) xy∈⟦r⟧ = rl (x , y) r (complete₂ K ⟨ r ⟩ (x , y) xy∈⟦r⟧) 186 | 187 | min-med : ∀ (K : KB Σ X) J → (J ⊨ K) → Mediated₀ (min K) J 188 | min-med K J J⊨K = (meet-lb excl-middle₁ (Interps K) J J⊨K , meet-uniq excl-middle₁ (Interps K) J J⊨K) 189 | 190 | min-init : ∀ (K : KB Σ X) → (K ⊫k K) → (min K ∈ Initial₀ K) 191 | min-init K K⊫K = ( ⊫-impl-min⊨ K K K⊫K , min-med K) 192 | 193 | min-uniq : ∀ (I : Interp Σ X) (K : KB Σ X) → (I ∈ Surjective) → (I ∈ Initial₀ K) → (I ≃ min K) 194 | min-uniq I K I∈Surj (I⊨K , I-med) = 195 | ( iso 196 | ≲⌊ meet-glb excl-middle₁ (Interps K) I I∈Surj lemma₁ ⌋ 197 | ≲⌊ meet-lb excl-middle₁ (Interps K) I I⊨K ⌋ 198 | (λ x → ≈-sym ⌊ I ⌋ (surj✓ I∈Surj x)) 199 | (λ x → is! (lemma₂ x)) 200 | , λ x → is! (lemma₂ x)) where 201 | 202 | lemma₁ : ∀ J J⊨K → I ≲ J 203 | lemma₁ J J⊨K with I-med J J⊨K 204 | lemma₁ J J⊨K | (I≲J , I≲J-uniq) = I≲J 205 | 206 | lemma₂ : ∀ x J J⊨K → ⌊ J ⌋ ⊨ ind J (ind⁻¹ I∈Surj (ind I x)) ≈ ind J x 207 | lemma₂ x J J⊨K = 208 | ≈-trans ⌊ J ⌋ (≈-sym ⌊ J ⌋ (≲-resp-ind (lemma₁ J J⊨K) (ind⁻¹ I∈Surj (ind I x)))) 209 | (≈-trans ⌊ J ⌋ (≲-resp-≈ ≲⌊ lemma₁ J J⊨K ⌋ (≈-sym ⌊ I ⌋ (surj✓ I∈Surj (ind I x)))) 210 | (≲-resp-ind (lemma₁ J J⊨K) x)) 211 | 212 | ⊫-impl-⊨₀ : ∀ (KB₁ KB₂ : KB Σ X) → (KB₁ ⊫k KB₁) → (KB₁ ⊫k KB₂) → (sur KB₁ ⊨₀ KB₂) 213 | ⊫-impl-⊨₀ KB₁ KB₂ KB₁⊫KB₁ KB₁⊫KB₂ = 214 | ( min KB₁ 215 | , meet-surj excl-middle₁ (Interps KB₁) 216 | , min-init KB₁ KB₁⊫KB₁ 217 | , ⊫-impl-min⊨ KB₁ KB₂ KB₁⊫KB₂ ) 218 | 219 | ⊨₀-impl-⊫₁ : ∀ (KB₁ KB₂ : KB Σ X) → (sur KB₁ ⊨₀ KB₂) → (KB₁ ⊫k KB₁) 220 | ⊨₀-impl-⊫₁ KB₁ KB₂ (I , I∈Surj , (I⊨KB₁ , I-med) , I⊨KB₂) = 221 | min⊨-impl-⊫ KB₁ KB₁ (⊨-resp-≃ (min-uniq I KB₁ I∈Surj (I⊨KB₁ , I-med)) KB₁ I⊨KB₁) 222 | 223 | ⊨₀-impl-⊫₂ : ∀ (KB₁ KB₂ : KB Σ X) → (sur KB₁ ⊨₀ KB₂) → (KB₁ ⊫k KB₂) 224 | ⊨₀-impl-⊫₂ KB₁ KB₂ (I , I∈Surj , I-init , I⊨KB₂) = 225 | min⊨-impl-⊫ KB₁ KB₂ (⊨-resp-≃ (min-uniq I KB₁ I∈Surj I-init) KB₂ I⊨KB₂) 226 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/KB.agda: -------------------------------------------------------------------------------- 1 | open import Web.Semantic.DL.ABox using ( ABox ) 2 | open import Web.Semantic.DL.Signature using ( Signature ) 3 | open import Web.Semantic.DL.TBox using ( TBox ) 4 | 5 | module Web.Semantic.DL.KB where 6 | 7 | infixr 4 _,_ 8 | 9 | data KB (Σ : Signature) (X : Set) : Set where 10 | _,_ : (T : TBox Σ) → (A : ABox Σ X) → KB Σ X 11 | 12 | tbox : ∀ {Σ X} → KB Σ X → TBox Σ 13 | tbox (T , A) = T 14 | 15 | abox : ∀ {Σ X} → KB Σ X → ABox Σ X 16 | abox (T , A) = A 17 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/KB/Model.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( _×_ ; _,_ ) 2 | open import Web.Semantic.DL.ABox.Model using ( _⊨a_ ; ⊨a-resp-≲ ) 3 | open import Web.Semantic.DL.ABox.Interp using ( Interp ; ⌊_⌋ ) 4 | open import Web.Semantic.DL.ABox.Interp.Morphism using ( _≃_ ; ≃⌊_⌋ ; ≃-impl-≲ ) 5 | open import Web.Semantic.DL.KB using ( KB ; tbox ; abox ) 6 | open import Web.Semantic.DL.Signature using ( Signature ) 7 | open import Web.Semantic.DL.TBox.Model using ( _⊨t_ ; ⊨t-resp-≃ ) 8 | 9 | module Web.Semantic.DL.KB.Model {Σ : Signature} {X : Set} where 10 | 11 | infixr 2 _⊨_ 12 | 13 | _⊨_ : Interp Σ X → KB Σ X → Set 14 | I ⊨ K = (⌊ I ⌋ ⊨t tbox K) × (I ⊨a abox K) 15 | 16 | Interps : KB Σ X → Interp Σ X → Set 17 | Interps K I = I ⊨ K 18 | 19 | ⊨-resp-≃ : ∀ {I J} → (I ≃ J) → ∀ K → (I ⊨ K) → (J ⊨ K) 20 | ⊨-resp-≃ I≃J K (I⊨T , I⊨A) = 21 | (⊨t-resp-≃ ≃⌊ I≃J ⌋ (tbox K) I⊨T , ⊨a-resp-≲ (≃-impl-≲ I≃J) (abox K) I⊨A) 22 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/KB/Skolemization.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( _×_ ; _,_ ) 2 | open import Relation.Unary using ( _∈_ ) 3 | open import Web.Semantic.DL.ABox.Interp using ( ⌊_⌋ ; ind ) 4 | open import Web.Semantic.DL.ABox.Skolemization using ( askolem ; askolem-sound ) 5 | open import Web.Semantic.DL.FOL using ( Formula ; _∧_ ) 6 | open import Web.Semantic.DL.FOL.Model using ( _⊨f_ ) 7 | open import Web.Semantic.DL.KB using ( KB ; tbox ; abox ) 8 | open import Web.Semantic.DL.KB.Model using ( _⊨_ ) 9 | open import Web.Semantic.DL.Signature using ( Signature ) 10 | open import Web.Semantic.DL.TBox.Skolemization using ( TSkolems ; tskolem ; tskolem-sound ) 11 | 12 | module Web.Semantic.DL.KB.Skolemization {Σ : Signature} {X : Set} where 13 | 14 | Skolems : Set → KB Σ X → Set 15 | Skolems Δ K = TSkolems Δ (tbox K) 16 | 17 | skolem : ∀ {Δ} → (X → Δ) → ∀ K → (Skolems Δ K) → Formula Σ Δ 18 | skolem i K Φ = tskolem (tbox K) Φ ∧ askolem i (abox K) 19 | 20 | skolem-sound : ∀ I K Φ → (⌊ I ⌋ ⊨f skolem (ind I) K Φ) → (I ⊨ K) 21 | skolem-sound I K Φ (I⊨T , I⊨A) = (tskolem-sound ⌊ I ⌋ (tbox K) Φ I⊨T , askolem-sound I (abox K) I⊨A) -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Role.agda: -------------------------------------------------------------------------------- 1 | open import Relation.Unary using ( ∅ ; _∪_ ) 2 | open import Web.Semantic.DL.Signature using ( Signature ; CN ; RN ) 3 | open import Web.Semantic.Util using ( Subset ; ⁅_⁆ ) 4 | 5 | module Web.Semantic.DL.Role where 6 | 7 | data Role (Σ : Signature) : Set where 8 | ⟨_⟩ : (r : RN Σ) → Role Σ 9 | ⟨_⟩⁻¹ : (r : RN Σ) → Role Σ 10 | 11 | inv : ∀ {Σ} → Role Σ → Role Σ 12 | inv ⟨ r ⟩ = ⟨ r ⟩⁻¹ 13 | inv ⟨ r ⟩⁻¹ = ⟨ r ⟩ 14 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Role/Model.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( _×_ ; _,_ ) 2 | open import Relation.Unary using ( _∈_ ) 3 | open import Web.Semantic.DL.Role using ( Role ; ⟨_⟩ ; ⟨_⟩⁻¹ ) 4 | open import Web.Semantic.DL.Signature using ( Signature ) 5 | open import Web.Semantic.DL.TBox.Interp using ( Interp ; Δ ; rol ; _⊨_≈_ ; rol-≈ ; ≈-refl ; ≈-sym ) 6 | open import Web.Semantic.DL.TBox.Interp.Morphism using ( _≲_ ; _≃_ ; ≲-resp-rol ; ≲-image² ; ≃-image ; ≃-image⁻¹ ; ≃-image² ; ≃-image²⁻¹ ; ≃-impl-≲ ; ≃-iso⁻¹ ) 7 | open import Web.Semantic.Util using ( Subset ; _⁻¹ ) 8 | 9 | module Web.Semantic.DL.Role.Model {Σ : Signature} where 10 | 11 | _⟦_⟧₂ : ∀ (I : Interp Σ) → Role Σ → Subset (Δ I × Δ I) 12 | I ⟦ ⟨ r ⟩ ⟧₂ = rol I r 13 | I ⟦ ⟨ r ⟩⁻¹ ⟧₂ = (rol I r)⁻¹ 14 | 15 | ⟦⟧₂-resp-≈ : ∀ I R {w x y z} → 16 | (I ⊨ w ≈ x) → ((x , y) ∈ I ⟦ R ⟧₂) → (I ⊨ y ≈ z) → ((w , z) ∈ I ⟦ R ⟧₂) 17 | ⟦⟧₂-resp-≈ I ⟨ r ⟩ w≈x xy∈⟦r⟧ y≈z = rol-≈ I r w≈x xy∈⟦r⟧ y≈z 18 | ⟦⟧₂-resp-≈ I ⟨ r ⟩⁻¹ w≈x yx∈⟦r⟧ y≈z = rol-≈ I r (≈-sym I y≈z) yx∈⟦r⟧ (≈-sym I w≈x) 19 | 20 | ⟦⟧₂-resp-≲ : ∀ {I J : Interp Σ} → (I≲J : I ≲ J) → ∀ {xy} R → 21 | (xy ∈ I ⟦ R ⟧₂) → (≲-image² I≲J xy ∈ J ⟦ R ⟧₂) 22 | ⟦⟧₂-resp-≲ I≲J {x , y} ⟨ r ⟩ xy∈⟦r⟧ = ≲-resp-rol I≲J xy∈⟦r⟧ 23 | ⟦⟧₂-resp-≲ I≲J {x , y} ⟨ r ⟩⁻¹ yx∈⟦r⟧ = ≲-resp-rol I≲J yx∈⟦r⟧ 24 | 25 | ⟦⟧₂-resp-≃ : ∀ {I J : Interp Σ} → (I≃J : I ≃ J) → ∀ {xy} R → 26 | (xy ∈ I ⟦ R ⟧₂) → (≃-image² I≃J xy ∈ J ⟦ R ⟧₂) 27 | ⟦⟧₂-resp-≃ I≃J = ⟦⟧₂-resp-≲ (≃-impl-≲ I≃J) 28 | 29 | ⟦⟧₂-refl-≃ : ∀ {I J : Interp Σ} → (I≃J : I ≃ J) → ∀ {xy} R → 30 | (≃-image²⁻¹ I≃J xy ∈ I ⟦ R ⟧₂) → (xy ∈ J ⟦ R ⟧₂) 31 | ⟦⟧₂-refl-≃ {I} {J} I≃J {x , y} R xy∈⟦R⟧ = 32 | ⟦⟧₂-resp-≈ J R (≈-sym J (≃-iso⁻¹ I≃J x)) 33 | (⟦⟧₂-resp-≃ I≃J R xy∈⟦R⟧) (≃-iso⁻¹ I≃J y) 34 | 35 | ⟦⟧₂-galois : ∀ {I J : Interp Σ} → (I≃J : I ≃ J) → ∀ {x y} R → 36 | ((≃-image⁻¹ I≃J x , y) ∈ I ⟦ R ⟧₂) → ((x , ≃-image I≃J y) ∈ J ⟦ R ⟧₂) 37 | ⟦⟧₂-galois {I} {J} I≃J {x} R xy∈⟦R⟧ = 38 | ⟦⟧₂-resp-≈ J R (≈-sym J (≃-iso⁻¹ I≃J x)) 39 | (⟦⟧₂-resp-≲ (≃-impl-≲ I≃J) R xy∈⟦R⟧) 40 | (≈-refl J) 41 | 42 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Role/Skolemization.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( _×_ ; _,_ ) 2 | open import Relation.Unary using ( _∈_ ) 3 | open import Web.Semantic.DL.FOL using ( Formula ; _∈₂_ ; _∈₂_⇒_ ) 4 | open import Web.Semantic.DL.FOL.Model using ( _⊨f_ ) 5 | open import Web.Semantic.DL.Role using ( Role ; ⟨_⟩ ; ⟨_⟩⁻¹ ) 6 | open import Web.Semantic.DL.Role.Model using ( _⟦_⟧₂ ; ⟦⟧₂-resp-≈ ) 7 | open import Web.Semantic.DL.Signature using ( Signature ) 8 | open import Web.Semantic.DL.TBox.Interp using ( Interp ; ≈-sym ) 9 | 10 | module Web.Semantic.DL.Role.Skolemization {Σ : Signature} where 11 | 12 | rskolem : ∀ {Δ} → Role Σ → Δ → Δ → Formula Σ Δ 13 | rskolem ⟨ r ⟩ x y = (x , y) ∈₂ r 14 | rskolem ⟨ r ⟩⁻¹ x y = (y , x) ∈₂ r 15 | 16 | rskolem⇒ : ∀ {Δ} → Role Σ → Δ → Δ → Formula Σ Δ → Formula Σ Δ 17 | rskolem⇒ ⟨ r ⟩ x y F = (x , y) ∈₂ r ⇒ F 18 | rskolem⇒ ⟨ r ⟩⁻¹ x y F = (y , x) ∈₂ r ⇒ F 19 | 20 | rskolem-sound : ∀ I R x y → (I ⊨f rskolem R x y) → ((x , y) ∈ I ⟦ R ⟧₂) 21 | rskolem-sound I ⟨ r ⟩ x y xy∈⟦r⟧ = xy∈⟦r⟧ 22 | rskolem-sound I ⟨ r ⟩⁻¹ x y yx∈⟦r⟧ = yx∈⟦r⟧ 23 | 24 | rskolem⇒-sound : ∀ I R x y F → (I ⊨f rskolem⇒ R x y F) → ((x , y) ∈ I ⟦ R ⟧₂) → (I ⊨f F) 25 | rskolem⇒-sound I ⟨ r ⟩ x y F xy∈r⇒F xy∈⟦r⟧ = xy∈r⇒F xy∈⟦r⟧ 26 | rskolem⇒-sound I ⟨ r ⟩⁻¹ x y F yx∈r⇒F yx∈⟦r⟧ = yx∈r⇒F yx∈⟦r⟧ 27 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Sequent.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( _×_ ; _,_ ) 2 | open import Relation.Nullary using ( ¬_ ) 3 | open import Relation.Unary using ( _∈_ ; ∅ ; _∪_ ) 4 | open import Web.Semantic.DL.ABox using 5 | ( ABox ; Assertions ; ε ; _,_ ; _∼_ ; _∈₁_ ; _∈₂_ ) 6 | open import Web.Semantic.DL.ABox.Interp using ( Interp ; ⌊_⌋ ; ind ) 7 | open import Web.Semantic.DL.Signature using ( Signature ) 8 | open import Web.Semantic.DL.Concept using 9 | ( Concept ; ⟨_⟩ ; ⊤ ; ⊥ ; _⊓_ ; _⊔_ ; ∀[_]_ ; ∃⟨_⟩_ ; ≤1 ) 10 | open import Web.Semantic.DL.KB using ( KB ; tbox ; abox ) 11 | open import Web.Semantic.DL.Role using ( Role ; ⟨_⟩ ; ⟨_⟩⁻¹ ) 12 | open import Web.Semantic.DL.TBox using 13 | ( TBox ; Axioms ; ε ; _,_ ;_⊑₁_ ; _⊑₂_ ; Ref ; Tra ) 14 | open import Web.Semantic.DL.TBox.Interp using ( Δ ; _⊨_≈_ ; con ; rol ) 15 | open import Web.Semantic.Util using ( _⊕_⊕_ ; inode ; bnode ; enode ) 16 | 17 | module Web.Semantic.DL.Sequent {Σ : Signature} {W X Y : Set} where 18 | 19 | infix 2 _⊕_⊢_∼_ _⊕_⊢_∈₁_ _⊕_⊢_∈₂_ 20 | 21 | Γ : (Interp Σ X) → (KB Σ (X ⊕ W ⊕ Y)) → Set 22 | Γ I K = Δ ⌊ I ⌋ ⊕ W ⊕ Y 23 | 24 | γ : ∀ I K → (X ⊕ W ⊕ Y) → Γ I K 25 | γ I K (inode x) = inode (ind I x) 26 | γ I K (bnode w) = bnode w 27 | γ I K (enode y) = enode y 28 | 29 | mutual 30 | 31 | data _⊕_⊢_∼_ I K : Γ I K → Γ I K → Set where 32 | ∼-assert : ∀ {x y} → ((x ∼ y) ∈ Assertions (abox K)) → (I ⊕ K ⊢ γ I K x ∼ γ I K y) 33 | ∼-import : ∀ {x y} → (⌊ I ⌋ ⊨ x ≈ y) → (I ⊕ K ⊢ inode x ∼ inode y) 34 | ∼-refl : ∀ {x} → (I ⊕ K ⊢ x ∼ x) 35 | ∼-sym : ∀ {x y} → (I ⊕ K ⊢ x ∼ y) → (I ⊕ K ⊢ y ∼ x) 36 | ∼-trans : ∀ {x y z} → (I ⊕ K ⊢ x ∼ y) → (I ⊕ K ⊢ y ∼ z) → (I ⊕ K ⊢ x ∼ z) 37 | ∼-≤1 : ∀ {x y z R} → (I ⊕ K ⊢ x ∈₁ (≤1 R)) → (I ⊕ K ⊢ (x , y) ∈₂ R) → (I ⊕ K ⊢ (x , z) ∈₂ R) → (I ⊕ K ⊢ y ∼ z) 38 | 39 | data _⊕_⊢_∈₁_ I K : Γ I K → Concept Σ → Set where 40 | ∈₁-assert : ∀ {x c} → ((x ∈₁ c) ∈ Assertions (abox K)) → (I ⊕ K ⊢ γ I K x ∈₁ ⟨ c ⟩) 41 | ∈₁-import : ∀ {x c} → (x ∈ con ⌊ I ⌋ c) → (I ⊕ K ⊢ inode x ∈₁ ⟨ c ⟩) 42 | ∈₁-resp-∼ : ∀ {x y C} → (I ⊕ K ⊢ x ∈₁ C) → (I ⊕ K ⊢ x ∼ y) → (I ⊕ K ⊢ y ∈₁ C) 43 | ∈₁-subsum : ∀ {x C D} → (I ⊕ K ⊢ x ∈₁ C) → ((C ⊑₁ D) ∈ Axioms (tbox K)) → (I ⊕ K ⊢ x ∈₁ D) 44 | ∈₁-⊤-I : ∀ {x} → (I ⊕ K ⊢ x ∈₁ ⊤) 45 | ∈₁-⊓-I : ∀ {x C D} → (I ⊕ K ⊢ x ∈₁ C) → (I ⊕ K ⊢ x ∈₁ D) → (I ⊕ K ⊢ x ∈₁ (C ⊓ D)) 46 | ∈₁-⊓-E₁ : ∀ {x C D} → (I ⊕ K ⊢ x ∈₁ (C ⊓ D)) → (I ⊕ K ⊢ x ∈₁ C) 47 | ∈₁-⊓-E₂ : ∀ {x C D} → (I ⊕ K ⊢ x ∈₁ (C ⊓ D)) → (I ⊕ K ⊢ x ∈₁ D) 48 | ∈₁-⊔-I₁ : ∀ {x C D} → (I ⊕ K ⊢ x ∈₁ C) → (I ⊕ K ⊢ x ∈₁ (C ⊔ D)) 49 | ∈₁-⊔-I₂ : ∀ {x C D} → (I ⊕ K ⊢ x ∈₁ D) → (I ⊕ K ⊢ x ∈₁ (C ⊔ D)) 50 | ∈₁-∀-E : ∀ {x y R C} → (I ⊕ K ⊢ x ∈₁ (∀[ R ] C)) → (I ⊕ K ⊢ (x , y) ∈₂ R) → (I ⊕ K ⊢ y ∈₁ C) 51 | ∈₁-∃-I : ∀ {x y R C} → (I ⊕ K ⊢ (x , y) ∈₂ R) → (I ⊕ K ⊢ y ∈₁ C) → (I ⊕ K ⊢ x ∈₁ (∃⟨ R ⟩ C)) 52 | 53 | data _⊕_⊢_∈₂_ I K : (Γ I K × Γ I K) → Role Σ → Set where 54 | ∈₂-assert : ∀ {x y r} → (((x , y) ∈₂ r) ∈ Assertions (abox K)) → (I ⊕ K ⊢ (γ I K x , γ I K y) ∈₂ ⟨ r ⟩) 55 | ∈₂-import : ∀ {x y r} → ((x , y) ∈ rol ⌊ I ⌋ r) → (I ⊕ K ⊢ (inode x , inode y) ∈₂ ⟨ r ⟩) 56 | ∈₂-resp-∼ : ∀ {w x y z R} → (I ⊕ K ⊢ w ∼ x) → (I ⊕ K ⊢ (x , y) ∈₂ R) → (I ⊕ K ⊢ y ∼ z) → (I ⊕ K ⊢ (w , z) ∈₂ R) 57 | ∈₂-subsum : ∀ {xy R S} → (I ⊕ K ⊢ xy ∈₂ R) → ((R ⊑₂ S) ∈ Axioms (tbox K)) → (I ⊕ K ⊢ xy ∈₂ S) 58 | ∈₂-refl : ∀ x {R} → ((Ref R) ∈ Axioms (tbox K)) → (I ⊕ K ⊢ (x , x) ∈₂ R) 59 | ∈₂-trans : ∀ {x y z R} → (I ⊕ K ⊢ (x , y) ∈₂ R) → (I ⊕ K ⊢ (y , z) ∈₂ R) → ((Tra R) ∈ Axioms (tbox K)) → (I ⊕ K ⊢ (x , z) ∈₂ R) 60 | ∈₂-inv-I : ∀ {x y r} → (I ⊕ K ⊢ (x , y) ∈₂ ⟨ r ⟩) → (I ⊕ K ⊢ (y , x) ∈₂ ⟨ r ⟩⁻¹) 61 | ∈₂-inv-E : ∀ {x y r} → (I ⊕ K ⊢ (x , y) ∈₂ ⟨ r ⟩⁻¹) → (I ⊕ K ⊢ (y , x) ∈₂ ⟨ r ⟩) 62 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Sequent/Model.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( _×_ ; _,_ ; proj₁ ; proj₂ ) 2 | open import Data.Sum using ( inj₁ ; inj₂ ) 3 | open import Data.Unit using ( tt ) 4 | open import Relation.Binary.PropositionalEquality using ( refl ) 5 | open import Relation.Unary using ( _∈_ ; _⊆_ ) 6 | open import Web.Semantic.DL.ABox using 7 | ( ABox ; Assertions ; ε ; _,_ ; _∼_ ; _∈₁_ ; _∈₂_ ) 8 | open import Web.Semantic.DL.ABox.Interp using 9 | ( Interp ; _,_ ; ⌊_⌋ ; ind ; _*_ ) 10 | open import Web.Semantic.DL.ABox.Interp.Morphism using 11 | ( _≲_ ; _,_ ; ≲⌊_⌋ ; ≲-resp-ind ; _≋_ ) 12 | open import Web.Semantic.DL.ABox.Model using 13 | ( _⟦_⟧₀ ; _⊨a_ ; Assertions✓ ) 14 | open import Web.Semantic.DL.Concept using 15 | ( ⟨_⟩ ; ¬⟨_⟩ ; ⊤ ; ⊥ ; _⊓_ ; _⊔_ ; ∀[_]_ ; ∃⟨_⟩_ ; ≤1 ; >1 ) 16 | open import Web.Semantic.DL.Concept.Model using ( _⟦_⟧₁ ; ⟦⟧₁-resp-≈ ) 17 | open import Web.Semantic.DL.Integrity using 18 | ( Unique ; Mediator ; Initial ; _>>_ ; _⊕_⊨_ ; _,_ ) 19 | open import Web.Semantic.DL.KB using ( KB ; tbox ; abox ) 20 | open import Web.Semantic.DL.KB.Model using ( _⊨_ ) 21 | open import Web.Semantic.DL.Role using ( ⟨_⟩ ; ⟨_⟩⁻¹ ) 22 | open import Web.Semantic.DL.Role.Model using ( _⟦_⟧₂ ; ⟦⟧₂-resp-≈ ) 23 | open import Web.Semantic.DL.Sequent using 24 | ( Γ ; γ ; _⊕_⊢_∼_ ; _⊕_⊢_∈₁_ ; _⊕_⊢_∈₂_ 25 | ; ∼-assert ; ∼-import ;∼-refl ; ∼-sym ; ∼-trans ; ∼-≤1 26 | ; ∈₂-assert ; ∈₂-import ; ∈₂-resp-∼ ; ∈₂-subsum 27 | ; ∈₂-refl ; ∈₂-trans ; ∈₂-inv-I ; ∈₂-inv-E 28 | ; ∈₁-assert ; ∈₁-import ; ∈₁-resp-∼ ; ∈₁-subsum 29 | ; ∈₁-⊤-I ; ∈₁-⊓-I ; ∈₁-⊓-E₁ ; ∈₁-⊓-E₂ 30 | ; ∈₁-⊔-I₁ ; ∈₁-⊔-I₂ ; ∈₁-∃-I ; ∈₁-∀-E ) 31 | open import Web.Semantic.DL.Signature using ( Signature ) 32 | open import Web.Semantic.DL.TBox using 33 | ( TBox ; Axioms ; ε ; _,_ ;_⊑₁_ ; _⊑₂_ ; Tra ) 34 | open import Web.Semantic.DL.TBox.Interp using 35 | ( interp ; Δ ; _⊨_≈_ ; ≈-refl ; ≈-sym ; ≈-trans ; con-≈ ; rol-≈ ) renaming 36 | ( Interp to Interp′ ) 37 | open import Web.Semantic.DL.TBox.Interp.Morphism using 38 | ( morph ; ≲-image ; ≲-resp-≈ ; ≲-resp-con ; ≲-resp-rol ) 39 | open import Web.Semantic.DL.TBox.Minimizable using 40 | ( LHS ; RHS ; μTBox 41 | ; ⟨_⟩ ; ⊤ ; ⊥ ; _⊓_ ; _⊔_ ; ∀[_]_ ; ∃⟨_⟩_ ; ≤1 42 | ; ε ; _,_ ;_⊑₁_ ; _⊑₂_ ; Ref ; Tra ) 43 | open import Web.Semantic.DL.TBox.Model using ( _⊨t_ ; Axioms✓ ) 44 | open import Web.Semantic.Util using 45 | ( Subset ; ⊆-refl ; id ; _∘_ ; _⊕_⊕_ ; _[⊕]_[⊕]_ ; inode ; bnode ; enode ) 46 | 47 | module Web.Semantic.DL.Sequent.Model {Σ : Signature} {W X Y : Set} where 48 | 49 | infix 5 _⊞_ 50 | 51 | minimal′ : Interp Σ X → KB Σ (X ⊕ W ⊕ Y) → Interp′ Σ 52 | minimal′ I KB = 53 | interp (Γ I KB) (λ x y → (I ⊕ KB ⊢ x ∼ y)) ∼-refl ∼-sym ∼-trans 54 | (λ c x → I ⊕ KB ⊢ x ∈₁ ⟨ c ⟩) (λ r xy → I ⊕ KB ⊢ xy ∈₂ ⟨ r ⟩) 55 | (λ c → ∈₁-resp-∼) (λ r → ∈₂-resp-∼) 56 | 57 | minimal : Interp Σ X → KB Σ (X ⊕ W ⊕ Y) → Interp Σ (X ⊕ W ⊕ Y) 58 | minimal I KB = (minimal′ I KB , γ I KB) 59 | 60 | complete₂ : ∀ I KB R {xy} → (xy ∈ minimal′ I KB ⟦ R ⟧₂) → (I ⊕ KB ⊢ xy ∈₂ R) 61 | complete₂ I KB ⟨ r ⟩ {(x , y)} xy∈⟦r⟧ = xy∈⟦r⟧ 62 | complete₂ I KB ⟨ r ⟩⁻¹ {(x , y)} yx∈⟦r⟧ = ∈₂-inv-I yx∈⟦r⟧ 63 | 64 | complete₁ : ∀ I KB {C x} → (C ∈ LHS) → (x ∈ minimal′ I KB ⟦ C ⟧₁) → 65 | (I ⊕ KB ⊢ x ∈₁ C) 66 | complete₁ I KB ⟨ c ⟩ x∈⟦c⟧ = 67 | x∈⟦c⟧ 68 | complete₁ I KB ⊤ _ = 69 | ∈₁-⊤-I 70 | complete₁ I KB (C ⊓ D) (x∈⟦C⟧ , x∈⟦D⟧) = 71 | ∈₁-⊓-I (complete₁ I KB C x∈⟦C⟧) (complete₁ I KB D x∈⟦D⟧) 72 | complete₁ I KB (C ⊔ D) (inj₁ x∈⟦C⟧) = 73 | ∈₁-⊔-I₁ (complete₁ I KB C x∈⟦C⟧) 74 | complete₁ I KB (C ⊔ D) (inj₂ x∈⟦D⟧) = 75 | ∈₁-⊔-I₂ (complete₁ I KB D x∈⟦D⟧) 76 | complete₁ I KB (∃⟨ R ⟩ C) (y , xy∈⟦R⟧ , y∈⟦C⟧) = 77 | ∈₁-∃-I (complete₂ I KB R xy∈⟦R⟧) (complete₁ I KB C y∈⟦C⟧) 78 | complete₁ I KB ⊥ () 79 | 80 | sound₂ : ∀ I KB R {xy} → (I ⊕ KB ⊢ xy ∈₂ R) → (xy ∈ minimal′ I KB ⟦ R ⟧₂) 81 | sound₂ I KB ⟨ r ⟩ {(x , y)} ⊢xy∈r = ⊢xy∈r 82 | sound₂ I KB ⟨ r ⟩⁻¹ {(x , y)} ⊢xy∈r⁻ = ∈₂-inv-E ⊢xy∈r⁻ 83 | 84 | sound₁ : ∀ I KB {C x} → (C ∈ RHS) → (I ⊕ KB ⊢ x ∈₁ C) → 85 | (x ∈ minimal′ I KB ⟦ C ⟧₁) 86 | sound₁ I KB ⟨ c ⟩ ⊢x∈c = 87 | ⊢x∈c 88 | sound₁ I KB ⊤ ⊢x∈⊤ = 89 | tt 90 | sound₁ I KB (C ⊓ D) ⊢x∈C⊓D = 91 | (sound₁ I KB C (∈₁-⊓-E₁ ⊢x∈C⊓D) , sound₁ I KB D (∈₁-⊓-E₂ ⊢x∈C⊓D)) 92 | sound₁ I KB (∀[ R ] C) ⊢x∈∀RC = λ y xy∈⟦R⟧ → 93 | sound₁ I KB C (∈₁-∀-E ⊢x∈∀RC (complete₂ I KB R xy∈⟦R⟧)) 94 | sound₁ I KB (≤1 R) ⊢x∈≤1R = λ y z xy∈⟦R⟧ xz∈⟦R⟧ → 95 | ∼-≤1 ⊢x∈≤1R (complete₂ I KB R xy∈⟦R⟧) (complete₂ I KB R xz∈⟦R⟧) 96 | 97 | minimal-⊨ : ∀ I KB → (tbox KB ∈ μTBox) → (minimal I KB ⊨ KB) 98 | minimal-⊨ I KB μT = 99 | ( minimal-tbox μT (⊆-refl (Axioms (tbox KB))) 100 | , minimal-abox (abox KB) (⊆-refl (Assertions (abox KB)))) where 101 | 102 | minimal-tbox : ∀ {T} → (T ∈ μTBox) → (Axioms T ⊆ Axioms (tbox KB)) → 103 | minimal′ I KB ⊨t T 104 | minimal-tbox ε ε⊆T = 105 | tt 106 | minimal-tbox (U , V) UV⊆T = 107 | ( minimal-tbox U (λ u → UV⊆T (inj₁ u)) 108 | , minimal-tbox V (λ v → UV⊆T (inj₂ v)) ) 109 | minimal-tbox (C ⊑₁ D) C⊑₁D∈T = λ x∈⟦C⟧ → 110 | sound₁ I KB D (∈₁-subsum (complete₁ I KB C x∈⟦C⟧) (C⊑₁D∈T refl)) 111 | minimal-tbox (Q ⊑₂ R) Q⊑₁R∈T = λ xy∈⟦Q⟧ → 112 | sound₂ I KB R (∈₂-subsum (complete₂ I KB Q xy∈⟦Q⟧) (Q⊑₁R∈T refl)) 113 | minimal-tbox (Ref R) RefR∈T = λ x → 114 | sound₂ I KB R (∈₂-refl x (RefR∈T refl)) 115 | minimal-tbox (Tra R) TraR∈T = λ xy∈⟦R⟧ yz∈⟦R⟧ → 116 | sound₂ I KB R (∈₂-trans 117 | (complete₂ I KB R xy∈⟦R⟧) (complete₂ I KB R yz∈⟦R⟧) (TraR∈T refl)) 118 | 119 | minimal-abox : ∀ A → (Assertions A ⊆ Assertions (abox KB)) → 120 | minimal I KB ⊨a A 121 | minimal-abox ε ε⊆A = 122 | tt 123 | minimal-abox (B , C) BC⊆A = 124 | ( minimal-abox B (λ b → BC⊆A (inj₁ b)) 125 | , minimal-abox C (λ c → BC⊆A (inj₂ c)) ) 126 | minimal-abox (x ∼ y) x∼y⊆A = 127 | ∼-assert (x∼y⊆A refl) 128 | minimal-abox (x ∈₁ c) x∈C⊆A = 129 | sound₁ I KB ⟨ c ⟩ (∈₁-assert (x∈C⊆A refl)) 130 | minimal-abox ((x , y) ∈₂ r) xy∈R⊆A = 131 | sound₂ I KB ⟨ r ⟩ (∈₂-assert (xy∈R⊆A refl)) 132 | 133 | minimal-≳ : ∀ I KB → (I ≲ inode * minimal I KB) 134 | minimal-≳ I KB = (morph inode ∼-import ∈₁-import ∈₂-import , λ x → ∼-refl) 135 | 136 | minimal-≲ : ∀ I KB J → (I ≲ inode * J) → (J ⊨ KB) → (minimal I KB ≲ J) 137 | minimal-≲ I KB J I≲J (J⊨T , J⊨A) = 138 | (morph f minimal-≈ minimal-con minimal-rol , fγ≈j) where 139 | 140 | f : Γ I KB → Δ ⌊ J ⌋ 141 | f = (≲-image ≲⌊ I≲J ⌋) [⊕] (ind J ∘ bnode) [⊕] (ind J ∘ enode) 142 | 143 | fγ≈j : ∀ x → ⌊ J ⌋ ⊨ f (γ I KB x) ≈ ind J x 144 | fγ≈j (inode x) = ≲-resp-ind I≲J x 145 | fγ≈j (bnode v) = ≈-refl ⌊ J ⌋ 146 | fγ≈j (enode y) = ≈-refl ⌊ J ⌋ 147 | 148 | mutual 149 | 150 | minimal-≈ : ∀ {x y} → (I ⊕ KB ⊢ x ∼ y) → (⌊ J ⌋ ⊨ f x ≈ f y) 151 | minimal-≈ (∼-assert x∼y∈A) = 152 | ≈-trans ⌊ J ⌋ (fγ≈j _) 153 | (≈-trans ⌊ J ⌋ (Assertions✓ J (abox KB) x∼y∈A J⊨A) 154 | (≈-sym ⌊ J ⌋ (fγ≈j _))) 155 | minimal-≈ (∼-import x≈y) = 156 | ≲-resp-≈ ≲⌊ I≲J ⌋ x≈y 157 | minimal-≈ ∼-refl = 158 | ≈-refl ⌊ J ⌋ 159 | minimal-≈ (∼-sym x∼y) = 160 | ≈-sym ⌊ J ⌋ (minimal-≈ x∼y) 161 | minimal-≈ (∼-trans x∼y y∼z) = 162 | ≈-trans ⌊ J ⌋ (minimal-≈ x∼y) (minimal-≈ y∼z) 163 | minimal-≈ (∼-≤1 x∈≤1R xy∈R xz∈R) = 164 | minimal-con x∈≤1R _ _ (minimal-rol xy∈R) (minimal-rol xz∈R) 165 | 166 | minimal-con : ∀ {x C} → (I ⊕ KB ⊢ x ∈₁ C) → (f x ∈ ⌊ J ⌋ ⟦ C ⟧₁) 167 | minimal-con (∈₁-assert x∈c∈A) = 168 | con-≈ ⌊ J ⌋ _ (Assertions✓ J (abox KB) x∈c∈A J⊨A) 169 | (≈-sym ⌊ J ⌋ (fγ≈j _)) 170 | minimal-con (∈₁-import x∈⟦c⟧) = 171 | ≲-resp-con ≲⌊ I≲J ⌋ x∈⟦c⟧ 172 | minimal-con {x} {C} (∈₁-resp-∼ x∈C x∼y) = 173 | ⟦⟧₁-resp-≈ ⌊ J ⌋ C (minimal-con x∈C) (minimal-≈ x∼y) 174 | minimal-con (∈₁-subsum x∈C C⊑D∈T) = 175 | Axioms✓ ⌊ J ⌋ (tbox KB) C⊑D∈T J⊨T (minimal-con x∈C) 176 | minimal-con ∈₁-⊤-I = 177 | tt 178 | minimal-con (∈₁-⊓-I x∈C x∈D) = 179 | (minimal-con x∈C , minimal-con x∈D) 180 | minimal-con (∈₁-⊓-E₁ x∈C⊓D) = 181 | proj₁ (minimal-con x∈C⊓D) 182 | minimal-con (∈₁-⊓-E₂ x∈C⊓D) = 183 | proj₂ (minimal-con x∈C⊓D) 184 | minimal-con (∈₁-⊔-I₁ x∈C) = 185 | inj₁ (minimal-con x∈C) 186 | minimal-con (∈₁-⊔-I₂ x∈D) = 187 | inj₂ (minimal-con x∈D) 188 | minimal-con (∈₁-∀-E x∈[R]C xy∈R) = 189 | minimal-con x∈[R]C _ (minimal-rol xy∈R) 190 | minimal-con (∈₁-∃-I xy∈R y∈C) = 191 | (_ , minimal-rol xy∈R , minimal-con y∈C) 192 | 193 | minimal-rol : ∀ {x y R} → (I ⊕ KB ⊢ (x , y) ∈₂ R) → 194 | ((f x , f y) ∈ ⌊ J ⌋ ⟦ R ⟧₂) 195 | minimal-rol (∈₂-assert xy∈r∈A) = 196 | rol-≈ ⌊ J ⌋ _ (fγ≈j _) (Assertions✓ J (abox KB) xy∈r∈A J⊨A) 197 | (≈-sym ⌊ J ⌋ (fγ≈j _)) 198 | minimal-rol (∈₂-import xy∈⟦r⟧) = 199 | ≲-resp-rol ≲⌊ I≲J ⌋ xy∈⟦r⟧ 200 | minimal-rol {x} {y} {R} (∈₂-resp-∼ w∼x xy∈R y∼z) = 201 | ⟦⟧₂-resp-≈ ⌊ J ⌋ R (minimal-≈ w∼x) (minimal-rol xy∈R) (minimal-≈ y∼z) 202 | minimal-rol (∈₂-subsum xy∈Q Q⊑R∈T) = 203 | Axioms✓ ⌊ J ⌋ (tbox KB) Q⊑R∈T J⊨T (minimal-rol xy∈Q) 204 | minimal-rol (∈₂-refl x RefR∈T) = 205 | Axioms✓ ⌊ J ⌋ (tbox KB) RefR∈T J⊨T _ 206 | minimal-rol (∈₂-trans xy∈R yz∈R TraR∈T) = 207 | Axioms✓ ⌊ J ⌋ (tbox KB) TraR∈T J⊨T 208 | (minimal-rol xy∈R) (minimal-rol yz∈R) 209 | minimal-rol (∈₂-inv-I xy∈r) = 210 | minimal-rol xy∈r 211 | minimal-rol (∈₂-inv-E xy∈r⁻) = 212 | minimal-rol xy∈r⁻ 213 | 214 | minimal-≋ : ∀ I KB K I≲K K⊨KB → 215 | (I≲K ≋ minimal-≳ I KB >> minimal-≲ I KB K I≲K K⊨KB) 216 | minimal-≋ I KB K I≲K (K⊨T , K⊨A) = λ x → ≈-refl ⌊ K ⌋ 217 | 218 | minimal-uniq : ∀ I KB K I≲K → 219 | Unique I (minimal I KB) K (minimal-≳ I KB) I≲K 220 | minimal-uniq I KB K I≲K J≲₁K J≲₂K I≲K≋I≲J≲₁K I≲K≋I≲J≲₂K (inode x) = 221 | ≈-trans ⌊ K ⌋ (≈-sym ⌊ K ⌋ (I≲K≋I≲J≲₁K x)) (I≲K≋I≲J≲₂K x) 222 | minimal-uniq I KB K I≲K J≲₁K J≲₂K I≲K≋I≲J≲₁K I≲K≋I≲J≲₂K (bnode v) = 223 | ≈-trans ⌊ K ⌋ (≲-resp-ind J≲₁K (bnode v)) 224 | (≈-sym ⌊ K ⌋ (≲-resp-ind J≲₂K (bnode v))) 225 | minimal-uniq I KB K I≲K J≲₁K J≲₂K I≲K≋I≲J≲₁K I≲K≋I≲J≲₂K (enode y) = 226 | ≈-trans ⌊ K ⌋ (≲-resp-ind J≲₁K (enode y)) 227 | (≈-sym ⌊ K ⌋ (≲-resp-ind J≲₂K (enode y))) 228 | 229 | minimal-med : ∀ I KB → Mediator I (minimal I KB) (minimal-≳ I KB) KB 230 | minimal-med I KB K I≲K K⊨KB = 231 | ( minimal-≲ I KB K I≲K K⊨KB 232 | , minimal-≋ I KB K I≲K K⊨KB 233 | , minimal-uniq I KB K I≲K) 234 | 235 | minimal-init : ∀ I KB → (tbox KB ∈ μTBox) → Initial I KB (minimal I KB) 236 | minimal-init I KB μKB = 237 | (minimal-≳ I KB 238 | , minimal-⊨ I KB μKB 239 | , minimal-med I KB) 240 | 241 | _⊞_ : Interp Σ X → KB Σ (X ⊕ W ⊕ Y) → Interp Σ Y 242 | I ⊞ KB = enode * minimal I KB 243 | 244 | ⊞-sound : ∀ I KB₁ KB₂ → (tbox KB₁ ∈ μTBox) → (I ⊞ KB₁ ⊨ KB₂) → (I ⊕ KB₁ ⊨ KB₂) 245 | ⊞-sound I KB₁ KB₂ μKB₁ ⊨KB₂ = (minimal I KB₁ , minimal-init I KB₁ μKB₁ , ⊨KB₂) -------------------------------------------------------------------------------- /src/Web/Semantic/DL/Signature.agda: -------------------------------------------------------------------------------- 1 | module Web.Semantic.DL.Signature where 2 | 3 | infixr 4 _,_ 4 | 5 | data Signature : Set₁ where 6 | _,_ : (CN RN : Set) → Signature 7 | 8 | CN : Signature → Set 9 | CN (CN , RN) = CN 10 | 11 | RN : Signature → Set 12 | RN (CN , RN) = RN 13 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/TBox.agda: -------------------------------------------------------------------------------- 1 | open import Relation.Unary using ( ∅ ; _∪_ ) 2 | open import Web.Semantic.DL.Concept using ( Concept ) 3 | open import Web.Semantic.DL.Role using ( Role ) 4 | open import Web.Semantic.DL.Signature using ( Signature ) 5 | open import Web.Semantic.Util using ( Subset ; ⁅_⁆ ) 6 | 7 | module Web.Semantic.DL.TBox where 8 | 9 | infixl 5 _⊑₁_ _⊑₂_ 10 | infixr 4 _,_ 11 | 12 | data TBox (Σ : Signature) : Set where 13 | ε : TBox Σ 14 | _,_ : (T U : TBox Σ) → TBox Σ 15 | _⊑₁_ : (C D : Concept Σ) → TBox Σ 16 | _⊑₂_ : (Q R : Role Σ) → TBox Σ 17 | Ref : (R : Role Σ) → TBox Σ 18 | Irr : (R : Role Σ) → TBox Σ 19 | Tra : (R : Role Σ) → TBox Σ 20 | Dis : (Q R : Role Σ) → TBox Σ 21 | 22 | Axioms : ∀ {Σ} → TBox Σ → Subset (TBox Σ) 23 | Axioms ε = ∅ 24 | Axioms (T , U) = (Axioms T) ∪ (Axioms U) 25 | Axioms (C ⊑₁ D) = ⁅ C ⊑₁ D ⁆ 26 | Axioms (Q ⊑₂ R) = ⁅ Q ⊑₂ R ⁆ 27 | Axioms (Ref R) = ⁅ Ref R ⁆ 28 | Axioms (Irr R) = ⁅ Irr R ⁆ 29 | Axioms (Tra R) = ⁅ Tra R ⁆ 30 | Axioms (Dis Q R) = ⁅ Dis Q R ⁆ 31 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/TBox/Interp.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( ∃ ; _×_ ; _,_ ) 2 | open import Relation.Binary.PropositionalEquality using ( _≡_ ; refl ) 3 | open import Relation.Nullary using ( ¬_ ) 4 | open import Relation.Unary using ( _∈_ ; ∅ ) 5 | open import Web.Semantic.DL.Signature using ( Signature ; CN ; RN ) 6 | open import Web.Semantic.Util using ( Setoid ; Subset ; _∘_ ; False ) 7 | 8 | module Web.Semantic.DL.TBox.Interp where 9 | 10 | data Interp (Σ : Signature) : Set₁ where 11 | interp : 12 | (Δ : Set) → 13 | (_≈_ : Δ → Δ → Set) → 14 | (ref : ∀ {x} → (x ≈ x)) → 15 | (sym : ∀ {x y} → (x ≈ y) → (y ≈ x)) → 16 | (trans : ∀ {x y z} → (x ≈ y) → (y ≈ z) → (x ≈ z)) → 17 | (con : CN Σ → Subset Δ) → 18 | (rol : RN Σ → Subset (Δ × Δ)) → 19 | (con-≈ : ∀ {x y} c → (x ∈ con c) → (x ≈ y) → (y ∈ con c)) → 20 | (rol-≈ : ∀ {w x y z} r → (w ≈ x) → ((x , y) ∈ rol r) → (y ≈ z) → ((w , z) ∈ rol r)) → 21 | Interp Σ 22 | 23 | Δ : ∀ {Σ} → Interp Σ → Set 24 | Δ (interp Δ _≈_ ref sym trans con rol con-≈ rol-≈) = Δ 25 | 26 | _⊨_≈_ : ∀ {Σ} → (I : Interp Σ) → (Δ I) → (Δ I) → Set 27 | _⊨_≈_ (interp Δ _≈_ ref sym trans con rol con-≈ rol-≈) = _≈_ 28 | 29 | ≈-refl : ∀ {Σ} → (I : Interp Σ) → ∀ {x} → (I ⊨ x ≈ x) 30 | ≈-refl (interp Δ _≈_ ref sym trans con rol con-≈ rol-≈) = ref 31 | 32 | ≈-sym : ∀ {Σ} → (I : Interp Σ) → ∀ {x y} → (I ⊨ x ≈ y) → (I ⊨ y ≈ x) 33 | ≈-sym (interp Δ _≈_ ref sym trans con rol con-≈ rol-≈) = sym 34 | 35 | ≈-trans : ∀ {Σ} → (I : Interp Σ) → ∀ {x y z} → (I ⊨ x ≈ y) → (I ⊨ y ≈ z) → (I ⊨ x ≈ z) 36 | ≈-trans (interp Δ _≈_ ref sym trans con rol con-≈ rol-≈) = trans 37 | 38 | con : ∀ {Σ} → (I : Interp Σ) → CN Σ → Subset (Δ I) 39 | con (interp Δ _≈_ ref sym trans con rol con-≈ rol-≈) = con 40 | 41 | rol : ∀ {Σ} → (I : Interp Σ) → RN Σ → Subset (Δ I × Δ I) 42 | rol (interp Δ _≈_ ref sym trans con rol con-≈ rol-≈) = rol 43 | 44 | con-≈ : ∀ {Σ} → (I : Interp Σ) → ∀ {x y} c → (x ∈ con I c) → (I ⊨ x ≈ y) → (y ∈ con I c) 45 | con-≈ (interp Δ _≈_ ref sym trans con rol con-≈ rol-≈) = con-≈ 46 | 47 | rol-≈ : ∀ {Σ} → (I : Interp Σ) → ∀ {w x y z} r → (I ⊨ w ≈ x) → ((x , y) ∈ rol I r) → (I ⊨ y ≈ z) → ((w , z) ∈ rol I r) 48 | rol-≈ (interp Δ _≈_ ref sym trans con rol con-≈ rol-≈) = rol-≈ 49 | 50 | _⊨_≉_ : ∀ {Σ} → (I : Interp Σ) → (Δ I) → (Δ I) → Set 51 | I ⊨ x ≉ y = ¬(I ⊨ x ≈ y) 52 | 53 | emp : ∀ {Σ} → Interp Σ 54 | emp = interp False (λ ()) (λ {}) (λ {}) (λ {}) (λ c → ∅) (λ r → ∅) (λ {}) (λ {}) 55 | 56 | ≈-refl′ : ∀ {Σ} (I : Interp Σ) → ∀ {x y} → (x ≡ y) → (I ⊨ x ≈ y) 57 | ≈-refl′ I refl = ≈-refl I 58 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/TBox/Interp/Morphism.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( ∃ ; _×_ ; _,_ ; proj₁ ; proj₂ ) 2 | open import Relation.Unary using ( _∈_ ) 3 | open import Web.Semantic.DL.TBox.Interp using 4 | ( Interp ; Δ ; _⊨_≈_ ; con ; rol ; ≈-refl ; ≈-sym ; ≈-trans ; con-≈ ; rol-≈ ) 5 | open import Web.Semantic.DL.Signature using ( Signature ) 6 | open import Web.Semantic.Util using ( id ; _∘_ ) 7 | 8 | module Web.Semantic.DL.TBox.Interp.Morphism {Σ : Signature} where 9 | 10 | infix 2 _≲_ _≃_ 11 | 12 | -- I ≲ J whenever J respects all the properties of I 13 | 14 | data _≲_ (I J : Interp Σ) : Set where 15 | morph : 16 | (f : Δ I → Δ J) → 17 | (≲-resp-≈ : ∀ {x y} → (I ⊨ x ≈ y) → (J ⊨ f x ≈ f y)) → 18 | (≲-resp-con : ∀ {c x} → (x ∈ con I c) → (f x ∈ con J c)) → 19 | (≲-resp-rol : ∀ {r x y} → ((x , y) ∈ rol I r) → ((f x , f y) ∈ rol J r)) → 20 | (I ≲ J) 21 | 22 | ≲-image : ∀ {I J} → (I ≲ J) → Δ I → Δ J 23 | ≲-image (morph f ≲-resp-≈ ≲-resp-con ≲-resp-rol) = f 24 | 25 | ≲-image² : ∀ {I J} → (I ≲ J) → (Δ I × Δ I) → (Δ J × Δ J) 26 | ≲-image² I≲J (x , y) = (≲-image I≲J x , ≲-image I≲J y) 27 | 28 | ≲-resp-≈ : ∀ {I J} → (I≲J : I ≲ J) → ∀ {x y} → (I ⊨ x ≈ y) → 29 | (J ⊨ ≲-image I≲J x ≈ ≲-image I≲J y) 30 | ≲-resp-≈ (morph f ≲-resp-≈ ≲-resp-con ≲-resp-rol) = ≲-resp-≈ 31 | 32 | ≲-resp-con : ∀ {I J} → (I≲J : I ≲ J) → ∀ {c x} → (x ∈ con I c) → 33 | (≲-image I≲J x ∈ con J c) 34 | ≲-resp-con (morph f ≲-resp-≈ ≲-resp-con ≲-resp-rol) = ≲-resp-con 35 | 36 | ≲-resp-rol : ∀ {I J} → (I≲J : I ≲ J) → ∀ {r xy} → (xy ∈ rol I r) → 37 | (≲-image² I≲J xy ∈ rol J r) 38 | ≲-resp-rol (morph f ≲-resp-≈ ≲-resp-con ≲-resp-rol) {r} {x , y} = ≲-resp-rol 39 | 40 | ≲-refl : ∀ I → (I ≲ I) 41 | ≲-refl I = morph id id id id 42 | 43 | ≲-trans : ∀ {I J K} → (I ≲ J) → (J ≲ K) → (I ≲ K) 44 | ≲-trans I≲J J≲K = morph 45 | (≲-image J≲K ∘ ≲-image I≲J) 46 | (≲-resp-≈ J≲K ∘ ≲-resp-≈ I≲J) 47 | (≲-resp-con J≲K ∘ ≲-resp-con I≲J) 48 | (≲-resp-rol J≲K ∘ ≲-resp-rol I≲J) 49 | 50 | -- I ≃ J whenever there is an isomprhism between I and J 51 | 52 | data _≃_ (I J : Interp Σ) : Set where 53 | iso : 54 | (≃-impl-≲ : I ≲ J) → 55 | (≃-impl-≳ : J ≲ I) → 56 | (≃-iso : ∀ x → I ⊨ (≲-image ≃-impl-≳ (≲-image ≃-impl-≲ x)) ≈ x) → 57 | (≃-iso⁻¹ : ∀ x → J ⊨ (≲-image ≃-impl-≲ (≲-image ≃-impl-≳ x)) ≈ x) → 58 | (I ≃ J) 59 | 60 | ≃-impl-≲ : ∀ {I J} → (I ≃ J) → (I ≲ J) 61 | ≃-impl-≲ (iso ≃-impl-≲ ≃-impl-≳ ≃-iso ≃-iso⁻¹) = ≃-impl-≲ 62 | 63 | ≃-impl-≳ : ∀ {I J} → (I ≃ J) → (J ≲ I) 64 | ≃-impl-≳ (iso ≃-impl-≲ ≃-impl-≳ ≃-iso ≃-iso⁻¹) = ≃-impl-≳ 65 | 66 | ≃-image : ∀ {I J} → (I ≃ J) → Δ I → Δ J 67 | ≃-image I≃J = ≲-image (≃-impl-≲ I≃J) 68 | 69 | ≃-image² : ∀ {I J} → (I ≃ J) → (Δ I × Δ I) → (Δ J × Δ J) 70 | ≃-image² I≃J = ≲-image² (≃-impl-≲ I≃J) 71 | 72 | ≃-image⁻¹ : ∀ {I J} → (I ≃ J) → Δ J → Δ I 73 | ≃-image⁻¹ I≃J = ≲-image (≃-impl-≳ I≃J) 74 | 75 | ≃-image²⁻¹ : ∀ {I J} → (I ≃ J) → (Δ J × Δ J) → (Δ I × Δ I) 76 | ≃-image²⁻¹ I≃J = ≲-image² (≃-impl-≳ I≃J) 77 | 78 | ≃-iso : ∀ {I J} → (I≃J : I ≃ J) → ∀ x → (I ⊨ (≃-image⁻¹ I≃J (≃-image I≃J x)) ≈ x) 79 | ≃-iso (iso ≃-impl-≲ ≃-impl-≳ ≃-iso ≃-iso⁻¹) = ≃-iso 80 | 81 | ≃-iso⁻¹ : ∀ {I J} → (I≃J : I ≃ J) → ∀ x → (J ⊨ (≃-image I≃J (≃-image⁻¹ I≃J x)) ≈ x) 82 | ≃-iso⁻¹ (iso ≃-impl-≲ ≃-impl-≳ ≃-iso ≃-iso⁻¹) = ≃-iso⁻¹ 83 | 84 | ≃-resp-≈ : ∀ {I J} → (I≃J : I ≃ J) → ∀ {x y} → 85 | (I ⊨ x ≈ y) → (J ⊨ ≃-image I≃J x ≈ ≃-image I≃J y) 86 | ≃-resp-≈ I≃J = ≲-resp-≈ (≃-impl-≲ I≃J) 87 | 88 | ≃-refl-≈ : ∀ {I J} → (I≃J : I ≃ J) → ∀ {x y} → 89 | (I ⊨ ≃-image⁻¹ I≃J x ≈ ≃-image⁻¹ I≃J y) → (J ⊨ x ≈ y) 90 | ≃-refl-≈ {I} {J} I≃J {x} {y} x≈y = 91 | ≈-trans J (≈-sym J (≃-iso⁻¹ I≃J x)) 92 | (≈-trans J (≃-resp-≈ I≃J x≈y) (≃-iso⁻¹ I≃J y)) 93 | 94 | ≃-refl : ∀ I → (I ≃ I) 95 | ≃-refl I = iso (≲-refl I) (≲-refl I) (λ x → ≈-refl I) (λ x → ≈-refl I) 96 | 97 | ≃-sym : ∀ {I J} → (I ≃ J) → (J ≃ I) 98 | ≃-sym I≃J = iso (≃-impl-≳ I≃J) (≃-impl-≲ I≃J) (≃-iso⁻¹ I≃J) (≃-iso I≃J) 99 | 100 | ≃-trans : ∀ {I J K} → (I ≃ J) → (J ≃ K) → (I ≃ K) 101 | ≃-trans {I} {J} {K} I≃J J≃K = iso 102 | (≲-trans (≃-impl-≲ I≃J) (≃-impl-≲ J≃K)) 103 | (≲-trans (≃-impl-≳ J≃K) (≃-impl-≳ I≃J)) 104 | (λ x → ≈-trans I (≲-resp-≈ (≃-impl-≳ I≃J) (≃-iso J≃K (≃-image I≃J x))) 105 | (≃-iso I≃J x)) 106 | (λ x → ≈-trans K (≲-resp-≈ (≃-impl-≲ J≃K) (≃-iso⁻¹ I≃J (≃-image⁻¹ J≃K x))) 107 | (≃-iso⁻¹ J≃K x)) 108 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/TBox/Minimizable.agda: -------------------------------------------------------------------------------- 1 | open import Data.Bool using ( Bool ; true ; false ; _∧_ ) 2 | open import Data.Product using ( _×_ ) 3 | open import Relation.Binary.PropositionalEquality using ( _≡_ ) 4 | open import Relation.Unary using ( _∈_ ) 5 | open import Web.Semantic.DL.Concept using 6 | ( Concept ; ⟨_⟩ ; ¬⟨_⟩ ; ⊤ ; ⊥ ; _⊓_ ; _⊔_ ; ∀[_]_ ; ∃⟨_⟩_ ; ≤1 ; >1 ) 7 | open import Web.Semantic.DL.Signature using ( Signature ) 8 | open import Web.Semantic.DL.TBox using 9 | ( TBox ; ε ; _,_ ;_⊑₁_ ; _⊑₂_ ; Dis ; Ref ; Irr ; Tra ) 10 | open import Web.Semantic.Util using ( Subset ; □ ; □-proj₁ ; □-proj₂ ) 11 | 12 | module Web.Semantic.DL.TBox.Minimizable {Σ : Signature} where 13 | 14 | data LHS : Subset (Concept Σ) where 15 | ⟨_⟩ : ∀ c → ⟨ c ⟩ ∈ LHS 16 | ⊤ : ⊤ ∈ LHS 17 | ⊥ : ⊥ ∈ LHS 18 | _⊓_ : ∀ {C D} → (C ∈ LHS) → (D ∈ LHS) → ((C ⊓ D) ∈ LHS) 19 | _⊔_ : ∀ {C D} → (C ∈ LHS) → (D ∈ LHS) → ((C ⊔ D) ∈ LHS) 20 | ∃⟨_⟩_ : ∀ R {C} → (C ∈ LHS) → ((∃⟨ R ⟩ C) ∈ LHS) 21 | 22 | data RHS : Subset (Concept Σ) where 23 | ⟨_⟩ : ∀ c → ⟨ c ⟩ ∈ RHS 24 | ⊤ : ⊤ ∈ RHS 25 | _⊓_ : ∀ {C D} → (C ∈ RHS) → (D ∈ RHS) → ((C ⊓ D) ∈ RHS) 26 | ∀[_]_ : ∀ R {C} → (C ∈ RHS) → ((∀[ R ] C) ∈ RHS) 27 | ≤1 : ∀ R → ((≤1 R) ∈ RHS) 28 | 29 | data μTBox : Subset (TBox Σ) where 30 | ε : μTBox ε 31 | _,_ : ∀ {T U} → (T ∈ μTBox) → (U ∈ μTBox) → ((T , U) ∈ μTBox) 32 | _⊑₁_ : ∀ {C D} → (C ∈ LHS) → (D ∈ RHS) → ((C ⊑₁ D) ∈ μTBox) 33 | _⊑₂_ : ∀ Q R → ((Q ⊑₂ R) ∈ μTBox) 34 | Ref : ∀ R → (Ref R ∈ μTBox) 35 | Tra : ∀ R → (Tra R ∈ μTBox) 36 | 37 | lhs? : Concept Σ → Bool 38 | lhs? ⟨ c ⟩ = true 39 | lhs? ¬⟨ c ⟩ = false 40 | lhs? ⊤ = true 41 | lhs? ⊥ = true 42 | lhs? (C ⊓ D) = lhs? C ∧ lhs? D 43 | lhs? (C ⊔ D) = lhs? C ∧ lhs? D 44 | lhs? (∀[ R ] C) = false 45 | lhs? (∃⟨ R ⟩ C) = lhs? C 46 | lhs? (≤1 R) = false 47 | lhs? (>1 R) = false 48 | 49 | lhs : ∀ C {C✓ : □(lhs? C)} → LHS C 50 | lhs ⟨ c ⟩ = ⟨ c ⟩ 51 | lhs ⊤ = ⊤ 52 | lhs ⊥ = ⊥ 53 | lhs (C ⊓ D) {C⊓D✓} = lhs C {□-proj₁ C⊓D✓} ⊓ lhs D {□-proj₂ {lhs? C} C⊓D✓} 54 | lhs (C ⊔ D) {C⊔D✓} = lhs C {□-proj₁ C⊔D✓} ⊔ lhs D {□-proj₂ {lhs? C} C⊔D✓} 55 | lhs (∃⟨ R ⟩ C) {C✓} = ∃⟨ R ⟩ (lhs C {C✓}) 56 | lhs ¬⟨ c ⟩ {} 57 | lhs (∀[ R ] C) {} 58 | lhs (≤1 R) {} 59 | lhs (>1 R) {} 60 | 61 | rhs? : Concept Σ → Bool 62 | rhs? ⟨ c ⟩ = true 63 | rhs? ¬⟨ c ⟩ = false 64 | rhs? ⊤ = true 65 | rhs? ⊥ = false 66 | rhs? (C ⊓ D) = rhs? C ∧ rhs? D 67 | rhs? (C ⊔ D) = false 68 | rhs? (∀[ R ] C) = rhs? C 69 | rhs? (∃⟨ R ⟩ C) = false 70 | rhs? (≤1 R) = true 71 | rhs? (>1 R) = false 72 | 73 | rhs : ∀ C {C✓ : □(rhs? C)} → RHS C 74 | rhs ⟨ c ⟩ = ⟨ c ⟩ 75 | rhs ⊤ = ⊤ 76 | rhs (C ⊓ D) {C⊓D✓} = rhs C {□-proj₁ C⊓D✓} ⊓ rhs D {□-proj₂ {rhs? C} C⊓D✓} 77 | rhs (∀[ R ] C) {C✓} = ∀[ R ] (rhs C {C✓}) 78 | rhs (≤1 R) = ≤1 R 79 | rhs ⊥ {} 80 | rhs ¬⟨ c ⟩ {} 81 | rhs (C ⊔ D) {} 82 | rhs (∃⟨ R ⟩ C) {} 83 | rhs (>1 R) {} 84 | 85 | μTBox? : TBox Σ → Bool 86 | μTBox? ε = true 87 | μTBox? (T , U) = μTBox? T ∧ μTBox? U 88 | μTBox? (C ⊑₁ D) = lhs? C ∧ rhs? D 89 | μTBox? (Q ⊑₂ R) = true 90 | μTBox? (Dis Q R) = false 91 | μTBox? (Ref R) = true 92 | μTBox? (Irr R) = false 93 | μTBox? (Tra R) = true 94 | 95 | μtBox : ∀ T {T✓ : □(μTBox? T)} → μTBox T 96 | μtBox ε = ε 97 | μtBox (T , U) {TU✓} = (μtBox T {□-proj₁ TU✓} , μtBox U {□-proj₂ {μTBox? T} TU✓}) 98 | μtBox (C ⊑₁ D) {C⊑D✓} = lhs C {□-proj₁ C⊑D✓} ⊑₁ rhs D {□-proj₂ {lhs? C} C⊑D✓} 99 | μtBox (Q ⊑₂ R) = Q ⊑₂ R 100 | μtBox (Ref R) = Ref R 101 | μtBox (Tra R) = Tra R 102 | μtBox (Dis Q R) {} 103 | μtBox (Irr R) {} 104 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/TBox/Model.agda: -------------------------------------------------------------------------------- 1 | open import Data.Product using ( _×_ ; _,_ ) 2 | open import Data.Sum using ( inj₁ ; inj₂ ) 3 | open import Relation.Binary.PropositionalEquality using ( refl ) 4 | open import Relation.Unary using ( _∈_ ; _∉_ ; _⊆_ ) 5 | open import Web.Semantic.DL.Concept.Model using 6 | ( _⟦_⟧₁ ; ⟦⟧₁-resp-≈ ; ⟦⟧₁-resp-≃; ⟦⟧₁-refl-≃ ) 7 | open import Web.Semantic.DL.Role.Model using 8 | ( _⟦_⟧₂ ; ⟦⟧₂-resp-≈ ; ⟦⟧₂-resp-≃ ; ⟦⟧₂-refl-≃ ) 9 | open import Web.Semantic.DL.Signature using ( Signature ) 10 | open import Web.Semantic.DL.TBox using 11 | ( TBox ; Axioms ; ε ; _,_ ;_⊑₁_ ; _⊑₂_ ; Dis ; Ref ; Irr ; Tra ) 12 | open import Web.Semantic.DL.TBox.Interp using ( Interp ) 13 | open import Web.Semantic.DL.TBox.Interp.Morphism using ( _≃_ ; ≃-sym ) 14 | open import Web.Semantic.Util using ( True ; tt ; _∘_ ) 15 | 16 | module Web.Semantic.DL.TBox.Model {Σ : Signature} where 17 | 18 | infixr 2 _⊨t_ 19 | 20 | _⊨t_ : Interp Σ → TBox Σ → Set 21 | I ⊨t ε = True 22 | I ⊨t (T , U) = (I ⊨t T) × (I ⊨t U) 23 | I ⊨t (C ⊑₁ D) = I ⟦ C ⟧₁ ⊆ I ⟦ D ⟧₁ 24 | I ⊨t (Q ⊑₂ R) = I ⟦ Q ⟧₂ ⊆ I ⟦ R ⟧₂ 25 | I ⊨t (Dis Q R) = ∀ {xy} → (xy ∈ I ⟦ Q ⟧₂) → (xy ∉ I ⟦ R ⟧₂) 26 | I ⊨t (Ref R) = ∀ x → ((x , x) ∈ I ⟦ R ⟧₂) 27 | I ⊨t (Irr R) = ∀ x → ((x , x) ∉ I ⟦ R ⟧₂) 28 | I ⊨t (Tra R) = ∀ {x y z} → 29 | ((x , y) ∈ I ⟦ R ⟧₂) → ((y , z) ∈ I ⟦ R ⟧₂) → ((x , z) ∈ I ⟦ R ⟧₂) 30 | 31 | Axioms✓ : ∀ I T {t} → (t ∈ Axioms T) → (I ⊨t T) → (I ⊨t t) 32 | Axioms✓ I ε () I⊨T 33 | Axioms✓ I (T , U) (inj₁ t∈T) (I⊨T , I⊨U) = Axioms✓ I T t∈T I⊨T 34 | Axioms✓ I (T , U) (inj₂ t∈U) (I⊨T , I⊨U) = Axioms✓ I U t∈U I⊨U 35 | Axioms✓ I (C ⊑₁ D) refl I⊨T = I⊨T 36 | Axioms✓ I (Q ⊑₂ R) refl I⊨T = I⊨T 37 | Axioms✓ I (Dis Q R) refl I⊨T = I⊨T 38 | Axioms✓ I (Ref R) refl I⊨T = I⊨T 39 | Axioms✓ I (Irr R) refl I⊨T = I⊨T 40 | Axioms✓ I (Tra R) refl I⊨T = I⊨T 41 | 42 | ⊨t-resp-≃ : ∀ {I J : Interp Σ} → (I ≃ J) → ∀ T → (I ⊨t T) → (J ⊨t T) 43 | ⊨t-resp-≃ {I} {J} I≃J ε _ = 44 | tt 45 | ⊨t-resp-≃ {I} {J} I≃J (T , U) (I⊨T , I⊨U) = 46 | (⊨t-resp-≃ I≃J T I⊨T , ⊨t-resp-≃ I≃J U I⊨U) 47 | ⊨t-resp-≃ {I} {J} I≃J (C ⊑₁ D) I⊨C⊑D = 48 | ⟦⟧₁-refl-≃ I≃J D ∘ I⊨C⊑D ∘ ⟦⟧₁-resp-≃ (≃-sym I≃J) C 49 | ⊨t-resp-≃ {I} {J} I≃J (Q ⊑₂ R) I⊨Q⊑R = ⟦⟧₂-refl-≃ I≃J R ∘ I⊨Q⊑R ∘ 50 | ⟦⟧₂-resp-≃ (≃-sym I≃J) Q 51 | ⊨t-resp-≃ {I} {J} I≃J (Dis Q R) I⊨DisQR = λ xy∈⟦Q⟧ xy∈⟦R⟧ → 52 | I⊨DisQR (⟦⟧₂-resp-≃ (≃-sym I≃J) Q xy∈⟦Q⟧) (⟦⟧₂-resp-≃ (≃-sym I≃J) R xy∈⟦R⟧) 53 | ⊨t-resp-≃ {I} {J} I≃J (Ref R) I⊨RefR = λ x → ⟦⟧₂-refl-≃ I≃J R (I⊨RefR _) 54 | ⊨t-resp-≃ {I} {J} I≃J (Irr R) I⊨IrrR = λ x xx∈⟦R⟧ → 55 | I⊨IrrR _ (⟦⟧₂-resp-≃ (≃-sym I≃J) R xx∈⟦R⟧) 56 | ⊨t-resp-≃ {I} {J} I≃J (Tra R) I⊨TraR = λ xy∈⟦R⟧ yz∈⟦R⟧ → 57 | ⟦⟧₂-refl-≃ I≃J R (I⊨TraR 58 | (⟦⟧₂-resp-≃ (≃-sym I≃J) R xy∈⟦R⟧) 59 | (⟦⟧₂-resp-≃ (≃-sym I≃J) R yz∈⟦R⟧)) 60 | -------------------------------------------------------------------------------- /src/Web/Semantic/DL/TBox/Skolemization.agda: -------------------------------------------------------------------------------- 1 | open import Data.Bool using ( Bool ; true ; false ; if_then_else_ ) 2 | open import Data.Empty using ( ⊥-elim ) 3 | open import Data.Product using ( _×_ ; _,_ ) 4 | open import Data.Sum using ( inj₁ ; inj₂ ) 5 | open import Relation.Unary using ( _∈_ ; _∉_ ) 6 | open import Web.Semantic.DL.Concept using ( neg ) 7 | open import Web.Semantic.DL.Concept.Model using ( _⟦_⟧₁ ; neg-sound ) 8 | open import Web.Semantic.DL.Concept.Skolemization using 9 | ( CSkolems ; cskolem ; cskolem-sound ) 10 | open import Web.Semantic.DL.FOL using 11 | ( Formula ; true ; false ; _∧_ ; _∈₁_ ; _∈₂_ ; _∼_ ; ∀₁ ) 12 | open import Web.Semantic.DL.FOL.Model using ( _⊨f_ ) 13 | open import Web.Semantic.DL.Role.Skolemization using 14 | ( rskolem ; rskolem⇒ ; rskolem-sound ; rskolem⇒-sound ) 15 | open import Web.Semantic.DL.Role.Model using ( _⟦_⟧₂ ) 16 | open import Web.Semantic.DL.Signature using ( Signature ) 17 | open import Web.Semantic.DL.TBox using 18 | ( TBox ; ε ; _,_ ;_⊑₁_ ; _⊑₂_ ; Dis ; Ref ; Irr ; Tra ) 19 | open import Web.Semantic.DL.TBox.Model using ( _⊨t_ ) 20 | open import Web.Semantic.Util using ( True ; tt ) 21 | 22 | module Web.Semantic.DL.TBox.Skolemization {Σ : Signature} where 23 | 24 | TSkolems : Set → TBox Σ → Set 25 | TSkolems Δ ε = True 26 | TSkolems Δ (T , U) = (TSkolems Δ T) × (TSkolems Δ U) 27 | TSkolems Δ (C ⊑₁ D) = (Δ → Bool) × (CSkolems Δ (neg C)) × (CSkolems Δ D) 28 | TSkolems Δ (Q ⊑₂ R) = True 29 | TSkolems Δ (Dis Q R) = True 30 | TSkolems Δ (Ref R) = True 31 | TSkolems Δ (Irr R) = True 32 | TSkolems Δ (Tra R) = True 33 | 34 | tskolem : ∀ {Δ} T → (TSkolems Δ T) → Formula Σ Δ 35 | tskolem ε Φ = true 36 | tskolem (T , U) (Φ , Ψ) = (tskolem T Φ) ∧ (tskolem U Ψ) 37 | tskolem (C ⊑₁ D) (φ , Φ , Ψ) = ∀₁ λ x → 38 | if (φ x) then (cskolem (neg C) Φ x) else (cskolem D Ψ x) 39 | tskolem (Q ⊑₂ R) Φ = ∀₁ λ x → ∀₁ λ y → 40 | rskolem⇒ Q x y (rskolem R x y) 41 | tskolem (Dis Q R) Φ = ∀₁ λ x → ∀₁ λ y → 42 | rskolem⇒ Q x y (rskolem⇒ R x y false) 43 | tskolem (Ref R) Φ = ∀₁ λ x → 44 | rskolem R x x 45 | tskolem (Irr R) Φ = ∀₁ λ x → 46 | rskolem⇒ R x x false 47 | tskolem (Tra R) Φ = ∀₁ λ x → ∀₁ λ y → ∀₁ λ z → 48 | rskolem⇒ R x y (rskolem⇒ R y z (rskolem R x z)) 49 | 50 | tskolem-sound : ∀ I T Φ → (I ⊨f tskolem T Φ) → (I ⊨t T) 51 | tskolem-sound I ε Φ _ = tt 52 | tskolem-sound I (T , U) (Φ , Ψ) (I⊨T , I⊨U) = 53 | (tskolem-sound I T Φ I⊨T , tskolem-sound I U Ψ I⊨U) 54 | tskolem-sound I (C ⊑₁ D) (φ , Φ , Ψ) I⊨C⊑D = lemma where 55 | lemma : ∀ {x} → (x ∈ I ⟦ C ⟧₁) → (x ∈ I ⟦ D ⟧₁) 56 | lemma {x} x∈⟦C⟧ with φ x | I⊨C⊑D x 57 | lemma {x} x∈⟦C⟧ | true | x∈⟦¬C⟧ = 58 | ⊥-elim (neg-sound I C (cskolem-sound I (neg C) Φ x x∈⟦¬C⟧) x∈⟦C⟧) 59 | lemma {x} x∈⟦C⟧ | false | x∈⟦D⟧ = 60 | cskolem-sound I D Ψ x x∈⟦D⟧ 61 | tskolem-sound I (Q ⊑₂ R) Φ I⊨Q⊑R = lemma where 62 | lemma : ∀ {xy} → (xy ∈ I ⟦ Q ⟧₂) → (xy ∈ I ⟦ R ⟧₂) 63 | lemma {x , y} xy∈⟦Q⟧ = 64 | rskolem-sound I R x y (rskolem⇒-sound I Q x y _ (I⊨Q⊑R x y) xy∈⟦Q⟧) 65 | tskolem-sound I (Dis Q R) Φ I⊨DisQR = lemma where 66 | lemma : ∀ {xy} → (xy ∈ I ⟦ Q ⟧₂) → (xy ∉ I ⟦ R ⟧₂) 67 | lemma {x , y} xy∈⟦Q⟧ = 68 | rskolem⇒-sound I R x y _ (rskolem⇒-sound I Q x y _ (I⊨DisQR x y) xy∈⟦Q⟧) 69 | tskolem-sound I (Ref R) Φ I⊨RefR = λ x → rskolem-sound I R x x (I⊨RefR x) 70 | tskolem-sound I (Irr R) Φ I⊨IrrR = λ x → rskolem⇒-sound I R x x _ (I⊨IrrR x) 71 | tskolem-sound I (Tra R) Φ I⊨TraR = 72 | λ {x} {y} {z} xy∈⟦R⟧ yz∈⟦R⟧ → rskolem-sound I R x z 73 | (rskolem⇒-sound I R y z _ (rskolem⇒-sound I R x y _ (I⊨TraR x y z) xy∈⟦R⟧) yz∈⟦R⟧) -------------------------------------------------------------------------------- /src/Web/Semantic/Everything.agda: -------------------------------------------------------------------------------- 1 | -- A module which imports every Web.Semantic module. 2 | 3 | module Web.Semantic.Everything where 4 | 5 | import Web.Semantic.DL.ABox 6 | import Web.Semantic.DL.ABox.Interp 7 | import Web.Semantic.DL.ABox.Interp.Meet 8 | import Web.Semantic.DL.ABox.Interp.Morphism 9 | import Web.Semantic.DL.ABox.Model 10 | import Web.Semantic.DL.ABox.Skolemization 11 | import Web.Semantic.DL.Category 12 | import Web.Semantic.DL.Category.Composition 13 | import Web.Semantic.DL.Category.Morphism 14 | import Web.Semantic.DL.Category.Object 15 | import Web.Semantic.DL.Category.Properties 16 | import Web.Semantic.DL.Category.Properties.Composition 17 | import Web.Semantic.DL.Category.Properties.Composition.Assoc 18 | import Web.Semantic.DL.Category.Properties.Composition.LeftUnit 19 | import Web.Semantic.DL.Category.Properties.Composition.Lemmas 20 | import Web.Semantic.DL.Category.Properties.Composition.RespectsEquiv 21 | import Web.Semantic.DL.Category.Properties.Composition.RespectsWiring 22 | import Web.Semantic.DL.Category.Properties.Composition.RightUnit 23 | import Web.Semantic.DL.Category.Properties.Equivalence 24 | import Web.Semantic.DL.Category.Properties.Tensor 25 | import Web.Semantic.DL.Category.Properties.Tensor.AssocNatural 26 | import Web.Semantic.DL.Category.Properties.Tensor.Coherence 27 | import Web.Semantic.DL.Category.Properties.Tensor.Functor 28 | import Web.Semantic.DL.Category.Properties.Tensor.Isomorphisms 29 | import Web.Semantic.DL.Category.Properties.Tensor.Lemmas 30 | import Web.Semantic.DL.Category.Properties.Tensor.RespectsEquiv 31 | import Web.Semantic.DL.Category.Properties.Tensor.RespectsWiring 32 | import Web.Semantic.DL.Category.Properties.Tensor.SymmNatural 33 | import Web.Semantic.DL.Category.Properties.Tensor.UnitNatural 34 | import Web.Semantic.DL.Category.Tensor 35 | import Web.Semantic.DL.Category.Unit 36 | import Web.Semantic.DL.Category.Wiring 37 | import Web.Semantic.DL.Concept 38 | import Web.Semantic.DL.Concept.Model 39 | import Web.Semantic.DL.Concept.Skolemization 40 | import Web.Semantic.DL.FOL 41 | import Web.Semantic.DL.FOL.Model 42 | import Web.Semantic.DL.Integrity 43 | import Web.Semantic.DL.Integrity.Closed 44 | import Web.Semantic.DL.Integrity.Closed.Alternate 45 | import Web.Semantic.DL.Integrity.Closed.Properties 46 | import Web.Semantic.DL.KB 47 | import Web.Semantic.DL.KB.Model 48 | import Web.Semantic.DL.KB.Skolemization 49 | import Web.Semantic.DL.Role 50 | import Web.Semantic.DL.Role.Model 51 | import Web.Semantic.DL.Role.Skolemization 52 | import Web.Semantic.DL.Sequent 53 | import Web.Semantic.DL.Sequent.Model 54 | import Web.Semantic.DL.Signature 55 | import Web.Semantic.DL.TBox 56 | import Web.Semantic.DL.TBox.Interp 57 | import Web.Semantic.DL.TBox.Interp.Morphism 58 | import Web.Semantic.DL.TBox.Minimizable 59 | import Web.Semantic.DL.TBox.Model 60 | import Web.Semantic.DL.TBox.Skolemization 61 | -------------------------------------------------------------------------------- /src/Web/Semantic/Util.agda: -------------------------------------------------------------------------------- 1 | {-# OPTIONS --universe-polymorphism #-} 2 | 3 | open import Data.Bool using ( Bool ; true ; false ; _∧_ ) 4 | open import Data.Empty using () 5 | open import Data.List using ( List ; [] ; _∷_ ; _++_ ; map ) 6 | open import Data.List.Any using ( here ; there ) 7 | open import Data.List.Membership.Propositional using () 8 | open import Data.Sum using ( _⊎_ ; inj₁ ; inj₂ ) 9 | open import Data.Product using ( ∃ ; ∄ ; _×_ ; _,_ ) 10 | open import Data.Unit using () 11 | open import Level using ( zero ) 12 | open import Relation.Binary using () 13 | open import Relation.Binary.PropositionalEquality using ( _≡_ ; refl ) 14 | open import Relation.Nullary using ( Dec ; yes ; no ) 15 | open import Relation.Unary using ( _∈_ ; _⊆_ ) 16 | 17 | module Web.Semantic.Util where 18 | 19 | infixr 9 _∘_ 20 | 21 | id : ∀ {X : Set} → X → X 22 | id x = x 23 | 24 | _∘_ : ∀ {X Y Z : Set} → (Y → Z) → (X → Y) → (X → Z) 25 | (f ∘ g) = λ x → f (g x) 26 | 27 | Setoid : Set₁ 28 | Setoid = Relation.Binary.Setoid zero zero 29 | 30 | Subset : Set → Set₁ 31 | Subset X = X → Set 32 | 33 | ⁅_⁆ : ∀ {X} → X → Subset X 34 | ⁅ x ⁆ y = x ≡ y 35 | 36 | ⊆-refl : ∀ {X} (P : Subset X) → (P ⊆ P) 37 | ⊆-refl P x∈P = x∈P 38 | 39 | _⁻¹ : ∀ {X Y : Set} → Subset (X × Y) → Subset (Y × X) 40 | (R ⁻¹) (y , x) = R (x , y) 41 | 42 | -- Some proofs are classical, and depend on excluded middle. 43 | 44 | ExclMiddle : Set₁ 45 | ExclMiddle = ∀ (X : Set) → Dec X 46 | 47 | ExclMiddle₁ : Set₂ 48 | ExclMiddle₁ = ∀ (X : Set₁) → Dec X 49 | 50 | smaller-excl-middle : ExclMiddle₁ → ExclMiddle 51 | smaller-excl-middle excl-middle₁ X = X? where 52 | 53 | data Large : Set₁ where 54 | large : X → Large 55 | 56 | X? : Dec X 57 | X? with excl-middle₁ Large 58 | X? | yes (large x) = yes x 59 | X? | no ¬x = no (λ x → ¬x (large x)) 60 | 61 | -- Some nameclashes between the standard library and semantic web terminology: 62 | -- ⊤ and ⊥ are used for concepts, and T is used to range over T-Boxes. 63 | 64 | open Data.Bool public using () renaming ( T to □ ) 65 | open Data.Empty public using () renaming ( ⊥ to False ; ⊥-elim to elim ) 66 | open Data.Unit public using ( tt ) renaming ( ⊤ to True ) 67 | 68 | □-proj₁ : ∀ {b c} → □(b ∧ c) → □ b 69 | □-proj₁ {true} □c = tt 70 | □-proj₁ {false} () 71 | 72 | □-proj₂ : ∀ {b c} → □(b ∧ c) → □ c 73 | □-proj₂ {true} □c = □c 74 | □-proj₂ {false} () 75 | 76 | -- Convert back and forth from Dec and Bool. 77 | 78 | is : ∀ {ℓ} {X : Set ℓ} → Dec X → Bool 79 | is (yes _) = true 80 | is (no _) = false 81 | 82 | is✓ : ∀ {ℓ} {X : Set ℓ} {x : Dec X} → □(is x) → X 83 | is✓ {ℓ} {X} {yes x} _ = x 84 | is✓ {ℓ} {X} {no _} () 85 | 86 | is! : ∀ {ℓ} {X : Set ℓ} {x : Dec X} → X → □(is x) 87 | is! {ℓ} {X} {yes _} x = tt 88 | is! {ℓ} {X} {no ¬x} x = ¬x x 89 | 90 | -- Finite sets are ones backed by a list 91 | 92 | open Data.List.Membership.Propositional public using () renaming ( _∈_ to _∈ˡ_ ) 93 | 94 | Finite : Set → Set 95 | Finite X = ∃ λ xs → ∀ (x : X) → (x ∈ˡ xs) 96 | 97 | False∈Fin : False ∈ Finite 98 | False∈Fin = ([] , λ ()) 99 | 100 | ⊎-resp-Fin : ∀ {X Y} → (X ∈ Finite) → (Y ∈ Finite) → ((X ⊎ Y) ∈ Finite) 101 | ⊎-resp-Fin {X} {Y} (xs , ∀x∙x∈xs) (ys , ∀y∙y∈ys) = 102 | ((map inj₁ xs ++ map inj₂ ys) , lemma) where 103 | 104 | lemma₁ : ∀ {x : X} {xs : List X} → 105 | (x ∈ˡ xs) → (inj₁ x ∈ˡ (map inj₁ xs ++ map inj₂ ys)) 106 | lemma₁ (here refl) = here refl 107 | lemma₁ (there x∈xs) = there (lemma₁ x∈xs) 108 | 109 | lemma₂ : ∀ (xs : List X) {y : Y} {ys : List Y} → 110 | (y ∈ˡ ys) → (inj₂ y ∈ˡ (map inj₁ xs ++ map inj₂ ys)) 111 | lemma₂ [] (here refl) = here refl 112 | lemma₂ [] (there y∈ys) = there (lemma₂ [] y∈ys) 113 | lemma₂ (x ∷ xs) y∈ys = there (lemma₂ xs y∈ys) 114 | 115 | lemma : ∀ x → (x ∈ˡ (map inj₁ xs ++ map inj₂ ys)) 116 | lemma (inj₁ x) = lemma₁ (∀x∙x∈xs x) 117 | lemma (inj₂ y) = lemma₂ xs (∀y∙y∈ys y) 118 | 119 | -- symmetric monoidal structure of sum 120 | 121 | _⟨⊎⟩_ : ∀ {W X Y Z : Set} → (W → X) → (Y → Z) → (W ⊎ Y) → (X ⊎ Z) 122 | _⟨⊎⟩_ f g (inj₁ x) = inj₁ (f x) 123 | _⟨⊎⟩_ f g (inj₂ y) = inj₂ (g y) 124 | 125 | _≡⊎≡_ : ∀ {X Y Z : Set} {f g : (X ⊎ Y) → Z} → 126 | (∀ x → (f (inj₁ x) ≡ g (inj₁ x))) → (∀ x → (f (inj₂ x) ≡ g (inj₂ x))) → 127 | ∀ x → (f x ≡ g x) 128 | (f₁≡g₁ ≡⊎≡ f₂≡g₂) (inj₁ x) = f₁≡g₁ x 129 | (f₁≡g₁ ≡⊎≡ f₂≡g₂) (inj₂ y) = f₂≡g₂ y 130 | 131 | ⊎-swap : ∀ {X Y : Set} → (X ⊎ Y) → (Y ⊎ X) 132 | ⊎-swap (inj₁ x) = inj₂ x 133 | ⊎-swap (inj₂ y) = inj₁ y 134 | 135 | ⊎-assoc : ∀ {X Y Z : Set} → ((X ⊎ Y) ⊎ Z) → (X ⊎ (Y ⊎ Z)) 136 | ⊎-assoc (inj₁ (inj₁ x)) = inj₁ x 137 | ⊎-assoc (inj₁ (inj₂ y)) = inj₂ (inj₁ y) 138 | ⊎-assoc (inj₂ z) = inj₂ (inj₂ z) 139 | 140 | ⊎-assoc⁻¹ : ∀ {X Y Z : Set} → (X ⊎ (Y ⊎ Z)) → ((X ⊎ Y) ⊎ Z) 141 | ⊎-assoc⁻¹ (inj₁ x) = inj₁ (inj₁ x) 142 | ⊎-assoc⁻¹ (inj₂ (inj₁ y)) = inj₁ (inj₂ y) 143 | ⊎-assoc⁻¹ (inj₂ (inj₂ z)) = inj₂ z 144 | 145 | ⊎-unit₁ : ∀ {X : Set} → (False ⊎ X) → X 146 | ⊎-unit₁ (inj₁ ()) 147 | ⊎-unit₁ (inj₂ x) = x 148 | 149 | ⊎-unit₂ : ∀ {X : Set} → (X ⊎ False) → X 150 | ⊎-unit₂ (inj₁ x) = x 151 | ⊎-unit₂ (inj₂ ()) 152 | 153 | inj⁻¹ : ∀ {X : Set} → (X ⊎ X) → X 154 | inj⁻¹ (inj₁ x) = x 155 | inj⁻¹ (inj₂ x) = x 156 | 157 | -- A set divided, like Gaul, into three parts 158 | 159 | infix 6 _⊕_⊕_ 160 | 161 | data _⊕_⊕_ (X V Y : Set) : Set where 162 | inode : (x : X) → (X ⊕ V ⊕ Y) -- Imported node 163 | bnode : (v : V) → (X ⊕ V ⊕ Y) -- Blank node 164 | enode : (y : Y) → (X ⊕ V ⊕ Y) -- Exported node 165 | 166 | _⟨⊕⟩_⟨⊕⟩_ : ∀ {U V W X Y Z} → (W → X) → (U → V) → (Y → Z) → 167 | (W ⊕ U ⊕ Y) → (X ⊕ V ⊕ Z) 168 | (f ⟨⊕⟩ g ⟨⊕⟩ h) (inode w) = inode (f w) 169 | (f ⟨⊕⟩ g ⟨⊕⟩ h) (bnode u) = bnode (g u) 170 | (f ⟨⊕⟩ g ⟨⊕⟩ h) (enode y) = enode (h y) 171 | 172 | _[⊕]_[⊕]_ : ∀ {X V Y Z : Set} → (X → Z) → (V → Z) → (Y → Z) → 173 | (X ⊕ V ⊕ Y) → Z 174 | (f [⊕] g [⊕] h) (inode x) = f x 175 | (f [⊕] g [⊕] h) (bnode v) = g v 176 | (f [⊕] g [⊕] h) (enode y) = h y 177 | 178 | left : ∀ {V W X Y Z} → (X ⊕ V ⊕ Y) → (X ⊕ (V ⊕ Y ⊕ W) ⊕ Z) 179 | left (inode x) = inode x 180 | left (bnode v) = bnode (inode v) 181 | left (enode y) = bnode (bnode y) 182 | 183 | right : ∀ {V W X Y Z} → (Y ⊕ W ⊕ Z) → (X ⊕ (V ⊕ Y ⊕ W) ⊕ Z) 184 | right (inode y) = bnode (bnode y) 185 | right (bnode w) = bnode (enode w) 186 | right (enode z) = enode z 187 | 188 | hmerge : ∀ {V W X Y Z A : Set} → 189 | ((X ⊕ V ⊕ Y) → A) → ((Y ⊕ W ⊕ Z) → A) → 190 | (X ⊕ (V ⊕ Y ⊕ W) ⊕ Z) → A 191 | hmerge f g (inode x) = f (inode x) 192 | hmerge f g (bnode (inode v)) = f (bnode v) 193 | hmerge f g (bnode (bnode y)) = g (inode y) 194 | hmerge f g (bnode (enode w)) = g (bnode w) 195 | hmerge f g (enode z) = g (enode z) 196 | 197 | →-dist-⊕ : ∀ {V X Y Z : Set} → ((X ⊕ V ⊕ Y) → Z) → 198 | ((X → Z) × (V → Z) × (Y → Z)) 199 | →-dist-⊕ i = ((i ∘ inode) , (i ∘ bnode) , (i ∘ enode)) 200 | 201 | up : ∀ {V₁ V₂ X₁ X₂ Y₁ Y₂} → (X₁ ⊕ V₁ ⊕ Y₁) → 202 | ((X₁ ⊎ X₂) ⊕ (V₁ ⊎ V₂) ⊕ (Y₁ ⊎ Y₂)) 203 | up (inode x) = inode (inj₁ x) 204 | up (bnode v) = bnode (inj₁ v) 205 | up (enode y) = enode (inj₁ y) 206 | 207 | down : ∀ {V₁ V₂ X₁ X₂ Y₁ Y₂} → (X₂ ⊕ V₂ ⊕ Y₂) → 208 | ((X₁ ⊎ X₂) ⊕ (V₁ ⊎ V₂) ⊕ (Y₁ ⊎ Y₂)) 209 | down (inode x) = inode (inj₂ x) 210 | down (bnode v) = bnode (inj₂ v) 211 | down (enode y) = enode (inj₂ y) 212 | 213 | vmerge : ∀ {V₁ V₂ X₁ X₂ Y₁ Y₂ A : Set} → 214 | ((X₁ ⊕ V₁ ⊕ Y₁) → A) → ((X₂ ⊕ V₂ ⊕ Y₂) → A) → 215 | ((X₁ ⊎ X₂) ⊕ (V₁ ⊎ V₂) ⊕ (Y₁ ⊎ Y₂)) → A 216 | vmerge j k (inode (inj₁ x)) = j (inode x) 217 | vmerge j k (inode (inj₂ x)) = k (inode x) 218 | vmerge j k (bnode (inj₁ v)) = j (bnode v) 219 | vmerge j k (bnode (inj₂ v)) = k (bnode v) 220 | vmerge j k (enode (inj₁ y)) = j (enode y) 221 | vmerge j k (enode (inj₂ y)) = k (enode y) 222 | --------------------------------------------------------------------------------